aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
-rw-r--r--CHANGES.md3
-rw-r--r--META.coq.in20
-rw-r--r--Makefile.common5
-rw-r--r--doc/sphinx/user-extensions/syntax-extensions.rst98
-rw-r--r--doc/stdlib/index-list.html.template2
-rw-r--r--interp/notation.ml285
-rw-r--r--interp/notation.mli24
-rw-r--r--plugins/extraction/ExtrHaskellString.v20
-rw-r--r--plugins/extraction/ExtrOcamlString.v16
-rw-r--r--plugins/syntax/ascii_syntax.ml100
-rw-r--r--plugins/syntax/ascii_syntax_plugin.mlpack1
-rw-r--r--plugins/syntax/g_string.mlg25
-rw-r--r--plugins/syntax/numeral.ml2
-rw-r--r--plugins/syntax/plugin_base.dune22
-rw-r--r--plugins/syntax/string_notation.ml98
-rw-r--r--plugins/syntax/string_notation.mli16
-rw-r--r--plugins/syntax/string_notation_plugin.mlpack2
-rw-r--r--plugins/syntax/string_syntax.ml81
-rw-r--r--plugins/syntax/string_syntax_plugin.mlpack1
-rw-r--r--proofs/logic.ml2
-rw-r--r--test-suite/output/Search.out16
-rw-r--r--test-suite/output/StringSyntax.out1089
-rw-r--r--test-suite/output/StringSyntax.v52
-rw-r--r--theories/Init/Byte.v830
-rw-r--r--theories/Init/Prelude.v5
-rw-r--r--theories/Strings/Ascii.v68
-rw-r--r--theories/Strings/Byte.v1214
-rw-r--r--theories/Strings/String.v86
-rw-r--r--vernac/explainErr.ml4
-rw-r--r--vernac/himsg.ml6
-rw-r--r--vernac/himsg.mli2
31 files changed, 3851 insertions, 344 deletions
diff --git a/CHANGES.md b/CHANGES.md
index 75a29de8e9..8cb71f19ea 100644
--- a/CHANGES.md
+++ b/CHANGES.md
@@ -46,6 +46,9 @@ Notations
`Bind Scope`, `Delimit Scope`, `Undelimit Scope`, or `Notation` is
deprecated.
+- New command `String Notation` to register string syntax for custom
+ inductive types.
+
Plugins
- The quote plugin (https://coq.inria.fr/distrib/V8.8.1/refman/proof-engine/detailed-tactic-examples.html#quote)
diff --git a/META.coq.in b/META.coq.in
index c2d3f85b9f..159984d87a 100644
--- a/META.coq.in
+++ b/META.coq.in
@@ -459,28 +459,16 @@ package "plugins" (
archive(native) = "int31_syntax_plugin.cmx"
)
- package "asciisyntax" (
+ package "string_notation" (
- description = "Coq asciisyntax plugin"
+ description = "Coq string_notation plugin"
version = "8.10"
requires = ""
directory = "syntax"
- archive(byte) = "ascii_syntax_plugin.cmo"
- archive(native) = "ascii_syntax_plugin.cmx"
- )
-
- package "stringsyntax" (
-
- description = "Coq stringsyntax plugin"
- version = "8.10"
-
- requires = "coq.plugins.asciisyntax"
- directory = "syntax"
-
- archive(byte) = "string_syntax_plugin.cmo"
- archive(native) = "string_syntax_plugin.cmx"
+ archive(byte) = "string_notation_plugin.cmo"
+ archive(native) = "string_notation_plugin.cmx"
)
package "derive" (
diff --git a/Makefile.common b/Makefile.common
index a59fbe676e..9f7ed9d46e 100644
--- a/Makefile.common
+++ b/Makefile.common
@@ -140,9 +140,8 @@ RTAUTOCMO:=plugins/rtauto/rtauto_plugin.cmo
SYNTAXCMO:=$(addprefix plugins/syntax/, \
r_syntax_plugin.cmo \
int31_syntax_plugin.cmo \
- ascii_syntax_plugin.cmo \
- string_syntax_plugin.cmo \
- numeral_notation_plugin.cmo)
+ numeral_notation_plugin.cmo \
+ string_notation_plugin.cmo)
DERIVECMO:=plugins/derive/derive_plugin.cmo
LTACCMO:=plugins/ltac/ltac_plugin.cmo plugins/ltac/tauto_plugin.cmo
SSRMATCHINGCMO:=plugins/ssrmatching/ssrmatching_plugin.cmo
diff --git a/doc/sphinx/user-extensions/syntax-extensions.rst b/doc/sphinx/user-extensions/syntax-extensions.rst
index a5869055fa..a54da6faf2 100644
--- a/doc/sphinx/user-extensions/syntax-extensions.rst
+++ b/doc/sphinx/user-extensions/syntax-extensions.rst
@@ -1583,6 +1583,104 @@ Numeral notations
As noted above, the :n:`(abstract after @num)` directive has no
effect when :n:`@ident__2` lands in an :g:`option` type.
+String notations
+-----------------
+
+.. cmd:: String Notation @ident__1 @ident__2 @ident__3 : @scope.
+ :name: String Notation
+
+ This command allows the user to customize the way strings are parsed
+ and printed.
+
+ The token :n:`@ident__1` should be the name of an inductive type,
+ while :n:`@ident__2` and :n:`@ident__3` should be the names of the
+ parsing and printing functions, respectively. The parsing function
+ :n:`@ident__2` should have one of the following types:
+
+ * :n:`Byte.byte -> @ident__1`
+ * :n:`Byte.byte -> option @ident__1`
+ * :n:`list Byte.byte -> @ident__1`
+ * :n:`list Byte.byte -> option @ident__1`
+
+ And the printing function :n:`@ident__3` should have one of the
+ following types:
+
+ * :n:`@ident__1 -> Byte.byte`
+ * :n:`@ident__1 -> option Byte.byte`
+ * :n:`@ident__1 -> list Byte.byte`
+ * :n:`@ident__1 -> option (list Byte.byte)`
+
+ When parsing, the application of the parsing function
+ :n:`@ident__2` to the string will be fully reduced, and universes
+ of the resulting term will be refreshed.
+
+ .. exn:: Cannot interpret this string as a value of type @type
+
+ The string notation registered for :token:`type` does not support
+ the given string. This error is given when the interpretation
+ function returns :g:`None`.
+
+ .. exn:: @ident should go from Byte.byte or (list Byte.byte) to @type or (option @type).
+
+ The parsing function given to the :cmd:`String Notation`
+ vernacular is not of the right type.
+
+ .. exn:: @ident should go from @type to Byte.byte or (option Byte.byte) or (list Byte.byte) or (option (list Byte.byte)).
+
+ The printing function given to the :cmd:`String Notation`
+ vernacular is not of the right type.
+
+ .. exn:: @type is not an inductive type.
+
+ String notations can only be declared for inductive types with no
+ arguments.
+
+ .. exn:: Unexpected term @term while parsing a string notation.
+
+ Parsing functions must always return ground terms, made up of
+ applications of constructors and inductive types. Parsing
+ functions may not return terms containing axioms, bare
+ (co)fixpoints, lambdas, etc.
+
+ .. exn:: Unexpected non-option term @term while parsing a string notation.
+
+ Parsing functions expected to return an :g:`option` must always
+ return a concrete :g:`Some` or :g:`None` when applied to a
+ concrete string expressed as a decimal. They may not return
+ opaque constants.
+
+ .. exn:: Cannot interpret in @scope because @ident could not be found in the current environment.
+
+ The inductive type used to register the string notation is no
+ longer available in the environment. Most likely, this is because
+ the string notation was declared inside a functor for an
+ inductive type inside the functor. This use case is not currently
+ supported.
+
+ Alternatively, you might be trying to use a primitive token
+ notation from a plugin which forgot to specify which module you
+ must :g:`Require` for access to that notation.
+
+ .. exn:: Syntax error: [prim:reference] expected after 'Notation' (in [vernac:command]).
+
+ The type passed to :cmd:`String Notation` must be a single
+ identifier.
+
+ .. exn:: Syntax error: [prim:reference] expected after [prim:reference] (in [vernac:command]).
+
+ Both functions passed to :cmd:`String Notation` must be single
+ identifiers.
+
+ .. exn:: The reference @ident was not found in the current environment.
+
+ Identifiers passed to :cmd:`String Notation` must exist in the
+ global environment.
+
+ .. exn:: @ident is bound to a notation that does not denote a reference.
+
+ Identifiers passed to :cmd:`String Notation` must be global
+ references, or notations which denote to single identifiers.
+
.. _TacticNotation:
Tactic Notations
diff --git a/doc/stdlib/index-list.html.template b/doc/stdlib/index-list.html.template
index 4fc9bf9e19..51f94d7e5a 100644
--- a/doc/stdlib/index-list.html.template
+++ b/doc/stdlib/index-list.html.template
@@ -17,6 +17,7 @@ through the <tt>Require Import</tt> command.</p>
theories/Init/Datatypes.v
theories/Init/Logic.v
theories/Init/Logic_Type.v
+ theories/Init/Byte.v
theories/Init/Nat.v
theories/Init/Decimal.v
theories/Init/Peano.v
@@ -497,6 +498,7 @@ through the <tt>Require Import</tt> command.</p>
Implementation of string as list of ascii characters
</dt>
<dd>
+ theories/Strings/Byte.v
theories/Strings/Ascii.v
theories/Strings/String.v
theories/Strings/BinaryString.v
diff --git a/interp/notation.ml b/interp/notation.ml
index e17cc65042..c866929234 100644
--- a/interp/notation.ml
+++ b/interp/notation.ml
@@ -530,11 +530,11 @@ let prim_token_uninterpreters =
(*******************************************************)
(* Numeral notation interpretation *)
-type numeral_notation_error =
+type prim_token_notation_error =
| UnexpectedTerm of Constr.t
| UnexpectedNonOptionTerm of Constr.t
-exception NumeralNotationError of Environ.env * Evd.evar_map * numeral_notation_error
+exception PrimTokenNotationError of string * Environ.env * Evd.evar_map * prim_token_notation_error
type numnot_option =
| Nop
@@ -554,20 +554,26 @@ type target_kind =
| UInt of Names.inductive (* Coq.Init.Decimal.uint *)
| Z of z_pos_ty (* Coq.Numbers.BinNums.Z and positive *)
+type string_target_kind =
+ | ListByte
+ | Byte
+
type option_kind = Option | Direct
-type conversion_kind = target_kind * option_kind
+type 'target conversion_kind = 'target * option_kind
-type numeral_notation_obj =
- { to_kind : conversion_kind;
+type ('target, 'warning) prim_token_notation_obj =
+ { to_kind : 'target conversion_kind;
to_ty : GlobRef.t;
- of_kind : conversion_kind;
+ of_kind : 'target conversion_kind;
of_ty : GlobRef.t;
- num_ty : Libnames.qualid; (* for warnings / error messages *)
- warning : numnot_option }
+ ty_name : Libnames.qualid; (* for warnings / error messages *)
+ warning : 'warning }
-module Numeral = struct
-(** * Numeral notation *)
+type numeral_notation_obj = (target_kind, numnot_option) prim_token_notation_obj
+type string_notation_obj = (string_target_kind, unit) prim_token_notation_obj
+module PrimTokenNotation = struct
+(** * Code shared between Numeral notation and String notation *)
(** Reduction
The constr [c] below isn't necessarily well-typed, since we
@@ -596,7 +602,69 @@ let eval_constr env sigma (c : Constr.t) =
let eval_constr_app env sigma c1 c2 =
eval_constr env sigma (mkApp (c1,[| c2 |]))
-exception NotANumber
+exception NotAValidPrimToken
+
+(** The uninterp function below work at the level of [glob_constr]
+ which is too low for us here. So here's a crude conversion back
+ to [constr] for the subset that concerns us. *)
+
+let rec constr_of_glob env sigma g = match DAst.get g with
+ | Glob_term.GRef (ConstructRef c, _) ->
+ let sigma,c = Evd.fresh_constructor_instance env sigma c in
+ sigma,mkConstructU c
+ | Glob_term.GApp (gc, gcl) ->
+ let sigma,c = constr_of_glob env sigma gc in
+ let sigma,cl = List.fold_left_map (constr_of_glob env) sigma gcl in
+ sigma,mkApp (c, Array.of_list cl)
+ | _ ->
+ raise NotAValidPrimToken
+
+let rec glob_of_constr token_kind ?loc env sigma c = match Constr.kind c with
+ | App (c, ca) ->
+ let c = glob_of_constr token_kind ?loc env sigma c in
+ let cel = List.map (glob_of_constr token_kind ?loc env sigma) (Array.to_list ca) in
+ DAst.make ?loc (Glob_term.GApp (c, cel))
+ | Construct (c, _) -> DAst.make ?loc (Glob_term.GRef (ConstructRef c, None))
+ | Const (c, _) -> DAst.make ?loc (Glob_term.GRef (ConstRef c, None))
+ | Ind (ind, _) -> DAst.make ?loc (Glob_term.GRef (IndRef ind, None))
+ | Var id -> DAst.make ?loc (Glob_term.GRef (VarRef id, None))
+ | _ -> Loc.raise ?loc (PrimTokenNotationError(token_kind,env,sigma,UnexpectedTerm c))
+
+let no_such_prim_token uninterpreted_token_kind ?loc ty =
+ CErrors.user_err ?loc
+ (str ("Cannot interpret this "^uninterpreted_token_kind^" as a value of type ") ++
+ pr_qualid ty)
+
+let interp_option uninterpreted_token_kind token_kind ty ?loc env sigma c =
+ match Constr.kind c with
+ | App (_Some, [| _; c |]) -> glob_of_constr token_kind ?loc env sigma c
+ | App (_None, [| _ |]) -> no_such_prim_token uninterpreted_token_kind ?loc ty
+ | x -> Loc.raise ?loc (PrimTokenNotationError(token_kind,env,sigma,UnexpectedNonOptionTerm c))
+
+let uninterp_option c =
+ match Constr.kind c with
+ | App (_Some, [| _; x |]) -> x
+ | _ -> raise NotAValidPrimToken
+
+let uninterp to_raw o (Glob_term.AnyGlobConstr n) =
+ let env = Global.env () in
+ let sigma = Evd.from_env env in
+ let sigma,of_ty = Evd.fresh_global env sigma o.of_ty in
+ let of_ty = EConstr.Unsafe.to_constr of_ty in
+ try
+ let sigma,n = constr_of_glob env sigma n in
+ let c = eval_constr_app env sigma of_ty n in
+ let c = if snd o.of_kind == Direct then c else uninterp_option c in
+ Some (to_raw (fst o.of_kind, c))
+ with
+ | Type_errors.TypeError _ | Pretype_errors.PretypeError _ -> None (* cf. eval_constr_app *)
+ | NotAValidPrimToken -> None (* all other functions except big2raw *)
+
+end
+
+module Numeral = struct
+(** * Numeral notation *)
+open PrimTokenNotation
let warn_large_num =
CWarnings.create ~name:"large-number" ~category:"numbers"
@@ -670,15 +738,15 @@ let rawnum_of_coquint c =
| Construct ((_,n), _) (* D0 to D9 *) ->
let () = Buffer.add_char buf (char_of_digit n) in
of_uint_loop a buf
- | _ -> raise NotANumber)
- | _ -> raise NotANumber
+ | _ -> raise NotAValidPrimToken)
+ | _ -> raise NotAValidPrimToken
in
let buf = Buffer.create 64 in
let () = of_uint_loop c buf in
if Int.equal (Buffer.length buf) 0 then
(* To avoid ambiguities between Nil and (D0 Nil), we choose
to not display Nil alone as "0" *)
- raise NotANumber
+ raise NotAValidPrimToken
else Buffer.contents buf
let rawnum_of_coqint c =
@@ -687,8 +755,8 @@ let rawnum_of_coqint c =
(match Constr.kind c with
| Construct ((_,1), _) (* Pos *) -> (rawnum_of_coquint c', true)
| Construct ((_,2), _) (* Neg *) -> (rawnum_of_coquint c', false)
- | _ -> raise NotANumber)
- | _ -> raise NotANumber
+ | _ -> raise NotAValidPrimToken)
+ | _ -> raise NotAValidPrimToken
(***********************************************************************)
@@ -718,9 +786,9 @@ let rec bigint_of_pos c = match Constr.kind c with
| 2 -> (* xO *) Bigint.mult_2 (bigint_of_pos d)
| n -> assert false (* no other constructor of type positive *)
end
- | x -> raise NotANumber
+ | x -> raise NotAValidPrimToken
end
- | x -> raise NotANumber
+ | x -> raise NotAValidPrimToken
(** Now, [Z] from/to bigint *)
@@ -745,51 +813,9 @@ let bigint_of_z z = match Constr.kind z with
| 3 -> (* Zneg *) Bigint.neg (bigint_of_pos d)
| n -> assert false (* no other constructor of type Z *)
end
- | _ -> raise NotANumber
+ | _ -> raise NotAValidPrimToken
end
- | _ -> raise NotANumber
-
-(** The uninterp function below work at the level of [glob_constr]
- which is too low for us here. So here's a crude conversion back
- to [constr] for the subset that concerns us. *)
-
-let rec constr_of_glob env sigma g = match DAst.get g with
- | Glob_term.GRef (ConstructRef c, _) ->
- let sigma,c = Evd.fresh_constructor_instance env sigma c in
- sigma,mkConstructU c
- | Glob_term.GApp (gc, gcl) ->
- let sigma,c = constr_of_glob env sigma gc in
- let sigma,cl = List.fold_left_map (constr_of_glob env) sigma gcl in
- sigma,mkApp (c, Array.of_list cl)
- | _ ->
- raise NotANumber
-
-let rec glob_of_constr ?loc env sigma c = match Constr.kind c with
- | App (c, ca) ->
- let c = glob_of_constr ?loc env sigma c in
- let cel = List.map (glob_of_constr ?loc env sigma) (Array.to_list ca) in
- DAst.make ?loc (Glob_term.GApp (c, cel))
- | Construct (c, _) -> DAst.make ?loc (Glob_term.GRef (ConstructRef c, None))
- | Const (c, _) -> DAst.make ?loc (Glob_term.GRef (ConstRef c, None))
- | Ind (ind, _) -> DAst.make ?loc (Glob_term.GRef (IndRef ind, None))
- | Var id -> DAst.make ?loc (Glob_term.GRef (VarRef id, None))
- | _ -> Loc.raise ?loc (NumeralNotationError(env,sigma,UnexpectedTerm c))
-
-let no_such_number ?loc ty =
- CErrors.user_err ?loc
- (str "Cannot interpret this number as a value of type " ++
- pr_qualid ty)
-
-let interp_option ty ?loc env sigma c =
- match Constr.kind c with
- | App (_Some, [| _; c |]) -> glob_of_constr ?loc env sigma c
- | App (_None, [| _ |]) -> no_such_number ?loc ty
- | x -> Loc.raise ?loc (NumeralNotationError(env,sigma,UnexpectedNonOptionTerm c))
-
-let uninterp_option c =
- match Constr.kind c with
- | App (_Some, [| _; x |]) -> x
- | _ -> raise NotANumber
+ | _ -> raise NotAValidPrimToken
let big2raw n =
if Bigint.is_pos_or_zero n then (Bigint.to_string n, true)
@@ -801,13 +827,13 @@ let raw2big (n,s) =
let interp o ?loc n =
begin match o.warning with
| Warning threshold when snd n && rawnum_compare (fst n) threshold >= 0 ->
- warn_large_num o.num_ty
+ warn_large_num o.ty_name
| _ -> ()
end;
let c = match fst o.to_kind with
| Int int_ty -> coqint_of_rawnum int_ty n
| UInt uint_ty when snd n -> coquint_of_rawnum uint_ty (fst n)
- | UInt _ (* n <= 0 *) -> no_such_number ?loc o.num_ty
+ | UInt _ (* n <= 0 *) -> no_such_prim_token "number" ?loc o.ty_name
| Z z_pos_ty -> z_of_bigint z_pos_ty (raw2big n)
in
let env = Global.env () in
@@ -816,30 +842,120 @@ let interp o ?loc n =
let to_ty = EConstr.Unsafe.to_constr to_ty in
match o.warning, snd o.to_kind with
| Abstract threshold, Direct when rawnum_compare (fst n) threshold >= 0 ->
- warn_abstract_large_num (o.num_ty,o.to_ty);
- glob_of_constr ?loc env sigma (mkApp (to_ty,[|c|]))
+ warn_abstract_large_num (o.ty_name,o.to_ty);
+ glob_of_constr "numeral" ?loc env sigma (mkApp (to_ty,[|c|]))
| _ ->
let res = eval_constr_app env sigma to_ty c in
match snd o.to_kind with
- | Direct -> glob_of_constr ?loc env sigma res
- | Option -> interp_option o.num_ty ?loc env sigma res
+ | Direct -> glob_of_constr "numeral" ?loc env sigma res
+ | Option -> interp_option "number" "numeral" o.ty_name ?loc env sigma res
+
+let uninterp o n =
+ PrimTokenNotation.uninterp
+ begin function
+ | (Int _, c) -> rawnum_of_coqint c
+ | (UInt _, c) -> (rawnum_of_coquint c, true)
+ | (Z _, c) -> big2raw (bigint_of_z c)
+ end o n
+end
+
+module Strings = struct
+(** * String notation *)
+open PrimTokenNotation
+
+let qualid_of_ref n =
+ n |> Coqlib.lib_ref |> Nametab.shortest_qualid_of_global Id.Set.empty
+
+let q_list () = qualid_of_ref "core.list.type"
+let q_byte () = qualid_of_ref "core.byte.type"
+
+let unsafe_locate_ind q =
+ match Nametab.locate q with
+ | IndRef i -> i
+ | _ -> raise Not_found
+
+let locate_list () = unsafe_locate_ind (q_list ())
+let locate_byte () = unsafe_locate_ind (q_byte ())
+
+(***********************************************************************)
+
+(** ** Conversion between Coq [list Byte.byte] and internal raw string *)
+
+let coqbyte_of_char_code byte c =
+ mkConstruct (byte, 1 + c)
+
+let coqbyte_of_string ?loc byte s =
+ let p =
+ if Int.equal (String.length s) 1 then int_of_char s.[0]
+ else
+ if Int.equal (String.length s) 3 && is_digit s.[0] && is_digit s.[1] && is_digit s.[2]
+ then int_of_string s
+ else
+ user_err ?loc ~hdr:"coqbyte_of_string"
+ (str "Expects a single character or a three-digits ascii code.") in
+ coqbyte_of_char_code byte p
+
+let coqbyte_of_char byte c = coqbyte_of_char_code byte (Char.code c)
+
+let make_ascii_string n =
+ if n>=32 && n<=126 then String.make 1 (char_of_int n)
+ else Printf.sprintf "%03d" n
+
+let char_code_of_coqbyte c =
+ match Constr.kind c with
+ | Construct ((_,c), _) -> c - 1
+ | _ -> raise NotAValidPrimToken
+
+let string_of_coqbyte c = make_ascii_string (char_code_of_coqbyte c)
+
+let coqlist_byte_of_string byte_ty list_ty str =
+ let cbyte = mkInd byte_ty in
+ let nil = mkApp (mkConstruct (list_ty, 1), [|cbyte|]) in
+ let cons x xs = mkApp (mkConstruct (list_ty, 2), [|cbyte; x; xs|]) in
+ let rec do_chars s i acc =
+ if i < 0 then acc
+ else
+ let b = coqbyte_of_char byte_ty s.[i] in
+ do_chars s (i-1) (cons b acc)
+ in
+ do_chars str (String.length str - 1) nil
+
+(* N.B. We rely on the fact that [nil] is the first constructor and [cons] is the second constructor, for [list] *)
+let string_of_coqlist_byte c =
+ let rec of_coqlist_byte_loop c buf =
+ match Constr.kind c with
+ | App (_nil, [|_ty|]) -> ()
+ | App (_cons, [|_ty;b;c|]) ->
+ let () = Buffer.add_char buf (Char.chr (char_code_of_coqbyte b)) in
+ of_coqlist_byte_loop c buf
+ | _ -> raise NotAValidPrimToken
+ in
+ let buf = Buffer.create 64 in
+ let () = of_coqlist_byte_loop c buf in
+ Buffer.contents buf
-let uninterp o (Glob_term.AnyGlobConstr n) =
+let interp o ?loc n =
+ let byte_ty = locate_byte () in
+ let list_ty = locate_list () in
+ let c = match fst o.to_kind with
+ | ListByte -> coqlist_byte_of_string byte_ty list_ty n
+ | Byte -> coqbyte_of_string ?loc byte_ty n
+ in
let env = Global.env () in
let sigma = Evd.from_env env in
- let sigma,of_ty = Evd.fresh_global env sigma o.of_ty in
- let of_ty = EConstr.Unsafe.to_constr of_ty in
- try
- let sigma,n = constr_of_glob env sigma n in
- let c = eval_constr_app env sigma of_ty n in
- let c = if snd o.of_kind == Direct then c else uninterp_option c in
- match fst o.of_kind with
- | Int _ -> Some (rawnum_of_coqint c)
- | UInt _ -> Some (rawnum_of_coquint c, true)
- | Z _ -> Some (big2raw (bigint_of_z c))
- with
- | Type_errors.TypeError _ | Pretype_errors.PretypeError _ -> None (* cf. eval_constr_app *)
- | NotANumber -> None (* all other functions except big2raw *)
+ let sigma,to_ty = Evd.fresh_global env sigma o.to_ty in
+ let to_ty = EConstr.Unsafe.to_constr to_ty in
+ let res = eval_constr_app env sigma to_ty c in
+ match snd o.to_kind with
+ | Direct -> glob_of_constr "string" ?loc env sigma res
+ | Option -> interp_option "string" "string" o.ty_name ?loc env sigma res
+
+let uninterp o n =
+ PrimTokenNotation.uninterp
+ begin function
+ | (ListByte, c) -> string_of_coqlist_byte c
+ | (Byte, c) -> string_of_coqbyte c
+ end o n
end
(* A [prim_token_infos], which is synchronized with the document
@@ -853,6 +969,7 @@ end
type prim_token_interp_info =
Uid of prim_token_uid
| NumeralNotation of numeral_notation_obj
+ | StringNotation of string_notation_obj
type prim_token_infos = {
pt_local : bool; (** Is this interpretation local? *)
@@ -1081,6 +1198,7 @@ let find_prim_token check_allowed ?loc p sc =
let interp = match info with
| Uid uid -> Hashtbl.find prim_token_interpreters uid
| NumeralNotation o -> InnerPrimToken.RawNumInterp (Numeral.interp o)
+ | StringNotation o -> InnerPrimToken.StringInterp (Strings.interp o)
in
let pat = InnerPrimToken.do_interp ?loc interp p in
check_allowed pat;
@@ -1270,6 +1388,7 @@ let uninterp_prim_token c =
let uninterp = match info with
| Uid uid -> Hashtbl.find prim_token_uninterpreters uid
| NumeralNotation o -> InnerPrimToken.RawNumUninterp (Numeral.uninterp o)
+ | StringNotation o -> InnerPrimToken.StringUninterp (Strings.uninterp o)
in
match InnerPrimToken.do_uninterp uninterp (AnyGlobConstr c) with
| None -> raise Notation_ops.No_match
@@ -1289,6 +1408,8 @@ let availability_of_prim_token n printer_scope local_scopes =
match n, uid with
| Numeral _, NumeralNotation _ -> true
| _, NumeralNotation _ -> false
+ | String _, StringNotation _ -> true
+ | _, StringNotation _ -> false
| _, Uid uid ->
let interp = Hashtbl.find prim_token_interpreters uid in
match n, interp with
diff --git a/interp/notation.mli b/interp/notation.mli
index ab0501a1e1..75034cad70 100644
--- a/interp/notation.mli
+++ b/interp/notation.mli
@@ -104,11 +104,11 @@ val register_string_interpretation :
(** * Numeral notation *)
-type numeral_notation_error =
+type prim_token_notation_error =
| UnexpectedTerm of Constr.t
| UnexpectedNonOptionTerm of Constr.t
-exception NumeralNotationError of Environ.env * Evd.evar_map * numeral_notation_error
+exception PrimTokenNotationError of string * Environ.env * Evd.evar_map * prim_token_notation_error
type numnot_option =
| Nop
@@ -128,20 +128,28 @@ type target_kind =
| UInt of Names.inductive (* Coq.Init.Decimal.uint *)
| Z of z_pos_ty (* Coq.Numbers.BinNums.Z and positive *)
+type string_target_kind =
+ | ListByte
+ | Byte
+
type option_kind = Option | Direct
-type conversion_kind = target_kind * option_kind
+type 'target conversion_kind = 'target * option_kind
-type numeral_notation_obj =
- { to_kind : conversion_kind;
+type ('target, 'warning) prim_token_notation_obj =
+ { to_kind : 'target conversion_kind;
to_ty : GlobRef.t;
- of_kind : conversion_kind;
+ of_kind : 'target conversion_kind;
of_ty : GlobRef.t;
- num_ty : Libnames.qualid; (* for warnings / error messages *)
- warning : numnot_option }
+ ty_name : Libnames.qualid; (* for warnings / error messages *)
+ warning : 'warning }
+
+type numeral_notation_obj = (target_kind, numnot_option) prim_token_notation_obj
+type string_notation_obj = (string_target_kind, unit) prim_token_notation_obj
type prim_token_interp_info =
Uid of prim_token_uid
| NumeralNotation of numeral_notation_obj
+ | StringNotation of string_notation_obj
type prim_token_infos = {
pt_local : bool; (** Is this interpretation local? *)
diff --git a/plugins/extraction/ExtrHaskellString.v b/plugins/extraction/ExtrHaskellString.v
index a4a40d3c5a..8c61f4e96b 100644
--- a/plugins/extraction/ExtrHaskellString.v
+++ b/plugins/extraction/ExtrHaskellString.v
@@ -6,6 +6,7 @@ Require Coq.extraction.Extraction.
Require Import Ascii.
Require Import String.
+Require Import Coq.Strings.Byte.
(**
* At the moment, Coq's extraction has no way to add extra import
@@ -40,3 +41,22 @@ Extract Inlined Constant Ascii.eqb => "(Prelude.==)".
Extract Inductive string => "Prelude.String" [ "([])" "(:)" ].
Extract Inlined Constant String.string_dec => "(Prelude.==)".
Extract Inlined Constant String.eqb => "(Prelude.==)".
+
+(* python -c 'print(" ".join(r""" "%s" """.strip() % (r"'"'\''"'" if chr(i) == "'"'"'" else repr(""" "" """.strip()) if chr(i) == """ " """.strip() else repr(chr(i))) for i in range(256)))' # " to satisfy Coq's comment parser *)
+Extract Inductive byte => "Prelude.Char"
+["'\x00'" "'\x01'" "'\x02'" "'\x03'" "'\x04'" "'\x05'" "'\x06'" "'\x07'" "'\x08'" "'\t'" "'\n'" "'\x0b'" "'\x0c'" "'\r'" "'\x0e'" "'\x0f'" "'\x10'" "'\x11'" "'\x12'" "'\x13'" "'\x14'" "'\x15'" "'\x16'" "'\x17'" "'\x18'" "'\x19'" "'\x1a'" "'\x1b'" "'\x1c'" "'\x1d'" "'\x1e'" "'\x1f'" "' '" "'!'" "'""'" "'#'" "'$'" "'%'" "'&'" "'\''" "'('" "')'" "'*'" "'+'" "','" "'-'" "'.'" "'/'" "'0'" "'1'" "'2'" "'3'" "'4'" "'5'" "'6'" "'7'" "'8'" "'9'" "':'" "';'" "'<'" "'='" "'>'" "'?'" "'@'" "'A'" "'B'" "'C'" "'D'" "'E'" "'F'" "'G'" "'H'" "'I'" "'J'" "'K'" "'L'" "'M'" "'N'" "'O'" "'P'" "'Q'" "'R'" "'S'" "'T'" "'U'" "'V'" "'W'" "'X'" "'Y'" "'Z'" "'['" "'\\'" "']'" "'^'" "'_'" "'`'" "'a'" "'b'" "'c'" "'d'" "'e'" "'f'" "'g'" "'h'" "'i'" "'j'" "'k'" "'l'" "'m'" "'n'" "'o'" "'p'" "'q'" "'r'" "'s'" "'t'" "'u'" "'v'" "'w'" "'x'" "'y'" "'z'" "'{'" "'|'" "'}'" "'~'" "'\x7f'" "'\x80'" "'\x81'" "'\x82'" "'\x83'" "'\x84'" "'\x85'" "'\x86'" "'\x87'" "'\x88'" "'\x89'" "'\x8a'" "'\x8b'" "'\x8c'" "'\x8d'" "'\x8e'" "'\x8f'" "'\x90'" "'\x91'" "'\x92'" "'\x93'" "'\x94'" "'\x95'" "'\x96'" "'\x97'" "'\x98'" "'\x99'" "'\x9a'" "'\x9b'" "'\x9c'" "'\x9d'" "'\x9e'" "'\x9f'" "'\xa0'" "'\xa1'" "'\xa2'" "'\xa3'" "'\xa4'" "'\xa5'" "'\xa6'" "'\xa7'" "'\xa8'" "'\xa9'" "'\xaa'" "'\xab'" "'\xac'" "'\xad'" "'\xae'" "'\xaf'" "'\xb0'" "'\xb1'" "'\xb2'" "'\xb3'" "'\xb4'" "'\xb5'" "'\xb6'" "'\xb7'" "'\xb8'" "'\xb9'" "'\xba'" "'\xbb'" "'\xbc'" "'\xbd'" "'\xbe'" "'\xbf'" "'\xc0'" "'\xc1'" "'\xc2'" "'\xc3'" "'\xc4'" "'\xc5'" "'\xc6'" "'\xc7'" "'\xc8'" "'\xc9'" "'\xca'" "'\xcb'" "'\xcc'" "'\xcd'" "'\xce'" "'\xcf'" "'\xd0'" "'\xd1'" "'\xd2'" "'\xd3'" "'\xd4'" "'\xd5'" "'\xd6'" "'\xd7'" "'\xd8'" "'\xd9'" "'\xda'" "'\xdb'" "'\xdc'" "'\xdd'" "'\xde'" "'\xdf'" "'\xe0'" "'\xe1'" "'\xe2'" "'\xe3'" "'\xe4'" "'\xe5'" "'\xe6'" "'\xe7'" "'\xe8'" "'\xe9'" "'\xea'" "'\xeb'" "'\xec'" "'\xed'" "'\xee'" "'\xef'" "'\xf0'" "'\xf1'" "'\xf2'" "'\xf3'" "'\xf4'" "'\xf5'" "'\xf6'" "'\xf7'" "'\xf8'" "'\xf9'" "'\xfa'" "'\xfb'" "'\xfc'" "'\xfd'" "'\xfe'" "'\xff'"].
+
+Extract Inlined Constant Byte.eqb => "(Prelude.==)".
+Extract Inlined Constant Byte.byte_eq_dec => "(Prelude.==)".
+Extract Inlined Constant Ascii.ascii_of_byte => "(\x -> x)".
+Extract Inlined Constant Ascii.byte_of_ascii => "(\x -> x)".
+
+(*
+Require Import ExtrHaskellBasic.
+Definition test := "ceci est un test"%string.
+Definition test2 := List.map (option_map Byte.to_nat) (List.map Byte.of_nat (List.seq 0 256)).
+Definition test3 := List.map ascii_of_nat (List.seq 0 256).
+
+Extraction Language Haskell.
+Recursive Extraction test Ascii.zero Ascii.one test2 test3 byte_rect.
+*)
diff --git a/plugins/extraction/ExtrOcamlString.v b/plugins/extraction/ExtrOcamlString.v
index a2a6a8fe67..f094d4860e 100644
--- a/plugins/extraction/ExtrOcamlString.v
+++ b/plugins/extraction/ExtrOcamlString.v
@@ -12,7 +12,7 @@
Require Coq.extraction.Extraction.
-Require Import Ascii String.
+Require Import Ascii String Coq.Strings.Byte.
Extract Inductive ascii => char
[
@@ -37,7 +37,19 @@ Extract Inlined Constant Ascii.eqb => "(=)".
Extract Inductive string => "char list" [ "[]" "(::)" ].
+(* python -c 'print(" ".join(r""" "%s" """.strip() % (r"'"'\''"'" if chr(i) == "'"'"'" else repr(""" "" """.strip()) if chr(i) == """ " """.strip() else repr(chr(i))) for i in range(256)))' # " to satisfy Coq's comment parser *)
+Extract Inductive byte => char
+["'\x00'" "'\x01'" "'\x02'" "'\x03'" "'\x04'" "'\x05'" "'\x06'" "'\x07'" "'\x08'" "'\t'" "'\n'" "'\x0b'" "'\x0c'" "'\r'" "'\x0e'" "'\x0f'" "'\x10'" "'\x11'" "'\x12'" "'\x13'" "'\x14'" "'\x15'" "'\x16'" "'\x17'" "'\x18'" "'\x19'" "'\x1a'" "'\x1b'" "'\x1c'" "'\x1d'" "'\x1e'" "'\x1f'" "' '" "'!'" "'""'" "'#'" "'$'" "'%'" "'&'" "'\''" "'('" "')'" "'*'" "'+'" "','" "'-'" "'.'" "'/'" "'0'" "'1'" "'2'" "'3'" "'4'" "'5'" "'6'" "'7'" "'8'" "'9'" "':'" "';'" "'<'" "'='" "'>'" "'?'" "'@'" "'A'" "'B'" "'C'" "'D'" "'E'" "'F'" "'G'" "'H'" "'I'" "'J'" "'K'" "'L'" "'M'" "'N'" "'O'" "'P'" "'Q'" "'R'" "'S'" "'T'" "'U'" "'V'" "'W'" "'X'" "'Y'" "'Z'" "'['" "'\\'" "']'" "'^'" "'_'" "'`'" "'a'" "'b'" "'c'" "'d'" "'e'" "'f'" "'g'" "'h'" "'i'" "'j'" "'k'" "'l'" "'m'" "'n'" "'o'" "'p'" "'q'" "'r'" "'s'" "'t'" "'u'" "'v'" "'w'" "'x'" "'y'" "'z'" "'{'" "'|'" "'}'" "'~'" "'\x7f'" "'\x80'" "'\x81'" "'\x82'" "'\x83'" "'\x84'" "'\x85'" "'\x86'" "'\x87'" "'\x88'" "'\x89'" "'\x8a'" "'\x8b'" "'\x8c'" "'\x8d'" "'\x8e'" "'\x8f'" "'\x90'" "'\x91'" "'\x92'" "'\x93'" "'\x94'" "'\x95'" "'\x96'" "'\x97'" "'\x98'" "'\x99'" "'\x9a'" "'\x9b'" "'\x9c'" "'\x9d'" "'\x9e'" "'\x9f'" "'\xa0'" "'\xa1'" "'\xa2'" "'\xa3'" "'\xa4'" "'\xa5'" "'\xa6'" "'\xa7'" "'\xa8'" "'\xa9'" "'\xaa'" "'\xab'" "'\xac'" "'\xad'" "'\xae'" "'\xaf'" "'\xb0'" "'\xb1'" "'\xb2'" "'\xb3'" "'\xb4'" "'\xb5'" "'\xb6'" "'\xb7'" "'\xb8'" "'\xb9'" "'\xba'" "'\xbb'" "'\xbc'" "'\xbd'" "'\xbe'" "'\xbf'" "'\xc0'" "'\xc1'" "'\xc2'" "'\xc3'" "'\xc4'" "'\xc5'" "'\xc6'" "'\xc7'" "'\xc8'" "'\xc9'" "'\xca'" "'\xcb'" "'\xcc'" "'\xcd'" "'\xce'" "'\xcf'" "'\xd0'" "'\xd1'" "'\xd2'" "'\xd3'" "'\xd4'" "'\xd5'" "'\xd6'" "'\xd7'" "'\xd8'" "'\xd9'" "'\xda'" "'\xdb'" "'\xdc'" "'\xdd'" "'\xde'" "'\xdf'" "'\xe0'" "'\xe1'" "'\xe2'" "'\xe3'" "'\xe4'" "'\xe5'" "'\xe6'" "'\xe7'" "'\xe8'" "'\xe9'" "'\xea'" "'\xeb'" "'\xec'" "'\xed'" "'\xee'" "'\xef'" "'\xf0'" "'\xf1'" "'\xf2'" "'\xf3'" "'\xf4'" "'\xf5'" "'\xf6'" "'\xf7'" "'\xf8'" "'\xf9'" "'\xfa'" "'\xfb'" "'\xfc'" "'\xfd'" "'\xfe'" "'\xff'"].
+
+Extract Inlined Constant Byte.eqb => "(=)".
+Extract Inlined Constant Byte.byte_eq_dec => "(=)".
+Extract Inlined Constant Ascii.ascii_of_byte => "(fun x -> x)".
+Extract Inlined Constant Ascii.byte_of_ascii => "(fun x -> x)".
+
(*
Definition test := "ceci est un test"%string.
-Recursive Extraction test Ascii.zero Ascii.one.
+Definition test2 := List.map (option_map Byte.to_nat) (List.map Byte.of_nat (List.seq 0 256)).
+Definition test3 := List.map ascii_of_nat (List.seq 0 256).
+
+Recursive Extraction test Ascii.zero Ascii.one test2 test3 byte_rect.
*)
diff --git a/plugins/syntax/ascii_syntax.ml b/plugins/syntax/ascii_syntax.ml
deleted file mode 100644
index 94255bab6c..0000000000
--- a/plugins/syntax/ascii_syntax.ml
+++ /dev/null
@@ -1,100 +0,0 @@
-(************************************************************************)
-(* * The Coq Proof Assistant / The Coq Development Team *)
-(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
-(* <O___,, * (see CREDITS file for the list of authors) *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(* * (see LICENSE file for the text of the license) *)
-(************************************************************************)
-
-
-(* Poor's man DECLARE PLUGIN *)
-let __coq_plugin_name = "ascii_syntax_plugin"
-let () = Mltop.add_known_module __coq_plugin_name
-
-open Pp
-open CErrors
-open Util
-open Names
-open Glob_term
-open Globnames
-open Coqlib
-
-exception Non_closed_ascii
-
-let make_dir l = DirPath.make (List.rev_map Id.of_string l)
-let make_path dir id = Libnames.make_path (make_dir dir) (Id.of_string id)
-
-let is_gr c gr = match DAst.get c with
-| GRef (r, _) -> GlobRef.equal r gr
-| _ -> false
-
-let ascii_module = ["Coq";"Strings";"Ascii"]
-let ascii_modpath = MPfile (make_dir ascii_module)
-
-let ascii_path = make_path ascii_module "ascii"
-
-let ascii_label = Label.make "ascii"
-let ascii_kn = MutInd.make2 ascii_modpath ascii_label
-let path_of_Ascii = ((ascii_kn,0),1)
-let static_glob_Ascii = ConstructRef path_of_Ascii
-
-let glob_Ascii = lazy (lib_ref "plugins.syntax.Ascii")
-
-open Lazy
-
-let interp_ascii ?loc p =
- let rec aux n p =
- if Int.equal n 0 then [] else
- let mp = p mod 2 in
- (DAst.make ?loc @@ GRef (lib_ref (if Int.equal mp 0 then "core.bool.false" else "core.bool.true"),None))
- :: (aux (n-1) (p/2)) in
- DAst.make ?loc @@ GApp (DAst.make ?loc @@ GRef(force glob_Ascii,None), aux 8 p)
-
-let interp_ascii_string ?loc s =
- let p =
- if Int.equal (String.length s) 1 then int_of_char s.[0]
- else
- if Int.equal (String.length s) 3 && is_digit s.[0] && is_digit s.[1] && is_digit s.[2]
- then int_of_string s
- else
- user_err ?loc ~hdr:"interp_ascii_string"
- (str "Expects a single character or a three-digits ascii code.") in
- interp_ascii ?loc p
-
-let uninterp_ascii r =
- let rec uninterp_bool_list n = function
- | [] when Int.equal n 0 -> 0
- | r::l when is_gr r (lib_ref "core.bool.true") -> 1+2*(uninterp_bool_list (n-1) l)
- | r::l when is_gr r (lib_ref "core.bool.false") -> 2*(uninterp_bool_list (n-1) l)
- | _ -> raise Non_closed_ascii in
- try
- let aux c = match DAst.get c with
- | GApp (r, l) when is_gr r (force glob_Ascii) -> uninterp_bool_list 8 l
- | _ -> raise Non_closed_ascii in
- Some (aux r)
- with
- Non_closed_ascii -> None
-
-let make_ascii_string n =
- if n>=32 && n<=126 then String.make 1 (char_of_int n)
- else Printf.sprintf "%03d" n
-
-let uninterp_ascii_string (AnyGlobConstr r) = Option.map make_ascii_string (uninterp_ascii r)
-
-open Notation
-
-let at_declare_ml_module f x =
- Mltop.declare_cache_obj (fun () -> f x) __coq_plugin_name
-
-let _ =
- let sc = "char_scope" in
- register_string_interpretation sc (interp_ascii_string,uninterp_ascii_string);
- at_declare_ml_module enable_prim_token_interpretation
- { pt_local = false;
- pt_scope = sc;
- pt_interp_info = Uid sc;
- pt_required = (ascii_path,ascii_module);
- pt_refs = [static_glob_Ascii];
- pt_in_match = true }
diff --git a/plugins/syntax/ascii_syntax_plugin.mlpack b/plugins/syntax/ascii_syntax_plugin.mlpack
deleted file mode 100644
index 7b9213a0e2..0000000000
--- a/plugins/syntax/ascii_syntax_plugin.mlpack
+++ /dev/null
@@ -1 +0,0 @@
-Ascii_syntax
diff --git a/plugins/syntax/g_string.mlg b/plugins/syntax/g_string.mlg
new file mode 100644
index 0000000000..1e06cd8ddb
--- /dev/null
+++ b/plugins/syntax/g_string.mlg
@@ -0,0 +1,25 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+DECLARE PLUGIN "string_notation_plugin"
+
+{
+
+open String_notation
+open Names
+open Stdarg
+
+}
+
+VERNAC COMMAND EXTEND StringNotation CLASSIFIED AS SIDEFF
+ | #[ locality = Attributes.locality; ] [ "String" "Notation" reference(ty) reference(f) reference(g) ":"
+ ident(sc) ] ->
+ { vernac_string_notation (Locality.make_module_locality locality) ty f g (Id.to_string sc) }
+END
diff --git a/plugins/syntax/numeral.ml b/plugins/syntax/numeral.ml
index 10a0af0b8f..470deb4a60 100644
--- a/plugins/syntax/numeral.ml
+++ b/plugins/syntax/numeral.ml
@@ -125,7 +125,7 @@ let vernac_numeral_notation local ty f g scope opts =
| None -> type_error_of g ty true
in
let o = { to_kind; to_ty; of_kind; of_ty;
- num_ty = ty;
+ ty_name = ty;
warning = opts }
in
(match opts, to_kind with
diff --git a/plugins/syntax/plugin_base.dune b/plugins/syntax/plugin_base.dune
index bfdd480fe9..1ab16c700d 100644
--- a/plugins/syntax/plugin_base.dune
+++ b/plugins/syntax/plugin_base.dune
@@ -6,6 +6,13 @@
(libraries coq.plugins.ltac))
(library
+ (name string_notation_plugin)
+ (public_name coq.plugins.string_notation)
+ (synopsis "Coq string notation plugin")
+ (modules g_string string_notation)
+ (libraries coq.vernac))
+
+(library
(name r_syntax_plugin)
(public_name coq.plugins.r_syntax)
(synopsis "Coq syntax plugin: reals")
@@ -13,23 +20,8 @@
(libraries coq.vernac))
(library
- (name ascii_syntax_plugin)
- (public_name coq.plugins.ascii_syntax)
- (synopsis "Coq syntax plugin: ASCII")
- (modules ascii_syntax)
- (libraries coq.vernac))
-
-(library
- (name string_syntax_plugin)
- (public_name coq.plugins.string_syntax)
- (synopsis "Coq syntax plugin: strings")
- (modules string_syntax)
- (libraries coq.plugins.ascii_syntax))
-
-(library
(name int31_syntax_plugin)
(public_name coq.plugins.int31_syntax)
(synopsis "Coq syntax plugin: int31")
(modules int31_syntax)
(libraries coq.vernac))
-
diff --git a/plugins/syntax/string_notation.ml b/plugins/syntax/string_notation.ml
new file mode 100644
index 0000000000..12ee4c6eda
--- /dev/null
+++ b/plugins/syntax/string_notation.ml
@@ -0,0 +1,98 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+open Pp
+open Util
+open Names
+open Libnames
+open Globnames
+open Constrexpr
+open Constrexpr_ops
+open Notation
+
+(** * String notation *)
+
+let get_constructors ind =
+ let mib,oib = Global.lookup_inductive ind in
+ let mc = oib.Declarations.mind_consnames in
+ Array.to_list
+ (Array.mapi (fun j c -> ConstructRef (ind, j + 1)) mc)
+
+let qualid_of_ref n =
+ n |> Coqlib.lib_ref |> Nametab.shortest_qualid_of_global Id.Set.empty
+
+let q_option () = qualid_of_ref "core.option.type"
+let q_list () = qualid_of_ref "core.list.type"
+let q_byte () = qualid_of_ref "core.byte.type"
+
+let has_type f ty =
+ let (sigma, env) = Pfedit.get_current_context () in
+ let c = mkCastC (mkRefC f, Glob_term.CastConv ty) in
+ try let _ = Constrintern.interp_constr env sigma c in true
+ with Pretype_errors.PretypeError _ -> false
+
+let type_error_to f ty =
+ CErrors.user_err
+ (pr_qualid f ++ str " should go from Byte.byte or (list Byte.byte) to " ++
+ pr_qualid ty ++ str " or (option " ++ pr_qualid ty ++ str ").")
+
+let type_error_of g ty =
+ CErrors.user_err
+ (pr_qualid g ++ str " should go from " ++ pr_qualid ty ++
+ str " to Byte.byte or (option Byte.byte) or (list Byte.byte) or (option (list Byte.byte)).")
+
+let vernac_string_notation local ty f g scope =
+ let app x y = mkAppC (x,[y]) in
+ let cref q = mkRefC q in
+ let cbyte = cref (q_byte ()) in
+ let clist = cref (q_list ()) in
+ let coption = cref (q_option ()) in
+ let opt r = app coption r in
+ let clist_byte = app clist cbyte in
+ let tyc = Smartlocate.global_inductive_with_alias ty in
+ let to_ty = Smartlocate.global_with_alias f in
+ let of_ty = Smartlocate.global_with_alias g in
+ let cty = cref ty in
+ let arrow x y =
+ mkProdC ([CAst.make Anonymous],Default Decl_kinds.Explicit, x, y)
+ in
+ let constructors = get_constructors tyc in
+ (* Check the type of f *)
+ let to_kind =
+ if has_type f (arrow clist_byte cty) then ListByte, Direct
+ else if has_type f (arrow clist_byte (opt cty)) then ListByte, Option
+ else if has_type f (arrow cbyte cty) then Byte, Direct
+ else if has_type f (arrow cbyte (opt cty)) then Byte, Option
+ else type_error_to f ty
+ in
+ (* Check the type of g *)
+ let of_kind =
+ if has_type g (arrow cty clist_byte) then ListByte, Direct
+ else if has_type g (arrow cty (opt clist_byte)) then ListByte, Option
+ else if has_type g (arrow cty cbyte) then Byte, Direct
+ else if has_type g (arrow cty (opt cbyte)) then Byte, Option
+ else type_error_of g ty
+ in
+ let o = { to_kind = to_kind;
+ to_ty = to_ty;
+ of_kind = of_kind;
+ of_ty = of_ty;
+ ty_name = ty;
+ warning = () }
+ in
+ let i =
+ { pt_local = local;
+ pt_scope = scope;
+ pt_interp_info = StringNotation o;
+ pt_required = Nametab.path_of_global (IndRef tyc),[];
+ pt_refs = constructors;
+ pt_in_match = true }
+ in
+ enable_prim_token_interpretation i
diff --git a/plugins/syntax/string_notation.mli b/plugins/syntax/string_notation.mli
new file mode 100644
index 0000000000..9a0174abf2
--- /dev/null
+++ b/plugins/syntax/string_notation.mli
@@ -0,0 +1,16 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+open Libnames
+open Vernacexpr
+
+(** * String notation *)
+
+val vernac_string_notation : locality_flag -> qualid -> qualid -> qualid -> Notation_term.scope_name -> unit
diff --git a/plugins/syntax/string_notation_plugin.mlpack b/plugins/syntax/string_notation_plugin.mlpack
new file mode 100644
index 0000000000..6aa081dab4
--- /dev/null
+++ b/plugins/syntax/string_notation_plugin.mlpack
@@ -0,0 +1,2 @@
+String_notation
+G_string
diff --git a/plugins/syntax/string_syntax.ml b/plugins/syntax/string_syntax.ml
deleted file mode 100644
index 59e65a0672..0000000000
--- a/plugins/syntax/string_syntax.ml
+++ /dev/null
@@ -1,81 +0,0 @@
-(************************************************************************)
-(* * The Coq Proof Assistant / The Coq Development Team *)
-(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
-(* <O___,, * (see CREDITS file for the list of authors) *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(* * (see LICENSE file for the text of the license) *)
-(************************************************************************)
-
-open Names
-open Globnames
-open Ascii_syntax_plugin.Ascii_syntax
-open Glob_term
-open Coqlib
-
-(* Poor's man DECLARE PLUGIN *)
-let __coq_plugin_name = "string_syntax_plugin"
-let () = Mltop.add_known_module __coq_plugin_name
-
-exception Non_closed_string
-
-(* make a string term from the string s *)
-
-let string_module = ["Coq";"Strings";"String"]
-
-let string_modpath = MPfile (make_dir string_module)
-let string_path = make_path string_module "string"
-
-let string_kn = MutInd.make2 string_modpath @@ Label.make "string"
-let static_glob_EmptyString = ConstructRef ((string_kn,0),1)
-let static_glob_String = ConstructRef ((string_kn,0),2)
-
-let glob_String = lazy (lib_ref "plugins.syntax.String")
-let glob_EmptyString = lazy (lib_ref "plugins.syntax.EmptyString")
-
-let is_gr c gr = match DAst.get c with
-| GRef (r, _) -> GlobRef.equal r gr
-| _ -> false
-
-open Lazy
-
-let interp_string ?loc s =
- let le = String.length s in
- let rec aux n =
- if n = le then DAst.make ?loc @@ GRef (force glob_EmptyString, None) else
- DAst.make ?loc @@ GApp (DAst.make ?loc @@ GRef (force glob_String, None),
- [interp_ascii ?loc (int_of_char s.[n]); aux (n+1)])
- in aux 0
-
-let uninterp_string (AnyGlobConstr r) =
- try
- let b = Buffer.create 16 in
- let rec aux c = match DAst.get c with
- | GApp (k,[a;s]) when is_gr k (force glob_String) ->
- (match uninterp_ascii a with
- | Some c -> Buffer.add_char b (Char.chr c); aux s
- | _ -> raise Non_closed_string)
- | GRef (z,_) when GlobRef.equal z (force glob_EmptyString) ->
- Some (Buffer.contents b)
- | _ ->
- raise Non_closed_string
- in aux r
- with
- Non_closed_string -> None
-
-open Notation
-
-let at_declare_ml_module f x =
- Mltop.declare_cache_obj (fun () -> f x) __coq_plugin_name
-
-let _ =
- let sc = "string_scope" in
- register_string_interpretation sc (interp_string,uninterp_string);
- at_declare_ml_module enable_prim_token_interpretation
- { pt_local = false;
- pt_scope = sc;
- pt_interp_info = Uid sc;
- pt_required = (string_path,["Coq";"Strings";"String"]);
- pt_refs = [static_glob_String; static_glob_EmptyString];
- pt_in_match = true }
diff --git a/plugins/syntax/string_syntax_plugin.mlpack b/plugins/syntax/string_syntax_plugin.mlpack
deleted file mode 100644
index 45d6e0fa23..0000000000
--- a/plugins/syntax/string_syntax_plugin.mlpack
+++ /dev/null
@@ -1 +0,0 @@
-String_syntax
diff --git a/proofs/logic.ml b/proofs/logic.ml
index 456ab879eb..3581e90b79 100644
--- a/proofs/logic.ml
+++ b/proofs/logic.ml
@@ -63,7 +63,7 @@ let catchable_exception = function
| CErrors.UserError _ | TypeError _
| Proof.OpenProof _
(* abstract will call close_proof inside a tactic *)
- | Notation.NumeralNotationError _
+ | Notation.PrimTokenNotationError _
| RefinerError _ | Indrec.RecursionSchemeError _
| Nametab.GlobalizationError _
(* reduction errors *)
diff --git a/test-suite/output/Search.out b/test-suite/output/Search.out
index 7446c17d98..f4544a0df3 100644
--- a/test-suite/output/Search.out
+++ b/test-suite/output/Search.out
@@ -34,17 +34,23 @@ bool_ind: forall P : bool -> Prop, P true -> P false -> forall b : bool, P b
bool_rec: forall P : bool -> Set, P true -> P false -> forall b : bool, P b
eq_true_rec:
forall P : bool -> Set, P true -> forall b : bool, eq_true b -> P b
-eq_true_ind:
- forall P : bool -> Prop, P true -> forall b : bool, eq_true b -> P b
+bool_rect: forall P : bool -> Type, P true -> P false -> forall b : bool, P b
eq_true_rect_r:
forall (P : bool -> Type) (b : bool), P b -> eq_true b -> P true
eq_true_rec_r:
forall (P : bool -> Set) (b : bool), P b -> eq_true b -> P true
eq_true_rect:
forall P : bool -> Type, P true -> forall b : bool, eq_true b -> P b
-bool_rect: forall P : bool -> Type, P true -> P false -> forall b : bool, P b
+eq_true_ind:
+ forall P : bool -> Prop, P true -> forall b : bool, eq_true b -> P b
eq_true_ind_r:
forall (P : bool -> Prop) (b : bool), P b -> eq_true b -> P true
+Byte.to_bits:
+ Byte.byte ->
+ bool * (bool * (bool * (bool * (bool * (bool * (bool * bool))))))
+Byte.of_bits:
+ bool * (bool * (bool * (bool * (bool * (bool * (bool * bool)))))) ->
+ Byte.byte
andb_true_intro:
forall b1 b2 : bool, b1 = true /\ b2 = true -> (b1 && b2)%bool = true
andb_prop: forall a b : bool, (a && b)%bool = true -> a = true /\ b = true
@@ -52,6 +58,10 @@ BoolSpec_ind:
forall (P Q : Prop) (P0 : bool -> Prop),
(P -> P0 true) ->
(Q -> P0 false) -> forall b : bool, BoolSpec P Q b -> P0 b
+Byte.to_bits_of_bits:
+ forall
+ b : bool * (bool * (bool * (bool * (bool * (bool * (bool * bool)))))),
+ Byte.to_bits (Byte.of_bits b) = b
bool_choice:
forall (S : Set) (R1 R2 : S -> Prop),
(forall x : S, {R1 x} + {R2 x}) ->
diff --git a/test-suite/output/StringSyntax.out b/test-suite/output/StringSyntax.out
new file mode 100644
index 0000000000..bbc936766d
--- /dev/null
+++ b/test-suite/output/StringSyntax.out
@@ -0,0 +1,1089 @@
+Monomorphic byte_rect =
+fun (P : byte -> Type) (f : P "000") (f0 : P "001") (f1 : P "002") (f2 : P "003") (f3 : P "004") (f4 : P "005") (f5 : P "006") (f6 : P "007") (f7 : P "008") (f8 : P "009") (f9 : P "010") (f10 : P "011") (f11 : P "012") (f12 : P "013") (f13 : P "014") (f14 : P "015") (f15 : P "016") (f16 : P "017") (f17 : P "018") (f18 : P "019") (f19 : P "020") (f20 : P "021") (f21 : P "022") (f22 : P "023") (f23 : P "024") (f24 : P "025") (f25 : P "026") (f26 : P "027") (f27 : P "028") (f28 : P "029") (f29 : P "030") (f30 : P "031") (f31 : P " ") (f32 : P "!") (f33 : P """") (f34 : P "#") (f35 : P "$") (f36 : P "%") (f37 : P "&") (f38 : P "'") (f39 : P "(") (f40 : P ")") (f41 : P "*") (f42 : P "+") (f43 : P ",") (f44 : P "-") (f45 : P ".") (f46 : P "/") (f47 : P "0") (f48 : P "1") (f49 : P "2") (f50 : P "3") (f51 : P "4") (f52 : P "5") (f53 : P "6") (f54 : P "7") (f55 : P "8") (f56 : P "9") (f57 : P ":") (f58 : P ";") (f59 : P "<") (f60 : P "=") (f61 : P ">") (f62 : P "?")
+ (f63 : P "@") (f64 : P "A") (f65 : P "B") (f66 : P "C") (f67 : P "D") (f68 : P "E") (f69 : P "F") (f70 : P "G") (f71 : P "H") (f72 : P "I") (f73 : P "J") (f74 : P "K") (f75 : P "L") (f76 : P "M") (f77 : P "N") (f78 : P "O") (f79 : P "P") (f80 : P "Q") (f81 : P "R") (f82 : P "S") (f83 : P "T") (f84 : P "U") (f85 : P "V") (f86 : P "W") (f87 : P "X") (f88 : P "Y") (f89 : P "Z") (f90 : P "[") (f91 : P "\") (f92 : P "]") (f93 : P "^") (f94 : P "_") (f95 : P "`") (f96 : P "a") (f97 : P "b") (f98 : P "c") (f99 : P "d") (f100 : P "e") (f101 : P "f") (f102 : P "g") (f103 : P "h") (f104 : P "i") (f105 : P "j") (f106 : P "k") (f107 : P "l") (f108 : P "m") (f109 : P "n") (f110 : P "o") (f111 : P "p") (f112 : P "q") (f113 : P "r") (f114 : P "s") (f115 : P "t") (f116 : P "u") (f117 : P "v") (f118 : P "w") (f119 : P "x") (f120 : P "y") (f121 : P "z") (f122 : P "{") (f123 : P "|") (f124 : P "}") (f125 : P "~") (f126 : P "127") (f127 : P "128") (f128 : P "129") (f129 : P "130")
+ (f130 : P "131") (f131 : P "132") (f132 : P "133") (f133 : P "134") (f134 : P "135") (f135 : P "136") (f136 : P "137") (f137 : P "138") (f138 : P "139") (f139 : P "140") (f140 : P "141") (f141 : P "142") (f142 : P "143") (f143 : P "144") (f144 : P "145") (f145 : P "146") (f146 : P "147") (f147 : P "148") (f148 : P "149") (f149 : P "150") (f150 : P "151") (f151 : P "152") (f152 : P "153") (f153 : P "154") (f154 : P "155") (f155 : P "156") (f156 : P "157") (f157 : P "158") (f158 : P "159") (f159 : P "160") (f160 : P "161") (f161 : P "162") (f162 : P "163") (f163 : P "164") (f164 : P "165") (f165 : P "166") (f166 : P "167") (f167 : P "168") (f168 : P "169") (f169 : P "170") (f170 : P "171") (f171 : P "172") (f172 : P "173") (f173 : P "174") (f174 : P "175") (f175 : P "176") (f176 : P "177") (f177 : P "178") (f178 : P "179") (f179 : P "180") (f180 : P "181") (f181 : P "182") (f182 : P "183") (f183 : P "184") (f184 : P "185") (f185 : P "186") (f186 : P "187")
+ (f187 : P "188") (f188 : P "189") (f189 : P "190") (f190 : P "191") (f191 : P "192") (f192 : P "193") (f193 : P "194") (f194 : P "195") (f195 : P "196") (f196 : P "197") (f197 : P "198") (f198 : P "199") (f199 : P "200") (f200 : P "201") (f201 : P "202") (f202 : P "203") (f203 : P "204") (f204 : P "205") (f205 : P "206") (f206 : P "207") (f207 : P "208") (f208 : P "209") (f209 : P "210") (f210 : P "211") (f211 : P "212") (f212 : P "213") (f213 : P "214") (f214 : P "215") (f215 : P "216") (f216 : P "217") (f217 : P "218") (f218 : P "219") (f219 : P "220") (f220 : P "221") (f221 : P "222") (f222 : P "223") (f223 : P "224") (f224 : P "225") (f225 : P "226") (f226 : P "227") (f227 : P "228") (f228 : P "229") (f229 : P "230") (f230 : P "231") (f231 : P "232") (f232 : P "233") (f233 : P "234") (f234 : P "235") (f235 : P "236") (f236 : P "237") (f237 : P "238") (f238 : P "239") (f239 : P "240") (f240 : P "241") (f241 : P "242") (f242 : P "243") (f243 : P "244")
+ (f244 : P "245") (f245 : P "246") (f246 : P "247") (f247 : P "248") (f248 : P "249") (f249 : P "250") (f250 : P "251") (f251 : P "252") (f252 : P "253") (f253 : P "254") (f254 : P "255") (b : byte) =>
+match b as b0 return (P b0) with
+| "000" => f
+| "001" => f0
+| "002" => f1
+| "003" => f2
+| "004" => f3
+| "005" => f4
+| "006" => f5
+| "007" => f6
+| "008" => f7
+| "009" => f8
+| "010" => f9
+| "011" => f10
+| "012" => f11
+| "013" => f12
+| "014" => f13
+| "015" => f14
+| "016" => f15
+| "017" => f16
+| "018" => f17
+| "019" => f18
+| "020" => f19
+| "021" => f20
+| "022" => f21
+| "023" => f22
+| "024" => f23
+| "025" => f24
+| "026" => f25
+| "027" => f26
+| "028" => f27
+| "029" => f28
+| "030" => f29
+| "031" => f30
+| " " => f31
+| "!" => f32
+| """" => f33
+| "#" => f34
+| "$" => f35
+| "%" => f36
+| "&" => f37
+| "'" => f38
+| "(" => f39
+| ")" => f40
+| "*" => f41
+| "+" => f42
+| "," => f43
+| "-" => f44
+| "." => f45
+| "/" => f46
+| "0" => f47
+| "1" => f48
+| "2" => f49
+| "3" => f50
+| "4" => f51
+| "5" => f52
+| "6" => f53
+| "7" => f54
+| "8" => f55
+| "9" => f56
+| ":" => f57
+| ";" => f58
+| "<" => f59
+| "=" => f60
+| ">" => f61
+| "?" => f62
+| "@" => f63
+| "A" => f64
+| "B" => f65
+| "C" => f66
+| "D" => f67
+| "E" => f68
+| "F" => f69
+| "G" => f70
+| "H" => f71
+| "I" => f72
+| "J" => f73
+| "K" => f74
+| "L" => f75
+| "M" => f76
+| "N" => f77
+| "O" => f78
+| "P" => f79
+| "Q" => f80
+| "R" => f81
+| "S" => f82
+| "T" => f83
+| "U" => f84
+| "V" => f85
+| "W" => f86
+| "X" => f87
+| "Y" => f88
+| "Z" => f89
+| "[" => f90
+| "\" => f91
+| "]" => f92
+| "^" => f93
+| "_" => f94
+| "`" => f95
+| "a" => f96
+| "b" => f97
+| "c" => f98
+| "d" => f99
+| "e" => f100
+| "f" => f101
+| "g" => f102
+| "h" => f103
+| "i" => f104
+| "j" => f105
+| "k" => f106
+| "l" => f107
+| "m" => f108
+| "n" => f109
+| "o" => f110
+| "p" => f111
+| "q" => f112
+| "r" => f113
+| "s" => f114
+| "t" => f115
+| "u" => f116
+| "v" => f117
+| "w" => f118
+| "x" => f119
+| "y" => f120
+| "z" => f121
+| "{" => f122
+| "|" => f123
+| "}" => f124
+| "~" => f125
+| "127" => f126
+| "128" => f127
+| "129" => f128
+| "130" => f129
+| "131" => f130
+| "132" => f131
+| "133" => f132
+| "134" => f133
+| "135" => f134
+| "136" => f135
+| "137" => f136
+| "138" => f137
+| "139" => f138
+| "140" => f139
+| "141" => f140
+| "142" => f141
+| "143" => f142
+| "144" => f143
+| "145" => f144
+| "146" => f145
+| "147" => f146
+| "148" => f147
+| "149" => f148
+| "150" => f149
+| "151" => f150
+| "152" => f151
+| "153" => f152
+| "154" => f153
+| "155" => f154
+| "156" => f155
+| "157" => f156
+| "158" => f157
+| "159" => f158
+| "160" => f159
+| "161" => f160
+| "162" => f161
+| "163" => f162
+| "164" => f163
+| "165" => f164
+| "166" => f165
+| "167" => f166
+| "168" => f167
+| "169" => f168
+| "170" => f169
+| "171" => f170
+| "172" => f171
+| "173" => f172
+| "174" => f173
+| "175" => f174
+| "176" => f175
+| "177" => f176
+| "178" => f177
+| "179" => f178
+| "180" => f179
+| "181" => f180
+| "182" => f181
+| "183" => f182
+| "184" => f183
+| "185" => f184
+| "186" => f185
+| "187" => f186
+| "188" => f187
+| "189" => f188
+| "190" => f189
+| "191" => f190
+| "192" => f191
+| "193" => f192
+| "194" => f193
+| "195" => f194
+| "196" => f195
+| "197" => f196
+| "198" => f197
+| "199" => f198
+| "200" => f199
+| "201" => f200
+| "202" => f201
+| "203" => f202
+| "204" => f203
+| "205" => f204
+| "206" => f205
+| "207" => f206
+| "208" => f207
+| "209" => f208
+| "210" => f209
+| "211" => f210
+| "212" => f211
+| "213" => f212
+| "214" => f213
+| "215" => f214
+| "216" => f215
+| "217" => f216
+| "218" => f217
+| "219" => f218
+| "220" => f219
+| "221" => f220
+| "222" => f221
+| "223" => f222
+| "224" => f223
+| "225" => f224
+| "226" => f225
+| "227" => f226
+| "228" => f227
+| "229" => f228
+| "230" => f229
+| "231" => f230
+| "232" => f231
+| "233" => f232
+| "234" => f233
+| "235" => f234
+| "236" => f235
+| "237" => f236
+| "238" => f237
+| "239" => f238
+| "240" => f239
+| "241" => f240
+| "242" => f241
+| "243" => f242
+| "244" => f243
+| "245" => f244
+| "246" => f245
+| "247" => f246
+| "248" => f247
+| "249" => f248
+| "250" => f249
+| "251" => f250
+| "252" => f251
+| "253" => f252
+| "254" => f253
+| "255" => f254
+end
+ : forall P : byte -> Type,
+ P "000" ->
+ P "001" ->
+ P "002" ->
+ P "003" ->
+ P "004" ->
+ P "005" ->
+ P "006" ->
+ P "007" ->
+ P "008" ->
+ P "009" ->
+ P "010" ->
+ P "011" ->
+ P "012" ->
+ P "013" ->
+ P "014" ->
+ P "015" ->
+ P "016" ->
+ P "017" ->
+ P "018" ->
+ P "019" ->
+ P "020" ->
+ P "021" ->
+ P "022" ->
+ P "023" ->
+ P "024" ->
+ P "025" ->
+ P "026" ->
+ P "027" ->
+ P "028" ->
+ P "029" ->
+ P "030" ->
+ P "031" ->
+ P " " ->
+ P "!" ->
+ P """" ->
+ P "#" ->
+ P "$" ->
+ P "%" ->
+ P "&" ->
+ P "'" ->
+ P "(" ->
+ P ")" ->
+ P "*" ->
+ P "+" ->
+ P "," ->
+ P "-" ->
+ P "." ->
+ P "/" ->
+ P "0" ->
+ P "1" ->
+ P "2" ->
+ P "3" ->
+ P "4" ->
+ P "5" ->
+ P "6" ->
+ P "7" ->
+ P "8" ->
+ P "9" ->
+ P ":" ->
+ P ";" ->
+ P "<" ->
+ P "=" ->
+ P ">" ->
+ P "?" ->
+ P "@" ->
+ P "A" ->
+ P "B" ->
+ P "C" ->
+ P "D" ->
+ P "E" ->
+ P "F" ->
+ P "G" ->
+ P "H" ->
+ P "I" ->
+ P "J" ->
+ P "K" ->
+ P "L" ->
+ P "M" ->
+ P "N" ->
+ P "O" ->
+ P "P" ->
+ P "Q" ->
+ P "R" ->
+ P "S" ->
+ P "T" ->
+ P "U" ->
+ P "V" ->
+ P "W" ->
+ P "X" ->
+ P "Y" ->
+ P "Z" ->
+ P "[" ->
+ P "\" ->
+ P "]" ->
+ P "^" ->
+ P "_" ->
+ P "`" ->
+ P "a" ->
+ P "b" ->
+ P "c" ->
+ P "d" ->
+ P "e" ->
+ P "f" ->
+ P "g" ->
+ P "h" ->
+ P "i" ->
+ P "j" ->
+ P "k" ->
+ P "l" ->
+ P "m" ->
+ P "n" ->
+ P "o" ->
+ P "p" ->
+ P "q" ->
+ P "r" ->
+ P "s" ->
+ P "t" ->
+ P "u" ->
+ P "v" ->
+ P "w" ->
+ P "x" ->
+ P "y" ->
+ P "z" ->
+ P "{" ->
+ P "|" ->
+ P "}" ->
+ P "~" ->
+ P "127" ->
+ P "128" ->
+ P "129" ->
+ P "130" ->
+ P "131" ->
+ P "132" ->
+ P "133" ->
+ P "134" ->
+ P "135" ->
+ P "136" ->
+ P "137" ->
+ P "138" ->
+ P "139" ->
+ P "140" ->
+ P "141" ->
+ P "142" ->
+ P "143" ->
+ P "144" ->
+ P "145" ->
+ P "146" ->
+ P "147" ->
+ P "148" ->
+ P "149" ->
+ P "150" ->
+ P "151" ->
+ P "152" ->
+ P "153" ->
+ P "154" ->
+ P "155" ->
+ P "156" ->
+ P "157" ->
+ P "158" ->
+ P "159" ->
+ P "160" ->
+ P "161" ->
+ P "162" ->
+ P "163" ->
+ P "164" ->
+ P "165" ->
+ P "166" ->
+ P "167" ->
+ P "168" -> P "169" -> P "170" -> P "171" -> P "172" -> P "173" -> P "174" -> P "175" -> P "176" -> P "177" -> P "178" -> P "179" -> P "180" -> P "181" -> P "182" -> P "183" -> P "184" -> P "185" -> P "186" -> P "187" -> P "188" -> P "189" -> P "190" -> P "191" -> P "192" -> P "193" -> P "194" -> P "195" -> P "196" -> P "197" -> P "198" -> P "199" -> P "200" -> P "201" -> P "202" -> P "203" -> P "204" -> P "205" -> P "206" -> P "207" -> P "208" -> P "209" -> P "210" -> P "211" -> P "212" -> P "213" -> P "214" -> P "215" -> P "216" -> P "217" -> P "218" -> P "219" -> P "220" -> P "221" -> P "222" -> P "223" -> P "224" -> P "225" -> P "226" -> P "227" -> P "228" -> P "229" -> P "230" -> P "231" -> P "232" -> P "233" -> P "234" -> P "235" -> P "236" -> P "237" -> P "238" -> P "239" -> P "240" -> P "241" -> P "242" -> P "243" -> P "244" -> P "245" -> P "246" -> P "247" -> P "248" -> P "249" -> P "250" -> P "251" -> P "252" -> P "253" -> P "254" -> P "255" -> forall b : byte, P b
+
+byte_rect is not universe polymorphic
+Argument scopes are [function_scope _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ byte_scope]
+Monomorphic byte_rec =
+fun P : byte -> Set => byte_rect P
+ : forall P : byte -> Set,
+ P "000" ->
+ P "001" ->
+ P "002" ->
+ P "003" ->
+ P "004" ->
+ P "005" ->
+ P "006" ->
+ P "007" ->
+ P "008" ->
+ P "009" ->
+ P "010" ->
+ P "011" ->
+ P "012" ->
+ P "013" ->
+ P "014" ->
+ P "015" ->
+ P "016" ->
+ P "017" ->
+ P "018" ->
+ P "019" ->
+ P "020" ->
+ P "021" ->
+ P "022" ->
+ P "023" ->
+ P "024" ->
+ P "025" ->
+ P "026" ->
+ P "027" ->
+ P "028" ->
+ P "029" ->
+ P "030" ->
+ P "031" ->
+ P " " ->
+ P "!" ->
+ P """" ->
+ P "#" ->
+ P "$" ->
+ P "%" ->
+ P "&" ->
+ P "'" ->
+ P "(" ->
+ P ")" ->
+ P "*" ->
+ P "+" ->
+ P "," ->
+ P "-" ->
+ P "." ->
+ P "/" ->
+ P "0" ->
+ P "1" ->
+ P "2" ->
+ P "3" ->
+ P "4" ->
+ P "5" ->
+ P "6" ->
+ P "7" ->
+ P "8" ->
+ P "9" ->
+ P ":" ->
+ P ";" ->
+ P "<" ->
+ P "=" ->
+ P ">" ->
+ P "?" ->
+ P "@" ->
+ P "A" ->
+ P "B" ->
+ P "C" ->
+ P "D" ->
+ P "E" ->
+ P "F" ->
+ P "G" ->
+ P "H" ->
+ P "I" ->
+ P "J" ->
+ P "K" ->
+ P "L" ->
+ P "M" ->
+ P "N" ->
+ P "O" ->
+ P "P" ->
+ P "Q" ->
+ P "R" ->
+ P "S" ->
+ P "T" ->
+ P "U" ->
+ P "V" ->
+ P "W" ->
+ P "X" ->
+ P "Y" ->
+ P "Z" ->
+ P "[" ->
+ P "\" ->
+ P "]" ->
+ P "^" ->
+ P "_" ->
+ P "`" ->
+ P "a" ->
+ P "b" ->
+ P "c" ->
+ P "d" ->
+ P "e" ->
+ P "f" ->
+ P "g" ->
+ P "h" ->
+ P "i" ->
+ P "j" ->
+ P "k" ->
+ P "l" ->
+ P "m" ->
+ P "n" ->
+ P "o" ->
+ P "p" ->
+ P "q" ->
+ P "r" ->
+ P "s" ->
+ P "t" ->
+ P "u" ->
+ P "v" ->
+ P "w" ->
+ P "x" ->
+ P "y" ->
+ P "z" ->
+ P "{" ->
+ P "|" ->
+ P "}" ->
+ P "~" ->
+ P "127" ->
+ P "128" ->
+ P "129" ->
+ P "130" ->
+ P "131" ->
+ P "132" ->
+ P "133" ->
+ P "134" ->
+ P "135" ->
+ P "136" ->
+ P "137" ->
+ P "138" ->
+ P "139" ->
+ P "140" ->
+ P "141" ->
+ P "142" ->
+ P "143" ->
+ P "144" ->
+ P "145" ->
+ P "146" ->
+ P "147" ->
+ P "148" ->
+ P "149" ->
+ P "150" ->
+ P "151" ->
+ P "152" ->
+ P "153" ->
+ P "154" ->
+ P "155" ->
+ P "156" ->
+ P "157" ->
+ P "158" ->
+ P "159" ->
+ P "160" ->
+ P "161" ->
+ P "162" ->
+ P "163" ->
+ P "164" ->
+ P "165" ->
+ P "166" ->
+ P "167" ->
+ P "168" -> P "169" -> P "170" -> P "171" -> P "172" -> P "173" -> P "174" -> P "175" -> P "176" -> P "177" -> P "178" -> P "179" -> P "180" -> P "181" -> P "182" -> P "183" -> P "184" -> P "185" -> P "186" -> P "187" -> P "188" -> P "189" -> P "190" -> P "191" -> P "192" -> P "193" -> P "194" -> P "195" -> P "196" -> P "197" -> P "198" -> P "199" -> P "200" -> P "201" -> P "202" -> P "203" -> P "204" -> P "205" -> P "206" -> P "207" -> P "208" -> P "209" -> P "210" -> P "211" -> P "212" -> P "213" -> P "214" -> P "215" -> P "216" -> P "217" -> P "218" -> P "219" -> P "220" -> P "221" -> P "222" -> P "223" -> P "224" -> P "225" -> P "226" -> P "227" -> P "228" -> P "229" -> P "230" -> P "231" -> P "232" -> P "233" -> P "234" -> P "235" -> P "236" -> P "237" -> P "238" -> P "239" -> P "240" -> P "241" -> P "242" -> P "243" -> P "244" -> P "245" -> P "246" -> P "247" -> P "248" -> P "249" -> P "250" -> P "251" -> P "252" -> P "253" -> P "254" -> P "255" -> forall b : byte, P b
+
+byte_rec is not universe polymorphic
+Argument scopes are [function_scope _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ byte_scope]
+Monomorphic byte_ind =
+fun (P : byte -> Prop) (f : P "000") (f0 : P "001") (f1 : P "002") (f2 : P "003") (f3 : P "004") (f4 : P "005") (f5 : P "006") (f6 : P "007") (f7 : P "008") (f8 : P "009") (f9 : P "010") (f10 : P "011") (f11 : P "012") (f12 : P "013") (f13 : P "014") (f14 : P "015") (f15 : P "016") (f16 : P "017") (f17 : P "018") (f18 : P "019") (f19 : P "020") (f20 : P "021") (f21 : P "022") (f22 : P "023") (f23 : P "024") (f24 : P "025") (f25 : P "026") (f26 : P "027") (f27 : P "028") (f28 : P "029") (f29 : P "030") (f30 : P "031") (f31 : P " ") (f32 : P "!") (f33 : P """") (f34 : P "#") (f35 : P "$") (f36 : P "%") (f37 : P "&") (f38 : P "'") (f39 : P "(") (f40 : P ")") (f41 : P "*") (f42 : P "+") (f43 : P ",") (f44 : P "-") (f45 : P ".") (f46 : P "/") (f47 : P "0") (f48 : P "1") (f49 : P "2") (f50 : P "3") (f51 : P "4") (f52 : P "5") (f53 : P "6") (f54 : P "7") (f55 : P "8") (f56 : P "9") (f57 : P ":") (f58 : P ";") (f59 : P "<") (f60 : P "=") (f61 : P ">") (f62 : P "?")
+ (f63 : P "@") (f64 : P "A") (f65 : P "B") (f66 : P "C") (f67 : P "D") (f68 : P "E") (f69 : P "F") (f70 : P "G") (f71 : P "H") (f72 : P "I") (f73 : P "J") (f74 : P "K") (f75 : P "L") (f76 : P "M") (f77 : P "N") (f78 : P "O") (f79 : P "P") (f80 : P "Q") (f81 : P "R") (f82 : P "S") (f83 : P "T") (f84 : P "U") (f85 : P "V") (f86 : P "W") (f87 : P "X") (f88 : P "Y") (f89 : P "Z") (f90 : P "[") (f91 : P "\") (f92 : P "]") (f93 : P "^") (f94 : P "_") (f95 : P "`") (f96 : P "a") (f97 : P "b") (f98 : P "c") (f99 : P "d") (f100 : P "e") (f101 : P "f") (f102 : P "g") (f103 : P "h") (f104 : P "i") (f105 : P "j") (f106 : P "k") (f107 : P "l") (f108 : P "m") (f109 : P "n") (f110 : P "o") (f111 : P "p") (f112 : P "q") (f113 : P "r") (f114 : P "s") (f115 : P "t") (f116 : P "u") (f117 : P "v") (f118 : P "w") (f119 : P "x") (f120 : P "y") (f121 : P "z") (f122 : P "{") (f123 : P "|") (f124 : P "}") (f125 : P "~") (f126 : P "127") (f127 : P "128") (f128 : P "129") (f129 : P "130")
+ (f130 : P "131") (f131 : P "132") (f132 : P "133") (f133 : P "134") (f134 : P "135") (f135 : P "136") (f136 : P "137") (f137 : P "138") (f138 : P "139") (f139 : P "140") (f140 : P "141") (f141 : P "142") (f142 : P "143") (f143 : P "144") (f144 : P "145") (f145 : P "146") (f146 : P "147") (f147 : P "148") (f148 : P "149") (f149 : P "150") (f150 : P "151") (f151 : P "152") (f152 : P "153") (f153 : P "154") (f154 : P "155") (f155 : P "156") (f156 : P "157") (f157 : P "158") (f158 : P "159") (f159 : P "160") (f160 : P "161") (f161 : P "162") (f162 : P "163") (f163 : P "164") (f164 : P "165") (f165 : P "166") (f166 : P "167") (f167 : P "168") (f168 : P "169") (f169 : P "170") (f170 : P "171") (f171 : P "172") (f172 : P "173") (f173 : P "174") (f174 : P "175") (f175 : P "176") (f176 : P "177") (f177 : P "178") (f178 : P "179") (f179 : P "180") (f180 : P "181") (f181 : P "182") (f182 : P "183") (f183 : P "184") (f184 : P "185") (f185 : P "186") (f186 : P "187")
+ (f187 : P "188") (f188 : P "189") (f189 : P "190") (f190 : P "191") (f191 : P "192") (f192 : P "193") (f193 : P "194") (f194 : P "195") (f195 : P "196") (f196 : P "197") (f197 : P "198") (f198 : P "199") (f199 : P "200") (f200 : P "201") (f201 : P "202") (f202 : P "203") (f203 : P "204") (f204 : P "205") (f205 : P "206") (f206 : P "207") (f207 : P "208") (f208 : P "209") (f209 : P "210") (f210 : P "211") (f211 : P "212") (f212 : P "213") (f213 : P "214") (f214 : P "215") (f215 : P "216") (f216 : P "217") (f217 : P "218") (f218 : P "219") (f219 : P "220") (f220 : P "221") (f221 : P "222") (f222 : P "223") (f223 : P "224") (f224 : P "225") (f225 : P "226") (f226 : P "227") (f227 : P "228") (f228 : P "229") (f229 : P "230") (f230 : P "231") (f231 : P "232") (f232 : P "233") (f233 : P "234") (f234 : P "235") (f235 : P "236") (f236 : P "237") (f237 : P "238") (f238 : P "239") (f239 : P "240") (f240 : P "241") (f241 : P "242") (f242 : P "243") (f243 : P "244")
+ (f244 : P "245") (f245 : P "246") (f246 : P "247") (f247 : P "248") (f248 : P "249") (f249 : P "250") (f250 : P "251") (f251 : P "252") (f252 : P "253") (f253 : P "254") (f254 : P "255") (b : byte) =>
+match b as b0 return (P b0) with
+| "000" => f
+| "001" => f0
+| "002" => f1
+| "003" => f2
+| "004" => f3
+| "005" => f4
+| "006" => f5
+| "007" => f6
+| "008" => f7
+| "009" => f8
+| "010" => f9
+| "011" => f10
+| "012" => f11
+| "013" => f12
+| "014" => f13
+| "015" => f14
+| "016" => f15
+| "017" => f16
+| "018" => f17
+| "019" => f18
+| "020" => f19
+| "021" => f20
+| "022" => f21
+| "023" => f22
+| "024" => f23
+| "025" => f24
+| "026" => f25
+| "027" => f26
+| "028" => f27
+| "029" => f28
+| "030" => f29
+| "031" => f30
+| " " => f31
+| "!" => f32
+| """" => f33
+| "#" => f34
+| "$" => f35
+| "%" => f36
+| "&" => f37
+| "'" => f38
+| "(" => f39
+| ")" => f40
+| "*" => f41
+| "+" => f42
+| "," => f43
+| "-" => f44
+| "." => f45
+| "/" => f46
+| "0" => f47
+| "1" => f48
+| "2" => f49
+| "3" => f50
+| "4" => f51
+| "5" => f52
+| "6" => f53
+| "7" => f54
+| "8" => f55
+| "9" => f56
+| ":" => f57
+| ";" => f58
+| "<" => f59
+| "=" => f60
+| ">" => f61
+| "?" => f62
+| "@" => f63
+| "A" => f64
+| "B" => f65
+| "C" => f66
+| "D" => f67
+| "E" => f68
+| "F" => f69
+| "G" => f70
+| "H" => f71
+| "I" => f72
+| "J" => f73
+| "K" => f74
+| "L" => f75
+| "M" => f76
+| "N" => f77
+| "O" => f78
+| "P" => f79
+| "Q" => f80
+| "R" => f81
+| "S" => f82
+| "T" => f83
+| "U" => f84
+| "V" => f85
+| "W" => f86
+| "X" => f87
+| "Y" => f88
+| "Z" => f89
+| "[" => f90
+| "\" => f91
+| "]" => f92
+| "^" => f93
+| "_" => f94
+| "`" => f95
+| "a" => f96
+| "b" => f97
+| "c" => f98
+| "d" => f99
+| "e" => f100
+| "f" => f101
+| "g" => f102
+| "h" => f103
+| "i" => f104
+| "j" => f105
+| "k" => f106
+| "l" => f107
+| "m" => f108
+| "n" => f109
+| "o" => f110
+| "p" => f111
+| "q" => f112
+| "r" => f113
+| "s" => f114
+| "t" => f115
+| "u" => f116
+| "v" => f117
+| "w" => f118
+| "x" => f119
+| "y" => f120
+| "z" => f121
+| "{" => f122
+| "|" => f123
+| "}" => f124
+| "~" => f125
+| "127" => f126
+| "128" => f127
+| "129" => f128
+| "130" => f129
+| "131" => f130
+| "132" => f131
+| "133" => f132
+| "134" => f133
+| "135" => f134
+| "136" => f135
+| "137" => f136
+| "138" => f137
+| "139" => f138
+| "140" => f139
+| "141" => f140
+| "142" => f141
+| "143" => f142
+| "144" => f143
+| "145" => f144
+| "146" => f145
+| "147" => f146
+| "148" => f147
+| "149" => f148
+| "150" => f149
+| "151" => f150
+| "152" => f151
+| "153" => f152
+| "154" => f153
+| "155" => f154
+| "156" => f155
+| "157" => f156
+| "158" => f157
+| "159" => f158
+| "160" => f159
+| "161" => f160
+| "162" => f161
+| "163" => f162
+| "164" => f163
+| "165" => f164
+| "166" => f165
+| "167" => f166
+| "168" => f167
+| "169" => f168
+| "170" => f169
+| "171" => f170
+| "172" => f171
+| "173" => f172
+| "174" => f173
+| "175" => f174
+| "176" => f175
+| "177" => f176
+| "178" => f177
+| "179" => f178
+| "180" => f179
+| "181" => f180
+| "182" => f181
+| "183" => f182
+| "184" => f183
+| "185" => f184
+| "186" => f185
+| "187" => f186
+| "188" => f187
+| "189" => f188
+| "190" => f189
+| "191" => f190
+| "192" => f191
+| "193" => f192
+| "194" => f193
+| "195" => f194
+| "196" => f195
+| "197" => f196
+| "198" => f197
+| "199" => f198
+| "200" => f199
+| "201" => f200
+| "202" => f201
+| "203" => f202
+| "204" => f203
+| "205" => f204
+| "206" => f205
+| "207" => f206
+| "208" => f207
+| "209" => f208
+| "210" => f209
+| "211" => f210
+| "212" => f211
+| "213" => f212
+| "214" => f213
+| "215" => f214
+| "216" => f215
+| "217" => f216
+| "218" => f217
+| "219" => f218
+| "220" => f219
+| "221" => f220
+| "222" => f221
+| "223" => f222
+| "224" => f223
+| "225" => f224
+| "226" => f225
+| "227" => f226
+| "228" => f227
+| "229" => f228
+| "230" => f229
+| "231" => f230
+| "232" => f231
+| "233" => f232
+| "234" => f233
+| "235" => f234
+| "236" => f235
+| "237" => f236
+| "238" => f237
+| "239" => f238
+| "240" => f239
+| "241" => f240
+| "242" => f241
+| "243" => f242
+| "244" => f243
+| "245" => f244
+| "246" => f245
+| "247" => f246
+| "248" => f247
+| "249" => f248
+| "250" => f249
+| "251" => f250
+| "252" => f251
+| "253" => f252
+| "254" => f253
+| "255" => f254
+end
+ : forall P : byte -> Prop,
+ P "000" ->
+ P "001" ->
+ P "002" ->
+ P "003" ->
+ P "004" ->
+ P "005" ->
+ P "006" ->
+ P "007" ->
+ P "008" ->
+ P "009" ->
+ P "010" ->
+ P "011" ->
+ P "012" ->
+ P "013" ->
+ P "014" ->
+ P "015" ->
+ P "016" ->
+ P "017" ->
+ P "018" ->
+ P "019" ->
+ P "020" ->
+ P "021" ->
+ P "022" ->
+ P "023" ->
+ P "024" ->
+ P "025" ->
+ P "026" ->
+ P "027" ->
+ P "028" ->
+ P "029" ->
+ P "030" ->
+ P "031" ->
+ P " " ->
+ P "!" ->
+ P """" ->
+ P "#" ->
+ P "$" ->
+ P "%" ->
+ P "&" ->
+ P "'" ->
+ P "(" ->
+ P ")" ->
+ P "*" ->
+ P "+" ->
+ P "," ->
+ P "-" ->
+ P "." ->
+ P "/" ->
+ P "0" ->
+ P "1" ->
+ P "2" ->
+ P "3" ->
+ P "4" ->
+ P "5" ->
+ P "6" ->
+ P "7" ->
+ P "8" ->
+ P "9" ->
+ P ":" ->
+ P ";" ->
+ P "<" ->
+ P "=" ->
+ P ">" ->
+ P "?" ->
+ P "@" ->
+ P "A" ->
+ P "B" ->
+ P "C" ->
+ P "D" ->
+ P "E" ->
+ P "F" ->
+ P "G" ->
+ P "H" ->
+ P "I" ->
+ P "J" ->
+ P "K" ->
+ P "L" ->
+ P "M" ->
+ P "N" ->
+ P "O" ->
+ P "P" ->
+ P "Q" ->
+ P "R" ->
+ P "S" ->
+ P "T" ->
+ P "U" ->
+ P "V" ->
+ P "W" ->
+ P "X" ->
+ P "Y" ->
+ P "Z" ->
+ P "[" ->
+ P "\" ->
+ P "]" ->
+ P "^" ->
+ P "_" ->
+ P "`" ->
+ P "a" ->
+ P "b" ->
+ P "c" ->
+ P "d" ->
+ P "e" ->
+ P "f" ->
+ P "g" ->
+ P "h" ->
+ P "i" ->
+ P "j" ->
+ P "k" ->
+ P "l" ->
+ P "m" ->
+ P "n" ->
+ P "o" ->
+ P "p" ->
+ P "q" ->
+ P "r" ->
+ P "s" ->
+ P "t" ->
+ P "u" ->
+ P "v" ->
+ P "w" ->
+ P "x" ->
+ P "y" ->
+ P "z" ->
+ P "{" ->
+ P "|" ->
+ P "}" ->
+ P "~" ->
+ P "127" ->
+ P "128" ->
+ P "129" ->
+ P "130" ->
+ P "131" ->
+ P "132" ->
+ P "133" ->
+ P "134" ->
+ P "135" ->
+ P "136" ->
+ P "137" ->
+ P "138" ->
+ P "139" ->
+ P "140" ->
+ P "141" ->
+ P "142" ->
+ P "143" ->
+ P "144" ->
+ P "145" ->
+ P "146" ->
+ P "147" ->
+ P "148" ->
+ P "149" ->
+ P "150" ->
+ P "151" ->
+ P "152" ->
+ P "153" ->
+ P "154" ->
+ P "155" ->
+ P "156" ->
+ P "157" ->
+ P "158" ->
+ P "159" ->
+ P "160" ->
+ P "161" ->
+ P "162" ->
+ P "163" ->
+ P "164" ->
+ P "165" ->
+ P "166" ->
+ P "167" ->
+ P "168" -> P "169" -> P "170" -> P "171" -> P "172" -> P "173" -> P "174" -> P "175" -> P "176" -> P "177" -> P "178" -> P "179" -> P "180" -> P "181" -> P "182" -> P "183" -> P "184" -> P "185" -> P "186" -> P "187" -> P "188" -> P "189" -> P "190" -> P "191" -> P "192" -> P "193" -> P "194" -> P "195" -> P "196" -> P "197" -> P "198" -> P "199" -> P "200" -> P "201" -> P "202" -> P "203" -> P "204" -> P "205" -> P "206" -> P "207" -> P "208" -> P "209" -> P "210" -> P "211" -> P "212" -> P "213" -> P "214" -> P "215" -> P "216" -> P "217" -> P "218" -> P "219" -> P "220" -> P "221" -> P "222" -> P "223" -> P "224" -> P "225" -> P "226" -> P "227" -> P "228" -> P "229" -> P "230" -> P "231" -> P "232" -> P "233" -> P "234" -> P "235" -> P "236" -> P "237" -> P "238" -> P "239" -> P "240" -> P "241" -> P "242" -> P "243" -> P "244" -> P "245" -> P "246" -> P "247" -> P "248" -> P "249" -> P "250" -> P "251" -> P "252" -> P "253" -> P "254" -> P "255" -> forall b : byte, P b
+
+byte_ind is not universe polymorphic
+Argument scopes are [function_scope _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ byte_scope]
+"000"
+ : byte
+"a"
+ : byte
+"127"
+ : byte
+The command has indeed failed with message:
+Expects a single character or a three-digits ascii code.
+"000"
+ : ascii
+"a"
+ : ascii
+"127"
+ : ascii
+The command has indeed failed with message:
+Expects a single character or a three-digits ascii code.
+"000"
+ : string
+"a"
+ : string
+"127"
+ : string
+"€"
+ : string
+""
+ : string
+ = "a"%char
+ : ascii
+ = "a"%byte
+ : byte
+ = "a"%string
+ : string
+ = ["a"%byte]
+ : list byte
+ = ["000"; "001"; "002"; "003"; "004"; "005"; "006"; "007"; "008"; "009"; "010"; "011"; "012"; "013"; "014"; "015"; "016"; "017"; "018"; "019"; "020"; "021"; "022"; "023"; "024"; "025"; "026"; "027"; "028"; "029"; "030"; "031"; " "; "!"; """"; "#"; "$"; "%"; "&"; "'"; "("; ")"; "*"; "+"; ","; "-"; "."; "/"; "0"; "1"; "2"; "3"; "4"; "5"; "6"; "7"; "8"; "9"; ":"; ";"; "<"; "="; ">"; "?"; "@"; "A"; "B"; "C"; "D"; "E"; "F"; "G"; "H"; "I"; "J"; "K"; "L"; "M"; "N"; "O"; "P"; "Q"; "R"; "S"; "T"; "U"; "V"; "W"; "X"; "Y"; "Z"; "["; "\"; "]"; "^"; "_"; "`"; "a"; "b"; "c"; "d"; "e"; "f"; "g"; "h"; "i"; "j"; "k"; "l"; "m"; "n"; "o"; "p"; "q"; "r"; "s"; "t"; "u"; "v"; "w"; "x"; "y"; "z"; "{"; "|"; "}"; "~"; "127"; "128"; "129"; "130"; "131"; "132"; "133"; "134"; "135"; "136"; "137"; "138"; "139"; "140"; "141"; "142"; "143"; "144"; "145"; "146"; "147"; "148"; "149"; "150"; "151"; "152"; "153"; "154"; "155"; "156"; "157"; "158"; "159"; "160"; "161"; "162"; "163"; "164"; "165"; "166"; "167";
+ "168"; "169"; "170"; "171"; "172"; "173"; "174"; "175"; "176"; "177"; "178"; "179"; "180"; "181"; "182"; "183"; "184"; "185"; "186"; "187"; "188"; "189"; "190"; "191"; "192"; "193"; "194"; "195"; "196"; "197"; "198"; "199"; "200"; "201"; "202"; "203"; "204"; "205"; "206"; "207"; "208"; "209"; "210"; "211"; "212"; "213"; "214"; "215"; "216"; "217"; "218"; "219"; "220"; "221"; "222"; "223"; "224"; "225"; "226"; "227"; "228"; "229"; "230"; "231"; "232"; "233"; "234"; "235"; "236"; "237"; "238"; "239"; "240"; "241"; "242"; "243"; "244"; "245"; "246"; "247"; "248"; "249"; "250"; "251"; "252"; "253"; "254"; "255"]
+ : list byte
+ = ["000"; "001"; "002"; "003"; "004"; "005"; "006"; "007"; "008"; "009"; "010"; "011"; "012"; "013"; "014"; "015"; "016"; "017"; "018"; "019"; "020"; "021"; "022"; "023"; "024"; "025"; "026"; "027"; "028"; "029"; "030"; "031"; " "; "!"; """"; "#"; "$"; "%"; "&"; "'"; "("; ")"; "*"; "+"; ","; "-"; "."; "/"; "0"; "1"; "2"; "3"; "4"; "5"; "6"; "7"; "8"; "9"; ":"; ";"; "<"; "="; ">"; "?"; "@"; "A"; "B"; "C"; "D"; "E"; "F"; "G"; "H"; "I"; "J"; "K"; "L"; "M"; "N"; "O"; "P"; "Q"; "R"; "S"; "T"; "U"; "V"; "W"; "X"; "Y"; "Z"; "["; "\"; "]"; "^"; "_"; "`"; "a"; "b"; "c"; "d"; "e"; "f"; "g"; "h"; "i"; "j"; "k"; "l"; "m"; "n"; "o"; "p"; "q"; "r"; "s"; "t"; "u"; "v"; "w"; "x"; "y"; "z"; "{"; "|"; "}"; "~"; "127"; "128"; "129"; "130"; "131"; "132"; "133"; "134"; "135"; "136"; "137"; "138"; "139"; "140"; "141"; "142"; "143"; "144"; "145"; "146"; "147"; "148"; "149"; "150"; "151"; "152"; "153"; "154"; "155"; "156"; "157"; "158"; "159"; "160"; "161"; "162"; "163"; "164"; "165"; "166"; "167";
+ "168"; "169"; "170"; "171"; "172"; "173"; "174"; "175"; "176"; "177"; "178"; "179"; "180"; "181"; "182"; "183"; "184"; "185"; "186"; "187"; "188"; "189"; "190"; "191"; "192"; "193"; "194"; "195"; "196"; "197"; "198"; "199"; "200"; "201"; "202"; "203"; "204"; "205"; "206"; "207"; "208"; "209"; "210"; "211"; "212"; "213"; "214"; "215"; "216"; "217"; "218"; "219"; "220"; "221"; "222"; "223"; "224"; "225"; "226"; "227"; "228"; "229"; "230"; "231"; "232"; "233"; "234"; "235"; "236"; "237"; "238"; "239"; "240"; "241"; "242"; "243"; "244"; "245"; "246"; "247"; "248"; "249"; "250"; "251"; "252"; "253"; "254"; "255"]
+ : list ascii
diff --git a/test-suite/output/StringSyntax.v b/test-suite/output/StringSyntax.v
new file mode 100644
index 0000000000..aab6e0bb03
--- /dev/null
+++ b/test-suite/output/StringSyntax.v
@@ -0,0 +1,52 @@
+Require Import Coq.Lists.List.
+Require Import Coq.Strings.String Coq.Strings.Byte Coq.Strings.Ascii.
+Import ListNotations.
+
+Set Printing Depth 100000.
+Set Printing Width 1000.
+
+Close Scope char_scope.
+Close Scope string_scope.
+
+Open Scope byte_scope.
+Print byte_rect.
+Print byte_rec.
+Print byte_ind.
+Check "000".
+Check "a".
+Check "127".
+Fail Check "€".
+Close Scope byte_scope.
+
+Open Scope char_scope.
+Check "000".
+Check "a".
+Check "127".
+Fail Check "€".
+Close Scope char_scope.
+
+Open Scope string_scope.
+Check "000".
+Check "a".
+Check "127".
+Check "€".
+Check String "001" EmptyString.
+Close Scope string_scope.
+
+Compute ascii_of_byte "a".
+Compute byte_of_ascii "a".
+Compute string_of_list_byte ("a"::nil)%byte.
+Compute list_byte_of_string "a".
+
+Local Open Scope byte_scope.
+Compute List.fold_right
+ (fun n ls => match Byte.of_nat n with
+ | Some b => cons b ls
+ | None => ls
+ end)
+ nil
+ (List.seq 0 256).
+Local Close Scope byte_scope.
+Local Open Scope char_scope.
+Compute List.map Ascii.ascii_of_nat (List.seq 0 256).
+Local Close Scope char_scope.
diff --git a/theories/Init/Byte.v b/theories/Init/Byte.v
new file mode 100644
index 0000000000..eede9d5028
--- /dev/null
+++ b/theories/Init/Byte.v
@@ -0,0 +1,830 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+(** * Bytes *)
+
+Require Import Coq.Init.Datatypes.
+Require Import Coq.Init.Logic.
+Require Import Coq.Init.Specif.
+Require Coq.Init.Nat.
+
+Declare ML Module "string_notation_plugin".
+
+(** We define an inductive for use with the [String Notation] command
+ which contains all ascii characters. We use 256 constructors for
+ efficiency and ease of conversion. *)
+
+Declare Scope byte_scope.
+Delimit Scope byte_scope with byte.
+
+Inductive byte :=
+| x00
+| x01
+| x02
+| x03
+| x04
+| x05
+| x06
+| x07
+| x08
+| x09
+| x0a
+| x0b
+| x0c
+| x0d
+| x0e
+| x0f
+| x10
+| x11
+| x12
+| x13
+| x14
+| x15
+| x16
+| x17
+| x18
+| x19
+| x1a
+| x1b
+| x1c
+| x1d
+| x1e
+| x1f
+| x20
+| x21
+| x22
+| x23
+| x24
+| x25
+| x26
+| x27
+| x28
+| x29
+| x2a
+| x2b
+| x2c
+| x2d
+| x2e
+| x2f
+| x30
+| x31
+| x32
+| x33
+| x34
+| x35
+| x36
+| x37
+| x38
+| x39
+| x3a
+| x3b
+| x3c
+| x3d
+| x3e
+| x3f
+| x40
+| x41
+| x42
+| x43
+| x44
+| x45
+| x46
+| x47
+| x48
+| x49
+| x4a
+| x4b
+| x4c
+| x4d
+| x4e
+| x4f
+| x50
+| x51
+| x52
+| x53
+| x54
+| x55
+| x56
+| x57
+| x58
+| x59
+| x5a
+| x5b
+| x5c
+| x5d
+| x5e
+| x5f
+| x60
+| x61
+| x62
+| x63
+| x64
+| x65
+| x66
+| x67
+| x68
+| x69
+| x6a
+| x6b
+| x6c
+| x6d
+| x6e
+| x6f
+| x70
+| x71
+| x72
+| x73
+| x74
+| x75
+| x76
+| x77
+| x78
+| x79
+| x7a
+| x7b
+| x7c
+| x7d
+| x7e
+| x7f
+| x80
+| x81
+| x82
+| x83
+| x84
+| x85
+| x86
+| x87
+| x88
+| x89
+| x8a
+| x8b
+| x8c
+| x8d
+| x8e
+| x8f
+| x90
+| x91
+| x92
+| x93
+| x94
+| x95
+| x96
+| x97
+| x98
+| x99
+| x9a
+| x9b
+| x9c
+| x9d
+| x9e
+| x9f
+| xa0
+| xa1
+| xa2
+| xa3
+| xa4
+| xa5
+| xa6
+| xa7
+| xa8
+| xa9
+| xaa
+| xab
+| xac
+| xad
+| xae
+| xaf
+| xb0
+| xb1
+| xb2
+| xb3
+| xb4
+| xb5
+| xb6
+| xb7
+| xb8
+| xb9
+| xba
+| xbb
+| xbc
+| xbd
+| xbe
+| xbf
+| xc0
+| xc1
+| xc2
+| xc3
+| xc4
+| xc5
+| xc6
+| xc7
+| xc8
+| xc9
+| xca
+| xcb
+| xcc
+| xcd
+| xce
+| xcf
+| xd0
+| xd1
+| xd2
+| xd3
+| xd4
+| xd5
+| xd6
+| xd7
+| xd8
+| xd9
+| xda
+| xdb
+| xdc
+| xdd
+| xde
+| xdf
+| xe0
+| xe1
+| xe2
+| xe3
+| xe4
+| xe5
+| xe6
+| xe7
+| xe8
+| xe9
+| xea
+| xeb
+| xec
+| xed
+| xee
+| xef
+| xf0
+| xf1
+| xf2
+| xf3
+| xf4
+| xf5
+| xf6
+| xf7
+| xf8
+| xf9
+| xfa
+| xfb
+| xfc
+| xfd
+| xfe
+| xff
+.
+
+Bind Scope byte_scope with byte.
+
+Register byte as core.byte.type.
+
+Local Notation "0" := false.
+Local Notation "1" := true.
+
+(** We pick a definition that matches with [Ascii.ascii] *)
+Definition of_bits (b : bool * (bool * (bool * (bool * (bool * (bool * (bool * bool))))))) : byte
+ := match b with
+ | (0,(0,(0,(0,(0,(0,(0,0))))))) => x00
+ | (1,(0,(0,(0,(0,(0,(0,0))))))) => x01
+ | (0,(1,(0,(0,(0,(0,(0,0))))))) => x02
+ | (1,(1,(0,(0,(0,(0,(0,0))))))) => x03
+ | (0,(0,(1,(0,(0,(0,(0,0))))))) => x04
+ | (1,(0,(1,(0,(0,(0,(0,0))))))) => x05
+ | (0,(1,(1,(0,(0,(0,(0,0))))))) => x06
+ | (1,(1,(1,(0,(0,(0,(0,0))))))) => x07
+ | (0,(0,(0,(1,(0,(0,(0,0))))))) => x08
+ | (1,(0,(0,(1,(0,(0,(0,0))))))) => x09
+ | (0,(1,(0,(1,(0,(0,(0,0))))))) => x0a
+ | (1,(1,(0,(1,(0,(0,(0,0))))))) => x0b
+ | (0,(0,(1,(1,(0,(0,(0,0))))))) => x0c
+ | (1,(0,(1,(1,(0,(0,(0,0))))))) => x0d
+ | (0,(1,(1,(1,(0,(0,(0,0))))))) => x0e
+ | (1,(1,(1,(1,(0,(0,(0,0))))))) => x0f
+ | (0,(0,(0,(0,(1,(0,(0,0))))))) => x10
+ | (1,(0,(0,(0,(1,(0,(0,0))))))) => x11
+ | (0,(1,(0,(0,(1,(0,(0,0))))))) => x12
+ | (1,(1,(0,(0,(1,(0,(0,0))))))) => x13
+ | (0,(0,(1,(0,(1,(0,(0,0))))))) => x14
+ | (1,(0,(1,(0,(1,(0,(0,0))))))) => x15
+ | (0,(1,(1,(0,(1,(0,(0,0))))))) => x16
+ | (1,(1,(1,(0,(1,(0,(0,0))))))) => x17
+ | (0,(0,(0,(1,(1,(0,(0,0))))))) => x18
+ | (1,(0,(0,(1,(1,(0,(0,0))))))) => x19
+ | (0,(1,(0,(1,(1,(0,(0,0))))))) => x1a
+ | (1,(1,(0,(1,(1,(0,(0,0))))))) => x1b
+ | (0,(0,(1,(1,(1,(0,(0,0))))))) => x1c
+ | (1,(0,(1,(1,(1,(0,(0,0))))))) => x1d
+ | (0,(1,(1,(1,(1,(0,(0,0))))))) => x1e
+ | (1,(1,(1,(1,(1,(0,(0,0))))))) => x1f
+ | (0,(0,(0,(0,(0,(1,(0,0))))))) => x20
+ | (1,(0,(0,(0,(0,(1,(0,0))))))) => x21
+ | (0,(1,(0,(0,(0,(1,(0,0))))))) => x22
+ | (1,(1,(0,(0,(0,(1,(0,0))))))) => x23
+ | (0,(0,(1,(0,(0,(1,(0,0))))))) => x24
+ | (1,(0,(1,(0,(0,(1,(0,0))))))) => x25
+ | (0,(1,(1,(0,(0,(1,(0,0))))))) => x26
+ | (1,(1,(1,(0,(0,(1,(0,0))))))) => x27
+ | (0,(0,(0,(1,(0,(1,(0,0))))))) => x28
+ | (1,(0,(0,(1,(0,(1,(0,0))))))) => x29
+ | (0,(1,(0,(1,(0,(1,(0,0))))))) => x2a
+ | (1,(1,(0,(1,(0,(1,(0,0))))))) => x2b
+ | (0,(0,(1,(1,(0,(1,(0,0))))))) => x2c
+ | (1,(0,(1,(1,(0,(1,(0,0))))))) => x2d
+ | (0,(1,(1,(1,(0,(1,(0,0))))))) => x2e
+ | (1,(1,(1,(1,(0,(1,(0,0))))))) => x2f
+ | (0,(0,(0,(0,(1,(1,(0,0))))))) => x30
+ | (1,(0,(0,(0,(1,(1,(0,0))))))) => x31
+ | (0,(1,(0,(0,(1,(1,(0,0))))))) => x32
+ | (1,(1,(0,(0,(1,(1,(0,0))))))) => x33
+ | (0,(0,(1,(0,(1,(1,(0,0))))))) => x34
+ | (1,(0,(1,(0,(1,(1,(0,0))))))) => x35
+ | (0,(1,(1,(0,(1,(1,(0,0))))))) => x36
+ | (1,(1,(1,(0,(1,(1,(0,0))))))) => x37
+ | (0,(0,(0,(1,(1,(1,(0,0))))))) => x38
+ | (1,(0,(0,(1,(1,(1,(0,0))))))) => x39
+ | (0,(1,(0,(1,(1,(1,(0,0))))))) => x3a
+ | (1,(1,(0,(1,(1,(1,(0,0))))))) => x3b
+ | (0,(0,(1,(1,(1,(1,(0,0))))))) => x3c
+ | (1,(0,(1,(1,(1,(1,(0,0))))))) => x3d
+ | (0,(1,(1,(1,(1,(1,(0,0))))))) => x3e
+ | (1,(1,(1,(1,(1,(1,(0,0))))))) => x3f
+ | (0,(0,(0,(0,(0,(0,(1,0))))))) => x40
+ | (1,(0,(0,(0,(0,(0,(1,0))))))) => x41
+ | (0,(1,(0,(0,(0,(0,(1,0))))))) => x42
+ | (1,(1,(0,(0,(0,(0,(1,0))))))) => x43
+ | (0,(0,(1,(0,(0,(0,(1,0))))))) => x44
+ | (1,(0,(1,(0,(0,(0,(1,0))))))) => x45
+ | (0,(1,(1,(0,(0,(0,(1,0))))))) => x46
+ | (1,(1,(1,(0,(0,(0,(1,0))))))) => x47
+ | (0,(0,(0,(1,(0,(0,(1,0))))))) => x48
+ | (1,(0,(0,(1,(0,(0,(1,0))))))) => x49
+ | (0,(1,(0,(1,(0,(0,(1,0))))))) => x4a
+ | (1,(1,(0,(1,(0,(0,(1,0))))))) => x4b
+ | (0,(0,(1,(1,(0,(0,(1,0))))))) => x4c
+ | (1,(0,(1,(1,(0,(0,(1,0))))))) => x4d
+ | (0,(1,(1,(1,(0,(0,(1,0))))))) => x4e
+ | (1,(1,(1,(1,(0,(0,(1,0))))))) => x4f
+ | (0,(0,(0,(0,(1,(0,(1,0))))))) => x50
+ | (1,(0,(0,(0,(1,(0,(1,0))))))) => x51
+ | (0,(1,(0,(0,(1,(0,(1,0))))))) => x52
+ | (1,(1,(0,(0,(1,(0,(1,0))))))) => x53
+ | (0,(0,(1,(0,(1,(0,(1,0))))))) => x54
+ | (1,(0,(1,(0,(1,(0,(1,0))))))) => x55
+ | (0,(1,(1,(0,(1,(0,(1,0))))))) => x56
+ | (1,(1,(1,(0,(1,(0,(1,0))))))) => x57
+ | (0,(0,(0,(1,(1,(0,(1,0))))))) => x58
+ | (1,(0,(0,(1,(1,(0,(1,0))))))) => x59
+ | (0,(1,(0,(1,(1,(0,(1,0))))))) => x5a
+ | (1,(1,(0,(1,(1,(0,(1,0))))))) => x5b
+ | (0,(0,(1,(1,(1,(0,(1,0))))))) => x5c
+ | (1,(0,(1,(1,(1,(0,(1,0))))))) => x5d
+ | (0,(1,(1,(1,(1,(0,(1,0))))))) => x5e
+ | (1,(1,(1,(1,(1,(0,(1,0))))))) => x5f
+ | (0,(0,(0,(0,(0,(1,(1,0))))))) => x60
+ | (1,(0,(0,(0,(0,(1,(1,0))))))) => x61
+ | (0,(1,(0,(0,(0,(1,(1,0))))))) => x62
+ | (1,(1,(0,(0,(0,(1,(1,0))))))) => x63
+ | (0,(0,(1,(0,(0,(1,(1,0))))))) => x64
+ | (1,(0,(1,(0,(0,(1,(1,0))))))) => x65
+ | (0,(1,(1,(0,(0,(1,(1,0))))))) => x66
+ | (1,(1,(1,(0,(0,(1,(1,0))))))) => x67
+ | (0,(0,(0,(1,(0,(1,(1,0))))))) => x68
+ | (1,(0,(0,(1,(0,(1,(1,0))))))) => x69
+ | (0,(1,(0,(1,(0,(1,(1,0))))))) => x6a
+ | (1,(1,(0,(1,(0,(1,(1,0))))))) => x6b
+ | (0,(0,(1,(1,(0,(1,(1,0))))))) => x6c
+ | (1,(0,(1,(1,(0,(1,(1,0))))))) => x6d
+ | (0,(1,(1,(1,(0,(1,(1,0))))))) => x6e
+ | (1,(1,(1,(1,(0,(1,(1,0))))))) => x6f
+ | (0,(0,(0,(0,(1,(1,(1,0))))))) => x70
+ | (1,(0,(0,(0,(1,(1,(1,0))))))) => x71
+ | (0,(1,(0,(0,(1,(1,(1,0))))))) => x72
+ | (1,(1,(0,(0,(1,(1,(1,0))))))) => x73
+ | (0,(0,(1,(0,(1,(1,(1,0))))))) => x74
+ | (1,(0,(1,(0,(1,(1,(1,0))))))) => x75
+ | (0,(1,(1,(0,(1,(1,(1,0))))))) => x76
+ | (1,(1,(1,(0,(1,(1,(1,0))))))) => x77
+ | (0,(0,(0,(1,(1,(1,(1,0))))))) => x78
+ | (1,(0,(0,(1,(1,(1,(1,0))))))) => x79
+ | (0,(1,(0,(1,(1,(1,(1,0))))))) => x7a
+ | (1,(1,(0,(1,(1,(1,(1,0))))))) => x7b
+ | (0,(0,(1,(1,(1,(1,(1,0))))))) => x7c
+ | (1,(0,(1,(1,(1,(1,(1,0))))))) => x7d
+ | (0,(1,(1,(1,(1,(1,(1,0))))))) => x7e
+ | (1,(1,(1,(1,(1,(1,(1,0))))))) => x7f
+ | (0,(0,(0,(0,(0,(0,(0,1))))))) => x80
+ | (1,(0,(0,(0,(0,(0,(0,1))))))) => x81
+ | (0,(1,(0,(0,(0,(0,(0,1))))))) => x82
+ | (1,(1,(0,(0,(0,(0,(0,1))))))) => x83
+ | (0,(0,(1,(0,(0,(0,(0,1))))))) => x84
+ | (1,(0,(1,(0,(0,(0,(0,1))))))) => x85
+ | (0,(1,(1,(0,(0,(0,(0,1))))))) => x86
+ | (1,(1,(1,(0,(0,(0,(0,1))))))) => x87
+ | (0,(0,(0,(1,(0,(0,(0,1))))))) => x88
+ | (1,(0,(0,(1,(0,(0,(0,1))))))) => x89
+ | (0,(1,(0,(1,(0,(0,(0,1))))))) => x8a
+ | (1,(1,(0,(1,(0,(0,(0,1))))))) => x8b
+ | (0,(0,(1,(1,(0,(0,(0,1))))))) => x8c
+ | (1,(0,(1,(1,(0,(0,(0,1))))))) => x8d
+ | (0,(1,(1,(1,(0,(0,(0,1))))))) => x8e
+ | (1,(1,(1,(1,(0,(0,(0,1))))))) => x8f
+ | (0,(0,(0,(0,(1,(0,(0,1))))))) => x90
+ | (1,(0,(0,(0,(1,(0,(0,1))))))) => x91
+ | (0,(1,(0,(0,(1,(0,(0,1))))))) => x92
+ | (1,(1,(0,(0,(1,(0,(0,1))))))) => x93
+ | (0,(0,(1,(0,(1,(0,(0,1))))))) => x94
+ | (1,(0,(1,(0,(1,(0,(0,1))))))) => x95
+ | (0,(1,(1,(0,(1,(0,(0,1))))))) => x96
+ | (1,(1,(1,(0,(1,(0,(0,1))))))) => x97
+ | (0,(0,(0,(1,(1,(0,(0,1))))))) => x98
+ | (1,(0,(0,(1,(1,(0,(0,1))))))) => x99
+ | (0,(1,(0,(1,(1,(0,(0,1))))))) => x9a
+ | (1,(1,(0,(1,(1,(0,(0,1))))))) => x9b
+ | (0,(0,(1,(1,(1,(0,(0,1))))))) => x9c
+ | (1,(0,(1,(1,(1,(0,(0,1))))))) => x9d
+ | (0,(1,(1,(1,(1,(0,(0,1))))))) => x9e
+ | (1,(1,(1,(1,(1,(0,(0,1))))))) => x9f
+ | (0,(0,(0,(0,(0,(1,(0,1))))))) => xa0
+ | (1,(0,(0,(0,(0,(1,(0,1))))))) => xa1
+ | (0,(1,(0,(0,(0,(1,(0,1))))))) => xa2
+ | (1,(1,(0,(0,(0,(1,(0,1))))))) => xa3
+ | (0,(0,(1,(0,(0,(1,(0,1))))))) => xa4
+ | (1,(0,(1,(0,(0,(1,(0,1))))))) => xa5
+ | (0,(1,(1,(0,(0,(1,(0,1))))))) => xa6
+ | (1,(1,(1,(0,(0,(1,(0,1))))))) => xa7
+ | (0,(0,(0,(1,(0,(1,(0,1))))))) => xa8
+ | (1,(0,(0,(1,(0,(1,(0,1))))))) => xa9
+ | (0,(1,(0,(1,(0,(1,(0,1))))))) => xaa
+ | (1,(1,(0,(1,(0,(1,(0,1))))))) => xab
+ | (0,(0,(1,(1,(0,(1,(0,1))))))) => xac
+ | (1,(0,(1,(1,(0,(1,(0,1))))))) => xad
+ | (0,(1,(1,(1,(0,(1,(0,1))))))) => xae
+ | (1,(1,(1,(1,(0,(1,(0,1))))))) => xaf
+ | (0,(0,(0,(0,(1,(1,(0,1))))))) => xb0
+ | (1,(0,(0,(0,(1,(1,(0,1))))))) => xb1
+ | (0,(1,(0,(0,(1,(1,(0,1))))))) => xb2
+ | (1,(1,(0,(0,(1,(1,(0,1))))))) => xb3
+ | (0,(0,(1,(0,(1,(1,(0,1))))))) => xb4
+ | (1,(0,(1,(0,(1,(1,(0,1))))))) => xb5
+ | (0,(1,(1,(0,(1,(1,(0,1))))))) => xb6
+ | (1,(1,(1,(0,(1,(1,(0,1))))))) => xb7
+ | (0,(0,(0,(1,(1,(1,(0,1))))))) => xb8
+ | (1,(0,(0,(1,(1,(1,(0,1))))))) => xb9
+ | (0,(1,(0,(1,(1,(1,(0,1))))))) => xba
+ | (1,(1,(0,(1,(1,(1,(0,1))))))) => xbb
+ | (0,(0,(1,(1,(1,(1,(0,1))))))) => xbc
+ | (1,(0,(1,(1,(1,(1,(0,1))))))) => xbd
+ | (0,(1,(1,(1,(1,(1,(0,1))))))) => xbe
+ | (1,(1,(1,(1,(1,(1,(0,1))))))) => xbf
+ | (0,(0,(0,(0,(0,(0,(1,1))))))) => xc0
+ | (1,(0,(0,(0,(0,(0,(1,1))))))) => xc1
+ | (0,(1,(0,(0,(0,(0,(1,1))))))) => xc2
+ | (1,(1,(0,(0,(0,(0,(1,1))))))) => xc3
+ | (0,(0,(1,(0,(0,(0,(1,1))))))) => xc4
+ | (1,(0,(1,(0,(0,(0,(1,1))))))) => xc5
+ | (0,(1,(1,(0,(0,(0,(1,1))))))) => xc6
+ | (1,(1,(1,(0,(0,(0,(1,1))))))) => xc7
+ | (0,(0,(0,(1,(0,(0,(1,1))))))) => xc8
+ | (1,(0,(0,(1,(0,(0,(1,1))))))) => xc9
+ | (0,(1,(0,(1,(0,(0,(1,1))))))) => xca
+ | (1,(1,(0,(1,(0,(0,(1,1))))))) => xcb
+ | (0,(0,(1,(1,(0,(0,(1,1))))))) => xcc
+ | (1,(0,(1,(1,(0,(0,(1,1))))))) => xcd
+ | (0,(1,(1,(1,(0,(0,(1,1))))))) => xce
+ | (1,(1,(1,(1,(0,(0,(1,1))))))) => xcf
+ | (0,(0,(0,(0,(1,(0,(1,1))))))) => xd0
+ | (1,(0,(0,(0,(1,(0,(1,1))))))) => xd1
+ | (0,(1,(0,(0,(1,(0,(1,1))))))) => xd2
+ | (1,(1,(0,(0,(1,(0,(1,1))))))) => xd3
+ | (0,(0,(1,(0,(1,(0,(1,1))))))) => xd4
+ | (1,(0,(1,(0,(1,(0,(1,1))))))) => xd5
+ | (0,(1,(1,(0,(1,(0,(1,1))))))) => xd6
+ | (1,(1,(1,(0,(1,(0,(1,1))))))) => xd7
+ | (0,(0,(0,(1,(1,(0,(1,1))))))) => xd8
+ | (1,(0,(0,(1,(1,(0,(1,1))))))) => xd9
+ | (0,(1,(0,(1,(1,(0,(1,1))))))) => xda
+ | (1,(1,(0,(1,(1,(0,(1,1))))))) => xdb
+ | (0,(0,(1,(1,(1,(0,(1,1))))))) => xdc
+ | (1,(0,(1,(1,(1,(0,(1,1))))))) => xdd
+ | (0,(1,(1,(1,(1,(0,(1,1))))))) => xde
+ | (1,(1,(1,(1,(1,(0,(1,1))))))) => xdf
+ | (0,(0,(0,(0,(0,(1,(1,1))))))) => xe0
+ | (1,(0,(0,(0,(0,(1,(1,1))))))) => xe1
+ | (0,(1,(0,(0,(0,(1,(1,1))))))) => xe2
+ | (1,(1,(0,(0,(0,(1,(1,1))))))) => xe3
+ | (0,(0,(1,(0,(0,(1,(1,1))))))) => xe4
+ | (1,(0,(1,(0,(0,(1,(1,1))))))) => xe5
+ | (0,(1,(1,(0,(0,(1,(1,1))))))) => xe6
+ | (1,(1,(1,(0,(0,(1,(1,1))))))) => xe7
+ | (0,(0,(0,(1,(0,(1,(1,1))))))) => xe8
+ | (1,(0,(0,(1,(0,(1,(1,1))))))) => xe9
+ | (0,(1,(0,(1,(0,(1,(1,1))))))) => xea
+ | (1,(1,(0,(1,(0,(1,(1,1))))))) => xeb
+ | (0,(0,(1,(1,(0,(1,(1,1))))))) => xec
+ | (1,(0,(1,(1,(0,(1,(1,1))))))) => xed
+ | (0,(1,(1,(1,(0,(1,(1,1))))))) => xee
+ | (1,(1,(1,(1,(0,(1,(1,1))))))) => xef
+ | (0,(0,(0,(0,(1,(1,(1,1))))))) => xf0
+ | (1,(0,(0,(0,(1,(1,(1,1))))))) => xf1
+ | (0,(1,(0,(0,(1,(1,(1,1))))))) => xf2
+ | (1,(1,(0,(0,(1,(1,(1,1))))))) => xf3
+ | (0,(0,(1,(0,(1,(1,(1,1))))))) => xf4
+ | (1,(0,(1,(0,(1,(1,(1,1))))))) => xf5
+ | (0,(1,(1,(0,(1,(1,(1,1))))))) => xf6
+ | (1,(1,(1,(0,(1,(1,(1,1))))))) => xf7
+ | (0,(0,(0,(1,(1,(1,(1,1))))))) => xf8
+ | (1,(0,(0,(1,(1,(1,(1,1))))))) => xf9
+ | (0,(1,(0,(1,(1,(1,(1,1))))))) => xfa
+ | (1,(1,(0,(1,(1,(1,(1,1))))))) => xfb
+ | (0,(0,(1,(1,(1,(1,(1,1))))))) => xfc
+ | (1,(0,(1,(1,(1,(1,(1,1))))))) => xfd
+ | (0,(1,(1,(1,(1,(1,(1,1))))))) => xfe
+ | (1,(1,(1,(1,(1,(1,(1,1))))))) => xff
+ end.
+
+Definition to_bits (b : byte) : bool * (bool * (bool * (bool * (bool * (bool * (bool * bool))))))
+ := match b with
+ | x00 => (0,(0,(0,(0,(0,(0,(0,0)))))))
+ | x01 => (1,(0,(0,(0,(0,(0,(0,0)))))))
+ | x02 => (0,(1,(0,(0,(0,(0,(0,0)))))))
+ | x03 => (1,(1,(0,(0,(0,(0,(0,0)))))))
+ | x04 => (0,(0,(1,(0,(0,(0,(0,0)))))))
+ | x05 => (1,(0,(1,(0,(0,(0,(0,0)))))))
+ | x06 => (0,(1,(1,(0,(0,(0,(0,0)))))))
+ | x07 => (1,(1,(1,(0,(0,(0,(0,0)))))))
+ | x08 => (0,(0,(0,(1,(0,(0,(0,0)))))))
+ | x09 => (1,(0,(0,(1,(0,(0,(0,0)))))))
+ | x0a => (0,(1,(0,(1,(0,(0,(0,0)))))))
+ | x0b => (1,(1,(0,(1,(0,(0,(0,0)))))))
+ | x0c => (0,(0,(1,(1,(0,(0,(0,0)))))))
+ | x0d => (1,(0,(1,(1,(0,(0,(0,0)))))))
+ | x0e => (0,(1,(1,(1,(0,(0,(0,0)))))))
+ | x0f => (1,(1,(1,(1,(0,(0,(0,0)))))))
+ | x10 => (0,(0,(0,(0,(1,(0,(0,0)))))))
+ | x11 => (1,(0,(0,(0,(1,(0,(0,0)))))))
+ | x12 => (0,(1,(0,(0,(1,(0,(0,0)))))))
+ | x13 => (1,(1,(0,(0,(1,(0,(0,0)))))))
+ | x14 => (0,(0,(1,(0,(1,(0,(0,0)))))))
+ | x15 => (1,(0,(1,(0,(1,(0,(0,0)))))))
+ | x16 => (0,(1,(1,(0,(1,(0,(0,0)))))))
+ | x17 => (1,(1,(1,(0,(1,(0,(0,0)))))))
+ | x18 => (0,(0,(0,(1,(1,(0,(0,0)))))))
+ | x19 => (1,(0,(0,(1,(1,(0,(0,0)))))))
+ | x1a => (0,(1,(0,(1,(1,(0,(0,0)))))))
+ | x1b => (1,(1,(0,(1,(1,(0,(0,0)))))))
+ | x1c => (0,(0,(1,(1,(1,(0,(0,0)))))))
+ | x1d => (1,(0,(1,(1,(1,(0,(0,0)))))))
+ | x1e => (0,(1,(1,(1,(1,(0,(0,0)))))))
+ | x1f => (1,(1,(1,(1,(1,(0,(0,0)))))))
+ | x20 => (0,(0,(0,(0,(0,(1,(0,0)))))))
+ | x21 => (1,(0,(0,(0,(0,(1,(0,0)))))))
+ | x22 => (0,(1,(0,(0,(0,(1,(0,0)))))))
+ | x23 => (1,(1,(0,(0,(0,(1,(0,0)))))))
+ | x24 => (0,(0,(1,(0,(0,(1,(0,0)))))))
+ | x25 => (1,(0,(1,(0,(0,(1,(0,0)))))))
+ | x26 => (0,(1,(1,(0,(0,(1,(0,0)))))))
+ | x27 => (1,(1,(1,(0,(0,(1,(0,0)))))))
+ | x28 => (0,(0,(0,(1,(0,(1,(0,0)))))))
+ | x29 => (1,(0,(0,(1,(0,(1,(0,0)))))))
+ | x2a => (0,(1,(0,(1,(0,(1,(0,0)))))))
+ | x2b => (1,(1,(0,(1,(0,(1,(0,0)))))))
+ | x2c => (0,(0,(1,(1,(0,(1,(0,0)))))))
+ | x2d => (1,(0,(1,(1,(0,(1,(0,0)))))))
+ | x2e => (0,(1,(1,(1,(0,(1,(0,0)))))))
+ | x2f => (1,(1,(1,(1,(0,(1,(0,0)))))))
+ | x30 => (0,(0,(0,(0,(1,(1,(0,0)))))))
+ | x31 => (1,(0,(0,(0,(1,(1,(0,0)))))))
+ | x32 => (0,(1,(0,(0,(1,(1,(0,0)))))))
+ | x33 => (1,(1,(0,(0,(1,(1,(0,0)))))))
+ | x34 => (0,(0,(1,(0,(1,(1,(0,0)))))))
+ | x35 => (1,(0,(1,(0,(1,(1,(0,0)))))))
+ | x36 => (0,(1,(1,(0,(1,(1,(0,0)))))))
+ | x37 => (1,(1,(1,(0,(1,(1,(0,0)))))))
+ | x38 => (0,(0,(0,(1,(1,(1,(0,0)))))))
+ | x39 => (1,(0,(0,(1,(1,(1,(0,0)))))))
+ | x3a => (0,(1,(0,(1,(1,(1,(0,0)))))))
+ | x3b => (1,(1,(0,(1,(1,(1,(0,0)))))))
+ | x3c => (0,(0,(1,(1,(1,(1,(0,0)))))))
+ | x3d => (1,(0,(1,(1,(1,(1,(0,0)))))))
+ | x3e => (0,(1,(1,(1,(1,(1,(0,0)))))))
+ | x3f => (1,(1,(1,(1,(1,(1,(0,0)))))))
+ | x40 => (0,(0,(0,(0,(0,(0,(1,0)))))))
+ | x41 => (1,(0,(0,(0,(0,(0,(1,0)))))))
+ | x42 => (0,(1,(0,(0,(0,(0,(1,0)))))))
+ | x43 => (1,(1,(0,(0,(0,(0,(1,0)))))))
+ | x44 => (0,(0,(1,(0,(0,(0,(1,0)))))))
+ | x45 => (1,(0,(1,(0,(0,(0,(1,0)))))))
+ | x46 => (0,(1,(1,(0,(0,(0,(1,0)))))))
+ | x47 => (1,(1,(1,(0,(0,(0,(1,0)))))))
+ | x48 => (0,(0,(0,(1,(0,(0,(1,0)))))))
+ | x49 => (1,(0,(0,(1,(0,(0,(1,0)))))))
+ | x4a => (0,(1,(0,(1,(0,(0,(1,0)))))))
+ | x4b => (1,(1,(0,(1,(0,(0,(1,0)))))))
+ | x4c => (0,(0,(1,(1,(0,(0,(1,0)))))))
+ | x4d => (1,(0,(1,(1,(0,(0,(1,0)))))))
+ | x4e => (0,(1,(1,(1,(0,(0,(1,0)))))))
+ | x4f => (1,(1,(1,(1,(0,(0,(1,0)))))))
+ | x50 => (0,(0,(0,(0,(1,(0,(1,0)))))))
+ | x51 => (1,(0,(0,(0,(1,(0,(1,0)))))))
+ | x52 => (0,(1,(0,(0,(1,(0,(1,0)))))))
+ | x53 => (1,(1,(0,(0,(1,(0,(1,0)))))))
+ | x54 => (0,(0,(1,(0,(1,(0,(1,0)))))))
+ | x55 => (1,(0,(1,(0,(1,(0,(1,0)))))))
+ | x56 => (0,(1,(1,(0,(1,(0,(1,0)))))))
+ | x57 => (1,(1,(1,(0,(1,(0,(1,0)))))))
+ | x58 => (0,(0,(0,(1,(1,(0,(1,0)))))))
+ | x59 => (1,(0,(0,(1,(1,(0,(1,0)))))))
+ | x5a => (0,(1,(0,(1,(1,(0,(1,0)))))))
+ | x5b => (1,(1,(0,(1,(1,(0,(1,0)))))))
+ | x5c => (0,(0,(1,(1,(1,(0,(1,0)))))))
+ | x5d => (1,(0,(1,(1,(1,(0,(1,0)))))))
+ | x5e => (0,(1,(1,(1,(1,(0,(1,0)))))))
+ | x5f => (1,(1,(1,(1,(1,(0,(1,0)))))))
+ | x60 => (0,(0,(0,(0,(0,(1,(1,0)))))))
+ | x61 => (1,(0,(0,(0,(0,(1,(1,0)))))))
+ | x62 => (0,(1,(0,(0,(0,(1,(1,0)))))))
+ | x63 => (1,(1,(0,(0,(0,(1,(1,0)))))))
+ | x64 => (0,(0,(1,(0,(0,(1,(1,0)))))))
+ | x65 => (1,(0,(1,(0,(0,(1,(1,0)))))))
+ | x66 => (0,(1,(1,(0,(0,(1,(1,0)))))))
+ | x67 => (1,(1,(1,(0,(0,(1,(1,0)))))))
+ | x68 => (0,(0,(0,(1,(0,(1,(1,0)))))))
+ | x69 => (1,(0,(0,(1,(0,(1,(1,0)))))))
+ | x6a => (0,(1,(0,(1,(0,(1,(1,0)))))))
+ | x6b => (1,(1,(0,(1,(0,(1,(1,0)))))))
+ | x6c => (0,(0,(1,(1,(0,(1,(1,0)))))))
+ | x6d => (1,(0,(1,(1,(0,(1,(1,0)))))))
+ | x6e => (0,(1,(1,(1,(0,(1,(1,0)))))))
+ | x6f => (1,(1,(1,(1,(0,(1,(1,0)))))))
+ | x70 => (0,(0,(0,(0,(1,(1,(1,0)))))))
+ | x71 => (1,(0,(0,(0,(1,(1,(1,0)))))))
+ | x72 => (0,(1,(0,(0,(1,(1,(1,0)))))))
+ | x73 => (1,(1,(0,(0,(1,(1,(1,0)))))))
+ | x74 => (0,(0,(1,(0,(1,(1,(1,0)))))))
+ | x75 => (1,(0,(1,(0,(1,(1,(1,0)))))))
+ | x76 => (0,(1,(1,(0,(1,(1,(1,0)))))))
+ | x77 => (1,(1,(1,(0,(1,(1,(1,0)))))))
+ | x78 => (0,(0,(0,(1,(1,(1,(1,0)))))))
+ | x79 => (1,(0,(0,(1,(1,(1,(1,0)))))))
+ | x7a => (0,(1,(0,(1,(1,(1,(1,0)))))))
+ | x7b => (1,(1,(0,(1,(1,(1,(1,0)))))))
+ | x7c => (0,(0,(1,(1,(1,(1,(1,0)))))))
+ | x7d => (1,(0,(1,(1,(1,(1,(1,0)))))))
+ | x7e => (0,(1,(1,(1,(1,(1,(1,0)))))))
+ | x7f => (1,(1,(1,(1,(1,(1,(1,0)))))))
+ | x80 => (0,(0,(0,(0,(0,(0,(0,1)))))))
+ | x81 => (1,(0,(0,(0,(0,(0,(0,1)))))))
+ | x82 => (0,(1,(0,(0,(0,(0,(0,1)))))))
+ | x83 => (1,(1,(0,(0,(0,(0,(0,1)))))))
+ | x84 => (0,(0,(1,(0,(0,(0,(0,1)))))))
+ | x85 => (1,(0,(1,(0,(0,(0,(0,1)))))))
+ | x86 => (0,(1,(1,(0,(0,(0,(0,1)))))))
+ | x87 => (1,(1,(1,(0,(0,(0,(0,1)))))))
+ | x88 => (0,(0,(0,(1,(0,(0,(0,1)))))))
+ | x89 => (1,(0,(0,(1,(0,(0,(0,1)))))))
+ | x8a => (0,(1,(0,(1,(0,(0,(0,1)))))))
+ | x8b => (1,(1,(0,(1,(0,(0,(0,1)))))))
+ | x8c => (0,(0,(1,(1,(0,(0,(0,1)))))))
+ | x8d => (1,(0,(1,(1,(0,(0,(0,1)))))))
+ | x8e => (0,(1,(1,(1,(0,(0,(0,1)))))))
+ | x8f => (1,(1,(1,(1,(0,(0,(0,1)))))))
+ | x90 => (0,(0,(0,(0,(1,(0,(0,1)))))))
+ | x91 => (1,(0,(0,(0,(1,(0,(0,1)))))))
+ | x92 => (0,(1,(0,(0,(1,(0,(0,1)))))))
+ | x93 => (1,(1,(0,(0,(1,(0,(0,1)))))))
+ | x94 => (0,(0,(1,(0,(1,(0,(0,1)))))))
+ | x95 => (1,(0,(1,(0,(1,(0,(0,1)))))))
+ | x96 => (0,(1,(1,(0,(1,(0,(0,1)))))))
+ | x97 => (1,(1,(1,(0,(1,(0,(0,1)))))))
+ | x98 => (0,(0,(0,(1,(1,(0,(0,1)))))))
+ | x99 => (1,(0,(0,(1,(1,(0,(0,1)))))))
+ | x9a => (0,(1,(0,(1,(1,(0,(0,1)))))))
+ | x9b => (1,(1,(0,(1,(1,(0,(0,1)))))))
+ | x9c => (0,(0,(1,(1,(1,(0,(0,1)))))))
+ | x9d => (1,(0,(1,(1,(1,(0,(0,1)))))))
+ | x9e => (0,(1,(1,(1,(1,(0,(0,1)))))))
+ | x9f => (1,(1,(1,(1,(1,(0,(0,1)))))))
+ | xa0 => (0,(0,(0,(0,(0,(1,(0,1)))))))
+ | xa1 => (1,(0,(0,(0,(0,(1,(0,1)))))))
+ | xa2 => (0,(1,(0,(0,(0,(1,(0,1)))))))
+ | xa3 => (1,(1,(0,(0,(0,(1,(0,1)))))))
+ | xa4 => (0,(0,(1,(0,(0,(1,(0,1)))))))
+ | xa5 => (1,(0,(1,(0,(0,(1,(0,1)))))))
+ | xa6 => (0,(1,(1,(0,(0,(1,(0,1)))))))
+ | xa7 => (1,(1,(1,(0,(0,(1,(0,1)))))))
+ | xa8 => (0,(0,(0,(1,(0,(1,(0,1)))))))
+ | xa9 => (1,(0,(0,(1,(0,(1,(0,1)))))))
+ | xaa => (0,(1,(0,(1,(0,(1,(0,1)))))))
+ | xab => (1,(1,(0,(1,(0,(1,(0,1)))))))
+ | xac => (0,(0,(1,(1,(0,(1,(0,1)))))))
+ | xad => (1,(0,(1,(1,(0,(1,(0,1)))))))
+ | xae => (0,(1,(1,(1,(0,(1,(0,1)))))))
+ | xaf => (1,(1,(1,(1,(0,(1,(0,1)))))))
+ | xb0 => (0,(0,(0,(0,(1,(1,(0,1)))))))
+ | xb1 => (1,(0,(0,(0,(1,(1,(0,1)))))))
+ | xb2 => (0,(1,(0,(0,(1,(1,(0,1)))))))
+ | xb3 => (1,(1,(0,(0,(1,(1,(0,1)))))))
+ | xb4 => (0,(0,(1,(0,(1,(1,(0,1)))))))
+ | xb5 => (1,(0,(1,(0,(1,(1,(0,1)))))))
+ | xb6 => (0,(1,(1,(0,(1,(1,(0,1)))))))
+ | xb7 => (1,(1,(1,(0,(1,(1,(0,1)))))))
+ | xb8 => (0,(0,(0,(1,(1,(1,(0,1)))))))
+ | xb9 => (1,(0,(0,(1,(1,(1,(0,1)))))))
+ | xba => (0,(1,(0,(1,(1,(1,(0,1)))))))
+ | xbb => (1,(1,(0,(1,(1,(1,(0,1)))))))
+ | xbc => (0,(0,(1,(1,(1,(1,(0,1)))))))
+ | xbd => (1,(0,(1,(1,(1,(1,(0,1)))))))
+ | xbe => (0,(1,(1,(1,(1,(1,(0,1)))))))
+ | xbf => (1,(1,(1,(1,(1,(1,(0,1)))))))
+ | xc0 => (0,(0,(0,(0,(0,(0,(1,1)))))))
+ | xc1 => (1,(0,(0,(0,(0,(0,(1,1)))))))
+ | xc2 => (0,(1,(0,(0,(0,(0,(1,1)))))))
+ | xc3 => (1,(1,(0,(0,(0,(0,(1,1)))))))
+ | xc4 => (0,(0,(1,(0,(0,(0,(1,1)))))))
+ | xc5 => (1,(0,(1,(0,(0,(0,(1,1)))))))
+ | xc6 => (0,(1,(1,(0,(0,(0,(1,1)))))))
+ | xc7 => (1,(1,(1,(0,(0,(0,(1,1)))))))
+ | xc8 => (0,(0,(0,(1,(0,(0,(1,1)))))))
+ | xc9 => (1,(0,(0,(1,(0,(0,(1,1)))))))
+ | xca => (0,(1,(0,(1,(0,(0,(1,1)))))))
+ | xcb => (1,(1,(0,(1,(0,(0,(1,1)))))))
+ | xcc => (0,(0,(1,(1,(0,(0,(1,1)))))))
+ | xcd => (1,(0,(1,(1,(0,(0,(1,1)))))))
+ | xce => (0,(1,(1,(1,(0,(0,(1,1)))))))
+ | xcf => (1,(1,(1,(1,(0,(0,(1,1)))))))
+ | xd0 => (0,(0,(0,(0,(1,(0,(1,1)))))))
+ | xd1 => (1,(0,(0,(0,(1,(0,(1,1)))))))
+ | xd2 => (0,(1,(0,(0,(1,(0,(1,1)))))))
+ | xd3 => (1,(1,(0,(0,(1,(0,(1,1)))))))
+ | xd4 => (0,(0,(1,(0,(1,(0,(1,1)))))))
+ | xd5 => (1,(0,(1,(0,(1,(0,(1,1)))))))
+ | xd6 => (0,(1,(1,(0,(1,(0,(1,1)))))))
+ | xd7 => (1,(1,(1,(0,(1,(0,(1,1)))))))
+ | xd8 => (0,(0,(0,(1,(1,(0,(1,1)))))))
+ | xd9 => (1,(0,(0,(1,(1,(0,(1,1)))))))
+ | xda => (0,(1,(0,(1,(1,(0,(1,1)))))))
+ | xdb => (1,(1,(0,(1,(1,(0,(1,1)))))))
+ | xdc => (0,(0,(1,(1,(1,(0,(1,1)))))))
+ | xdd => (1,(0,(1,(1,(1,(0,(1,1)))))))
+ | xde => (0,(1,(1,(1,(1,(0,(1,1)))))))
+ | xdf => (1,(1,(1,(1,(1,(0,(1,1)))))))
+ | xe0 => (0,(0,(0,(0,(0,(1,(1,1)))))))
+ | xe1 => (1,(0,(0,(0,(0,(1,(1,1)))))))
+ | xe2 => (0,(1,(0,(0,(0,(1,(1,1)))))))
+ | xe3 => (1,(1,(0,(0,(0,(1,(1,1)))))))
+ | xe4 => (0,(0,(1,(0,(0,(1,(1,1)))))))
+ | xe5 => (1,(0,(1,(0,(0,(1,(1,1)))))))
+ | xe6 => (0,(1,(1,(0,(0,(1,(1,1)))))))
+ | xe7 => (1,(1,(1,(0,(0,(1,(1,1)))))))
+ | xe8 => (0,(0,(0,(1,(0,(1,(1,1)))))))
+ | xe9 => (1,(0,(0,(1,(0,(1,(1,1)))))))
+ | xea => (0,(1,(0,(1,(0,(1,(1,1)))))))
+ | xeb => (1,(1,(0,(1,(0,(1,(1,1)))))))
+ | xec => (0,(0,(1,(1,(0,(1,(1,1)))))))
+ | xed => (1,(0,(1,(1,(0,(1,(1,1)))))))
+ | xee => (0,(1,(1,(1,(0,(1,(1,1)))))))
+ | xef => (1,(1,(1,(1,(0,(1,(1,1)))))))
+ | xf0 => (0,(0,(0,(0,(1,(1,(1,1)))))))
+ | xf1 => (1,(0,(0,(0,(1,(1,(1,1)))))))
+ | xf2 => (0,(1,(0,(0,(1,(1,(1,1)))))))
+ | xf3 => (1,(1,(0,(0,(1,(1,(1,1)))))))
+ | xf4 => (0,(0,(1,(0,(1,(1,(1,1)))))))
+ | xf5 => (1,(0,(1,(0,(1,(1,(1,1)))))))
+ | xf6 => (0,(1,(1,(0,(1,(1,(1,1)))))))
+ | xf7 => (1,(1,(1,(0,(1,(1,(1,1)))))))
+ | xf8 => (0,(0,(0,(1,(1,(1,(1,1)))))))
+ | xf9 => (1,(0,(0,(1,(1,(1,(1,1)))))))
+ | xfa => (0,(1,(0,(1,(1,(1,(1,1)))))))
+ | xfb => (1,(1,(0,(1,(1,(1,(1,1)))))))
+ | xfc => (0,(0,(1,(1,(1,(1,(1,1)))))))
+ | xfd => (1,(0,(1,(1,(1,(1,(1,1)))))))
+ | xfe => (0,(1,(1,(1,(1,(1,(1,1)))))))
+ | xff => (1,(1,(1,(1,(1,(1,(1,1)))))))
+ end.
+
+Lemma of_bits_to_bits (b : byte) : of_bits (to_bits b) = b.
+Proof. destruct b; exact eq_refl. Qed.
+
+Lemma to_bits_of_bits (b : _) : to_bits (of_bits b) = b.
+Proof.
+ repeat match goal with
+ | p : prod _ _ |- _ => destruct p
+ | b : bool |- _ => destruct b
+ end;
+ exact eq_refl.
+Qed.
+
+Definition byte_of_byte (b : byte) : byte := b.
+
+Module Export ByteSyntaxNotations.
+ String Notation byte byte_of_byte byte_of_byte : byte_scope.
+End ByteSyntaxNotations.
diff --git a/theories/Init/Prelude.v b/theories/Init/Prelude.v
index 6d98bcb34a..5e29f854e8 100644
--- a/theories/Init/Prelude.v
+++ b/theories/Init/Prelude.v
@@ -13,6 +13,7 @@ Require Export Logic.
Require Export Logic_Type.
Require Export Datatypes.
Require Export Specif.
+Require Coq.Init.Byte.
Require Coq.Init.Decimal.
Require Coq.Init.Nat.
Require Export Peano.
@@ -26,6 +27,7 @@ Require Export Coq.Init.Tauto.
Declare ML Module "cc_plugin".
Declare ML Module "ground_plugin".
Declare ML Module "numeral_notation_plugin".
+Declare ML Module "string_notation_plugin".
(* Parsing / printing of decimal numbers *)
Arguments Nat.of_uint d%dec_uint_scope.
@@ -38,5 +40,8 @@ Numeral Notation Decimal.int Decimal.int_of_int Decimal.int_of_int
(* Parsing / printing of [nat] numbers *)
Numeral Notation nat Nat.of_uint Nat.to_uint : nat_scope (abstract after 5000).
+(* Printing/Parsing of bytes *)
+Export Byte.ByteSyntaxNotations.
+
(* Default substrings not considered by queries like SearchAbout *)
Add Search Blacklist "_subproof" "_subterm" "Private_".
diff --git a/theories/Strings/Ascii.v b/theories/Strings/Ascii.v
index b7c1eaa788..6a0c5f066e 100644
--- a/theories/Strings/Ascii.v
+++ b/theories/Strings/Ascii.v
@@ -12,7 +12,7 @@
(** Contributed by Laurent Théry (INRIA);
Adapted to Coq V8 by the Coq Development Team *)
-Require Import Bool BinPos BinNat PeanoNat Nnat.
+Require Import Bool BinPos BinNat PeanoNat Nnat Coq.Strings.Byte.
(** * Definition of ascii characters *)
@@ -20,10 +20,7 @@ Require Import Bool BinPos BinNat PeanoNat Nnat.
Inductive ascii : Set := Ascii (_ _ _ _ _ _ _ _ : bool).
-Register Ascii as plugins.syntax.Ascii.
-
Declare Scope char_scope.
-Module Export AsciiSyntax. Declare ML Module "ascii_syntax_plugin". End AsciiSyntax.
Delimit Scope char_scope with char.
Bind Scope char_scope with ascii.
@@ -140,6 +137,12 @@ do 8 (destruct p; [ | | intros; vm_compute; reflexivity ]);
intro H; vm_compute in H; destruct p; discriminate.
Qed.
+Theorem N_ascii_bounded :
+ forall a : ascii, (N_of_ascii a < 256)%N.
+Proof.
+ destruct a as [[|][|][|][|][|][|][|][|]]; vm_compute; reflexivity.
+Qed.
+
Theorem ascii_nat_embedding :
forall a : ascii, ascii_of_nat (nat_of_ascii a) = a.
Proof.
@@ -158,6 +161,15 @@ Proof.
now apply Nat.compare_lt_iff.
Qed.
+Theorem nat_ascii_bounded :
+ forall a : ascii, nat_of_ascii a < 256.
+Proof.
+ intro a; unfold nat_of_ascii.
+ change 256 with (N.to_nat 256).
+ rewrite <- Nat.compare_lt_iff, <- N2Nat.inj_compare, N.compare_lt_iff.
+ apply N_ascii_bounded.
+Qed.
+
(** * Concrete syntax *)
@@ -175,7 +187,53 @@ Qed.
stand-alone utf8 characters so that only the notation "nnn" is
available for them (unless your terminal is able to represent them,
which is typically not the case in coqide).
-*)
+ *)
+
+Definition ascii_of_byte (b : byte) : ascii
+ := let '(b0, (b1, (b2, (b3, (b4, (b5, (b6, b7))))))) := Byte.to_bits b in
+ Ascii b0 b1 b2 b3 b4 b5 b6 b7.
+
+Definition byte_of_ascii (a : ascii) : byte
+ := let (b0, b1, b2, b3, b4, b5, b6, b7) := a in
+ Byte.of_bits (b0, (b1, (b2, (b3, (b4, (b5, (b6, b7))))))).
+
+Lemma ascii_of_byte_of_ascii x : ascii_of_byte (byte_of_ascii x) = x.
+Proof.
+ cbv [ascii_of_byte byte_of_ascii].
+ destruct x; rewrite to_bits_of_bits; reflexivity.
+Qed.
+
+Lemma byte_of_ascii_of_byte x : byte_of_ascii (ascii_of_byte x) = x.
+Proof.
+ cbv [ascii_of_byte byte_of_ascii].
+ repeat match goal with
+ | [ |- context[match ?x with pair _ _ => _ end] ]
+ => rewrite (surjective_pairing x)
+ | [ |- context[(fst ?x, snd ?x)] ]
+ => rewrite <- (surjective_pairing x)
+ end.
+ rewrite of_bits_to_bits; reflexivity.
+Qed.
+
+Lemma ascii_of_byte_via_N x : ascii_of_byte x = ascii_of_N (Byte.to_N x).
+Proof. destruct x; reflexivity. Qed.
+
+Lemma ascii_of_byte_via_nat x : ascii_of_byte x = ascii_of_nat (Byte.to_nat x).
+Proof. destruct x; reflexivity. Qed.
+
+Lemma byte_of_ascii_via_N x : Some (byte_of_ascii x) = Byte.of_N (N_of_ascii x).
+Proof.
+ rewrite <- (ascii_of_byte_of_ascii x); destruct (byte_of_ascii x); reflexivity.
+Qed.
+
+Lemma byte_of_ascii_via_nat x : Some (byte_of_ascii x) = Byte.of_nat (nat_of_ascii x).
+Proof.
+ rewrite <- (ascii_of_byte_of_ascii x); destruct (byte_of_ascii x); reflexivity.
+Qed.
+
+Module Export AsciiSyntax.
+ String Notation ascii ascii_of_byte byte_of_ascii : char_scope.
+End AsciiSyntax.
Local Open Scope char_scope.
diff --git a/theories/Strings/Byte.v b/theories/Strings/Byte.v
new file mode 100644
index 0000000000..2759ea60cb
--- /dev/null
+++ b/theories/Strings/Byte.v
@@ -0,0 +1,1214 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+Require Import Coq.Arith.EqNat.
+Require Import Coq.NArith.BinNat.
+Require Import Coq.NArith.Nnat.
+Require Export Coq.Init.Byte.
+
+Local Set Implicit Arguments.
+
+Definition eqb (a b : byte) : bool
+ := let '(a0, (a1, (a2, (a3, (a4, (a5, (a6, a7))))))) := to_bits a in
+ let '(b0, (b1, (b2, (b3, (b4, (b5, (b6, b7))))))) := to_bits b in
+ (Bool.eqb a0 b0 && Bool.eqb a1 b1 && Bool.eqb a2 b2 && Bool.eqb a3 b3 &&
+ Bool.eqb a4 b4 && Bool.eqb a5 b5 && Bool.eqb a6 b6 && Bool.eqb a7 b7)%bool.
+
+Module Export ByteNotations.
+ Export ByteSyntaxNotations.
+ Infix "=?" := eqb (at level 70) : byte_scope.
+End ByteNotations.
+
+Lemma byte_dec_lb x y : x = y -> eqb x y = true.
+Proof. intro; subst y; destruct x; reflexivity. Defined.
+
+Lemma byte_dec_bl x y (H : eqb x y = true) : x = y.
+Proof.
+ rewrite <- (of_bits_to_bits x), <- (of_bits_to_bits y).
+ cbv [eqb] in H; revert H.
+ generalize (to_bits x) (to_bits y); clear x y; intros x y H.
+ repeat match goal with
+ | [ H : and _ _ |- _ ] => destruct H
+ | [ H : prod _ _ |- _ ] => destruct H
+ | [ H : context[andb _ _ = true] |- _ ] => rewrite Bool.andb_true_iff in H
+ | [ H : context[Bool.eqb _ _ = true] |- _ ] => rewrite Bool.eqb_true_iff in H
+ | _ => progress subst
+ | _ => reflexivity
+ end.
+Qed.
+
+Lemma eqb_false x y : eqb x y = false -> x <> y.
+Proof. intros H H'; pose proof (byte_dec_lb H'); congruence. Qed.
+
+Definition byte_eq_dec (x y : byte) : {x = y} + {x <> y}
+ := (if eqb x y as beq return eqb x y = beq -> _
+ then fun pf => left (byte_dec_bl x y pf)
+ else fun pf => right (eqb_false pf))
+ eq_refl.
+
+Section nat.
+ Definition to_nat (x : byte) : nat
+ := match x with
+ | x00 => 0
+ | x01 => 1
+ | x02 => 2
+ | x03 => 3
+ | x04 => 4
+ | x05 => 5
+ | x06 => 6
+ | x07 => 7
+ | x08 => 8
+ | x09 => 9
+ | x0a => 10
+ | x0b => 11
+ | x0c => 12
+ | x0d => 13
+ | x0e => 14
+ | x0f => 15
+ | x10 => 16
+ | x11 => 17
+ | x12 => 18
+ | x13 => 19
+ | x14 => 20
+ | x15 => 21
+ | x16 => 22
+ | x17 => 23
+ | x18 => 24
+ | x19 => 25
+ | x1a => 26
+ | x1b => 27
+ | x1c => 28
+ | x1d => 29
+ | x1e => 30
+ | x1f => 31
+ | x20 => 32
+ | x21 => 33
+ | x22 => 34
+ | x23 => 35
+ | x24 => 36
+ | x25 => 37
+ | x26 => 38
+ | x27 => 39
+ | x28 => 40
+ | x29 => 41
+ | x2a => 42
+ | x2b => 43
+ | x2c => 44
+ | x2d => 45
+ | x2e => 46
+ | x2f => 47
+ | x30 => 48
+ | x31 => 49
+ | x32 => 50
+ | x33 => 51
+ | x34 => 52
+ | x35 => 53
+ | x36 => 54
+ | x37 => 55
+ | x38 => 56
+ | x39 => 57
+ | x3a => 58
+ | x3b => 59
+ | x3c => 60
+ | x3d => 61
+ | x3e => 62
+ | x3f => 63
+ | x40 => 64
+ | x41 => 65
+ | x42 => 66
+ | x43 => 67
+ | x44 => 68
+ | x45 => 69
+ | x46 => 70
+ | x47 => 71
+ | x48 => 72
+ | x49 => 73
+ | x4a => 74
+ | x4b => 75
+ | x4c => 76
+ | x4d => 77
+ | x4e => 78
+ | x4f => 79
+ | x50 => 80
+ | x51 => 81
+ | x52 => 82
+ | x53 => 83
+ | x54 => 84
+ | x55 => 85
+ | x56 => 86
+ | x57 => 87
+ | x58 => 88
+ | x59 => 89
+ | x5a => 90
+ | x5b => 91
+ | x5c => 92
+ | x5d => 93
+ | x5e => 94
+ | x5f => 95
+ | x60 => 96
+ | x61 => 97
+ | x62 => 98
+ | x63 => 99
+ | x64 => 100
+ | x65 => 101
+ | x66 => 102
+ | x67 => 103
+ | x68 => 104
+ | x69 => 105
+ | x6a => 106
+ | x6b => 107
+ | x6c => 108
+ | x6d => 109
+ | x6e => 110
+ | x6f => 111
+ | x70 => 112
+ | x71 => 113
+ | x72 => 114
+ | x73 => 115
+ | x74 => 116
+ | x75 => 117
+ | x76 => 118
+ | x77 => 119
+ | x78 => 120
+ | x79 => 121
+ | x7a => 122
+ | x7b => 123
+ | x7c => 124
+ | x7d => 125
+ | x7e => 126
+ | x7f => 127
+ | x80 => 128
+ | x81 => 129
+ | x82 => 130
+ | x83 => 131
+ | x84 => 132
+ | x85 => 133
+ | x86 => 134
+ | x87 => 135
+ | x88 => 136
+ | x89 => 137
+ | x8a => 138
+ | x8b => 139
+ | x8c => 140
+ | x8d => 141
+ | x8e => 142
+ | x8f => 143
+ | x90 => 144
+ | x91 => 145
+ | x92 => 146
+ | x93 => 147
+ | x94 => 148
+ | x95 => 149
+ | x96 => 150
+ | x97 => 151
+ | x98 => 152
+ | x99 => 153
+ | x9a => 154
+ | x9b => 155
+ | x9c => 156
+ | x9d => 157
+ | x9e => 158
+ | x9f => 159
+ | xa0 => 160
+ | xa1 => 161
+ | xa2 => 162
+ | xa3 => 163
+ | xa4 => 164
+ | xa5 => 165
+ | xa6 => 166
+ | xa7 => 167
+ | xa8 => 168
+ | xa9 => 169
+ | xaa => 170
+ | xab => 171
+ | xac => 172
+ | xad => 173
+ | xae => 174
+ | xaf => 175
+ | xb0 => 176
+ | xb1 => 177
+ | xb2 => 178
+ | xb3 => 179
+ | xb4 => 180
+ | xb5 => 181
+ | xb6 => 182
+ | xb7 => 183
+ | xb8 => 184
+ | xb9 => 185
+ | xba => 186
+ | xbb => 187
+ | xbc => 188
+ | xbd => 189
+ | xbe => 190
+ | xbf => 191
+ | xc0 => 192
+ | xc1 => 193
+ | xc2 => 194
+ | xc3 => 195
+ | xc4 => 196
+ | xc5 => 197
+ | xc6 => 198
+ | xc7 => 199
+ | xc8 => 200
+ | xc9 => 201
+ | xca => 202
+ | xcb => 203
+ | xcc => 204
+ | xcd => 205
+ | xce => 206
+ | xcf => 207
+ | xd0 => 208
+ | xd1 => 209
+ | xd2 => 210
+ | xd3 => 211
+ | xd4 => 212
+ | xd5 => 213
+ | xd6 => 214
+ | xd7 => 215
+ | xd8 => 216
+ | xd9 => 217
+ | xda => 218
+ | xdb => 219
+ | xdc => 220
+ | xdd => 221
+ | xde => 222
+ | xdf => 223
+ | xe0 => 224
+ | xe1 => 225
+ | xe2 => 226
+ | xe3 => 227
+ | xe4 => 228
+ | xe5 => 229
+ | xe6 => 230
+ | xe7 => 231
+ | xe8 => 232
+ | xe9 => 233
+ | xea => 234
+ | xeb => 235
+ | xec => 236
+ | xed => 237
+ | xee => 238
+ | xef => 239
+ | xf0 => 240
+ | xf1 => 241
+ | xf2 => 242
+ | xf3 => 243
+ | xf4 => 244
+ | xf5 => 245
+ | xf6 => 246
+ | xf7 => 247
+ | xf8 => 248
+ | xf9 => 249
+ | xfa => 250
+ | xfb => 251
+ | xfc => 252
+ | xfd => 253
+ | xfe => 254
+ | xff => 255
+ end.
+
+ Definition of_nat (x : nat) : option byte
+ := match x with
+ | 0 => Some x00
+ | 1 => Some x01
+ | 2 => Some x02
+ | 3 => Some x03
+ | 4 => Some x04
+ | 5 => Some x05
+ | 6 => Some x06
+ | 7 => Some x07
+ | 8 => Some x08
+ | 9 => Some x09
+ | 10 => Some x0a
+ | 11 => Some x0b
+ | 12 => Some x0c
+ | 13 => Some x0d
+ | 14 => Some x0e
+ | 15 => Some x0f
+ | 16 => Some x10
+ | 17 => Some x11
+ | 18 => Some x12
+ | 19 => Some x13
+ | 20 => Some x14
+ | 21 => Some x15
+ | 22 => Some x16
+ | 23 => Some x17
+ | 24 => Some x18
+ | 25 => Some x19
+ | 26 => Some x1a
+ | 27 => Some x1b
+ | 28 => Some x1c
+ | 29 => Some x1d
+ | 30 => Some x1e
+ | 31 => Some x1f
+ | 32 => Some x20
+ | 33 => Some x21
+ | 34 => Some x22
+ | 35 => Some x23
+ | 36 => Some x24
+ | 37 => Some x25
+ | 38 => Some x26
+ | 39 => Some x27
+ | 40 => Some x28
+ | 41 => Some x29
+ | 42 => Some x2a
+ | 43 => Some x2b
+ | 44 => Some x2c
+ | 45 => Some x2d
+ | 46 => Some x2e
+ | 47 => Some x2f
+ | 48 => Some x30
+ | 49 => Some x31
+ | 50 => Some x32
+ | 51 => Some x33
+ | 52 => Some x34
+ | 53 => Some x35
+ | 54 => Some x36
+ | 55 => Some x37
+ | 56 => Some x38
+ | 57 => Some x39
+ | 58 => Some x3a
+ | 59 => Some x3b
+ | 60 => Some x3c
+ | 61 => Some x3d
+ | 62 => Some x3e
+ | 63 => Some x3f
+ | 64 => Some x40
+ | 65 => Some x41
+ | 66 => Some x42
+ | 67 => Some x43
+ | 68 => Some x44
+ | 69 => Some x45
+ | 70 => Some x46
+ | 71 => Some x47
+ | 72 => Some x48
+ | 73 => Some x49
+ | 74 => Some x4a
+ | 75 => Some x4b
+ | 76 => Some x4c
+ | 77 => Some x4d
+ | 78 => Some x4e
+ | 79 => Some x4f
+ | 80 => Some x50
+ | 81 => Some x51
+ | 82 => Some x52
+ | 83 => Some x53
+ | 84 => Some x54
+ | 85 => Some x55
+ | 86 => Some x56
+ | 87 => Some x57
+ | 88 => Some x58
+ | 89 => Some x59
+ | 90 => Some x5a
+ | 91 => Some x5b
+ | 92 => Some x5c
+ | 93 => Some x5d
+ | 94 => Some x5e
+ | 95 => Some x5f
+ | 96 => Some x60
+ | 97 => Some x61
+ | 98 => Some x62
+ | 99 => Some x63
+ | 100 => Some x64
+ | 101 => Some x65
+ | 102 => Some x66
+ | 103 => Some x67
+ | 104 => Some x68
+ | 105 => Some x69
+ | 106 => Some x6a
+ | 107 => Some x6b
+ | 108 => Some x6c
+ | 109 => Some x6d
+ | 110 => Some x6e
+ | 111 => Some x6f
+ | 112 => Some x70
+ | 113 => Some x71
+ | 114 => Some x72
+ | 115 => Some x73
+ | 116 => Some x74
+ | 117 => Some x75
+ | 118 => Some x76
+ | 119 => Some x77
+ | 120 => Some x78
+ | 121 => Some x79
+ | 122 => Some x7a
+ | 123 => Some x7b
+ | 124 => Some x7c
+ | 125 => Some x7d
+ | 126 => Some x7e
+ | 127 => Some x7f
+ | 128 => Some x80
+ | 129 => Some x81
+ | 130 => Some x82
+ | 131 => Some x83
+ | 132 => Some x84
+ | 133 => Some x85
+ | 134 => Some x86
+ | 135 => Some x87
+ | 136 => Some x88
+ | 137 => Some x89
+ | 138 => Some x8a
+ | 139 => Some x8b
+ | 140 => Some x8c
+ | 141 => Some x8d
+ | 142 => Some x8e
+ | 143 => Some x8f
+ | 144 => Some x90
+ | 145 => Some x91
+ | 146 => Some x92
+ | 147 => Some x93
+ | 148 => Some x94
+ | 149 => Some x95
+ | 150 => Some x96
+ | 151 => Some x97
+ | 152 => Some x98
+ | 153 => Some x99
+ | 154 => Some x9a
+ | 155 => Some x9b
+ | 156 => Some x9c
+ | 157 => Some x9d
+ | 158 => Some x9e
+ | 159 => Some x9f
+ | 160 => Some xa0
+ | 161 => Some xa1
+ | 162 => Some xa2
+ | 163 => Some xa3
+ | 164 => Some xa4
+ | 165 => Some xa5
+ | 166 => Some xa6
+ | 167 => Some xa7
+ | 168 => Some xa8
+ | 169 => Some xa9
+ | 170 => Some xaa
+ | 171 => Some xab
+ | 172 => Some xac
+ | 173 => Some xad
+ | 174 => Some xae
+ | 175 => Some xaf
+ | 176 => Some xb0
+ | 177 => Some xb1
+ | 178 => Some xb2
+ | 179 => Some xb3
+ | 180 => Some xb4
+ | 181 => Some xb5
+ | 182 => Some xb6
+ | 183 => Some xb7
+ | 184 => Some xb8
+ | 185 => Some xb9
+ | 186 => Some xba
+ | 187 => Some xbb
+ | 188 => Some xbc
+ | 189 => Some xbd
+ | 190 => Some xbe
+ | 191 => Some xbf
+ | 192 => Some xc0
+ | 193 => Some xc1
+ | 194 => Some xc2
+ | 195 => Some xc3
+ | 196 => Some xc4
+ | 197 => Some xc5
+ | 198 => Some xc6
+ | 199 => Some xc7
+ | 200 => Some xc8
+ | 201 => Some xc9
+ | 202 => Some xca
+ | 203 => Some xcb
+ | 204 => Some xcc
+ | 205 => Some xcd
+ | 206 => Some xce
+ | 207 => Some xcf
+ | 208 => Some xd0
+ | 209 => Some xd1
+ | 210 => Some xd2
+ | 211 => Some xd3
+ | 212 => Some xd4
+ | 213 => Some xd5
+ | 214 => Some xd6
+ | 215 => Some xd7
+ | 216 => Some xd8
+ | 217 => Some xd9
+ | 218 => Some xda
+ | 219 => Some xdb
+ | 220 => Some xdc
+ | 221 => Some xdd
+ | 222 => Some xde
+ | 223 => Some xdf
+ | 224 => Some xe0
+ | 225 => Some xe1
+ | 226 => Some xe2
+ | 227 => Some xe3
+ | 228 => Some xe4
+ | 229 => Some xe5
+ | 230 => Some xe6
+ | 231 => Some xe7
+ | 232 => Some xe8
+ | 233 => Some xe9
+ | 234 => Some xea
+ | 235 => Some xeb
+ | 236 => Some xec
+ | 237 => Some xed
+ | 238 => Some xee
+ | 239 => Some xef
+ | 240 => Some xf0
+ | 241 => Some xf1
+ | 242 => Some xf2
+ | 243 => Some xf3
+ | 244 => Some xf4
+ | 245 => Some xf5
+ | 246 => Some xf6
+ | 247 => Some xf7
+ | 248 => Some xf8
+ | 249 => Some xf9
+ | 250 => Some xfa
+ | 251 => Some xfb
+ | 252 => Some xfc
+ | 253 => Some xfd
+ | 254 => Some xfe
+ | 255 => Some xff
+ | _ => None
+ end.
+
+ Lemma of_to_nat x : of_nat (to_nat x) = Some x.
+ Proof. destruct x; reflexivity. Qed.
+
+ Lemma to_of_nat x y : of_nat x = Some y -> to_nat y = x.
+ Proof.
+ do 256 try destruct x as [|x]; cbv [of_nat]; intro.
+ all: repeat match goal with
+ | _ => reflexivity
+ | _ => progress subst
+ | [ H : Some ?a = Some ?b |- _ ] => assert (a = b) by refine match H with eq_refl => eq_refl end; clear H
+ | [ H : None = Some _ |- _ ] => solve [ inversion H ]
+ end.
+ Qed.
+
+ Lemma to_of_nat_iff x y : of_nat x = Some y <-> to_nat y = x.
+ Proof. split; intro; subst; (apply of_to_nat || apply to_of_nat); assumption. Qed.
+
+ Lemma to_of_nat_option_map x : option_map to_nat (of_nat x) = if Nat.leb x 255 then Some x else None.
+ Proof. do 256 try destruct x as [|x]; reflexivity. Qed.
+
+ Lemma to_nat_bounded x : to_nat x <= 255.
+ Proof.
+ generalize (to_of_nat_option_map (to_nat x)).
+ rewrite of_to_nat; cbn [option_map].
+ destruct (Nat.leb (to_nat x) 255) eqn:H; [ | congruence ].
+ rewrite (PeanoNat.Nat.leb_le (to_nat x) 255) in H.
+ intro; assumption.
+ Qed.
+
+ Lemma of_nat_None_iff x : of_nat x = None <-> 255 < x.
+ Proof.
+ generalize (to_of_nat_option_map x).
+ destruct (of_nat x), (Nat.leb x 255) eqn:H; cbn [option_map]; try congruence.
+ { rewrite PeanoNat.Nat.leb_le in H; split; [ congruence | ].
+ rewrite PeanoNat.Nat.lt_nge; intro H'; exfalso; apply H'; assumption. }
+ { rewrite PeanoNat.Nat.leb_nle in H; split; [ | reflexivity ].
+ rewrite PeanoNat.Nat.lt_nge; intro; assumption. }
+ Qed.
+End nat.
+
+Section N.
+ Local Open Scope N_scope.
+
+ Definition to_N (x : byte) : N
+ := match x with
+ | x00 => 0
+ | x01 => 1
+ | x02 => 2
+ | x03 => 3
+ | x04 => 4
+ | x05 => 5
+ | x06 => 6
+ | x07 => 7
+ | x08 => 8
+ | x09 => 9
+ | x0a => 10
+ | x0b => 11
+ | x0c => 12
+ | x0d => 13
+ | x0e => 14
+ | x0f => 15
+ | x10 => 16
+ | x11 => 17
+ | x12 => 18
+ | x13 => 19
+ | x14 => 20
+ | x15 => 21
+ | x16 => 22
+ | x17 => 23
+ | x18 => 24
+ | x19 => 25
+ | x1a => 26
+ | x1b => 27
+ | x1c => 28
+ | x1d => 29
+ | x1e => 30
+ | x1f => 31
+ | x20 => 32
+ | x21 => 33
+ | x22 => 34
+ | x23 => 35
+ | x24 => 36
+ | x25 => 37
+ | x26 => 38
+ | x27 => 39
+ | x28 => 40
+ | x29 => 41
+ | x2a => 42
+ | x2b => 43
+ | x2c => 44
+ | x2d => 45
+ | x2e => 46
+ | x2f => 47
+ | x30 => 48
+ | x31 => 49
+ | x32 => 50
+ | x33 => 51
+ | x34 => 52
+ | x35 => 53
+ | x36 => 54
+ | x37 => 55
+ | x38 => 56
+ | x39 => 57
+ | x3a => 58
+ | x3b => 59
+ | x3c => 60
+ | x3d => 61
+ | x3e => 62
+ | x3f => 63
+ | x40 => 64
+ | x41 => 65
+ | x42 => 66
+ | x43 => 67
+ | x44 => 68
+ | x45 => 69
+ | x46 => 70
+ | x47 => 71
+ | x48 => 72
+ | x49 => 73
+ | x4a => 74
+ | x4b => 75
+ | x4c => 76
+ | x4d => 77
+ | x4e => 78
+ | x4f => 79
+ | x50 => 80
+ | x51 => 81
+ | x52 => 82
+ | x53 => 83
+ | x54 => 84
+ | x55 => 85
+ | x56 => 86
+ | x57 => 87
+ | x58 => 88
+ | x59 => 89
+ | x5a => 90
+ | x5b => 91
+ | x5c => 92
+ | x5d => 93
+ | x5e => 94
+ | x5f => 95
+ | x60 => 96
+ | x61 => 97
+ | x62 => 98
+ | x63 => 99
+ | x64 => 100
+ | x65 => 101
+ | x66 => 102
+ | x67 => 103
+ | x68 => 104
+ | x69 => 105
+ | x6a => 106
+ | x6b => 107
+ | x6c => 108
+ | x6d => 109
+ | x6e => 110
+ | x6f => 111
+ | x70 => 112
+ | x71 => 113
+ | x72 => 114
+ | x73 => 115
+ | x74 => 116
+ | x75 => 117
+ | x76 => 118
+ | x77 => 119
+ | x78 => 120
+ | x79 => 121
+ | x7a => 122
+ | x7b => 123
+ | x7c => 124
+ | x7d => 125
+ | x7e => 126
+ | x7f => 127
+ | x80 => 128
+ | x81 => 129
+ | x82 => 130
+ | x83 => 131
+ | x84 => 132
+ | x85 => 133
+ | x86 => 134
+ | x87 => 135
+ | x88 => 136
+ | x89 => 137
+ | x8a => 138
+ | x8b => 139
+ | x8c => 140
+ | x8d => 141
+ | x8e => 142
+ | x8f => 143
+ | x90 => 144
+ | x91 => 145
+ | x92 => 146
+ | x93 => 147
+ | x94 => 148
+ | x95 => 149
+ | x96 => 150
+ | x97 => 151
+ | x98 => 152
+ | x99 => 153
+ | x9a => 154
+ | x9b => 155
+ | x9c => 156
+ | x9d => 157
+ | x9e => 158
+ | x9f => 159
+ | xa0 => 160
+ | xa1 => 161
+ | xa2 => 162
+ | xa3 => 163
+ | xa4 => 164
+ | xa5 => 165
+ | xa6 => 166
+ | xa7 => 167
+ | xa8 => 168
+ | xa9 => 169
+ | xaa => 170
+ | xab => 171
+ | xac => 172
+ | xad => 173
+ | xae => 174
+ | xaf => 175
+ | xb0 => 176
+ | xb1 => 177
+ | xb2 => 178
+ | xb3 => 179
+ | xb4 => 180
+ | xb5 => 181
+ | xb6 => 182
+ | xb7 => 183
+ | xb8 => 184
+ | xb9 => 185
+ | xba => 186
+ | xbb => 187
+ | xbc => 188
+ | xbd => 189
+ | xbe => 190
+ | xbf => 191
+ | xc0 => 192
+ | xc1 => 193
+ | xc2 => 194
+ | xc3 => 195
+ | xc4 => 196
+ | xc5 => 197
+ | xc6 => 198
+ | xc7 => 199
+ | xc8 => 200
+ | xc9 => 201
+ | xca => 202
+ | xcb => 203
+ | xcc => 204
+ | xcd => 205
+ | xce => 206
+ | xcf => 207
+ | xd0 => 208
+ | xd1 => 209
+ | xd2 => 210
+ | xd3 => 211
+ | xd4 => 212
+ | xd5 => 213
+ | xd6 => 214
+ | xd7 => 215
+ | xd8 => 216
+ | xd9 => 217
+ | xda => 218
+ | xdb => 219
+ | xdc => 220
+ | xdd => 221
+ | xde => 222
+ | xdf => 223
+ | xe0 => 224
+ | xe1 => 225
+ | xe2 => 226
+ | xe3 => 227
+ | xe4 => 228
+ | xe5 => 229
+ | xe6 => 230
+ | xe7 => 231
+ | xe8 => 232
+ | xe9 => 233
+ | xea => 234
+ | xeb => 235
+ | xec => 236
+ | xed => 237
+ | xee => 238
+ | xef => 239
+ | xf0 => 240
+ | xf1 => 241
+ | xf2 => 242
+ | xf3 => 243
+ | xf4 => 244
+ | xf5 => 245
+ | xf6 => 246
+ | xf7 => 247
+ | xf8 => 248
+ | xf9 => 249
+ | xfa => 250
+ | xfb => 251
+ | xfc => 252
+ | xfd => 253
+ | xfe => 254
+ | xff => 255
+ end.
+
+ Definition of_N (x : N) : option byte
+ := match x with
+ | 0 => Some x00
+ | 1 => Some x01
+ | 2 => Some x02
+ | 3 => Some x03
+ | 4 => Some x04
+ | 5 => Some x05
+ | 6 => Some x06
+ | 7 => Some x07
+ | 8 => Some x08
+ | 9 => Some x09
+ | 10 => Some x0a
+ | 11 => Some x0b
+ | 12 => Some x0c
+ | 13 => Some x0d
+ | 14 => Some x0e
+ | 15 => Some x0f
+ | 16 => Some x10
+ | 17 => Some x11
+ | 18 => Some x12
+ | 19 => Some x13
+ | 20 => Some x14
+ | 21 => Some x15
+ | 22 => Some x16
+ | 23 => Some x17
+ | 24 => Some x18
+ | 25 => Some x19
+ | 26 => Some x1a
+ | 27 => Some x1b
+ | 28 => Some x1c
+ | 29 => Some x1d
+ | 30 => Some x1e
+ | 31 => Some x1f
+ | 32 => Some x20
+ | 33 => Some x21
+ | 34 => Some x22
+ | 35 => Some x23
+ | 36 => Some x24
+ | 37 => Some x25
+ | 38 => Some x26
+ | 39 => Some x27
+ | 40 => Some x28
+ | 41 => Some x29
+ | 42 => Some x2a
+ | 43 => Some x2b
+ | 44 => Some x2c
+ | 45 => Some x2d
+ | 46 => Some x2e
+ | 47 => Some x2f
+ | 48 => Some x30
+ | 49 => Some x31
+ | 50 => Some x32
+ | 51 => Some x33
+ | 52 => Some x34
+ | 53 => Some x35
+ | 54 => Some x36
+ | 55 => Some x37
+ | 56 => Some x38
+ | 57 => Some x39
+ | 58 => Some x3a
+ | 59 => Some x3b
+ | 60 => Some x3c
+ | 61 => Some x3d
+ | 62 => Some x3e
+ | 63 => Some x3f
+ | 64 => Some x40
+ | 65 => Some x41
+ | 66 => Some x42
+ | 67 => Some x43
+ | 68 => Some x44
+ | 69 => Some x45
+ | 70 => Some x46
+ | 71 => Some x47
+ | 72 => Some x48
+ | 73 => Some x49
+ | 74 => Some x4a
+ | 75 => Some x4b
+ | 76 => Some x4c
+ | 77 => Some x4d
+ | 78 => Some x4e
+ | 79 => Some x4f
+ | 80 => Some x50
+ | 81 => Some x51
+ | 82 => Some x52
+ | 83 => Some x53
+ | 84 => Some x54
+ | 85 => Some x55
+ | 86 => Some x56
+ | 87 => Some x57
+ | 88 => Some x58
+ | 89 => Some x59
+ | 90 => Some x5a
+ | 91 => Some x5b
+ | 92 => Some x5c
+ | 93 => Some x5d
+ | 94 => Some x5e
+ | 95 => Some x5f
+ | 96 => Some x60
+ | 97 => Some x61
+ | 98 => Some x62
+ | 99 => Some x63
+ | 100 => Some x64
+ | 101 => Some x65
+ | 102 => Some x66
+ | 103 => Some x67
+ | 104 => Some x68
+ | 105 => Some x69
+ | 106 => Some x6a
+ | 107 => Some x6b
+ | 108 => Some x6c
+ | 109 => Some x6d
+ | 110 => Some x6e
+ | 111 => Some x6f
+ | 112 => Some x70
+ | 113 => Some x71
+ | 114 => Some x72
+ | 115 => Some x73
+ | 116 => Some x74
+ | 117 => Some x75
+ | 118 => Some x76
+ | 119 => Some x77
+ | 120 => Some x78
+ | 121 => Some x79
+ | 122 => Some x7a
+ | 123 => Some x7b
+ | 124 => Some x7c
+ | 125 => Some x7d
+ | 126 => Some x7e
+ | 127 => Some x7f
+ | 128 => Some x80
+ | 129 => Some x81
+ | 130 => Some x82
+ | 131 => Some x83
+ | 132 => Some x84
+ | 133 => Some x85
+ | 134 => Some x86
+ | 135 => Some x87
+ | 136 => Some x88
+ | 137 => Some x89
+ | 138 => Some x8a
+ | 139 => Some x8b
+ | 140 => Some x8c
+ | 141 => Some x8d
+ | 142 => Some x8e
+ | 143 => Some x8f
+ | 144 => Some x90
+ | 145 => Some x91
+ | 146 => Some x92
+ | 147 => Some x93
+ | 148 => Some x94
+ | 149 => Some x95
+ | 150 => Some x96
+ | 151 => Some x97
+ | 152 => Some x98
+ | 153 => Some x99
+ | 154 => Some x9a
+ | 155 => Some x9b
+ | 156 => Some x9c
+ | 157 => Some x9d
+ | 158 => Some x9e
+ | 159 => Some x9f
+ | 160 => Some xa0
+ | 161 => Some xa1
+ | 162 => Some xa2
+ | 163 => Some xa3
+ | 164 => Some xa4
+ | 165 => Some xa5
+ | 166 => Some xa6
+ | 167 => Some xa7
+ | 168 => Some xa8
+ | 169 => Some xa9
+ | 170 => Some xaa
+ | 171 => Some xab
+ | 172 => Some xac
+ | 173 => Some xad
+ | 174 => Some xae
+ | 175 => Some xaf
+ | 176 => Some xb0
+ | 177 => Some xb1
+ | 178 => Some xb2
+ | 179 => Some xb3
+ | 180 => Some xb4
+ | 181 => Some xb5
+ | 182 => Some xb6
+ | 183 => Some xb7
+ | 184 => Some xb8
+ | 185 => Some xb9
+ | 186 => Some xba
+ | 187 => Some xbb
+ | 188 => Some xbc
+ | 189 => Some xbd
+ | 190 => Some xbe
+ | 191 => Some xbf
+ | 192 => Some xc0
+ | 193 => Some xc1
+ | 194 => Some xc2
+ | 195 => Some xc3
+ | 196 => Some xc4
+ | 197 => Some xc5
+ | 198 => Some xc6
+ | 199 => Some xc7
+ | 200 => Some xc8
+ | 201 => Some xc9
+ | 202 => Some xca
+ | 203 => Some xcb
+ | 204 => Some xcc
+ | 205 => Some xcd
+ | 206 => Some xce
+ | 207 => Some xcf
+ | 208 => Some xd0
+ | 209 => Some xd1
+ | 210 => Some xd2
+ | 211 => Some xd3
+ | 212 => Some xd4
+ | 213 => Some xd5
+ | 214 => Some xd6
+ | 215 => Some xd7
+ | 216 => Some xd8
+ | 217 => Some xd9
+ | 218 => Some xda
+ | 219 => Some xdb
+ | 220 => Some xdc
+ | 221 => Some xdd
+ | 222 => Some xde
+ | 223 => Some xdf
+ | 224 => Some xe0
+ | 225 => Some xe1
+ | 226 => Some xe2
+ | 227 => Some xe3
+ | 228 => Some xe4
+ | 229 => Some xe5
+ | 230 => Some xe6
+ | 231 => Some xe7
+ | 232 => Some xe8
+ | 233 => Some xe9
+ | 234 => Some xea
+ | 235 => Some xeb
+ | 236 => Some xec
+ | 237 => Some xed
+ | 238 => Some xee
+ | 239 => Some xef
+ | 240 => Some xf0
+ | 241 => Some xf1
+ | 242 => Some xf2
+ | 243 => Some xf3
+ | 244 => Some xf4
+ | 245 => Some xf5
+ | 246 => Some xf6
+ | 247 => Some xf7
+ | 248 => Some xf8
+ | 249 => Some xf9
+ | 250 => Some xfa
+ | 251 => Some xfb
+ | 252 => Some xfc
+ | 253 => Some xfd
+ | 254 => Some xfe
+ | 255 => Some xff
+ | _ => None
+ end.
+
+ Lemma of_to_N x : of_N (to_N x) = Some x.
+ Proof. destruct x; reflexivity. Qed.
+
+ Lemma to_of_N x y : of_N x = Some y -> to_N y = x.
+ Proof.
+ cbv [of_N];
+ repeat match goal with
+ | [ |- context[match ?x with _ => _ end] ] => is_var x; destruct x
+ | _ => intro
+ | _ => reflexivity
+ | _ => progress subst
+ | [ H : Some ?a = Some ?b |- _ ] => assert (a = b) by refine match H with eq_refl => eq_refl end; clear H
+ | [ H : None = Some _ |- _ ] => solve [ inversion H ]
+ end.
+ Qed.
+
+ Lemma to_of_N_iff x y : of_N x = Some y <-> to_N y = x.
+ Proof. split; intro; subst; (apply of_to_N || apply to_of_N); assumption. Qed.
+
+ Lemma to_of_N_option_map x : option_map to_N (of_N x) = if N.leb x 255 then Some x else None.
+ Proof.
+ cbv [of_N];
+ repeat match goal with
+ | [ |- context[match ?x with _ => _ end] ] => is_var x; destruct x
+ end;
+ reflexivity.
+ Qed.
+
+ Lemma to_N_bounded x : to_N x <= 255.
+ Proof.
+ generalize (to_of_N_option_map (to_N x)).
+ rewrite of_to_N; cbn [option_map].
+ destruct (N.leb (to_N x) 255) eqn:H; [ | congruence ].
+ rewrite (N.leb_le (to_N x) 255) in H.
+ intro; assumption.
+ Qed.
+
+ Lemma of_N_None_iff x : of_N x = None <-> 255 < x.
+ Proof.
+ generalize (to_of_N_option_map x).
+ destruct (of_N x), (N.leb x 255) eqn:H; cbn [option_map]; try congruence.
+ { rewrite N.leb_le in H; split; [ congruence | ].
+ rewrite N.lt_nge; intro H'; exfalso; apply H'; assumption. }
+ { rewrite N.leb_nle in H; split; [ | reflexivity ].
+ rewrite N.lt_nge; intro; assumption. }
+ Qed.
+
+ Lemma to_N_via_nat x : to_N x = N.of_nat (to_nat x).
+ Proof. destruct x; reflexivity. Qed.
+
+ Lemma to_nat_via_N x : to_nat x = N.to_nat (to_N x).
+ Proof. destruct x; reflexivity. Qed.
+
+ Lemma of_N_via_nat x : of_N x = of_nat (N.to_nat x).
+ Proof.
+ destruct (of_N x) as [b|] eqn:H1.
+ { rewrite to_of_N_iff in H1; subst.
+ destruct b; reflexivity. }
+ { rewrite of_N_None_iff, <- N.compare_lt_iff in H1.
+ symmetry; rewrite of_nat_None_iff, <- PeanoNat.Nat.compare_lt_iff.
+ rewrite Nat2N.inj_compare, N2Nat.id; assumption. }
+ Qed.
+
+ Lemma of_nat_via_N x : of_nat x = of_N (N.of_nat x).
+ Proof.
+ destruct (of_nat x) as [b|] eqn:H1.
+ { rewrite to_of_nat_iff in H1; subst.
+ destruct b; reflexivity. }
+ { rewrite of_nat_None_iff, <- PeanoNat.Nat.compare_lt_iff in H1.
+ symmetry; rewrite of_N_None_iff, <- N.compare_lt_iff.
+ rewrite N2Nat.inj_compare, Nat2N.id; assumption. }
+ Qed.
+End N.
diff --git a/theories/Strings/String.v b/theories/Strings/String.v
index a09d518892..08ccfac877 100644
--- a/theories/Strings/String.v
+++ b/theories/Strings/String.v
@@ -15,6 +15,7 @@
Require Import Arith.
Require Import Ascii.
Require Import Bool.
+Require Import Coq.Strings.Byte.
(** *** Definition of strings *)
@@ -25,7 +26,6 @@ Inductive string : Set :=
| String : ascii -> string -> string.
Declare Scope string_scope.
-Module Export StringSyntax. Declare ML Module "string_syntax_plugin". End StringSyntax.
Delimit Scope string_scope with string.
Bind Scope string_scope with string.
Local Open Scope string_scope.
@@ -114,12 +114,12 @@ Theorem get_correct :
Proof.
intros s1; elim s1; simpl.
intros s2; case s2; simpl; split; auto.
-intros H; generalize (H 0); intros H1; inversion H1.
+intros H; generalize (H O); intros H1; inversion H1.
intros; discriminate.
intros a s1' Rec s2; case s2; simpl; split; auto.
-intros H; generalize (H 0); intros H1; inversion H1.
+intros H; generalize (H O); intros H1; inversion H1.
intros; discriminate.
-intros H; generalize (H 0); simpl; intros H1; inversion H1.
+intros H; generalize (H O); simpl; intros H1; inversion H1.
case (Rec s).
intros H0; rewrite H0; auto.
intros n; exact (H (S n)).
@@ -150,7 +150,7 @@ Proof.
intros s1; elim s1; simpl; auto.
intros s2 n; rewrite plus_comm; simpl; auto.
intros a s1' Rec s2 n; case n; simpl; auto.
-generalize (Rec s2 0); simpl; auto. intros.
+generalize (Rec s2 O); simpl; auto. intros.
rewrite <- Plus.plus_Snm_nSm; auto.
Qed.
@@ -162,9 +162,9 @@ Qed.
Fixpoint substring (n m : nat) (s : string) : string :=
match n, m, s with
- | 0, 0, _ => EmptyString
- | 0, S m', EmptyString => s
- | 0, S m', String c s' => String c (substring 0 m' s')
+ | O, O, _ => EmptyString
+ | O, S m', EmptyString => s
+ | O, S m', String c s' => String c (substring 0 m' s')
| S n', _, EmptyString => s
| S n', _, String c s' => substring n' m s'
end.
@@ -257,16 +257,16 @@ Qed.
Fixpoint index (n : nat) (s1 s2 : string) : option nat :=
match s2, n with
- | EmptyString, 0 =>
+ | EmptyString, O =>
match s1 with
- | EmptyString => Some 0
+ | EmptyString => Some O
| String a s1' => None
end
| EmptyString, S n' => None
- | String b s2', 0 =>
- if prefix s1 s2 then Some 0
+ | String b s2', O =>
+ if prefix s1 s2 then Some O
else
- match index 0 s1 s2' with
+ match index O s1 s2' with
| Some n => Some (S n)
| None => None
end
@@ -300,8 +300,8 @@ generalize (prefix_correct s1 (String b s2'));
intros H0 H; injection H as <-; auto.
case H0; simpl; auto.
case m; simpl; auto.
-case (index 0 s1 s2'); intros; discriminate.
-intros m'; generalize (Rec 0 m' s1); case (index 0 s1 s2'); auto.
+case (index O s1 s2'); intros; discriminate.
+intros m'; generalize (Rec O m' s1); case (index O s1 s2'); auto.
intros x H H0 H1; apply H; injection H1; auto.
intros; discriminate.
intros n'; case m; simpl; auto.
@@ -335,7 +335,7 @@ intros H0 H; injection H as <-; auto.
intros p H2 H3; inversion H3.
case m; simpl; auto.
case (index 0 s1 s2'); intros; discriminate.
-intros m'; generalize (Rec 0 m' s1); case (index 0 s1 s2'); auto.
+intros m'; generalize (Rec O m' s1); case (index 0 s1 s2'); auto.
intros x H H0 H1 p; try case p; simpl; auto.
intros H2 H3; red; intros H4; case H0.
intros H5 H6; absurd (false = true); auto with bool.
@@ -383,7 +383,7 @@ intros H4 H5; absurd (false = true); auto with bool.
case s1; simpl; auto.
intros a s n0 H H0 H1 H2;
change (substring n0 (length (String a s)) s2' <> String a s);
- apply (Rec 0); auto.
+ apply (Rec O); auto.
generalize H0; case (index 0 (String a s) s2'); simpl; auto; intros;
discriminate.
apply Le.le_O_n.
@@ -423,9 +423,53 @@ Qed.
Definition findex n s1 s2 :=
match index n s1 s2 with
| Some n => n
- | None => 0
+ | None => O
end.
+(** *** Conversion to/from [list ascii] and [list byte] *)
+
+Fixpoint string_of_list_ascii (s : list ascii) : string
+ := match s with
+ | nil => EmptyString
+ | cons ch s => String ch (string_of_list_ascii s)
+ end.
+
+Fixpoint list_ascii_of_string (s : string) : list ascii
+ := match s with
+ | EmptyString => nil
+ | String ch s => cons ch (list_ascii_of_string s)
+ end.
+
+Lemma string_of_list_ascii_of_string s : string_of_list_ascii (list_ascii_of_string s) = s.
+Proof.
+ induction s as [|? ? IHs]; [ reflexivity | cbn; apply f_equal, IHs ].
+Defined.
+
+Lemma list_ascii_of_string_of_list_ascii s : list_ascii_of_string (string_of_list_ascii s) = s.
+Proof.
+ induction s as [|? ? IHs]; [ reflexivity | cbn; apply f_equal, IHs ].
+Defined.
+
+Definition string_of_list_byte (s : list byte) : string
+ := string_of_list_ascii (List.map ascii_of_byte s).
+
+Definition list_byte_of_string (s : string) : list byte
+ := List.map byte_of_ascii (list_ascii_of_string s).
+
+Lemma string_of_list_byte_of_string s : string_of_list_byte (list_byte_of_string s) = s.
+Proof.
+ cbv [string_of_list_byte list_byte_of_string].
+ erewrite List.map_map, List.map_ext, List.map_id, string_of_list_ascii_of_string; [ reflexivity | intro ].
+ apply ascii_of_byte_of_ascii.
+Qed.
+
+Lemma list_byte_of_string_of_list_byte s : list_byte_of_string (string_of_list_byte s) = s.
+Proof.
+ cbv [string_of_list_byte list_byte_of_string].
+ erewrite list_ascii_of_string_of_list_ascii, List.map_map, List.map_ext, List.map_id; [ reflexivity | intro ].
+ apply byte_of_ascii_of_byte.
+Qed.
+
(** *** Concrete syntax *)
(**
@@ -438,7 +482,11 @@ Definition findex n s1 s2 :=
part of a valid utf8 sequence of characters are not representable
using the Coq string notation (use explicitly the String constructor
with the ascii codes of the characters).
-*)
+ *)
+
+Module Export StringSyntax.
+ String Notation string string_of_list_byte list_byte_of_string : string_scope.
+End StringSyntax.
Example HelloWorld := " ""Hello world!""
".
diff --git a/vernac/explainErr.ml b/vernac/explainErr.ml
index befb4d7ccf..e1496e58d7 100644
--- a/vernac/explainErr.ml
+++ b/vernac/explainErr.ml
@@ -64,8 +64,8 @@ let process_vernac_interp_error exn = match fst exn with
wrap_vernac_error exn (Himsg.explain_type_error ctx Evd.empty te)
| PretypeError(ctx,sigma,te) ->
wrap_vernac_error exn (Himsg.explain_pretype_error ctx sigma te)
- | Notation.NumeralNotationError(ctx,sigma,te) ->
- wrap_vernac_error exn (Himsg.explain_numeral_notation_error ctx sigma te)
+ | Notation.PrimTokenNotationError(kind,ctx,sigma,te) ->
+ wrap_vernac_error exn (Himsg.explain_prim_token_notation_error kind ctx sigma te)
| Typeclasses_errors.TypeClassError(env, te) ->
wrap_vernac_error exn (Himsg.explain_typeclass_error env te)
| Implicit_quantifiers.MismatchedContextInstance(e,c,l,x) ->
diff --git a/vernac/himsg.ml b/vernac/himsg.ml
index 68d6db75ee..a2b5c8d70a 100644
--- a/vernac/himsg.ml
+++ b/vernac/himsg.ml
@@ -1326,12 +1326,12 @@ let explain_reduction_tactic_error = function
spc () ++ str "is not well typed." ++ fnl () ++
explain_type_error env' (Evd.from_env env') e
-let explain_numeral_notation_error env sigma = function
+let explain_prim_token_notation_error kind env sigma = function
| Notation.UnexpectedTerm c ->
(strbrk "Unexpected term " ++
pr_constr_env env sigma c ++
- strbrk " while parsing a numeral notation.")
+ strbrk (" while parsing a "^kind^" notation."))
| Notation.UnexpectedNonOptionTerm c ->
(strbrk "Unexpected non-option term " ++
pr_constr_env env sigma c ++
- strbrk " while parsing a numeral notation.")
+ strbrk (" while parsing a "^kind^" notation."))
diff --git a/vernac/himsg.mli b/vernac/himsg.mli
index db05aaa125..bab66b2af4 100644
--- a/vernac/himsg.mli
+++ b/vernac/himsg.mli
@@ -47,4 +47,4 @@ val explain_module_internalization_error :
val map_pguard_error : ('c -> 'd) -> 'c pguard_error -> 'd pguard_error
val map_ptype_error : ('c -> 'd) -> ('c, 'c) ptype_error -> ('d, 'd) ptype_error
-val explain_numeral_notation_error : env -> Evd.evar_map -> Notation.numeral_notation_error -> Pp.t
+val explain_prim_token_notation_error : string -> env -> Evd.evar_map -> Notation.prim_token_notation_error -> Pp.t