diff --git a/.gitignore b/.gitignore index fdfa49c5f..370616ad7 100644 --- a/.gitignore +++ b/.gitignore @@ -475,3 +475,5 @@ erasure/src/envMap.ml erasure/src/envMap.mli erasure/src/eEtaExpanded.ml erasure/src/eEtaExpanded.mli +erasure/src/pCUICEqualityDec.ml +erasure/src/pCUICEqualityDec.mli diff --git a/TODO.md b/TODO.md index 11fc50edc..2f55b18ff 100644 --- a/TODO.md +++ b/TODO.md @@ -1,9 +1,12 @@ +# Performance + +- change eq_kername to use comparison rather than eq_dec, check we're not using + eq_dec unnecessarily so that erasure/typechecking runs faster inside Coq + (and we avoid useless sumbool to bool conversions in ML/extracted code) # Small stuffs - `assumption_context` should be a boolean function. -- remove duplication of eq_context / eq_context_upto and eq_decl / eq_decl_upto - - Make `wf Σ` `wf_ext Σ` some typeclasses (as at the begining of PCUICCanonicity) et changer les : wf Σ -> en {wfΣ : wf Σ} partout, ce qui éviterait bien des conditions de bord triviales diff --git a/erasure/_PluginProject.in b/erasure/_PluginProject.in index 088e36f3c..4372ee6cc 100644 --- a/erasure/_PluginProject.in +++ b/erasure/_PluginProject.in @@ -5,22 +5,6 @@ src/META.coq-metacoq-erasure # src/classes0.mli # src/classes0.ml -# From Coq's stdlib -src/orderedTypeAlt.mli -src/orderedTypeAlt.ml -src/orderedType0.mli -src/orderedType0.ml -src/fMapList.mli -src/fMapList.ml -src/decidableType.mli -src/decidableType.ml -src/fMapInterface.mli -src/fMapInterface.ml -src/fMapAVL.mli -src/fMapAVL.ml -src/fMapFacts.mli -src/fMapFacts.ml - # src/string2pos.mli # src/string2pos.ml # src/canonicalTries.mli @@ -31,8 +15,6 @@ src/uGraph0.ml src/uGraph0.mli src/wGraph.ml src/wGraph.mli -src/kernames.mli -src/kernames.ml src/envMap.mli src/envMap.ml diff --git a/erasure/src/g_metacoq_erasure.mlg b/erasure/src/g_metacoq_erasure.mlg index ffbd66ce7..3e0795120 100644 --- a/erasure/src/g_metacoq_erasure.mlg +++ b/erasure/src/g_metacoq_erasure.mlg @@ -23,11 +23,11 @@ let bytes_of_list l = let pr_char_list l = (* We allow utf8 encoding *) - str (Bytes.to_string (bytes_of_list l)) + str (Caml_bytestring.caml_string_of_bytestring l) let check ~bypass ~fast env evm c = debug (fun () -> str"Quoting"); - let term = time (str"Quoting") (Ast_quoter.quote_term_rec bypass env) (EConstr.to_constr evm c) in + let term = time (str"Quoting") (Ast_quoter.quote_term_rec ~bypass env) (EConstr.to_constr evm c) in let checker_flags = Config0.extraction_checker_flags in let erase = time (str"Erasure") (if fast then Erasure.erase_fast_and_print_template_program checker_flags diff --git a/erasure/src/metacoq_erasure_plugin.mlpack b/erasure/src/metacoq_erasure_plugin.mlpack index 6f43af023..e99df7065 100644 --- a/erasure/src/metacoq_erasure_plugin.mlpack +++ b/erasure/src/metacoq_erasure_plugin.mlpack @@ -5,14 +5,6 @@ Utils WGraph UGraph0 -OrderedTypeAlt -OrderedType0 -Kernames -FMapList -DecidableType -FMapInterface -FMapAVL -FMapFacts EnvMap WcbvEval diff --git a/erasure/src/pCUICEqualityDec.ml b/erasure/src/pCUICEqualityDec.ml deleted file mode 100644 index ca3e06a15..000000000 --- a/erasure/src/pCUICEqualityDec.ml +++ /dev/null @@ -1,52 +0,0 @@ -open All_Forall -open BasicAst -open Datatypes -open List0 -open PCUICAst -open PCUICEquality -open Universes0 - -(** val compare_universe_variance : - (Universe.t -> Universe.t -> bool) -> (Universe.t -> Universe.t -> bool) - -> Variance.t -> Level.t -> Level.t -> bool **) - -let compare_universe_variance equ lequ v u u' = - match v with - | Variance.Irrelevant -> true - | Variance.Covariant -> lequ (Universe.make u) (Universe.make u') - | Variance.Invariant -> equ (Universe.make u) (Universe.make u') - -(** val compare_universe_instance : - (Universe.t -> Universe.t -> bool) -> Level.t list -> Level.t list -> bool **) - -let compare_universe_instance equ u u' = - forallb2 equ (map Universe.make u) (map Universe.make u') - -(** val compare_universe_instance_variance : - (Universe.t -> Universe.t -> bool) -> (Universe.t -> Universe.t -> bool) - -> Variance.t list -> Level.t list -> Level.t list -> bool **) - -let rec compare_universe_instance_variance equ lequ v u u' = - match u with - | [] -> (match u' with - | [] -> true - | _ :: _ -> false) - | u0 :: us -> - (match u' with - | [] -> false - | u'0 :: us' -> - (match v with - | [] -> compare_universe_instance_variance equ lequ v us us' - | v0 :: vs -> - (&&) (compare_universe_variance equ lequ v0 u0 u'0) - (compare_universe_instance_variance equ lequ vs us us'))) - -(** val compare_global_instance : - PCUICEnvironment.global_env -> (Universe.t -> Universe.t -> bool) -> - (Universe.t -> Universe.t -> bool) -> global_reference -> nat -> Level.t - list -> Level.t list -> bool **) - -let compare_global_instance _UU03a3_ equ lequ gr napp = - match global_variance _UU03a3_ gr napp with - | Some v -> compare_universe_instance_variance equ lequ v - | None -> compare_universe_instance equ diff --git a/erasure/src/pCUICEqualityDec.mli b/erasure/src/pCUICEqualityDec.mli deleted file mode 100644 index f83a9a332..000000000 --- a/erasure/src/pCUICEqualityDec.mli +++ /dev/null @@ -1,23 +0,0 @@ -open All_Forall -open BasicAst -open Datatypes -open List0 -open PCUICAst -open PCUICEquality -open Universes0 - -val compare_universe_variance : - (Universe.t -> Universe.t -> bool) -> (Universe.t -> Universe.t -> bool) -> - Variance.t -> Level.t -> Level.t -> bool - -val compare_universe_instance : - (Universe.t -> Universe.t -> bool) -> Level.t list -> Level.t list -> bool - -val compare_universe_instance_variance : - (Universe.t -> Universe.t -> bool) -> (Universe.t -> Universe.t -> bool) -> - Variance.t list -> Level.t list -> Level.t list -> bool - -val compare_global_instance : - PCUICEnvironment.global_env -> (Universe.t -> Universe.t -> bool) -> - (Universe.t -> Universe.t -> bool) -> global_reference -> nat -> Level.t - list -> Level.t list -> bool diff --git a/erasure/theories/EDeps.v b/erasure/theories/EDeps.v index 26d2947ee..c5bf924fe 100644 --- a/erasure/theories/EDeps.v +++ b/erasure/theories/EDeps.v @@ -461,14 +461,12 @@ Proof. apply PCUICWeakeningEnvConv.lookup_env_Some_fresh in H as not_fresh. econstructor. - unfold PCUICAst.declared_constant in *; cbn. - unfold eq_kername. inversion wfΣ; subst. - destruct kername_eq_dec as [<-|]; [congruence|]. + destruct (eqb_spec kn0 kn) as [<-|]; [congruence|]. eassumption. - - unfold ETyping.declared_constant in *. cbn -[Reflect.eqb]. + - unfold ETyping.declared_constant in *. cbn -[ReflectEq.eqb]. inversion wfΣ; subst. - change (eq_kername kn0 kn) with (Reflect.eqb kn0 kn). - destruct (Reflect.eqb_spec kn0 kn); [congruence|]. + destruct (ReflectEq.eqb_spec kn0 kn); [congruence|]. eassumption. - unfold erases_constant_body in *. destruct PCUICAst.PCUICEnvironment.cst_body eqn:body. @@ -476,9 +474,8 @@ Proof. assert (PCUICAst.declared_constant (add_global_decl Σ (kn, decl)) kn0 cb). { unfold PCUICAst.declared_constant. cbn. - unfold eq_kername. inversion wfΣ; subst. - destruct kername_eq_dec as [<-|]; [congruence|]. + destruct (eqb_spec kn0 kn) as [<-|]; [congruence|]. easy. } inversion wfΣ; subst. eapply declared_constant_inv in H4; eauto. @@ -503,21 +500,21 @@ Proof. invs wfΣ. destruct H0. split. 2: eauto. destruct d. split; eauto. - red. cbn. unfold eq_kername. destruct kername_eq_dec; eauto. + red. cbn. cbn in *. + destruct (eqb_spec (inductive_mind ind) kn). cbn in *. subst. - eapply PCUICWeakeningEnvConv.lookup_env_Some_fresh in H5. eauto. eapply H. + eapply PCUICWeakeningEnvConv.lookup_env_Some_fresh in H5. eauto. eapply H. exact H0. - econstructor; eauto. destruct H as [H H']. split; eauto. red in H |- *. inv wfΣ. unfold PCUICEnvironment.lookup_env. - simpl. unfold eq_kername. - destruct (kername_eq_dec (inductive_mind p.1) kn); auto. subst. + simpl. destruct (eqb_spec (inductive_mind p.1) kn); auto. subst. eapply PCUICWeakeningEnvConv.lookup_env_Some_fresh in H; eauto. contradiction. destruct H0 as [H0 H0']. split; eauto. red in H0 |- *. - inv wfΣ. cbn. change (eq_kername (inductive_mind p.1) kn) with (Reflect.eqb (inductive_mind p.1) kn). - destruct (Reflect.eqb_spec (inductive_mind p.1) kn); auto. subst. + inv wfΣ. cbn. change (eq_kername (inductive_mind p.1) kn) with (ReflectEq.eqb (inductive_mind p.1) kn). + destruct (ReflectEq.eqb_spec (inductive_mind p.1) kn); auto. subst. destruct H as [H _]. eapply PCUICWeakeningEnvConv.lookup_env_Some_fresh in H. eauto. contradiction. - econstructor; eauto. @@ -525,15 +522,15 @@ Proof. split; eauto. red in H |- *. inv wfΣ. unfold PCUICEnvironment.lookup_env. simpl. - change (eq_kername (inductive_mind p.1.1) kn) with (Reflect.eqb (inductive_mind p.1.1) kn); auto. - destruct (Reflect.eqb_spec (inductive_mind p.1.1) kn). subst. + change (eq_kername (inductive_mind p.1.1) kn) with (ReflectEq.eqb (inductive_mind p.1.1) kn); auto. + destruct (ReflectEq.eqb_spec (inductive_mind p.1.1) kn). subst. eapply PCUICWeakeningEnvConv.lookup_env_Some_fresh in H; eauto. contradiction. apply H. destruct H0 as [H0 H0']. split; eauto. red in H0 |- *. inv wfΣ. simpl. - change (eq_kername (inductive_mind p.1.1) kn) with (Reflect.eqb (inductive_mind p.1.1) kn); auto. - destruct (Reflect.eqb_spec (inductive_mind p.1.1) kn); auto. subst. + change (eq_kername (inductive_mind p.1.1) kn) with (ReflectEq.eqb (inductive_mind p.1.1) kn); auto. + destruct (ReflectEq.eqb_spec (inductive_mind p.1.1) kn); auto. subst. destruct H as [H _]. eapply PCUICWeakeningEnvConv.lookup_env_Some_fresh in H. eauto. contradiction. Qed. @@ -716,14 +713,13 @@ Proof. apply IH in erg'. 2:{ depelim wf. now depelim o0. } assert (decl_ext: PCUICAst.declared_constant Σu kn' cst'). { unfold PCUICAst.declared_constant in *; cbn in *. - unfold eq_kername in *. - now destruct kername_eq_dec; [|congruence]. } + destruct (eqb_spec kn' kn); [|congruence]. subst. contradiction. } specialize (proj1 erg' kn' cst' decl_ext) as (cst & decl'' & ? & ?). exists cst. split; [|split]. * unfold declared_constant in *; cbn. rewrite decl''. - change (eq_kername kn' kn) with (Reflect.eqb kn' kn). - destruct (Reflect.eqb_spec kn' kn); auto. congruence. + change (eq_kername kn' kn) with (ReflectEq.eqb kn' kn). + destruct (ReflectEq.eqb_spec kn' kn); auto. congruence. * inversion wf; subst. eapply declared_constant_inv in decl_ext; eauto. 2: exact weaken_env_prop_typing. @@ -745,14 +741,13 @@ Proof. destruct decli as [decli ?]. split; auto. red in decli |- *. simpl in decli |- *. unfold PCUICEnvironment.lookup_env in decli |- *. simpl in *. - unfold eq_kername in decli |- *. - destruct kername_eq_dec. subst. discriminate. auto. + destruct (eqb_spec (inductive_mind k) kn). subst. discriminate. auto. destruct IH as [mdecl' [idecl' [decli' er]]]. exists mdecl', idecl'. split; auto. red. destruct decli'; split; auto. red in decli. unfold declared_minductive in *. - simpl. unfold eq_kername. destruct kername_eq_dec; subst; auto. + simpl. destruct (eqb_spec (inductive_mind k) kn); subst; auto. unfold PCUICAst.declared_minductive in decli. unfold PCUICEnvironment.lookup_env in decli. simpl in decli. rewrite eq_kername_refl in decli. intuition discriminate. @@ -762,16 +757,15 @@ Proof. simpl in decli |- *. unfold PCUICAst.declared_minductive, PCUICEnvironment.lookup_env in decli. simpl in decli. - unfold eq_kername in decli |- *. simpl in *. - destruct kername_eq_dec. subst. noconf decli. + destruct (eqb_specT (inductive_mind k) kn). simpl in *. subst. noconf decli. destruct (Forall2_nth_error_left (proj1 H) _ _ H3); eauto. eexists _, _; intuition eauto. split; eauto. red. - simpl. unfold eq_kername. destruct kername_eq_dec; try congruence. + simpl. rewrite eqb_refl. congruence. destruct (proj2 IH _ _ _ (conj decli H3)) as [m' [i' [decli' ei]]]. eexists _, _; intuition eauto. destruct decli'; red; split; eauto. red in d |- *. simpl. - unfold eq_kername; destruct kername_eq_dec; subst; try congruence. + apply neqb in n. destruct eqb; cbn in n; try congruence. Qed. Lemma erases_global_erases_deps Σ Γ t T et Σ' : diff --git a/erasure/theories/EExtends.v b/erasure/theories/EExtends.v index 2bde68073..98ddc3476 100644 --- a/erasure/theories/EExtends.v +++ b/erasure/theories/EExtends.v @@ -18,7 +18,7 @@ Lemma lookup_env_Some_fresh {Σ c decl} : lookup_env Σ c = Some decl -> ~ (fresh_global c Σ). Proof. induction Σ; cbn. 1: congruence. - unfold BasicAst.eq_kername. destruct BasicAst.kername_eq_dec; subst. + destruct (eqb_spec c a.1); subst. - intros [= <-] H2. invs H2. contradiction. - intros H1 H2. apply IHΣ; tas. @@ -37,8 +37,8 @@ Proof. - specialize (IHΣ'' c decl). forward IHΣ''. + now inv wfΣ'. + intros HΣ. specialize (IHΣ'' HΣ). - inv wfΣ'. simpl in *. unfold BasicAst.eq_kername. - destruct BasicAst.kername_eq_dec; subst; auto. + inv wfΣ'. simpl in *. + destruct (eqb_spec c kn); subst; auto. apply lookup_env_Some_fresh in IHΣ''; contradiction. Qed. diff --git a/erasure/theories/EOptimizePropDiscr.v b/erasure/theories/EOptimizePropDiscr.v index 0dc50a694..b9869010b 100644 --- a/erasure/theories/EOptimizePropDiscr.v +++ b/erasure/theories/EOptimizePropDiscr.v @@ -455,7 +455,7 @@ Lemma lookup_env_optimize Σ kn : Proof. unfold optimize_env. induction Σ at 2 4; simpl; auto. - unfold eq_kername; destruct kername_eq_dec => //. + case: eqb_spec => //. Qed. Lemma is_propositional_optimize Σ ind : diff --git a/erasure/theories/EPretty.v b/erasure/theories/EPretty.v index feeebeb13..2daba3972 100644 --- a/erasure/theories/EPretty.v +++ b/erasure/theories/EPretty.v @@ -6,17 +6,9 @@ From MetaCoq.PCUIC Require Import PCUICPrimitive. (** * Pretty printing *) -Section print_term. +Section freshnames. Context (Σ : global_context). - Definition print_def {A : Set} (f : A -> string) (def : def A) := - string_of_name (dname def) ^ " { struct " ^ string_of_nat (rarg def) ^ " }" - ^ " := " ^ nl ^ f (dbody def). - - Definition print_defs (print_term : context -> bool -> bool -> term -> string) Γ (defs : mfixpoint term) := - let ctx' := List.map (fun d => {| decl_name := dname d; decl_body := None |}) defs in - print_list (print_def (print_term (ctx' ++ Γ) true false)) (nl ^ " with ") defs. - Definition lookup_ind_decl ind i := match lookup_env Σ ind with | Some (InductiveDecl {| ind_bodies := l |}) => @@ -42,7 +34,7 @@ Section print_term. List.forallb (fun decl => match decl.(decl_name) with - | nNamed id' => negb (ident_eq id id') + | nNamed id' => negb (eqb id id') | nAnon => true end) Γ. @@ -61,7 +53,7 @@ Section print_term. match i with | 0 => id | S i' => - let id' := id ^ string_of_nat (n - i) in + let id' := id ^ (string_of_nat (n - i)) in if is_fresh Γ id' then id' else aux i' end @@ -75,128 +67,156 @@ Section print_term. in if is_fresh Γ id then nNamed id else nNamed (fresh_id_from Γ 10 id). +End freshnames. + +Module PrintTermTree. + Import Tree. + Infix "^" := append. + + Section print_term. + Context (Σ : global_context). - Fixpoint print_term (Γ : context) (top : bool) (inapp : bool) (t : term) {struct t} := - match t with - | tBox => "∎" - | tRel n => - match nth_error Γ n with - | Some {| decl_name := na |} => - match na with - | nAnon => "Anonymous (" ^ string_of_nat n ^ ")" - | nNamed id => id + Definition print_def {A : Set} (f : A -> t) (def : def A) := + string_of_name (dname def) ^ " { struct " ^ string_of_nat (rarg def) ^ " }" + ^ " := " ^ nl ^ f (dbody def). + + Definition print_defs (print_term : context -> bool -> bool -> term -> t) Γ (defs : mfixpoint term) := + let ctx' := List.map (fun d => {| decl_name := dname d; decl_body := None |}) defs in + print_list (print_def (print_term (ctx' ++ Γ)%list true false)) (nl ^ " with ") defs. + + + Fixpoint print_term (Γ : context) (top : bool) (inapp : bool) (t : term) {struct t} : Tree.t := + match t with + | tBox => "∎" + | tRel n => + match nth_error Γ n with + | Some {| decl_name := na |} => + match na with + | nAnon => "Anonymous (" ^ string_of_nat n ^ ")" + | nNamed id => id + end + | None => "UnboundRel(" ^ string_of_nat n ^ ")" end - | None => "UnboundRel(" ^ string_of_nat n ^ ")" - end - | tVar n => "Var(" ^ n ^ ")" - | tEvar ev args => "Evar(" ^ string_of_nat ev ^ "[]" (* TODO *) ^ ")" - | tLambda na body => - let na' := fresh_name Γ na t in - parens top ("fun " ^ string_of_name na' - ^ " => " ^ print_term (vass na' :: Γ) true false body) - | tLetIn na def body => - let na' := fresh_name Γ na t in - parens top ("let " ^ string_of_name na' ^ - " := " ^ print_term Γ true false def ^ " in " ^ nl ^ - print_term (vdef na' def :: Γ) true false body) - | tApp f l => - parens (top || inapp) (print_term Γ false true f ^ " " ^ print_term Γ false false l) - | tConst c => string_of_kername c - | tConstruct (mkInd i k as ind) l => - match lookup_ind_decl i k with - | Some oib => - match nth_error oib.(ind_ctors) l with - | Some (na, _) => na + | tVar n => "Var(" ^ n ^ ")" + | tEvar ev args => "Evar(" ^ string_of_nat ev ^ "[]" (* TODO *) ^ ")" + | tLambda na body => + let na' := fresh_name Γ na t in + parens top ("fun " ^ string_of_name na' + ^ " => " ^ print_term (vass na' :: Γ) true false body) + | tLetIn na def body => + let na' := fresh_name Γ na t in + parens top ("let " ^ string_of_name na' ^ + " := " ^ print_term Γ true false def ^ " in " ^ nl ^ + print_term (vdef na' def :: Γ) true false body) + | tApp f l => + parens (top || inapp) (print_term Γ false true f ^ " " ^ print_term Γ false false l) + | tConst c => string_of_kername c + | tConstruct (mkInd i k as ind) l => + match lookup_ind_decl Σ i k with + | Some oib => + match nth_error oib.(ind_ctors) l with + | Some (na, _) => na + | None => + "UnboundConstruct(" ^ string_of_inductive ind ^ "," ^ string_of_nat l ^ ")" + end | None => "UnboundConstruct(" ^ string_of_inductive ind ^ "," ^ string_of_nat l ^ ")" end - | None => - "UnboundConstruct(" ^ string_of_inductive ind ^ "," ^ string_of_nat l ^ ")" - end - | tCase (mkInd mind i as ind, pars) t brs => - match lookup_ind_decl mind i with - | Some oib => - let fix print_args Γ nas br {struct nas} := - match nas with - | [] => "=>" ^ " " ^ br Γ - | na :: nas => - string_of_name na ^ " " ^ print_args (vass na :: Γ) nas br + | tCase (mkInd mind i as ind, pars) t brs => + match lookup_ind_decl Σ mind i with + | Some oib => + let fix print_args Γ nas br {struct nas} := + match nas with + | [] => "=>" ^ " " ^ br Γ + | na :: nas => + string_of_name na ^ " " ^ print_args (vass na :: Γ) nas br + end + in + let brs := map (fun '(nas, br) => print_args Γ (List.rev nas) (fun Γ => print_term Γ true false br)) brs in + let brs := combine brs oib.(ind_ctors) in + parens top ("match " ^ print_term Γ true false t ^ + " with " ^ nl ^ + print_list (fun '(b, (na, _)) => (na : String.t) ^ " " ^ b) + (nl ^ " | ") brs ^ nl ^ "end" ^ nl) + | None => + "Case(" ^ string_of_inductive ind ^ "," ^ string_of_nat i ^ "," ^ string_of_term t ^ "," + ^ string_of_list (fun b => string_of_term (snd b)) brs ^ ")" + end + | tProj (mkInd mind i as ind, pars, k) c => + match lookup_ind_decl Σ mind i with + | Some oib => + match nth_error oib.(ind_projs) k with + | Some na => print_term Γ false false c ^ ".(" ^ na ^ ")" + | None => + "UnboundProj(" ^ string_of_inductive ind ^ "," ^ string_of_nat i ^ "," ^ string_of_nat k ^ "," + ^ print_term Γ true false c ^ ")" end - in - let brs := map (fun '(nas, br) => print_args Γ (List.rev nas) (fun Γ => print_term Γ true false br)) brs in - let brs := combine brs oib.(ind_ctors) in - parens top ("match " ^ print_term Γ true false t ^ - " with " ^ nl ^ - print_list (fun '(b, (na, _)) => na ^ " " ^ b) - (nl ^ " | ") brs ^ nl ^ "end" ^ nl) - | None => - "Case(" ^ string_of_inductive ind ^ "," ^ string_of_nat i ^ "," ^ string_of_term t ^ "," - ^ string_of_list (fun b => string_of_term (snd b)) brs ^ ")" - end - | tProj (mkInd mind i as ind, pars, k) c => - match lookup_ind_decl mind i with - | Some oib => - match nth_error oib.(ind_projs) k with - | Some na => print_term Γ false false c ^ ".(" ^ na ^ ")" | None => "UnboundProj(" ^ string_of_inductive ind ^ "," ^ string_of_nat i ^ "," ^ string_of_nat k ^ "," - ^ print_term Γ true false c ^ ")" + ^ print_term Γ true false c ^ ")" end - | None => - "UnboundProj(" ^ string_of_inductive ind ^ "," ^ string_of_nat i ^ "," ^ string_of_nat k ^ "," - ^ print_term Γ true false c ^ ")" + + + | tFix l n => + parens top ("let fix " ^ print_defs print_term Γ l ^ nl ^ + " in " ^ List.nth_default (string_of_nat n) (map (string_of_name ∘ dname) l) n) + | tCoFix l n => + parens top ("let cofix " ^ print_defs print_term Γ l ^ nl ^ + " in " ^ List.nth_default (string_of_nat n) (map (string_of_name ∘ dname) l) n) + (* | tPrim p => *) + (* parens top (string_of_prim (print_term Γ false false) p) *) + end. + End print_term. + + Definition pr Σ (t : term) := print_term Σ [] true false t. + + Definition print_constant_body Σ kn decl := + match decl.(cst_body) with + | Some b => "Definition " ^ string_of_kername kn ^ " := " ^ pr Σ b + | None => "Axiom " ^ string_of_kername kn + end. + + Definition pr_allowed_elim (elims : Universes.allowed_eliminations) := + match elims with + | Universes.IntoSProp => "into sprop" + | Universes.IntoPropSProp => "into prop or sprop" + | Universes.IntoSetPropSProp => "into set, prop or sprop" + | Universes.IntoAny => "into any sort" + end. + + Definition print_one_inductive_body npars body : Tree.t := + let params := string_of_nat npars ^ " parameters" in + let prop := if body.(ind_propositional) then "propositional" else "computational" in + let kelim := pr_allowed_elim body.(ind_kelim) in + let ctors := print_list (fun idn => "| " ^ (idn.1 : ident) ^ " " ^ string_of_nat idn.2 ^ " arguments") nl body.(ind_ctors) in + let projs := + match body.(ind_projs) return Tree.t with + | [] => "" + | _ => nl ^ "projections: " ^ print_list (fun x => x : ident) ", " body.(ind_projs) end + in + "Inductive " ^ body.(ind_name) ^ "(" ^ params ^ "," ^ prop ^ ", elimination " ^ kelim ^ ") := " ^ nl ^ ctors ^ projs. + + Definition print_inductive_body decl := + print_list (print_one_inductive_body decl.(ind_npars)) nl decl.(ind_bodies). + + Definition print_decl Σ '(kn, d) := + match d with + | ConstantDecl body => print_constant_body Σ kn body + | InductiveDecl mind => print_inductive_body mind + end. + + Definition print_global_context (g : global_context) := + print_list (print_decl g) nl (List.rev g). + Notation print_env := print_global_context. + + Definition print_program (p : program) : t := + pr p.1 p.2 ^ nl ^ "in" ^ print_env p.1. + +End PrintTermTree. + +Definition pr Σ := Tree.to_string ∘ PrintTermTree.pr Σ. +Definition print_global_context := Tree.to_string ∘ PrintTermTree.print_global_context. - | tFix l n => - parens top ("let fix " ^ print_defs print_term Γ l ^ nl ^ - " in " ^ List.nth_default (string_of_nat n) (map (string_of_name ∘ dname) l) n) - | tCoFix l n => - parens top ("let cofix " ^ print_defs print_term Γ l ^ nl ^ - " in " ^ List.nth_default (string_of_nat n) (map (string_of_name ∘ dname) l) n) - (* | tPrim p => *) - (* parens top (string_of_prim (print_term Γ false false) p) *) - end. - -End print_term. - -Definition pr Σ (t : term) := print_term Σ [] true false t. - -Definition print_constant_body Σ kn decl := - match decl.(cst_body) with - | Some b => "Definition " ^ string_of_kername kn ^ " := " ^ pr Σ b - | None => "Axiom " ^ string_of_kername kn - end. - -Definition pr_allowed_elim (elims : Universes.allowed_eliminations) := - match elims with - | Universes.IntoSProp => "into sprop" - | Universes.IntoPropSProp => "into prop or sprop" - | Universes.IntoSetPropSProp => "into set, prop or sprop" - | Universes.IntoAny => "into any sort" - end. - -Definition print_one_inductive_body npars body := - let params := string_of_nat npars ^ " parameters" in - let prop := if body.(ind_propositional) then "propositional" else "computational" in - let kelim := pr_allowed_elim body.(ind_kelim) in - let ctors := print_list (fun idn => "| " ^ idn.1 ^ " " ^ string_of_nat idn.2 ^ " arguments") nl body.(ind_ctors) in - let projs := - match body.(ind_projs) with - | [] => "" - | _ => nl ^ "projections: " ^ print_list (fun x => x) ", " body.(ind_projs) - end - in - "Inductive " ^ body.(ind_name) ^ "(" ^ params ^ "," ^ prop ^ ", elimination " ^ kelim ^ ") := " ^ nl ^ ctors ^ projs. - -Definition print_inductive_body decl := - print_list (print_one_inductive_body decl.(ind_npars)) nl decl.(ind_bodies). - -Definition print_decl Σ '(kn, d) := - match d with - | ConstantDecl body => print_constant_body Σ kn body - | InductiveDecl mind => print_inductive_body mind - end. - -Definition print_global_context (g : global_context) := - print_list (print_decl g) nl (List.rev g). +Definition print_program := Tree.to_string ∘ PrintTermTree.print_program. diff --git a/erasure/theories/EReflect.v b/erasure/theories/EReflect.v index 5b8f7a148..265e3590e 100644 --- a/erasure/theories/EReflect.v +++ b/erasure/theories/EReflect.v @@ -19,9 +19,9 @@ Local Ltac term_dec_tac term_dec := repeat match goal with | t : term, u : term |- _ => fcase (term_dec t u) | n : nat, m : nat |- _ => fcase (Nat.eq_dec n m) - | i : ident, i' : ident |- _ => fcase (string_dec i i') - | i : kername, i' : kername |- _ => fcase (kername_eq_dec i i') - | i : string, i' : kername |- _ => fcase (string_dec i i') + | i : ident, i' : ident |- _ => fcase (eq_dec i i') + | i : kername, i' : kername |- _ => fcase (eq_dec i i') + | i : string, i' : kername |- _ => fcase (eq_dec i i') | n : name, n' : name |- _ => fcase (eq_dec n n') | i : prim_val _, i' : prim_val _ |- _ => fcase (eq_dec i i') | i : inductive, i' : inductive |- _ => fcase (eq_dec i i') @@ -105,7 +105,7 @@ Proof. Defined. #[global] -Instance ReflectEq_term : Reflect.ReflectEq _ := +Instance ReflectEq_term : ReflectEq.ReflectEq _ := @EqDec_ReflectEq _ EqDec_term. Definition eqb_constant_body (x y : constant_body) := diff --git a/erasure/theories/ERemoveParams.v b/erasure/theories/ERemoveParams.v index 6c5691d1d..7755c60a4 100644 --- a/erasure/theories/ERemoveParams.v +++ b/erasure/theories/ERemoveParams.v @@ -406,7 +406,7 @@ Proof. unfold strip_env. destruct Σ as [Σ map repr wf]; cbn. induction Σ at 2 4; simpl; auto. - rewrite /eq_kername; destruct kername_eq_dec => //. + case: eqb_spec => //. Qed. Lemma is_propositional_strip (Σ : GlobalContextMap.t) ind : @@ -557,10 +557,11 @@ Proof. - move=> _; rewrite /declared_constant /lookup_env /= //. - move=> /andP[] etaa etaΣ. destruct a as [kn' d']; cbn in *. - rewrite /declared_constant /=; rewrite /eq_kername; destruct kername_eq_dec. - * subst kn'. move=> [=]. intros ->. + rewrite /declared_constant /=. + case:eqb_specT => //. + * intros e; subst kn'. move=> [=]. intros ->. exists Σ. split => //. now exists [(kn, d)]. - * move=> Hl. destruct (IHΣ etaΣ Hl) as [Σ' [ext eta]]. + * intros ne. move=> Hl. destruct (IHΣ etaΣ Hl) as [Σ' [ext eta]]. exists Σ'; split => //. destruct ext as [Σ'' ->]. now exists ((kn', d')::Σ''). diff --git a/erasure/theories/ETyping.v b/erasure/theories/ETyping.v index 8916a98b7..5b260db70 100644 --- a/erasure/theories/ETyping.v +++ b/erasure/theories/ETyping.v @@ -42,7 +42,7 @@ Lemma elookup_env_cons_fresh {kn d Σ kn'} : ETyping.lookup_env ((kn, d) :: Σ) kn' = ETyping.lookup_env Σ kn'. Proof. simpl. change (eq_kername kn' kn) with (eqb kn' kn). - destruct (Reflect.eqb_spec kn' kn). subst => //. auto. + destruct (eqb_spec kn' kn). subst => //. auto. Qed. (** Knowledge of propositionality status of an inductive type and parameters *) @@ -87,7 +87,7 @@ Lemma lookup_env_Some_fresh {Σ c decl} : lookup_env Σ c = Some decl -> ~ (fresh_global c Σ). Proof. induction Σ; cbn. 1: congruence. - unfold eq_kername; destruct kername_eq_dec; subst. + case: eqb_spec; intros e; subst. - intros [= <-] H2. inv H2. contradiction. - intros H1 H2. apply IHΣ; tas. @@ -107,7 +107,7 @@ Proof. + now inv wfΣ'. + intros HΣ. specialize (IHΣ'' HΣ). inv wfΣ'. simpl in *. - unfold eq_kername; destruct kername_eq_dec; subst; auto. + case: eqb_spec; intros e; subst; auto. apply lookup_env_Some_fresh in IHΣ''; contradiction. Qed. @@ -182,7 +182,7 @@ Proof. induction mfix; simpl; auto. Qed. -Definition tDummy := tVar ""%string. +Definition tDummy := tVar "". Definition iota_red npar args (br : list name * term) := substl (List.rev (List.skipn npar args)) br.2. diff --git a/erasure/theories/EWtAst.v b/erasure/theories/EWtAst.v new file mode 100644 index 000000000..43a34124c --- /dev/null +++ b/erasure/theories/EWtAst.v @@ -0,0 +1,440 @@ +(* Distributed under the terms of the MIT license. *) +From Coq Require Import ssreflect ssrbool. +From MetaCoq.Template Require Import utils BasicAst Universes. +From MetaCoq.PCUIC Require Import PCUICPrimitive. +From MetaCoq.Erasure Require Import EAst EAstUtils EInduction ECSubst ELiftSubst ETyping. +(* EEtaExpanded.*) +From Equations Require Import Equations. +Set Equations Transparent. + +Import MCMonadNotation. + +(** * Well-scoped and well-formed extracted terms *) + +Definition isSome {A} (o : option A) : bool := + match o with + | Some _ => true + | None => false + end. + +Section WellScoped. + Context (Σ : global_context). + + Definition lookup_constant kn : option constant_body := + decl <- ETyping.lookup_env Σ kn;; + match decl with + | ConstantDecl decl => Some decl + | InductiveDecl mdecl => None + end. + + Definition lookup_minductive kn : option mutual_inductive_body := + decl <- ETyping.lookup_env Σ kn;; + match decl with + | ConstantDecl _ => None + | InductiveDecl mdecl => ret mdecl + end. + + Definition lookup_inductive kn : option (mutual_inductive_body * one_inductive_body) := + mdecl <- lookup_minductive (inductive_mind kn) ;; + idecl <- nth_error mdecl.(ind_bodies) (inductive_ind kn) ;; + ret (mdecl, idecl). + + Definition lookup_constructor kn c : option (mutual_inductive_body * one_inductive_body * (ident × nat)) := + '(mdecl, idecl) <- lookup_inductive kn ;; + cdecl <- nth_error idecl.(ind_ctors) c ;; + ret (mdecl, idecl, cdecl). + + Definition lookup_projection (proj : projection) := + '(mdecl, idecl) <- lookup_inductive (fst (fst proj)) ;; + pdecl <- List.nth_error idecl.(ind_projs) (snd proj) ;; + ret (mdecl, idecl, pdecl). + + Definition declared_constant id : bool := + isSome (lookup_constant id). + + Definition declared_minductive mind := + isSome (lookup_minductive mind). + + Definition declared_inductive ind := + isSome (lookup_inductive ind). + + Definition declared_constructor kn c := + isSome (lookup_constructor kn c). + + Definition declared_projection kn := + isSome (lookup_projection kn). + + Equations well_scoped (n : nat) (t : EAst.term) : bool := + { | n, tBox => true + | n, tRel i := i closedn n t. + Proof. + revert n t. + apply (well_scoped_elim + (fun n t e => e -> closedn n t) + (fun n l e => e -> forallb (closedn n) l) + (fun n l e => e -> forallb (fun br => closedn (#|br.1| + n) (snd br)) l) + (fun n l e => e -> forallb (closedn n ∘ dbody) l)); cbn => //. + all:intros *; intros; simp well_scoped. + all:rtoProp; intuition eauto. + Qed. + + Definition eterm n := { t : EAst.term | well_scoped n t }. + Definition eterm_term {n} (e : eterm n) := proj1_sig e. + Definition well_scoped_eterm {n} (e : eterm n) : well_scoped n (eterm_term e) := proj2_sig e. +End WellScoped. + +Arguments eterm_term {Σ n}. + +Coercion eterm_term : eterm >-> term. +Coercion well_scoped_eterm : eterm >-> is_true. + +Module Constructors. + Section Constructors. + Context {Σ : global_context} {n : nat}. + + Obligation Tactic := idtac. + Program Definition tBox : eterm Σ n := tBox. + Next Obligation. cbn. exact eq_refl. Defined. + + + Program Definition tRel (i : nat) (Hi : i < n) : eterm Σ n := tRel i. + Next Obligation. cbn; intros. eapply Nat.leb_le, Hi. Defined. + + Lemma andP (a b : bool) : a -> b -> a && b. + Proof. destruct a, b; cbn; intros; try exact eq_refl; try discriminate. Defined. + + Program Definition tEvar k (l : list (eterm Σ n)) : eterm Σ n := tEvar k (map eterm_term l). + Next Obligation. + cbn; intros. + induction l; cbn. exact eq_refl. + apply andP. exact a. exact IHl. + Defined. + + Program Definition tLambda na (b : eterm Σ (S n)) : eterm Σ n := tLambda na b. + Next Obligation. + intros. exact b. + Defined. + + Program Definition tLetIn na (b : eterm Σ n) (b' : eterm Σ (S n)) : eterm Σ n := tLetIn na b b'. + Next Obligation. + intros. cbn. apply andP. exact b. exact b'. + Defined. + + Lemma ap {A B : Type} (f : A -> B) (x y : A) : x = y -> f x = f y. + Proof. destruct 1. exact eq_refl. Defined. + + Lemma map_app {A B : Type} (f : A -> B) (l l' : list A) : map f (l ++ l') = map f l ++ map f l'. + Proof. + induction l; cbn. exact eq_refl. + now apply ap. + Defined. + + Program Definition tApp (f : eterm Σ n) (l : list (eterm Σ n)) (napp : ~~ isApp f) (nnil : l <> nil) : eterm Σ n := + mkApps f (map eterm_term l). + Next Obligation. + induction l using rev_ind => //. + intros Hf Hl. + specialize (IHl Hf). + rewrite map_app EAstUtils.mkApps_app /=. + apply andP. + - destruct l; cbn; [exact f|]. + apply IHl. intros Heq. congruence. + - exact x. + Defined. + + Program Definition tConst kn (isdecl : declared_constant Σ kn) : eterm Σ n := tConst kn. + + Program Definition tConstruct ind k (isdecl : declared_constructor Σ ind k) : eterm Σ n := tConstruct ind k. + + Program Definition tCase ci (c : eterm Σ n) (brs : list (∑ args : list name, eterm Σ (#|args| + n))) + (isdecl : declared_inductive Σ ci.1) : eterm Σ n := + tCase ci c (map (fun br : ∑ args, eterm Σ (#|args| + n) => (br.π1, proj1_sig br.π2)) brs). + Next Obligation. + intros. cbn. + destruct ci. apply andP. + - apply andP => //. exact c. + - induction brs; simp well_scoped => //. + cbn. rewrite IHbrs andb_true_r. exact a.π2. + Qed. + + Program Definition tProj p (c : eterm Σ n) (isdecl : declared_projection Σ p) : eterm Σ n := + tProj p c. + Next Obligation. + now cbn; intros. + Defined. + Next Obligation. + now cbn. + Defined. + Next Obligation. + cbn. + intros. apply andP => //. exact c. + Defined. + + Definition edefs := ∑ mfix : mfixpoint term, well_scoped_mfix Σ (#|mfix| + n) mfix. + + Program Definition tFix (mfix : edefs) idx : eterm Σ n := (tFix mfix.π1 idx). + Next Obligation. + cbn; intros. exact mfix.π2. + Defined. + + Program Definition tCoFix (mfix : edefs) idx : eterm Σ n := (tCoFix mfix.π1 idx). + Next Obligation. + cbn; intros; exact mfix.π2. + Defined. + End Constructors. +End Constructors. + +Definition ebr_br Σ n := (fun br : ∑ args : list name, eterm Σ (#|args| + n) => (br.π1, proj1_sig br.π2)). + +Module View. + Section view. + Context {Σ : global_context} {n : nat}. + + Inductive t : eterm Σ n -> Set := + | tBox : t Constructors.tBox + | tRel (i : nat) (le : i < n) : t (Constructors.tRel i le) + | tEvar (k : nat) l : t (Constructors.tEvar k l) + | tLambda na b : t (Constructors.tLambda na b) + | tLetIn na b b' : t (Constructors.tLetIn na b b') + | tApp (f : eterm Σ n) l (napp : ~~ isApp f) (nnil : l <> nil) : t (Constructors.tApp f l napp nnil) + | tConst kn (isdecl : declared_constant Σ kn): t (Constructors.tConst kn isdecl) + | tConstruct i k isdecl : t (Constructors.tConstruct i k isdecl) + | tCase ci c brs isdecl : t (Constructors.tCase ci c brs isdecl) + | tProj p c isdecl : t (Constructors.tProj p c isdecl) + | tFix mfix idx : t (Constructors.tFix mfix idx) + | tCoFix mfix idx : t (Constructors.tCoFix mfix idx). + Derive Signature for t. + + Equations view_term {e} (v : t e) : term := + | tBox => EAst.tBox + | tRel i le => EAst.tRel i + | tEvar k l => EAst.tEvar k (map eterm_term l) + | tLambda na b => EAst.tLambda na b + | tLetIn na b b' => EAst.tLetIn na b b' + | tApp f l napp nnil => EAst.mkApps f (map eterm_term l) + | tConst kn isdecl => EAst.tConst kn + | tConstruct i k isdecl => EAst.tConstruct i k + | tCase ci c brs isdecl => EAst.tCase ci c (map (ebr_br Σ n) brs) + | tProj p c isdecl => EAst.tProj p c + | tFix mfix idx => EAst.tFix mfix.π1 idx + | tCoFix mfix idx => EAst.tCoFix mfix.π1 idx. + + End view. + + Section ViewSizes. + Context {Σ : global_context} {n : nat}. + + Lemma view_size_let_def {na} {b : eterm Σ n} {b' : eterm Σ (S n)} : size b < size (Constructors.tLetIn na b b'). + Proof. cbn. lia. Qed. + + Lemma view_size_let_body {na} {b : eterm Σ n} {b' : eterm Σ (S n)} : size b' < size (Constructors.tLetIn na b b'). + Proof. cbn. lia. Qed. + + Lemma view_size_lambda {na} {b : eterm Σ (S n)} : size b < size (Constructors.tLambda na b). + Proof. cbn. lia. Qed. + + End ViewSizes. + +End View. + + +Lemma well_scoped_irr {Σ n t} (ws ws' : well_scoped Σ n t) : ws = ws'. +Proof. apply uip. Defined. + +Section view. + Context {Σ : global_context} {n : nat}. + + Equations? ws_list (l : list term) (ws : well_scoped_terms Σ n l) : list (eterm Σ n) := + | [], _ => [] + | t :: ts, ws => (exist _ t _) :: ws_list ts _. + Proof. + all:move/andP: ws => [] //. + Qed. + + Lemma ws_list_eq l ws : map eterm_term (ws_list l ws) = l. + Proof. funelim (ws_list l ws); cbn; auto. now rewrite H. Qed. + + Equations? ws_brs (l : list (list name × term)) (ws : well_scoped_brs Σ n l) : list (∑ args, eterm Σ (#|args| + n)) := + | [], _ => [] + | t :: ts, ws => (t.1; exist _ t.2 _) :: ws_brs ts _. + Proof. + all:move/andP: ws => [] //. + Qed. + + Lemma ws_brs_eq l ws : map (ebr_br Σ n) (ws_brs l ws) = l. + Proof. funelim (ws_brs l ws); cbn; auto. rewrite H. destruct t => //. Qed. + + Lemma andb_left {a b} : a && b -> a. + Proof. + move/andP=>[] //. + Qed. + + Lemma andb_right {a b} : a && b -> b. + Proof. + move/andP=>[] //. + Qed. + + Lemma well_scoped_terms_forallb {l} : well_scoped_terms Σ n l = forallb (well_scoped Σ n) l. + Proof. + induction l; simp well_scoped; auto. + now rewrite IHl. + Qed. + + Lemma well_scoped_mkApps {f l} : well_scoped Σ n (mkApps f l) -> well_scoped Σ n f && well_scoped_terms Σ n l. + Proof. + induction l using rev_ind; cbn; auto. + - now intros ->. + - rewrite mkApps_app /=. + move/andP=> [] wfl wx. + move/andP: (IHl wfl) => [] => -> wfnl /=. + now rewrite well_scoped_terms_forallb forallb_app -well_scoped_terms_forallb /= wfnl wx. + Qed. + + Notation " ( x ⧆ y ) " := (exist _ x y). + + Definition view (x : eterm Σ n) : View.t x. + Proof. + destruct x as [t ws]. + set (t' := t). + destruct t; + try solve [evar (ws' : well_scoped Σ n t' = true); rewrite (well_scoped_irr ws ws'); subst ws' t'; + unshelve econstructor; eauto]. + - evar (ws' : well_scoped Σ n t' = true); rewrite (well_scoped_irr ws ws'); subst ws' t'. + unshelve econstructor. now eapply Nat.leb_le. + - cbn in ws; discriminate. + - pose proof (ws_list_eq l ws). set(l' := ws_list l ws) in *. clearbody l'. + subst t'. + revert ws. destruct H. + intros ws. set (t' := tEvar _ _). + evar (ws' : well_scoped Σ n t' = true); rewrite (well_scoped_irr ws ws'); subst ws' t'. + econstructor. + - cbn in ws. + change t with (eterm_term (exist _ t ws)) in t'. + evar (ws' : well_scoped Σ n t' = true); rewrite (well_scoped_irr ws ws'); subst ws' t'. + econstructor. + - change t2 with (eterm_term (t2 ⧆ andb_right ws)) in t'. + change t1 with (eterm_term (t1 ⧆ andb_left ws)) in t'. + revert t'. + set (prf := andb_left ws). + set (prf' := andb_right _). clearbody prf prf'. + intros t'. + evar (ws' : well_scoped Σ n t' = true). + rewrite (well_scoped_irr ws ws'); subst ws' t'. + econstructor. + - destruct (decompose_app (tApp t1 t2)) eqn:da. + move: ws. revert t'. + rewrite (decompose_app_inv da). + intros t' ws. + move: (decompose_app_notApp _ _ _ da). + move: (EInduction.decompose_app_app _ _ _ _ da). + clear da t1 t2. + pose proof (wfa := well_scoped_mkApps ws). + pose proof (ws_list_eq l (andb_right wfa)). + set(l'' := ws_list l _) in *. clearbody l''. + move: ws. revert t'. + destruct H. + intros t'. + change t with (eterm_term (t ⧆ andb_left wfa)) in t'. + set (prfa := andb_left wfa) in *. clearbody prfa. + intros ws hl ht. + evar (ws' : well_scoped Σ n t' = true). + rewrite (well_scoped_irr ws ws'); subst ws' t'. + unshelve econstructor. exact ht. + intros hl'. subst l''. cbn in hl. congruence. + + - destruct p as [ind k]. + change t with (eterm_term (t ⧆ andb_right (andb_left ws))) in t'. + set (prf1 := andb_right _) in t'. clearbody prf1. + pose proof (ws_brs_eq _ (andb_right ws)). + set (prf2 := andb_right _) in *. clearbody prf2. + revert t' ws. destruct H. + intros t' ws. + evar (ws' : well_scoped Σ n t' = true). + rewrite (well_scoped_irr ws ws'); subst ws' t'. + unshelve econstructor. + cbn. apply (andb_left (andb_left ws)). + + - change t with (eterm_term (t ⧆ andb_right ws)) in t'. + set (prf := andb_right ws) in *. clearbody prf. + evar (ws' : well_scoped Σ n t' = true). + rewrite (well_scoped_irr ws ws'); subst ws' t'. + unshelve econstructor. apply (andb_left ws). + - change m with (((m; ws) : Constructors.edefs).π1) in t'. + evar (ws' : well_scoped Σ n t' = true). + rewrite (well_scoped_irr ws ws'); subst ws' t'. + econstructor. + - change m with (((m; ws) : Constructors.edefs).π1) in t'. + evar (ws' : well_scoped Σ n t' = true). + rewrite (well_scoped_irr ws ws'); subst ws' t'. + econstructor. + Defined. + + Lemma view_term (t : eterm Σ n) (v : View.t t) : View.view_term v = t. + Proof. + induction v; cbn; reflexivity. + Qed. + + Lemma view_view_term (t : eterm Σ n) : View.view_term (view t) = t. + Proof. apply view_term. Qed. + + Lemma view_view_size (t : eterm Σ n) : size (View.view_term (view t)) = size t. + Proof. now rewrite view_view_term. Qed. + +End view. + +From Coq Require Import Relation_Definitions. + +Section test_view. + Context (Σ : global_context). + Import View. + + Definition eterm_size : relation (∑ n, eterm Σ n) := + MR lt (fun x : ∑ n, eterm Σ n => size x.π2). + + Instance wf_size : WellFounded eterm_size. + Proof. unfold eterm_size. tc. Defined. + Obligation Tactic := idtac. + + #[derive(eliminator=no)] + Equations? test_view (n : nat) (t : eterm Σ n) : bool + by wf (n; t) eterm_size := + test_view n t with view t := { + | tBox => true + | tRel i lei => false + | tApp f args notapp notnil => test_view n f + | _ => false + }. + Proof. + do 2 red. cbn. rewrite size_mkApps. + destruct args; cbn; try congruence. lia. + Qed. + +End test_view. + + diff --git a/erasure/theories/Erasure.v b/erasure/theories/Erasure.v index 12e8d7552..ae833caf9 100644 --- a/erasure/theories/Erasure.v +++ b/erasure/theories/Erasure.v @@ -18,10 +18,11 @@ From MetaCoq.Erasure Require ErasureFunction EOptimizePropDiscr ERemoveParams EW (* Used to show timings of the ML execution *) -Definition time : forall {A}, string -> (unit -> A) -> unit -> A := - fun A s f x => f x. +Definition time : forall {A B}, string -> (A -> B) -> A -> B := + fun A B s f x => f x. -Extract Constant time => "(fun c f x -> let s = Tm_util.list_to_string c in Tm_util.time (Pp.str s) f x)". +Extract Constant time => + "(fun c f x -> let s = Caml_bytestring.caml_string_of_bytestring c in Tm_util.time (Pp.str s) f x)". (* This is the total erasure function + let-expansion of constructor arguments and case branches + @@ -437,10 +438,10 @@ Local Open Scope string_scope. Coq definition). *) Program Definition erase_and_print_template_program {cf : checker_flags} (p : Ast.Env.program) : string := - let (Σ', t) := erase_program p (todo "wf_env and welltyped term") in - time "Pretty printing" (fun _ => print_term Σ' [] true false t ^ nl ^ "in:" ^ nl ^ print_global_context Σ') tt. + let p' := erase_program p (todo "wf_env and welltyped term") in + time "Pretty printing" print_program p'. Program Definition erase_fast_and_print_template_program {cf : checker_flags} (p : Ast.Env.program) : string := - let (Σ', t) := erase_program_fast p (todo "wf_env and welltyped term") in - time "pretty-printing" (fun _ => print_term Σ' [] true false t ^ nl ^ "in:" ^ nl ^ print_global_context Σ') tt. + let p' := erase_program_fast p (todo "wf_env and welltyped term") in + time "pretty-printing" print_program p'. diff --git a/erasure/theories/ErasureFunction.v b/erasure/theories/ErasureFunction.v index 6252ecd31..da91f9aff 100644 --- a/erasure/theories/ErasureFunction.v +++ b/erasure/theories/ErasureFunction.v @@ -1237,14 +1237,13 @@ Proof. red. destruct H. split; eauto. red in H. red. inv wfΣ. inv X. - rewrite -H. simpl. unfold lookup_env; simpl; unfold eq_kername. destruct kername_eq_dec; try congruence. + rewrite -H. simpl. unfold lookup_env; simpl. destruct (eqb_spec (inductive_mind p.1) kn); try congruence. eapply lookup_env_Some_fresh in H. subst kn; contradiction. - econstructor; eauto. red. destruct H. split; eauto. inv wfΣ. inv X. red in H |- *. - rewrite -H. simpl. unfold lookup_env; simpl; unfold eq_kername. - unfold lookup_env in H; simpl in H. destruct kername_eq_dec; try congruence. + rewrite -H. simpl. unfold lookup_env; simpl; destruct (eqb_spec (inductive_mind p.1.1) kn); try congruence. eapply lookup_env_Some_fresh in H. subst kn. contradiction. Qed. @@ -1256,23 +1255,23 @@ Proof. intros wf hl. eapply lookup_env_Some_fresh in hl. inv wf. inv X. - destruct (kername_eq_dec kn kn'); subst; congruence. + destruct (eqb_spec kn kn'); subst; congruence. Qed. Lemma lookup_env_cons_disc {Σ kn kn' d} : kn <> kn' -> lookup_env (add_global_decl Σ (kn', d)) kn = lookup_env Σ kn. Proof. - intros Hk. simpl; unfold lookup_env; simpl; unfold eq_kername; simpl. - destruct kername_eq_dec; congruence. + intros Hk. simpl; unfold lookup_env; simpl. + destruct (eqb_spec kn kn'); congruence. Qed. Lemma elookup_env_cons_disc {Σ kn kn' d} : kn <> kn' -> ETyping.lookup_env ((kn', d) :: Σ) kn = ETyping.lookup_env Σ kn. Proof. - intros Hk. simpl; unfold eq_kername. - destruct kername_eq_dec; congruence. + intros Hk. simpl. + destruct (eqb_spec kn kn'); congruence. Qed. Lemma global_erases_with_deps_cons kn kn' d d' Σ Σ' : @@ -1404,9 +1403,9 @@ Lemma lookup_env_closed {Σ kn decl} : ETyping.closed_env Σ -> ETyping.lookup_e Proof. induction Σ; cbn => //. move/andP => [] cla cle. - unfold eq_kername; destruct kername_eq_dec. + case: eqb_spec. intros e; subst. move=> [= <-]. apply cla. - now eapply IHΣ. + intros hne; now eapply IHΣ. Qed. Lemma erases_closed Σ Γ t t' : Σ;;; Γ |- t ⇝ℇ t' -> PCUICAst.closedn #|Γ| t -> ELiftSubst.closedn #|Γ| t'. @@ -1450,7 +1449,7 @@ Proof. eapply KernameSet.subset_spec in sub. destruct (H i hin) as [[decl Hdecl]]. unfold lookup_env in Hdecl. cbn in Hdecl. - pose proof (eqb_spec i kn). unfold eqb in H0; cbn in H0. + pose proof (eqb_spec i kn). rewrite e in Hdecl. move: Hdecl. cbn -[erase_global_decls]. elim: H0. intros -> [= <-]. * destruct d as [|]; [left|right]. @@ -1737,19 +1736,19 @@ Proof. - move=> /= //. - intros e. destruct a as [kn' d']. cbn -[erase_global_decls]. - unfold eq_kername. destruct kername_eq_dec. subst kn'. + case: (eqb_spec kn kn'); intros e'; subst. intros [= ->]. unfold erase_global_decls. eapply KernameSet.mem_spec in hin. rewrite hin /=. now rewrite eq_kername_refl. intros hl. destruct d'. simpl. destruct KernameSet.mem. cbn. - unfold eq_kername. destruct kername_eq_dec => //. + rewrite (negbTE (proj2 (neqb _ _) e')). eapply IHg => //. eapply KernameSet.union_spec. left => //. eapply IHg => //. simpl. destruct KernameSet.mem. cbn. - unfold eq_kername. destruct kername_eq_dec => //. + rewrite (negbTE (proj2 (neqb _ _) e')). eapply IHg => //. eapply IHg => //. Qed. diff --git a/erasure/theories/Extraction.v b/erasure/theories/Extraction.v index 0d992f5fe..6a5a9c29f 100644 --- a/erasure/theories/Extraction.v +++ b/erasure/theories/Extraction.v @@ -1,5 +1,5 @@ (* Distributed under the terms of the MIT license. *) -From Coq Require Import Ascii FSets ExtrOcamlBasic ExtrOcamlString ExtrOcamlZInt ExtrOCamlFloats ExtrOCamlInt63. +From Coq Require Import Ascii FSets ExtrOcamlBasic ExtrOcamlZInt ExtrOCamlFloats ExtrOCamlInt63. From MetaCoq.Template Require Import utils. (** * Extraction setup for the erasure phase of template-coq. @@ -7,11 +7,6 @@ From MetaCoq.Template Require Import utils. Any extracted code planning to link with the plugin should use these same directives for consistency. *) - -Extract Constant ascii_compare => - "fun x y -> match Char.compare x y with 0 -> Eq | x when x < 0 -> Lt | _ -> Gt". -Extract Constant Ascii.compare => - "fun x y -> match Char.compare x y with 0 -> Eq | x when x < 0 -> Lt | _ -> Gt". Extraction Blacklist Classes config uGraph Universes Ast String List Nat Int UnivSubst Typing Checker Retyping OrderedType Logic Common ws_cumul_pb Classes Numeral @@ -38,10 +33,8 @@ Extract Constant Equations.Init.pr1 => "fst". Extract Constant Equations.Init.pr2 => "snd". Extraction Inline Equations.Init.pr1 Equations.Init.pr2. -Extract Constant PCUICWfEnv.guard_impl => - "(fun _ _ _ _ -> true)". - -Extract Constant PCUICTyping.guard_checking => "{ fix_guard = (fun _ _ _ -> true); cofix_guard = (fun _ _ _ -> true) }". +Extract Constant PCUICWfEnv.guard_impl => "(fun _ _ _ _ -> true)". +Extract Constant PCUICTyping.guard_checking => "(fun _ _ _ _ -> true)". Cd "src". diff --git a/examples/add_constructor.v b/examples/add_constructor.v index ac9da6bc3..03b28d6ed 100644 --- a/examples/add_constructor.v +++ b/examples/add_constructor.v @@ -91,7 +91,7 @@ Polymorphic Definition add_constructor (tm : Ast.term) (** * Examples *) -Local Open Scope string. +Local Open Scope bs_scope. (** Here we add a silly constructor to bool. *) MetaCoq Run ( add_constructor <% bool %> "foo" <% (fun x : Type => nat -> x -> bool -> x) %>). diff --git a/examples/demo.v b/examples/demo.v index cb4843b07..c05d427f4 100644 --- a/examples/demo.v +++ b/examples/demo.v @@ -83,9 +83,9 @@ Print eo_from_syntax. Local Notation Nat_module := (MPfile ["Datatypes"; "Init"; "Coq"], "nat"). -MetaCoq Unquote Definition two_from_syntax := (Ast.tApp (Ast.tConstruct (BasicAst.mkInd Nat_module 0) 1 nil) - (Ast.tApp (Ast.tConstruct (BasicAst.mkInd Nat_module 0) 1 nil) - (Ast.tConstruct (BasicAst.mkInd Nat_module 0) 0 nil :: nil) :: nil)). +MetaCoq Unquote Definition two_from_syntax := (Ast.tApp (Ast.tConstruct (Kernames.mkInd Nat_module 0) 1 nil) + (Ast.tApp (Ast.tConstruct (Kernames.mkInd Nat_module 0) 1 nil) + (Ast.tConstruct (Kernames.mkInd Nat_module 0) 0 nil :: nil) :: nil)). MetaCoq Quote Recursively Definition plus_syntax := plus. diff --git a/examples/metacoq_tour.v b/examples/metacoq_tour.v index 15508a1de..b6dacff61 100644 --- a/examples/metacoq_tour.v +++ b/examples/metacoq_tour.v @@ -1,10 +1,9 @@ -From Coq Require Import String. From MetaCoq.Template Require config utils All. From MetaCoq.Template Require Import TemplateMonad. From MetaCoq.PCUIC Require Import PCUICAst PCUICReduction PCUICCumulativity PCUICTyping PCUICSafeLemmata. Import MCMonadNotation. -Local Open Scope string_scope. +Local Open Scope bs_scope. (** MetaCoq is: diff --git a/examples/tauto.v b/examples/tauto.v index 89ec9e182..e6ab44e56 100644 --- a/examples/tauto.v +++ b/examples/tauto.v @@ -859,36 +859,36 @@ Equations reify (Σ : global_env_ext) (Γ : context) (P : term) : option form | None => None } ; | tInd ind [] - with kername_eq_dec ind.(inductive_mind) q_and := { - | left e2 with args := { + with eqb ind.(inductive_mind) q_and := { + | true with args := { | [ A ; B ] => af <- reify Σ Γ A ;; bf <- reify Σ Γ B ;; ret (And af bf) ; | _ => None } ; - | right _ - with kername_eq_dec ind.(inductive_mind) q_or := { - | left e2 with args := { + | false + with eqb ind.(inductive_mind) q_or := { + | true with args := { | [ A ; B ] => af <- reify Σ Γ A ;; bf <- reify Σ Γ B ;; ret (Or af bf) ; | _ => None } ; - | right _ - with kername_eq_dec ind.(inductive_mind) q_False := { - | left e2 with args := { + | false + with eqb ind.(inductive_mind) q_False := { + | true with args := { | [] => ret Fa ; | _ => None } ; - | right _ - with kername_eq_dec ind.(inductive_mind) q_True := { - | left e2 with args := { + | false + with eqb ind.(inductive_mind) q_True := { + | true with args := { | [] => ret Tr ; | _ => None } ; - | right _ => None + | false => None } } }} ; diff --git a/examples/typing_correctness.v b/examples/typing_correctness.v index e4834a990..8886f19c5 100644 --- a/examples/typing_correctness.v +++ b/examples/typing_correctness.v @@ -40,12 +40,10 @@ Proof. let wf_proof := eval hnf in (make_wf_env_ext gctx) in match wf_proof with | CorrectDecl ?x => exact x - | _ => fail "Couldn't prove the global environment is well-formed" + | _ => idtac wf_proof; fail "Couldn't prove the global environment is well-formed" end. Defined. - - (** There is always a proof of `forall x : Sort s, x -> x` *) Definition inh {cf:checker_flags} (Σ : wf_env_ext) Γ T := ∑ t, forall Σ0 : global_env_ext, abstract_env_rel' Σ Σ0 -> ∥ typing Σ0 Γ t T ∥. diff --git a/pcuic/theories/Conversion/PCUICNamelessConv.v b/pcuic/theories/Conversion/PCUICNamelessConv.v index c8fdac30b..aa6ed976b 100644 --- a/pcuic/theories/Conversion/PCUICNamelessConv.v +++ b/pcuic/theories/Conversion/PCUICNamelessConv.v @@ -237,7 +237,7 @@ Proof. rewrite /lookup_env /=. induction Σ. 1: reflexivity. simpl. - unfold eq_kername; destruct kername_eq_dec; subst. + case: eqb_spec; intros e; subst. - reflexivity. - assumption. Qed. diff --git a/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v b/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v index 78808dd01..d1c106735 100644 --- a/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v +++ b/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v @@ -1641,14 +1641,14 @@ Lemma wf_universe_subst_instance {cf : checker_flags} (Σ : global_env_ext) univ Proof. destruct l; simpl; auto. rename n into t. intros wfΣ Hl Hu e [[l n] [inl ->]]%In_subst_instance. - destruct l; simpl; auto. + destruct l as [|s|n']; simpl; auto. - unfold global_ext_levels. apply LS.union_spec. right. apply global_levels_Set. - specialize (Hl (Level.Level s, n) inl). simpl in Hl. apply monomorphic_level_in_global_ext in Hl. eapply LS.union_spec. now right. - - specialize (Hl (Level.Var n0, n) inl). + - specialize (Hl (Level.Var n', n) inl). eapply LS.union_spec in Hl as [Hl|Hl]. + red in Hu. unfold levels_of_udecl in Hl. @@ -1657,7 +1657,7 @@ Proof. destruct u; try discriminate. lsets. * destruct Hu as [declu [us vc]]. unfold subst_instance. simpl. - destruct (nth_error u n0) eqn:hnth. + destruct (nth_error u n') eqn:hnth. 2:{ simpl. eapply LS.union_spec; right; apply global_levels_Set. } eapply forallb_Forall in declu. eapply nth_error_forall in declu; eauto. diff --git a/pcuic/theories/Conversion/PCUICWeakeningEnvConv.v b/pcuic/theories/Conversion/PCUICWeakeningEnvConv.v index 314e0359e..c1bffc0ca 100644 --- a/pcuic/theories/Conversion/PCUICWeakeningEnvConv.v +++ b/pcuic/theories/Conversion/PCUICWeakeningEnvConv.v @@ -145,7 +145,7 @@ Lemma lookup_global_Some_fresh `{checker_flags} Σ c decl : lookup_global Σ c = Some decl -> ~ (fresh_global c Σ). Proof. induction Σ; cbn. 1: congruence. - unfold eq_kername; destruct kername_eq_dec; subst. + destruct (eqb_spec c a.1); subst. - intros [= <-] H2. inv H2. contradiction. - intros H1 H2. apply IHΣ; tas. @@ -171,7 +171,7 @@ Proof. - simpl. auto. - intros hl. depelim hΣ. specialize (IHΣ'' c decl hΣ hl). simpl in *. - unfold eq_kername; destruct kername_eq_dec; subst; auto. + destruct (eqb_spec c kn); subst; auto. apply lookup_global_Some_fresh in IHΣ''; contradiction. Qed. #[global] @@ -550,7 +550,7 @@ Proof. destruct Σ as [univs Σ]; cbn in *. induction wfΣ; simpl. 1: discriminate. cbn in HH. subst udecl. - unfold eq_kername in HH; destruct kername_eq_dec; subst. + destruct (eqb_spec c kn); subst. - apply some_inj in HH; destruct HH. subst. clear -o. unfold on_udecl, on_udecl_prop in *. destruct o as [H1 [H2 [H3 H4]]]. repeat split. diff --git a/pcuic/theories/Extraction.v b/pcuic/theories/Extraction.v index f29cd4135..febd24001 100644 --- a/pcuic/theories/Extraction.v +++ b/pcuic/theories/Extraction.v @@ -1,6 +1,6 @@ (* Distributed under the terms of the MIT license. *) Require Import MetaCoq.Template.utils. -Require Import FSets ssreflect ExtrOcamlBasic ExtrOcamlString ExtrOcamlZInt. +Require Import FSets ssreflect ExtrOcamlBasic ExtrOcamlZInt. (** * Extraction setup for the pcuic phase of MetaCoq. diff --git a/pcuic/theories/PCUICCanonicity.v b/pcuic/theories/PCUICCanonicity.v index 523f6e63d..2c831fa8d 100644 --- a/pcuic/theories/PCUICCanonicity.v +++ b/pcuic/theories/PCUICCanonicity.v @@ -155,7 +155,7 @@ Section Spines. eapply forallb_All in eqhd. destruct idx; simpl in Hl; [congruence|]. eapply All_nth_error in eqhd; eauto. - now eapply Reflect.eqb_eq in eqhd. + now eapply ReflectEq.eqb_eq in eqhd. Qed. Lemma wf_cofixpoint_inv mfix idx decl : @@ -180,7 +180,7 @@ Section Spines. eapply forallb_All in eqhd. destruct idx; simpl in Hl; [congruence|]. eapply All_nth_error in eqhd; eauto. - now eapply Reflect.eqb_eq in eqhd. + now eapply ReflectEq.eqb_eq in eqhd. Qed. Lemma on_free_vars_it_mkLambda_or_LetIn {P Δ t} : @@ -837,7 +837,7 @@ Section WeakNormalization. - move=> [hargs ccum']. rewrite expand_lets_mkApps subst_mkApps /= in ccum'. eapply invert_cumul_ind_ind in ccum' as ((? & ?) & ?). - len in r. eapply Reflect.eqb_eq in i0. now subst ind'. + len in r. eapply ReflectEq.eqb_eq in i0. now subst ind'. Qed. Lemma check_recursivity_kind_inj {mind rk rk'} : @@ -846,8 +846,8 @@ Section WeakNormalization. Proof. rewrite /check_recursivity_kind. case: lookup_env => //; case => // m. - elim: Reflect.eqb_spec; - elim: Reflect.eqb_spec; congruence. + elim: ReflectEq.eqb_spec; + elim: ReflectEq.eqb_spec; congruence. Qed. Lemma wt_closed_red_refl {Γ t T} : Σ ;;; Γ |- t : T -> Σ ;;; Γ ⊢ t ⇝ t. diff --git a/pcuic/theories/PCUICConversion.v b/pcuic/theories/PCUICConversion.v index 96630c23b..ebbf7457f 100644 --- a/pcuic/theories/PCUICConversion.v +++ b/pcuic/theories/PCUICConversion.v @@ -1441,7 +1441,7 @@ Section Inversions. Lemma invert_cumul_ind_ind {Γ ind ind' u u' args args'} : Σ ;;; Γ ⊢ mkApps (tInd ind u) args ≤ mkApps (tInd ind' u') args' -> - (Reflect.eqb ind ind' * PCUICEquality.R_global_instance Σ (eq_universe Σ) (leq_universe Σ) (IndRef ind) #|args| u u' * + (eqb ind ind' * PCUICEquality.R_global_instance Σ (eq_universe Σ) (leq_universe Σ) (IndRef ind) #|args| u u' * ws_cumul_pb_terms Σ Γ args args'). Proof. intros ht; eapply ws_cumul_pb_Ind_l_inv in ht as (? & ? & [? ? ?]); auto. diff --git a/pcuic/theories/PCUICExpandLetsCorrectness.v b/pcuic/theories/PCUICExpandLetsCorrectness.v index 41f9335d1..99140df12 100644 --- a/pcuic/theories/PCUICExpandLetsCorrectness.v +++ b/pcuic/theories/PCUICExpandLetsCorrectness.v @@ -281,8 +281,7 @@ Proof. destruct Σ as [univs Σ]. induction Σ. - reflexivity. - - cbn. - unfold eq_kername in *; destruct kername_eq_dec; subst. + - cbn. case: eqb_spec; intros e; subst. + destruct a; auto. + now rewrite IHΣ. Qed. diff --git a/pcuic/theories/PCUICInductiveInversion.v b/pcuic/theories/PCUICInductiveInversion.v index 0db1caf1b..2221fb902 100644 --- a/pcuic/theories/PCUICInductiveInversion.v +++ b/pcuic/theories/PCUICInductiveInversion.v @@ -238,8 +238,8 @@ Proof. apply decompose_app_inv in dapp => //. rewrite dapp. do 3 f_equal. symmetry. - change (eq_kername ind k) with (Reflect.eqb ind k) in eqk. - destruct (Reflect.eqb_spec ind k); auto. discriminate. + change (eq_kername ind k) with (ReflectEq.eqb ind k) in eqk. + destruct (ReflectEq.eqb_spec ind k); auto. discriminate. discriminate. - discriminate. - discriminate. diff --git a/pcuic/theories/PCUICParallelReductionConfluence.v b/pcuic/theories/PCUICParallelReductionConfluence.v index cd301000f..ca7043742 100644 --- a/pcuic/theories/PCUICParallelReductionConfluence.v +++ b/pcuic/theories/PCUICParallelReductionConfluence.v @@ -2051,7 +2051,7 @@ Section Rho. destruct (inspect_nth_error_rename (r:=r) eqbr) as [prf ->]; simp rho. cbv zeta. simp rho. cbn -[eqb]. autorewrite with len. - case: eqb_spec => // hlen. + case: eqb_specT => // hlen. + erewrite rename_rho_iota_red; tea => //. f_equal. pose proof (decompose_app_inv decapp). subst t. @@ -3680,7 +3680,7 @@ Section Rho. rewrite rho_app_case. rewrite decompose_app_mkApps; auto. change eq_inductive with (@eqb inductive _). - destruct (eqb_spec ci.(ci_ind) ci.(ci_ind)); try discriminate. + destruct (eqb_specT ci.(ci_ind) ci.(ci_ind)); try discriminate. 2:{ congruence. } unfold iota_red. eapply forallb_All in p4. eapply All2_All_mix_left in X3; tea. eapply All2_nth_error_Some_right in X3 as [br0 [hnth [hcl [IHbr [[predbod IHbod] breq]]]]]; tea. @@ -3957,7 +3957,7 @@ Section Rho. - simpl in *. inv_on_free_vars. rewrite rho_app_proj. rewrite decompose_app_mkApps; auto. change eq_inductive with (@eqb inductive _). - destruct (eqb_spec i i) => //. + destruct (eqb_specT i i) => //. eapply All2_nth_error_Some_right in heq_nth_error as [t' [? ?]]; eauto. simpl in y. rewrite e0. simpl. auto. eapply y => //. @@ -4128,9 +4128,9 @@ Section Rho. subst c0. cbn -[eqb]. simp rho. change eq_inductive with (@eqb inductive _). - destruct (eqb_spec ci.(ci_ind) ind). subst ind. + destruct (eqb_specT ci.(ci_ind) ind). subst ind. destruct (nth_error brs0 n) eqn:hbr => //. - case: eqb_spec => [eq|neq]. + case: eqb_specT => [eq|neq]. eapply pred1_mkApps_tConstruct in predc0 as [args' [? ?]]; pcuic. subst c1. intuition auto. rewrite rho_app_construct in Hc. @@ -4277,7 +4277,7 @@ Section Rho. eapply pred1_mkApps_refl_tConstruct in Hc. destruct nth_error eqn:Heq. change eq_inductive with (@eqb inductive _). - destruct (eqb_spec ind ind0); subst. + destruct (eqb_specT ind ind0); subst. econstructor; eauto. now rewrite nth_error_map Heq. eapply pred_proj_congr, pred_mkApps; auto with pcuic. diff --git a/pcuic/theories/PCUICSR.v b/pcuic/theories/PCUICSR.v index 3ab0d25e0..4c35e9d4d 100644 --- a/pcuic/theories/PCUICSR.v +++ b/pcuic/theories/PCUICSR.v @@ -2066,7 +2066,7 @@ Proof. eapply wf_cofixpoint_typing_spine in t; eauto. unfold check_recursivity_kind in t. rewrite isdecl.p1 in t. - apply Reflect.eqb_eq in t. rewrite t /= in heq_isCoFinite. + apply ReflectEq.eqb_eq in t. rewrite t /= in heq_isCoFinite. discriminate. - (* Case congruence on a parameter *) diff --git a/pcuic/theories/PCUICToTemplateCorrectness.v b/pcuic/theories/PCUICToTemplateCorrectness.v index 9d6c93b69..316ace2e0 100644 --- a/pcuic/theories/PCUICToTemplateCorrectness.v +++ b/pcuic/theories/PCUICToTemplateCorrectness.v @@ -141,8 +141,7 @@ Proof. destruct Σ as [univs Σ]. induction Σ. - reflexivity. - - cbn. - unfold eq_kername in *; destruct kername_eq_dec; subst. + - cbn. case: eqb_spec => intros; subst. + destruct a; auto. + now rewrite IHΣ. Qed. diff --git a/pcuic/theories/PCUICTyping.v b/pcuic/theories/PCUICTyping.v index 79d496175..7ecd373c5 100644 --- a/pcuic/theories/PCUICTyping.v +++ b/pcuic/theories/PCUICTyping.v @@ -82,7 +82,7 @@ Definition isCoFinite (r : recursivity_kind) := Definition check_recursivity_kind (lookup: kername -> option global_decl) ind r := match lookup ind with - | Some (InductiveDecl mib) => Reflect.eqb mib.(ind_finite) r + | Some (InductiveDecl mib) => ReflectEq.eqb mib.(ind_finite) r | _ => false end. @@ -109,7 +109,7 @@ Definition wf_fixpoint_gen | Some (ind :: inds) => (* Check that mutually recursive fixpoints are all on the same mututal inductive block *) - forallb (Reflect.eqb ind) inds && + forallb (eqb ind) inds && check_recursivity_kind lookup ind Finite | _ => false end. @@ -135,7 +135,7 @@ Definition wf_cofixpoint_gen | Some (ind :: inds) => (* Check that mutually recursive cofixpoints are all producing coinductives in the same mututal coinductive block *) - forallb (Reflect.eqb ind) inds && + forallb (eqb ind) inds && check_recursivity_kind lookup ind CoFinite | _ => false end. @@ -1252,14 +1252,14 @@ Section All_local_env. destruct Σ as [univs Σ]; rewrite /on_global_env /lookup_env; cbn. intros [cu Σp]. induction Σp; simpl. congruence. - unfold eq_kername; destruct kername_eq_dec; subst. - intros [= ->]. - exists ({| universes := univs; declarations := Σ |}). - split. - - red; cbn. split; [split;[lsets|csets]|]. - exists [(kn, decl)] => //. - - split => //. - - apply o0. + destruct (eqb_specT c kn); subst. + - intros [= ->]. + exists ({| universes := univs; declarations := Σ |}). + split. + * red; cbn. split; [split;[lsets|csets]|]. + exists [(kn, decl)] => //. + * split => //. + * apply o0. - intros hl. destruct (IHΣp hl) as [Σ' []]. exists Σ'. split=> //. diff --git a/pcuic/theories/PCUICWfUniverses.v b/pcuic/theories/PCUICWfUniverses.v index 8a24fe315..d6103759e 100644 --- a/pcuic/theories/PCUICWfUniverses.v +++ b/pcuic/theories/PCUICWfUniverses.v @@ -90,7 +90,7 @@ Section CheckerFlags. Proof. destruct l; simpl; auto. rename n into t. intros wfΣ Hl Hu e [[l n] [inl ->]]%In_subst_instance. - destruct l; simpl; auto. + destruct l as [|s|n']; simpl; auto. - unfold global_ext_levels. apply LS.union_spec. right. apply global_levels_Set. @@ -98,7 +98,7 @@ Section CheckerFlags. simpl in Hl. apply monomorphic_level_in_global_ext in Hl. eapply LS.union_spec. now right. - - specialize (Hl (Level.Var n0, n) inl). + - specialize (Hl (Level.Var n', n) inl). eapply LS.union_spec in Hl as [Hl|Hl]. + red in Hu. unfold levels_of_udecl in Hl. @@ -109,9 +109,10 @@ Section CheckerFlags. eapply nth_error_forall in Hu; eauto. eapply LS.union_spec; right. eapply global_levels_Set. * unfold subst_instance. simpl. - destruct (nth_error u n0) eqn:hnth. + destruct (nth_error u n') eqn:hnth. 2:{ simpl. rewrite hnth. eapply LS.union_spec; right; apply global_levels_Set. } - eapply nth_error_forall in Hu. 2:eauto. change (nth_error u n0) with (nth_error u n0) in *. + eapply nth_error_forall in Hu. 2:eauto. + change (nth_error u n') with (nth_error u n') in *. rewrite -> hnth. simpl. apply Hu. + now apply not_var_global_levels in Hl. Qed. diff --git a/pcuic/theories/Syntax/PCUICCases.v b/pcuic/theories/Syntax/PCUICCases.v index 8f9ffb898..35a40f355 100644 --- a/pcuic/theories/Syntax/PCUICCases.v +++ b/pcuic/theories/Syntax/PCUICCases.v @@ -209,10 +209,10 @@ Definition wf_predicateb mdecl idecl (p : predicate term) : bool := && forallb2 (fun na decl => eqb_binder_annot na decl.(decl_name)) (forget_types p.(pcontext)) (decl :: idecl.(ind_indices)). -Lemma wf_predicateP mdecl idecl p : reflect (wf_predicate mdecl idecl p) (wf_predicateb mdecl idecl p). +Lemma wf_predicateP mdecl idecl p : reflectProp (wf_predicate mdecl idecl p) (wf_predicateb mdecl idecl p). Proof. rewrite /wf_predicate /wf_predicate_gen /wf_predicateb. - case: Reflect.eqb_spec => eqpars /= //. + case: ReflectEq.eqb_spec => eqpars /= //. * case: (forallb2P _ _ (forget_types (pcontext p)) (idecl_binder idecl :: ind_indices idecl) (fun na decl => eqb_annot_reflect na decl.(decl_name))); constructor => //. intros [_ Hi]; contradiction. diff --git a/pcuic/theories/Syntax/PCUICReflect.v b/pcuic/theories/Syntax/PCUICReflect.v index e48db61d4..5ae4a841f 100644 --- a/pcuic/theories/Syntax/PCUICReflect.v +++ b/pcuic/theories/Syntax/PCUICReflect.v @@ -31,9 +31,9 @@ Local Ltac term_dec_tac term_dec := fcase (eq_dec x y) | x : list aname, y : list aname |- _ => fcase (eq_dec x y) | n : nat, m : nat |- _ => fcase (Nat.eq_dec n m) - | i : ident, i' : ident |- _ => fcase (string_dec i i') - | i : kername, i' : kername |- _ => fcase (kername_eq_dec i i') - | i : string, i' : kername |- _ => fcase (string_dec i i') + | i : ident, i' : ident |- _ => fcase (eq_dec i i') + | i : kername, i' : kername |- _ => fcase (eq_dec i i') + | i : string, i' : kername |- _ => fcase (eq_dec i i') | n : name, n' : name |- _ => fcase (eq_dec n n') | n : aname, n' : aname |- _ => fcase (eq_dec n n') | i : prim_val, j : prim_val |- _ => fcase (eq_dec i j) diff --git a/pcuic/theories/Typing/PCUICWeakeningEnvTyp.v b/pcuic/theories/Typing/PCUICWeakeningEnvTyp.v index 02ce3229f..0eb7555fd 100644 --- a/pcuic/theories/Typing/PCUICWeakeningEnvTyp.v +++ b/pcuic/theories/Typing/PCUICWeakeningEnvTyp.v @@ -365,7 +365,7 @@ Proof. destruct Hext as [univs' [Σ'' HΣ'']]. split; eauto. exists (Σ'' ++ [(kn, d)]). now rewrite <- app_assoc. } - unfold eq_kername; destruct kername_eq_dec; subst. + case: eqb_specT; intro eq; subst. - intros [= ->]. subst. clear Hext; eapply weakening_on_global_decl_ext. 3:tea. all:eauto. - apply IHonΣ; auto. @@ -386,7 +386,7 @@ Proof. destruct Hext as [univs' [Σ'' HΣ'']]. split; eauto. exists (Σ'' ++ [(kn, d)]). now rewrite <- app_assoc. } - unfold eq_kername; destruct kername_eq_dec; subst. + case: eqb_specT; intro e; subst. - intros [= ->]. subst. clear Hext; eapply weakening_on_global_decl. 5:tea. all:eauto. destruct wfΣ. split => //. now depelim o2. diff --git a/pcuic/theories/utils/PCUICAstUtils.v b/pcuic/theories/utils/PCUICAstUtils.v index 71ed8ba2c..2b57706de 100644 --- a/pcuic/theories/utils/PCUICAstUtils.v +++ b/pcuic/theories/utils/PCUICAstUtils.v @@ -92,7 +92,7 @@ Lemma lookup_env_cons {kn d Σ kn' d'} : lookup_global ((kn, d) :: Σ) kn' = Som (kn = kn' /\ d = d') \/ (kn <> kn' /\ lookup_global Σ kn' = Some d'). Proof. simpl. - epose proof (Reflect.eqb_spec (A:=kername) kn' kn). simpl in H. + epose proof (ReflectEq.eqb_spec (A:=kername) kn' kn). simpl in H. elim: H. intros -> [= <-]; intuition auto. intros diff look. intuition auto. Qed. @@ -102,7 +102,7 @@ Lemma lookup_env_cons_fresh {kn d Σ kn'} : lookup_global ((kn, d) :: Σ) kn' = lookup_global Σ kn'. Proof. simpl. - epose proof (Reflect.eqb_spec (A:=kername) kn' kn). simpl in H. + epose proof (ReflectEq.eqb_spec (A:=kername) kn' kn). simpl in H. elim: H. intros -> => //. auto. Qed. diff --git a/pcuic/theories/utils/PCUICPretty.v b/pcuic/theories/utils/PCUICPretty.v index 5791a9c4b..d98f3da56 100644 --- a/pcuic/theories/utils/PCUICPretty.v +++ b/pcuic/theories/utils/PCUICPretty.v @@ -18,7 +18,7 @@ Section lookups. end. End lookups. -Section print_term. +Section fresh. Context (Σ : global_env_ext). Fixpoint decompose_lam (t : term) (n : nat) : (list aname) * (list term) * term := @@ -34,7 +34,7 @@ Section print_term. end. Definition is_fresh (Γ : list ident) (id : ident) := - List.forallb (fun id' => negb (ident_eq id id')) Γ. + List.forallb (fun id' => negb (eqb id id')) Γ. Fixpoint name_from_term (t : term) := match t with @@ -47,7 +47,7 @@ Section print_term. | tConst c u => "x" | tInd (mkInd i k) u => match lookup_ind_decl Σ i k with - | Some (_, body) => substring 0 1 (body.(ind_name)) + | Some (_, body) => String.substring 0 1 (body.(ind_name)) | None => "X" end | _ => "U" @@ -109,201 +109,232 @@ Section print_term. Γ end in aux Γ (MCList.rev Γ'). +End fresh. - Section Aux. - Context (print_term : list ident -> bool -> bool -> term -> string). - Definition print_defs Γ (defs : mfixpoint term) := - let ctx' := fix_context defs in - print_list (print_def (print_term Γ true false) (print_term (fresh_names Γ ctx') true false)) - (nl ^ " with ") defs. +Module PrintTermTree. + Section env. + Context (Σ : global_env_ext). + Import bytestring.Tree. + Infix "^" := append. - Definition pr_context_decl Γ (c : context_decl) : ident * string := - match c with - | {| decl_name := na; decl_type := ty; decl_body := None |} => - let na' := (fresh_name Γ na.(binder_name) (Some ty)) in - (na', ("(" ++ na' ++ " : " ++ print_term Γ true false ty ++ ")")%string) - | {| decl_name := na; decl_type := ty; decl_body := Some b |} => - let na' := (fresh_name Γ na.(binder_name) (Some ty)) in - (na', ("(" ++ na' ++ " : " ++ print_term Γ true false ty ++ " := " ++ - print_term Γ true false b ++ ")")%string) - end. - - Fixpoint print_context_gen Γ Δ := - match Δ with - | [] => (Γ, ""%string) - | d :: decls => - let '(Γ, s) := print_context_gen Γ decls in - let '(na, s') := pr_context_decl Γ d in - match decls with - | [] => (na :: Γ, (s ++ s')%string) - | _ => (na :: Γ, (s ++ " " ++ s')%string) - end - end. - - Fixpoint print_context_names Γ Δ := - match Δ with - | [] => (Γ, ""%string) - | d :: decls => - let '(Γ, s) := print_context_names Γ decls in - let na := (fresh_name Γ d.(decl_name).(binder_name) (Some d.(decl_type))) in - match decls with - | [] => (na :: Γ, (s ++ na)%string) - | _ => (na :: Γ, (s ++ " " ++ na)%string) - end - end. + Section Aux. + Context (print_term : list ident -> bool -> bool -> term -> t). - End Aux. + Definition print_def {A} (f : A -> t) (g : A -> t) (def : def A) := + string_of_name (binder_name (dname def)) ^ " { struct " ^ string_of_nat (rarg def) ^ " }" ^ + " : " ^ f (dtype def) ^ " := " ^ nl ^ g (dbody def). + + Definition print_defs Γ (defs : mfixpoint term) := + let ctx' := fix_context defs in + print_list (print_def (print_term Γ true false) (print_term (fresh_names Σ Γ ctx') true false)) + (nl ^ " with ") defs. + + Definition pr_context_decl Γ (c : context_decl) : ident * t := + match c with + | {| decl_name := na; decl_type := ty; decl_body := None |} => + let na' := (fresh_name Σ Γ na.(binder_name) (Some ty)) in + (na', ("(" ^ na' ^ " : " ^ print_term Γ true false ty ^ ")")%bs) + | {| decl_name := na; decl_type := ty; decl_body := Some b |} => + let na' := (fresh_name Σ Γ na.(binder_name) (Some ty)) in + (na', ("(" ^ na' ^ " : " ^ print_term Γ true false ty ^ " := " ^ + print_term Γ true false b ^ ")")%bs) + end. + + Fixpoint print_context_gen Γ Δ := + match Δ with + | [] => (Γ, "" : t) + | d :: decls => + let '(Γ, s) := print_context_gen Γ decls in + let '(na, s') := pr_context_decl Γ d in + match decls with + | [] => (na :: Γ, s ^ s') + | _ => (na :: Γ, s ^ " " ^ s') + end + end. + + Fixpoint print_context_names Γ Δ := + match Δ with + | [] => (Γ, "" : t) + | d :: decls => + let '(Γ, s) := print_context_names Γ decls in + let na := (fresh_name Σ Γ d.(decl_name).(binder_name) (Some d.(decl_type))) in + match decls with + | [] => (na :: Γ, (s ^ na)) + | _ => (na :: Γ, (s ^ " " ^ na)) + end + end. + + End Aux. - Context (all : bool). + Context (all : bool). - Fixpoint print_term (Γ : list ident) (top : bool)(inapp : bool) (t : term) {struct t} := - match t with - | tRel n => - match nth_error Γ n with - | Some id => id - | None => "UnboundRel(" ^ string_of_nat n ^ ")" - end - | tVar n => "Var(" ^ n ^ ")" - | tEvar ev args => "Evar(" ^ string_of_nat ev ^ "[]" (* TODO *) ^ ")" - | tSort s => string_of_sort s - | tProd na dom codom => - let na' := fresh_name Γ na.(binder_name) (Some dom) in - parens top - ("∀ " ^ na' ^ " : " ^ - print_term Γ true false dom ^ ", " ^ print_term (na':: Γ) true false codom) - | tLambda na dom body => - let na' := fresh_name Γ na.(binder_name) (Some dom) in - parens top ("fun " ^ na' ^ " : " ^ print_term Γ true false dom - ^ " => " ^ print_term (na':: Γ) true false body) - | tLetIn na def dom body => - let na' := fresh_name Γ na.(binder_name) (Some dom) in - parens top ("let" ^ na' ^ " : " ^ print_term Γ true false dom ^ - " := " ^ print_term Γ true false def ^ " in " ^ nl ^ - print_term (na' :: Γ) true false body) - | tApp f l => - parens (top || inapp) (print_term Γ false true f ^ " " ^ print_term Γ false false l) - | tConst c u => string_of_kername c ^ print_universe_instance u - | tInd (mkInd i k) u => - match lookup_ind_decl Σ i k with - | Some (_, oib) => oib.(ind_name) ^ print_universe_instance u - | None => - "UnboundInd(" ^ string_of_inductive (mkInd i k) ^ "," ^ string_of_universe_instance u ^ ")" - end - | tConstruct (mkInd i k as ind) l u => - match lookup_ind_decl Σ i k with - | Some (_, oib) => - match nth_error oib.(ind_ctors) l with - | Some cdecl => cdecl.(cstr_name) ^ print_universe_instance u + Fixpoint print_term (Γ : list ident) (top : bool)(inapp : bool) (t : term) {struct t} : Tree.t := + match t with + | tRel n => + match nth_error Γ n with + | Some id => id + | None => "UnboundRel(" ^ string_of_nat n ^ ")" + end + | tVar n => "Var(" ^ n ^ ")" + | tEvar ev args => "Evar(" ^ string_of_nat ev ^ "[]" (* TODO *) ^ ")" + | tSort s => string_of_sort s + | tProd na dom codom => + let na' := fresh_name Σ Γ na.(binder_name) (Some dom) in + parens top + ("∀ " ^ na' ^ " : " ^ + print_term Γ true false dom ^ ", " ^ print_term (na':: Γ) true false codom) + | tLambda na dom body => + let na' := fresh_name Σ Γ na.(binder_name) (Some dom) in + parens top ("fun " ^ na' ^ " : " ^ print_term Γ true false dom + ^ " => " ^ print_term (na':: Γ) true false body) + | tLetIn na def dom body => + let na' := fresh_name Σ Γ na.(binder_name) (Some dom) in + parens top ("let" ^ na' ^ " : " ^ print_term Γ true false dom ^ + " := " ^ print_term Γ true false def ^ " in " ^ nl ^ + print_term (na' :: Γ) true false body) + | tApp f l => + parens (top || inapp) (print_term Γ false true f ^ " " ^ print_term Γ false false l) + | tConst c u => string_of_kername c ^ print_universe_instance u + | tInd (mkInd i k) u => + match lookup_ind_decl Σ i k with + | Some (_, oib) => oib.(ind_name) ^ print_universe_instance u + | None => + "UnboundInd(" ^ string_of_inductive (mkInd i k) ^ "," ^ string_of_universe_instance u ^ ")" + end + | tConstruct (mkInd i k as ind) l u => + match lookup_ind_decl Σ i k with + | Some (_, oib) => + match nth_error oib.(ind_ctors) l with + | Some cdecl => cdecl.(cstr_name) ^ print_universe_instance u + | None => + "UnboundConstruct(" ^ string_of_inductive ind ^ "," ^ string_of_nat l ^ "," + ^ string_of_universe_instance u ^ ")" + end | None => "UnboundConstruct(" ^ string_of_inductive ind ^ "," ^ string_of_nat l ^ "," ^ string_of_universe_instance u ^ ")" end - | None => - "UnboundConstruct(" ^ string_of_inductive ind ^ "," ^ string_of_nat l ^ "," - ^ string_of_universe_instance u ^ ")" - end - | tCase {| ci_ind := mkInd mind i as ind; ci_npar := pars |} p t brs => - match lookup_ind_decl Σ mind i with - | Some (_, oib) => - let Γret := p.(pcontext) in - let Γret := fresh_names Γ Γret in - let ret_binders := firstn #|pcontext p| Γret in - let (as_name, indices) := (hd "_" ret_binders, MCList.rev (tail ret_binders)) in - let in_args := repeat "_" #|pparams p| ++ indices in - let in_str := oib.(ind_name) ^ String.concat "" (map (fun a => " " ^ a) in_args) in + | tCase {| ci_ind := mkInd mind i as ind; ci_npar := pars |} p t brs => + match lookup_ind_decl Σ mind i with + | Some (_, oib) => + let Γret := p.(pcontext) in + let Γret := fresh_names Σ Γ Γret in + let ret_binders := firstn #|pcontext p| Γret in + let (as_name, indices) := (hd "_" ret_binders, MCList.rev (tail ret_binders)) in + let in_args := (repeat "_" #|pparams p| ++ indices)%list in + let in_str := oib.(ind_name) ^ concat "" (map (fun a : String.t => " " ^ a) in_args) in - let brs := map (fun br => - let (Γctx, pctx) := - if all then print_context_gen print_term Γ br.(bcontext) - else print_context_names Γ br.(bcontext) - in - pctx ^ " ⇒ " ^ print_term Γctx true false br.(bbody)) brs in - let brs := combine brs oib.(ind_ctors) in - - parens top ("match " ^ print_term Γ true false t ^ - " as " ^ as_name ^ - " in " ^ in_str ^ - " return " ^ print_term Γret true false (preturn p) ^ - " with " ^ nl ^ - print_list (fun '(b, cdecl) => cdecl.(cstr_name) ^ " " ^ b) - (nl ^ " | ") brs ^ nl ^ "end" ^ nl) - | None => - "Case(" ^ string_of_inductive ind ^ "," ^ string_of_nat i ^ "," ^ string_of_term t ^ "," - ^ string_of_predicate string_of_term p ^ "," ^ - string_of_list (pretty_string_of_branch string_of_term) brs ^ ")" - end - | tProj (mkInd mind i as ind, pars, k) c => - match lookup_ind_decl Σ mind i with - | Some (_, oib) => - match nth_error oib.(ind_projs) k with - | Some (na, _) => print_term Γ false false c ^ ".(" ^ na ^ ")" + let brs := map (fun br => + let (Γctx, pctx) := + if all then print_context_gen print_term Γ br.(bcontext) + else print_context_names Γ br.(bcontext) + in + pctx ^ " ⇒ " ^ print_term Γctx true false br.(bbody)) brs in + let brs := combine brs oib.(ind_ctors) in + + parens top ("match " ^ print_term Γ true false t ^ + " as " ^ as_name ^ + " in " ^ in_str ^ + " return " ^ print_term Γret true false (preturn p) ^ + " with " ^ nl ^ + print_list (fun '(b, cdecl) => cdecl.(cstr_name) ^ " " ^ b) + (nl ^ " | ") brs ^ nl ^ "end" ^ nl) + | None => + "Case(" ^ string_of_inductive ind ^ "," ^ string_of_nat i ^ "," ^ string_of_term t ^ "," + ^ string_of_predicate string_of_term p ^ "," ^ + string_of_list (pretty_string_of_branch string_of_term) brs ^ ")" + end + | tProj (mkInd mind i as ind, pars, k) c => + match lookup_ind_decl Σ mind i with + | Some (_, oib) => + match nth_error oib.(ind_projs) k with + | Some (na, _) => print_term Γ false false c ^ ".(" ^ na ^ ")" + | None => + "UnboundProj(" ^ string_of_inductive ind ^ "," ^ string_of_nat i ^ "," ^ string_of_nat k ^ "," + ^ print_term Γ true false c ^ ")" + end | None => "UnboundProj(" ^ string_of_inductive ind ^ "," ^ string_of_nat i ^ "," ^ string_of_nat k ^ "," - ^ print_term Γ true false c ^ ")" + ^ print_term Γ true false c ^ ")" end - | None => - "UnboundProj(" ^ string_of_inductive ind ^ "," ^ string_of_nat i ^ "," ^ string_of_nat k ^ "," - ^ print_term Γ true false c ^ ")" - end - | tFix l n => - parens top ("let fix " ^ print_defs print_term Γ l ^ nl ^ - " in " ^ List.nth_default (string_of_nat n) (map (string_of_aname ∘ dname) l) n) - | tCoFix l n => - parens top ("let cofix " ^ print_defs print_term Γ l ^ nl ^ - " in " ^ List.nth_default (string_of_nat n) (map (string_of_aname ∘ dname) l) n) - (* | tPrim i => parens top (string_of_prim (print_term Γ true false) i) *) - end. + | tFix l n => + parens top ("let fix " ^ print_defs print_term Γ l ^ nl ^ + " in " ^ List.nth_default (string_of_nat n) (map (string_of_aname ∘ dname) l) n) + | tCoFix l n => + parens top ("let cofix " ^ print_defs print_term Γ l ^ nl ^ + " in " ^ List.nth_default (string_of_nat n) (map (string_of_aname ∘ dname) l) n) + (* | tPrim i => parens top (string_of_prim (print_term Γ true false) i) *) + end. + End env. -End print_term. + Import bytestring.Tree. + Infix "^" := append. -Notation pr_term Σ Γ top := (print_term Σ true Γ top false). + Section env. + Context (Σ : global_env_ext). + Notation print_context := (print_context_gen Σ (print_term Σ true)). + Notation pr_term Γ top := (print_term Σ true Γ top false). -Notation print_context Σ := (print_context_gen Σ (print_term Σ true)). + Definition print_one_cstr Γ (mib : mutual_inductive_body) (c : constructor_body) : t := + let '(Γargs, s) := print_context Γ c.(cstr_args) in + c.(cstr_name) ^ " : " ^ s ^ "_" ^ print_list (pr_term Γargs true) " " c.(cstr_indices). + + Definition print_one_ind (short : bool) Γ (mib : mutual_inductive_body) (oib : one_inductive_body) : t := + let '(Γpars, spars) := print_context Γ mib.(ind_params) in + let '(Γinds, sinds) := print_context Γpars oib.(ind_indices) in + oib.(ind_name) ^ spars ^ " : " ^ sinds ^ " " ^ pr_term Γinds true (tSort oib.(ind_sort)) ^ ":=" ^ nl ^ + if short then "..." + else print_list (print_one_cstr Γpars mib) nl oib.(ind_ctors). + End env. -Definition print_one_cstr Σ Γ (mib : mutual_inductive_body) (c : constructor_body) : string := - let '(Γargs, s) := print_context Σ Γ c.(cstr_args) in - c.(cstr_name) ++ " : " ++ s ++ "_" ++ print_list (pr_term Σ Γargs true) " " c.(cstr_indices). - -Definition print_one_ind (short : bool) Σ Γ (mib : mutual_inductive_body) (oib : one_inductive_body) : string := - let '(Γpars, spars) := print_context Σ Γ mib.(ind_params) in - let '(Γinds, sinds) := print_context Σ Γpars oib.(ind_indices) in - oib.(ind_name) ++ spars ++ sinds ++ pr_term Σ Γinds true (tSort oib.(ind_sort)) ++ ":=" ++ nl ++ - if short then "..." - else print_list (print_one_cstr Σ Γpars mib) nl oib.(ind_ctors). + Fixpoint print_env_aux (short : bool) (prefix : nat) (Σ : global_env) (acc : t) : t := + match prefix with + | 0 => match Σ.(declarations) with [] => acc | _ => ("..." ^ nl ^ acc) end + | S n => + let univs := Σ.(universes) in + match Σ.(declarations) with + | [] => acc + | (kn, InductiveDecl mib) :: Σ => + let Σ' := ({| universes := univs; declarations := Σ |}, mib.(ind_universes)) in + let names := fresh_names Σ' [] (arities_context mib.(ind_bodies)) in + print_env_aux short n Σ'.1 + ("Inductive " ^ + print_list (print_one_ind Σ' short names mib) nl mib.(ind_bodies) ^ "." ^ + nl ^ acc) + | (kn, ConstantDecl cb) :: Σ => + let Σ' := ({| universes := univs; declarations := Σ |}, cb.(cst_universes)) in + print_env_aux short n Σ'.1 + ((match cb.(cst_body) with + | Some _ => "Definition " + | None => "Axiom " + end) ^ string_of_kername kn ^ " : " ^ print_term Σ' true nil true false cb.(cst_type) ^ + match cb.(cst_body) with + | Some b => + if short then ("..." ^ nl) + else (" := " ^ nl ^ print_term Σ' true nil true false b ^ "." ^ nl) + | None => "." + end ^ acc) + end + end. + + Definition print_env (short : bool) (prefix : nat) Σ := + print_env_aux short prefix Σ (Tree.string ""). -Fixpoint print_env_aux (short : bool) (prefix : nat) univs (Σ : global_declarations) (acc : string) := - match prefix with - | 0 => match Σ with [] => acc | _ => ("..." ++ nl ++ acc)%string end - | S n => - match Σ with - | [] => acc - | (kn, InductiveDecl mib) :: Σ => - let Σ' := ({| universes := univs; declarations := Σ |}, mib.(ind_universes)) in - let names := fresh_names Σ' [] (arities_context mib.(ind_bodies)) in - print_env_aux short n univs Σ - ("Inductive " ++ - print_list (print_one_ind short Σ' names mib) nl mib.(ind_bodies) ++ "." ++ - nl ++ acc)%string - | (kn, ConstantDecl cb) :: Σ => - let Σ' := ({| universes := univs; declarations := Σ |}, cb.(cst_universes)) in - print_env_aux short n univs Σ - ((match cb.(cst_body) with - | Some _ => "Definition " - | None => "Axiom " - end) ++ string_of_kername kn ++ " : " ++ pr_term Σ' nil true cb.(cst_type) ++ - match cb.(cst_body) with - | Some b => - if short then ("..." ++ nl)%string - else (" := " ++ nl ++ pr_term Σ' nil true b ++ "." ++ nl) - | None => "." - end ++ acc)%string - end - end. + Definition print_program (short : bool) (prefix : nat) (p : program) : t := + print_env short prefix (fst p) ^ nl ^ print_term (empty_ext (fst p)) true nil true false (snd p). -Definition print_env (short : bool) (prefix : nat) Σ := print_env_aux short prefix Σ.(universes) Σ.(declarations) EmptyString. +End PrintTermTree. +Definition print_term Σ all Γ top inapp t := + Tree.to_string (PrintTermTree.print_term Σ all Γ top inapp t). + +Definition print_context Σ Γ Δ : string := + Tree.to_string (PrintTermTree.print_context_gen Σ (PrintTermTree.print_term Σ true) Γ Δ).2. + +Definition print_env (short : bool) (prefix : nat) Σ := + Tree.to_string (PrintTermTree.print_env short prefix Σ). + Definition print_program (short : bool) (prefix : nat) (p : program) : string := - print_env short prefix (fst p) ++ nl ++ - pr_term (empty_ext (fst p)) nil true (snd p). + Tree.to_string (PrintTermTree.print_program short prefix p). \ No newline at end of file diff --git a/safechecker/_PluginProject.in b/safechecker/_PluginProject.in index bf97d347d..1ca874069 100644 --- a/safechecker/_PluginProject.in +++ b/safechecker/_PluginProject.in @@ -9,10 +9,6 @@ src/META.coq-metacoq-safechecker # From Coq's stdlib src/eqdepFacts.mli src/eqdepFacts.ml -src/orderedTypeAlt.mli -src/orderedTypeAlt.ml -src/orderedType0.mli -src/orderedType0.ml # From template src/ssrbool.ml @@ -83,12 +79,6 @@ src/pCUICPretty.ml # src/string2pos.ml # src/canonicalTries.mli # src/canonicalTries.ml -src/kernames.mli -src/kernames.ml -src/fMapList.mli -src/fMapList.ml -src/fMapAVL.mli -src/fMapAVL.ml src/pCUICErrors.mli src/pCUICErrors.ml src/pCUICWfEnv.mli diff --git a/safechecker/src/g_metacoq_safechecker.mlg b/safechecker/src/g_metacoq_safechecker.mlg index 4e99b2958..9ddb8aec0 100644 --- a/safechecker/src/g_metacoq_safechecker.mlg +++ b/safechecker/src/g_metacoq_safechecker.mlg @@ -9,25 +9,17 @@ open PeanoNat.Nat open Datatypes open PCUICSafeChecker open Tm_util +open Caml_bytestring -let bytes_of_list l = - let bytes = Bytes.create (List.length l) in - let rec fill acc = function - | [] -> bytes - | c :: cs -> - Bytes.set bytes acc c; - fill (1 + acc) cs - in fill 0 l - -let pr_char_list l = +let pr_string l = (* We allow utf8 encoding *) - str (Bytes.to_string (bytes_of_list l)) + str (caml_string_of_bytestring l) let check env evm poly (c, ustate) = debug (fun () -> str"Quoting"); let uctx = Evd.universe_context_set evm in let env = if poly then env else Environ.push_context_set ~strict:true uctx env in - let prog = time (str"Quoting") (Ast_quoter.quote_term_rec true env) (EConstr.to_constr evm c) in + let prog = time (str"Quoting") (Ast_quoter.quote_term_rec ~bypass:true ~with_universes:true env) (EConstr.to_constr evm c) in let uctx = if poly then let uctx = Evd.to_universe_context evm in @@ -44,8 +36,8 @@ let check env evm poly (c, ustate) = uctx in match check with - | Coq_inl s -> Feedback.msg_info (pr_char_list s) - | Coq_inr s -> CErrors.user_err (pr_char_list s) + | Coq_inl s -> Feedback.msg_info (pr_string s) + | Coq_inr s -> CErrors.user_err (pr_string s) } VERNAC COMMAND EXTEND MetaCoqSafeCheck CLASSIFIED AS QUERY diff --git a/safechecker/src/metacoq_safechecker_plugin.mlpack b/safechecker/src/metacoq_safechecker_plugin.mlpack index e51f0c2b8..963e90883 100644 --- a/safechecker/src/metacoq_safechecker_plugin.mlpack +++ b/safechecker/src/metacoq_safechecker_plugin.mlpack @@ -2,11 +2,6 @@ Int0 MSetAVL EqdepFacts Utils -OrderedTypeAlt -OrderedType0 -Kernames -FMapList -FMapAVL Ssrbool WGraph UGraph0 diff --git a/safechecker/theories/Extraction.v b/safechecker/theories/Extraction.v index 18143f3eb..08a7d41ab 100644 --- a/safechecker/theories/Extraction.v +++ b/safechecker/theories/Extraction.v @@ -1,5 +1,5 @@ (* Distributed under the terms of the MIT license. *) -From Coq Require Import OrdersTac Ascii ExtrOcamlBasic ExtrOcamlString ExtrOcamlZInt ExtrOCamlInt63 ExtrOCamlFloats. +From Coq Require Import OrdersTac Ascii ExtrOcamlBasic ExtrOcamlZInt ExtrOCamlInt63 ExtrOCamlFloats. From MetaCoq.Template Require Import utils MC_ExtrOCamlZPosInt. From MetaCoq.SafeChecker Require Import PCUICSafeChecker PCUICSafeConversion SafeTemplateChecker. @@ -12,11 +12,6 @@ From MetaCoq.SafeChecker Require Import PCUICSafeChecker PCUICSafeConversion (** Here we could extract uint63_from/to_model to the identity *) -Extract Constant ascii_compare => - "fun x y -> match Char.compare x y with 0 -> Eq | x when x < 0 -> Lt | _ -> Gt". -Extract Constant Ascii.compare => - "fun x y -> match Char.compare x y with 0 -> Eq | x when x < 0 -> Lt | _ -> Gt". - Extraction Blacklist Classes config uGraph Universes Ast String List Nat Int Init UnivSubst Typing Checker Retyping OrderedType Logic Common Equality Classes Uint63. diff --git a/safechecker/theories/PCUICConsistency.v b/safechecker/theories/PCUICConsistency.v index aa5049b5d..7f1610d43 100644 --- a/safechecker/theories/PCUICConsistency.v +++ b/safechecker/theories/PCUICConsistency.v @@ -6,7 +6,6 @@ provided term to build an inhabitant and then canonicity to show that this is a contradiction. *) -From Coq Require Import Ascii String. From Equations Require Import Equations. From MetaCoq.PCUIC Require Import PCUICAst. From MetaCoq.PCUIC Require Import PCUICAstUtils. @@ -31,7 +30,7 @@ Local Opaque hnf. Fixpoint string_repeat c (n : nat) : string := match n with | 0 => "" - | S n => String c (string_repeat c n) + | S n => String.String c (string_repeat c n) end. Lemma string_repeat_length c n : @@ -55,7 +54,7 @@ Proof. Qed. Definition make_fresh_name (Σ : global_env) : kername := - (MPfile [], string_repeat "a"%char (S (max_name_length Σ.(declarations)))). + (MPfile [], string_repeat "a"%byte (S (max_name_length Σ.(declarations)))). Lemma make_fresh_name_fresh Σ : fresh_global (make_fresh_name Σ) Σ.(declarations). diff --git a/safechecker/theories/PCUICEqualityDec.v b/safechecker/theories/PCUICEqualityDec.v index d2cd61e52..5ec87174f 100644 --- a/safechecker/theories/PCUICEqualityDec.v +++ b/safechecker/theories/PCUICEqualityDec.v @@ -168,7 +168,7 @@ Notation eqb_term_upto_univ eq leq gen_compare_global_instance := Ltac eqspec := lazymatch goal with | |- context [ eqb ?u ?v ] => - destruct (eqb_spec u v) ; nodec ; subst + destruct (eqb_specT u v) ; nodec ; subst end. Ltac eqspecs := @@ -307,7 +307,7 @@ Proof. now constructor. Qed. -Definition eqb_univ_reflect : forall u u' : Universe.t, reflect (u = u') (eqb u u'). +Definition eqb_univ_reflect : forall u u' : Universe.t, reflectProp (u = u') (eqb u u'). Proof. intros u u'. destruct (eqb_spec u u'); constructor; auto. @@ -333,16 +333,16 @@ Proof. destruct X; cbn ; subst. all: destruct (eqb_annot_reflect na na') => /= //. 2: apply/andb_and; split. - all: apply eq_dec_to_bool_refl. + all: eapply eqb_refl. - intros hcc'. eapply forallb2_bcompare_decl_All2_fold in hcc'; tea. eapply All2_fold_impl in hcc'; tea; simpl; intuition eauto. move: H. destruct d as [na [bod|] ty], d' as [na' [bod'|] ty']; cbn in * => //. + + destruct (eqb_annot_reflect na na') => // /=. + repeat case: eqb_specT => //; intros; subst; cbn; auto; constructor; auto. + destruct (eqb_annot_reflect na na') => //. - unfold eq_dec_to_bool. repeat destruct eq_dec => //; subst; cbn; auto; constructor; auto. - + destruct (eqb_annot_reflect na na') => //. - unfold eq_dec_to_bool. repeat destruct eq_dec => //; subst; cbn; auto; constructor; auto. + repeat case: eqb_specT => //; intros; subst; cbn; auto; constructor; auto. Qed. Lemma forallb_true {A : Type} (l : list A) : forallb xpredT l. @@ -586,7 +586,7 @@ Proof. intros H; now inv H. * constructor. intros H; now inv H. - cbn - [eqb]. - destruct (eqb_spec s k) ; nodec. subst. + destruct (eqb_specT s k) ; nodec. subst. induction u in ui, ht, ht' |- *. + destruct ui. * constructor. constructor. constructor. diff --git a/safechecker/theories/PCUICErrors.v b/safechecker/theories/PCUICErrors.v index 9b0706181..a4b3b95e3 100644 --- a/safechecker/theories/PCUICErrors.v +++ b/safechecker/theories/PCUICErrors.v @@ -309,11 +309,11 @@ Definition string_of_type_error Σ (e : type_error) : string := print_term Σ Γ t' ^ nl ^ "and:" ^ nl ^ print_term Σ Γ u' ^ nl ^ "error:" ^ nl ^ string_of_conv_error Σ e ^ (* nl ^ "in universe graph:" ^ nl ^ print_universes_graph G ^ nl ^ *) - " and context: " ^ nl ^ snd (print_context Σ [] Γ) + " and context: " ^ nl ^ print_context Σ [] Γ | NotConvertible Γ t u => "Terms are not convertible:" ^ nl ^ print_term Σ Γ t ^ nl ^ "and:" ^ nl ^ print_term Σ Γ u ^ (* nl ^ "in universe graph:" ^ nl ^ print_universes_graph G ^ nl ^ *) - " and context: " ^ nl ^ snd (print_context Σ [] Γ) + " and context: " ^ nl ^ print_context Σ [] Γ | NotASort t => "Not a sort: " ^ print_term Σ [] t | NotAProduct t t' => "Not a product" ^ print_term Σ [] t ^ nl ^ "(after reducing to " ^ print_term Σ [] t' diff --git a/safechecker/theories/PCUICSafeChecker.v b/safechecker/theories/PCUICSafeChecker.v index e903dafef..fc37abd99 100644 --- a/safechecker/theories/PCUICSafeChecker.v +++ b/safechecker/theories/PCUICSafeChecker.v @@ -867,8 +867,8 @@ Section CheckEnv. rename Heq_anonymous2 into dc. *) symmetry in Heq_anonymous0. eapply decompose_prod_n_assum_inv in Heq_anonymous0; simpl in Heq_anonymous0; subst. - destruct (eqb_spec params (ind_params mdecl)) => //. subst params. - destruct (eqb_spec args (cstr_args cdecl)) => //. subst args. + destruct (eqb_specT params (ind_params mdecl)) => //. subst params. + destruct (eqb_specT args (cstr_args cdecl)) => //. subst args. eapply eqb_eq in eqarglen0. eapply eqb_eq in eqbindices. eapply eqb_eq in eqbpars. @@ -1367,7 +1367,7 @@ Section CheckEnv. sq. rename Heq_anonymous into eqa. symmetry in eqa; eapply decompose_app_inv in eqa. subst t0. - move: eqhd; case: eqb_spec => // -> _. + move: eqhd; case: eqb_specT => // -> _. constructor. now eapply forallb_All in closedargs. Qed. @@ -2103,7 +2103,7 @@ Section CheckEnv. Qed. Next Obligation. - destruct (eqb_spec params (ind_params mdecl)); [|discriminate]. subst params. + destruct (eqb_specT params (ind_params mdecl)); [|discriminate]. subst params. red. red. eapply nth_error_all in wfars; eauto; simpl in wfars. destruct wfars as [s Hs]. now exists s. @@ -2399,12 +2399,11 @@ Section CheckEnv. let Σ := fst p in '(exist wfΣ eq) <- check_wf_ext (Σ, φ) ;; inft <- infer_term wfΣ (snd p) ;; - ret _. + ret (inft.π1; _). Next Obligation. - exists inft. assert (wfx : abstract_env_rel' x x) by reflexivity. - specialize_Σ wfx. - have [wfΣ] := (x.(wf_env_ext_wf)); sq. split; eauto. + have [wfΣ] := (x.(wf_env_ext_wf)). + cbn. specialize_Σ wfx. sq. split; eauto. Qed. End CheckEnv. diff --git a/safechecker/theories/PCUICSafeConversion.v b/safechecker/theories/PCUICSafeConversion.v index d648bde04..76fd15a54 100644 --- a/safechecker/theories/PCUICSafeConversion.v +++ b/safechecker/theories/PCUICSafeConversion.v @@ -2296,7 +2296,7 @@ Qed. apply All2_app. 1: assumption. constructor. 2: constructor. change (true = eqb u.(rarg) v.(rarg)) in eq1. - destruct (eqb_spec u.(rarg) v.(rarg)). 2: discriminate. + destruct (eqb_specT u.(rarg) v.(rarg)). 2: discriminate. symmetry in eqann. apply eqb_binder_annot_spec in eqann. split; auto. @@ -2332,7 +2332,7 @@ Qed. constructor. 2: eauto. change (true = eqb u.(rarg) v.(rarg)) in eq1. - destruct (eqb_spec u.(rarg) v.(rarg)). 2: discriminate. + destruct (eqb_specT u.(rarg) v.(rarg)). 2: discriminate. symmetry in eqann. apply eqb_binder_annot_spec in eqann. clear eq1. @@ -2367,7 +2367,7 @@ Qed. specialize_Σ wfΣ. destruct H as [H]; inversion H; destruct X0 as [_ [eq_uv eqann]]. change (?ru =? ?rv) with (eqb ru rv) in eq1. - destruct (eqb_spec (rarg u) (rarg v)) as [|neq_uv]; [discriminate|]. + destruct (eqb_specT (rarg u) (rarg v)) as [|neq_uv]; [discriminate|]. exact (neq_uv eq_uv). Qed. @@ -3771,9 +3771,7 @@ Qed. eexists. eassumption. Qed. Next Obligation. - change (eq_dec_to_bool ci ci') with (eqb ci ci') in eq5. - destruct (eqb_spec ci ci'). 2: discriminate. - assumption. + now apply eqb_eq. Qed. Next Obligation. eapply R_positionR. all: simpl. @@ -3782,9 +3780,7 @@ Qed. eapply positionR_poscat. constructor. Qed. Next Obligation. - change (eq_dec_to_bool ci ci') with (eqb ci ci') in eq5. - destruct (eqb_spec ci ci'). 2: discriminate. - assumption. + now apply eqb_eq. Qed. Next Obligation. unshelve eapply R_stateR. @@ -3802,7 +3798,7 @@ Qed. sq. apply ws_cumul_pb_eq_le_gen. change (eq_dec_to_bool ci ci') with (eqb ci ci') in eq5. - destruct (eqb_spec ci ci'). 2: discriminate. + destruct (eqb_specT ci ci'). 2: discriminate. subst. eapply ws_cumul_pb_Case. all: tas. * clear aux eq2 eq4. eapply welltyped_zipc_zipp in h1; eauto. eapply welltyped_is_open_term in h1. @@ -4536,7 +4532,7 @@ Qed. destruct convfix as [h1']. eapply conv_conv_cum_l. change (true = eqb idx idx') in eq4. - destruct (eqb_spec idx idx'). 2: discriminate. + destruct (eqb_specT idx idx'). 2: discriminate. subst. eapply ws_cumul_pb_Fix; eauto. - clear aux eee. eapply welltyped_zipc_zipp in h1; fvs. @@ -4584,7 +4580,7 @@ Qed. destruct convfix as [h1']. eapply conv_conv_cum_l. change (true = eqb idx idx') in eq4. - destruct (eqb_spec idx idx'). 2: discriminate. + destruct (eqb_specT idx idx'). 2: discriminate. subst. eapply ws_cumul_pb_CoFix; eauto. - clear aux h1'. eapply welltyped_zipc_zipp in h1; fvs. diff --git a/safechecker/theories/PCUICTypeChecker.v b/safechecker/theories/PCUICTypeChecker.v index 58ace43ee..59ca06d6a 100644 --- a/safechecker/theories/PCUICTypeChecker.v +++ b/safechecker/theories/PCUICTypeChecker.v @@ -2202,7 +2202,7 @@ Section Typecheck. eapply declared_inductive_inj in isdecl as []; tea. subst. apply absurd. - now apply/eqb_spec. + now apply/eqb_specT. Qed. Next Obligation. @@ -2250,7 +2250,7 @@ Section Typecheck. eapply infering_ind_ind in X0 as [? []]. 2-3: now auto. 2: now econstructor ; tea ; apply closed_red_red. - now apply/eqb_spec. + now apply/eqb_specT. Qed. Next Obligation. @@ -2308,20 +2308,20 @@ Section Typecheck. eapply infer_Proj with (pdecl := (i1, t)). - split. split. eassumption. cbn. rewrite hctors. reflexivity. split. symmetry; eassumption. cbn in *. - now apply beq_nat_true. - - cbn. destruct (ssrbool.elimT (eqb_spec ind I)); [assumption|]. + now apply Nat.eqb_eq. + - cbn. destruct (ssrbool.elimT (eqb_specT ind I)); [assumption|]. econstructor ; tea. now apply closed_red_red. - eapply type_reduction_closed in X2; eauto. 2: now apply infering_typing. eapply validity in X2; eauto. - destruct (ssrbool.elimT (eqb_spec ind I)); auto. + destruct (ssrbool.elimT (eqb_specT ind I)); auto. unshelve eapply (PCUICInductives.isType_mkApps_Ind_inv _ decl _) in X2 as [parsubst [argsubst [sp sp' cu]]]; eauto. - pose proof (PCUICContextSubst.context_subst_length2 (PCUICSpine.inst_ctx_subst sp)). - pose proof (PCUICContextSubst.context_subst_length2 (PCUICSpine.inst_ctx_subst sp')). - autorewrite with len in H, H0. + pose proof (Hl := PCUICContextSubst.context_subst_length2 (PCUICSpine.inst_ctx_subst sp)). + pose proof (Hr := PCUICContextSubst.context_subst_length2 (PCUICSpine.inst_ctx_subst sp')). + autorewrite with len in Hl, Hr. destruct (on_declared_inductive decl) eqn:ond. - rewrite -o.(onNpars) -H. + rewrite -o.(onNpars) -Hl. forward (o0.(onProjections)). intros H'; rewrite H' nth_error_nil // in HH. destruct ind_ctors as [|cs []]; auto. @@ -2329,7 +2329,7 @@ Section Typecheck. unshelve epose proof (onps.(on_projs_noidx _ _ _ _ _ _)). destruct (ind_indices idecl) => //. simpl in *. - rewrite List.skipn_length in H0. + rewrite List.skipn_length in e. rewrite List.firstn_length. lia. - destruct ind_projs => //. rewrite nth_error_nil in HH; congruence. Defined. @@ -2350,7 +2350,7 @@ Section Typecheck. pose proof (heΣ _ wfΣ) as [heΣ]. cbn in *. specialize_Σ wfΣ ; sq. apply absurd. - apply/eqb_spec. + apply/eqb_specT. cbn in *. sq. inversion X1. diff --git a/template-coq/_CoqProject b/template-coq/_CoqProject index 887d0440f..0444a18ec 100644 --- a/template-coq/_CoqProject +++ b/template-coq/_CoqProject @@ -7,6 +7,9 @@ theories/utils/canonicaltries/String2pos.v theories/utils/canonicaltries/CanonicalTries.v # utils +theories/utils/ByteCompare.v +theories/utils/ByteCompareSpec.v +theories/utils/bytestring.v theories/utils/MCPrelude.v theories/utils/MCReflect.v theories/utils/All_Forall.v @@ -23,9 +26,12 @@ theories/utils/MCSquash.v theories/utils/MCString.v theories/utils/wGraph.v theories/utils/MCUtils.v +# theories/utils/ByteCompare.v +# theories/utils/ByteCompareSpec.v #theories/utils/MC_ExtrOCamlInt63.v theories/utils/MC_ExtrOCamlNatInt.v theories/utils/MC_ExtrOCamlZPosInt.v +theories/utils/ReflectEq.v # common theories/common/uGraph.v @@ -33,6 +39,7 @@ theories/common/uGraph.v # Basic Ast files theories/utils.v theories/config.v +theories/Kernames.v theories/Universes.v theories/BasicAst.v theories/Environment.v @@ -40,7 +47,6 @@ theories/Ast.v theories/AstUtils.v theories/Reflect.v theories/ReflectAst.v -theories/Kernames.v theories/EnvMap.v theories/Induction.v theories/EnvironmentTyping.v diff --git a/template-coq/_PluginProject b/template-coq/_PluginProject index bccc5ca02..8ec2dbe9d 100644 --- a/template-coq/_PluginProject +++ b/template-coq/_PluginProject @@ -15,6 +15,8 @@ gen-src/all_Forall.ml gen-src/all_Forall.mli gen-src/ascii.ml gen-src/ascii.mli +gen-src/caml_nat.mli +gen-src/caml_nat.ml gen-src/ast0.ml gen-src/ast0.mli gen-src/ast_denoter.ml @@ -116,8 +118,10 @@ gen-src/mCReflect.mli gen-src/mCReflect.ml gen-src/primFloat.mli gen-src/primFloat.ml -gen-src/decimalString.mli -gen-src/decimalString.ml +#gen-src/decimalString.mli +#gen-src/decimalString.ml +gen-src/reflectEq.mli +gen-src/reflectEq.ml gen-src/mCString.ml gen-src/mCString.mli gen-src/mCUtils.ml @@ -146,6 +150,20 @@ gen-src/orders.ml gen-src/orders.mli gen-src/ordersTac.ml gen-src/ordersTac.mli +gen-src/orderedType0.mli +gen-src/orderedType0.ml +gen-src/orderedTypeAlt.mli +gen-src/orderedTypeAlt.ml +gen-src/decidableType.mli +gen-src/decidableType.ml +gen-src/fMapList.mli +gen-src/fMapList.ml +gen-src/fMapInterface.mli +gen-src/fMapInterface.ml +gen-src/fMapAVL.mli +gen-src/fMapAVL.ml +gen-src/fMapFacts.mli +gen-src/fMapFacts.ml gen-src/peanoNat.ml gen-src/peanoNat.mli gen-src/plugin_core.ml @@ -166,7 +184,22 @@ gen-src/specif.ml gen-src/specif.mli gen-src/string0.ml gen-src/string0.mli +gen-src/byte0.mli +gen-src/byte0.ml +gen-src/byte.mli +gen-src/byte.ml +gen-src/caml_byte.mli +gen-src/caml_byte.ml +gen-src/byteCompare.mli +gen-src/byteCompare.ml +gen-src/byteCompareSpec.mli +gen-src/byteCompareSpec.ml +gen-src/bytestring.mli +gen-src/bytestring.ml +gen-src/caml_bytestring.ml gen-src/tm_util.ml +gen-src/kernames.mli +gen-src/kernames.ml gen-src/universes0.ml gen-src/universes0.mli diff --git a/template-coq/gen-src/metacoq_template_plugin.mlpack b/template-coq/gen-src/metacoq_template_plugin.mlpack index 646267f16..77effcede 100644 --- a/template-coq/gen-src/metacoq_template_plugin.mlpack +++ b/template-coq/gen-src/metacoq_template_plugin.mlpack @@ -5,6 +5,7 @@ Decimal Hexadecimal Numeral0 Nat0 +Caml_nat List0 PeanoNat Specif @@ -15,15 +16,26 @@ BinNat BinInt PrimInt63 Uint0 -Uint63 +Int63 +Byte Ascii -Number +Byte0 +Caml_byte +ByteCompare +ByteCompareSpec String0 Orders OrdersTac OrdersFacts OrdersLists +OrderedType0 +OrderedTypeAlt Int0 +FMapList +FMapAVL +DecidableType +FMapInterface +FMapFacts MSetInterface MSetFacts MSetDecide @@ -47,6 +59,9 @@ MCRelations MCOption MCProd MCCompare +ReflectEq +Bytestring +Caml_bytestring Monad_utils Sumbool Zeven @@ -62,6 +77,7 @@ MCUtils Signature All_Forall Config0 +Kernames Universes0 BasicAst Environment diff --git a/template-coq/src/ast_denoter.ml b/template-coq/src/ast_denoter.ml index 23ca64c8b..d51f6c9de 100644 --- a/template-coq/src/ast_denoter.ml +++ b/template-coq/src/ast_denoter.ml @@ -1,5 +1,6 @@ open Names open Constr +open Kernames open BasicAst open Ast0 open Env @@ -8,7 +9,7 @@ open Tm_util module BaseExtractionDenoter = struct type t = Ast0.term - type quoted_ident = char list + type quoted_ident = Bytestring.String.t type quoted_int = Datatypes.nat type quoted_int63 = Uint63.t type quoted_float64 = Float64.t @@ -18,7 +19,7 @@ struct type quoted_relevance = relevance type quoted_sort = Universes0.Universe.t type quoted_cast_kind = cast_kind - type quoted_kernel_name = BasicAst.kername + type quoted_kernel_name = kername type quoted_inductive = inductive type quoted_proj = projection type quoted_global_reference = global_reference @@ -124,8 +125,10 @@ struct (* | Coq_tInt i -> ACoq_tInt i *) (* | Coq_tFloat f -> ACoq_tFloat f *) + let unquote_string = Caml_bytestring.caml_string_of_bytestring + let unquote_ident (qi: quoted_ident) : Id.t = - let s = list_to_string qi in + let s = unquote_string qi in Id.of_string s let unquote_relevance (r : relevance) : Sorts.relevance = @@ -142,10 +145,7 @@ struct {Context.binder_name = unquote_name q.binder_name; Context.binder_relevance = unquote_relevance q.binder_relevance} - let rec unquote_int (q: quoted_int) : int = - match q with - | Datatypes.O -> 0 - | Datatypes.S x -> succ (unquote_int x) + let unquote_int (q: quoted_int) : int = Caml_nat.caml_int_of_nat q let unquote_evar env evm n l = let id = Evar.unsafe_of_int (unquote_int n) in @@ -192,7 +192,7 @@ struct match trm with | Universes0.Level.Coq_lzero -> Univ.Level.set | Universes0.Level.Level s -> - let s = list_to_string s in + let s = unquote_string s in let comps = CString.split_on_char '.' s in let last, dp = CList.sep_last comps in let dp = DirPath.make (List.map Id.of_string comps) in @@ -212,7 +212,7 @@ struct | Universes0.Universe.Coq_lSProp -> evd, Sorts.sprop | Universes0.Universe.Coq_lProp -> evd, Sorts.prop | Universes0.Universe.Coq_lType u -> - (* let u = Universes0.Universe.t_set l in *) + let u = Universes0.Universe.t_set u in let ux_list = Universes0.UnivExprSet.elements u in let l = List.map unquote_level_expr ux_list in let u = List.fold_left Univ.Universe.sup (List.hd l) (List.tl l) in @@ -222,7 +222,7 @@ struct = (evm, Univ.Instance.of_array (Array.of_list (List0.map unquote_level l))) - let unquote_global_reference (trm : BasicAst.global_reference) : GlobRef.t = + let unquote_global_reference (trm : Kernames.global_reference) : GlobRef.t = let open GlobRef in match trm with | VarRef id -> VarRef (unquote_ident id) diff --git a/template-coq/src/ast_quoter.ml b/template-coq/src/ast_quoter.ml index 5812f6378..2ac1c1369 100644 --- a/template-coq/src/ast_quoter.ml +++ b/template-coq/src/ast_quoter.ml @@ -1,14 +1,16 @@ open Names open Datatypes +open Kernames open BasicAst open Ast0 open Ast0.Env open Tm_util +open Caml_bytestring module ExtractedASTBaseQuoter = struct type t = Ast0.term - type quoted_ident = char list + type quoted_ident = Bytestring.String.t type quoted_int = Datatypes.nat type quoted_int63 = Uint63.t type quoted_float64 = Float64.t @@ -18,7 +20,7 @@ struct type quoted_relevance = BasicAst.relevance type quoted_sort = Universes0.Universe.t type quoted_cast_kind = cast_kind - type quoted_kernel_name = BasicAst.kername + type quoted_kernel_name = Kernames.kername type quoted_inductive = inductive type quoted_proj = projection type quoted_global_reference = global_reference @@ -53,8 +55,8 @@ struct type quoted_global_env = global_env type quoted_program = program - let quote_ident id = - string_to_list (Id.to_string id) + let quote_string = bytestring_of_caml_string + let quote_ident id = quote_string (Id.to_string id) let quote_relevance = function | Sorts.Relevant -> BasicAst.Relevant @@ -68,12 +70,7 @@ struct let {Context.binder_name = n; Context.binder_relevance = relevance} = ann_n in { BasicAst.binder_name = quote_name n; BasicAst.binder_relevance = quote_relevance relevance } - let quote_int i = - let rec aux acc i = - if i < 0 then acc - else aux (Datatypes.S acc) (i - 1) - in aux Datatypes.O (i - 1) - + let quote_int = Caml_nat.nat_of_caml_int let quote_bool x = x let quote_int63 x = x @@ -85,7 +82,7 @@ struct if Univ.Level.is_set l then Universes0.Level.Coq_lzero else match Univ.Level.var_index l with | Some x -> Universes0.Level.Var (quote_int x) - | None -> Universes0.Level.Level (string_to_list (Univ.Level.to_string l)) + | None -> Universes0.Level.Level (quote_string (Univ.Level.to_string l)) let quote_level (l : Univ.Level.t) : (Universes0.PropLevel.t, Universes0.Level.t) Datatypes.sum = try Coq_inr (quote_nonprop_level l) @@ -120,18 +117,18 @@ struct | Constr.VMcast -> VmCast - let quote_dirpath (dp : DirPath.t) : BasicAst.dirpath = + let quote_dirpath (dp : DirPath.t) : Kernames.dirpath = let l = DirPath.repr dp in List.map quote_ident l - let rec quote_modpath (mp : ModPath.t) : BasicAst.modpath = + let rec quote_modpath (mp : ModPath.t) : Kernames.modpath = match mp with | MPfile dp -> MPfile (quote_dirpath dp) | MPbound mbid -> let (i, id, dp) = MBId.repr mbid in MPbound (quote_dirpath dp, quote_ident id, quote_int i) | MPdot (mp, id) -> MPdot (quote_modpath mp, quote_ident (Label.to_id id)) - let quote_kn (kn : KerName.t) : BasicAst.kername = + let quote_kn (kn : KerName.t) : Kernames.kername = (quote_modpath (KerName.modpath kn), quote_ident (Label.to_id (KerName.label kn))) diff --git a/template-coq/src/caml_byte.ml b/template-coq/src/caml_byte.ml new file mode 100644 index 000000000..4f1f093ef --- /dev/null +++ b/template-coq/src/caml_byte.ml @@ -0,0 +1,515 @@ +open Byte +let char_of_byte = function +| Coq_x00 -> '\000' +| Coq_x01 -> '\001' +| Coq_x02 -> '\002' +| Coq_x03 -> '\003' +| Coq_x04 -> '\004' +| Coq_x05 -> '\005' +| Coq_x06 -> '\006' +| Coq_x07 -> '\007' +| Coq_x08 -> '\008' +| Coq_x09 -> '\009' +| Coq_x0a -> '\010' +| Coq_x0b -> '\011' +| Coq_x0c -> '\012' +| Coq_x0d -> '\013' +| Coq_x0e -> '\014' +| Coq_x0f -> '\015' +| Coq_x10 -> '\016' +| Coq_x11 -> '\017' +| Coq_x12 -> '\018' +| Coq_x13 -> '\019' +| Coq_x14 -> '\020' +| Coq_x15 -> '\021' +| Coq_x16 -> '\022' +| Coq_x17 -> '\023' +| Coq_x18 -> '\024' +| Coq_x19 -> '\025' +| Coq_x1a -> '\026' +| Coq_x1b -> '\027' +| Coq_x1c -> '\028' +| Coq_x1d -> '\029' +| Coq_x1e -> '\030' +| Coq_x1f -> '\031' +| Coq_x20 -> '\032' +| Coq_x21 -> '\033' +| Coq_x22 -> '\034' +| Coq_x23 -> '\035' +| Coq_x24 -> '\036' +| Coq_x25 -> '\037' +| Coq_x26 -> '\038' +| Coq_x27 -> '\039' +| Coq_x28 -> '\040' +| Coq_x29 -> '\041' +| Coq_x2a -> '\042' +| Coq_x2b -> '\043' +| Coq_x2c -> '\044' +| Coq_x2d -> '\045' +| Coq_x2e -> '\046' +| Coq_x2f -> '\047' +| Coq_x30 -> '\048' +| Coq_x31 -> '\049' +| Coq_x32 -> '\050' +| Coq_x33 -> '\051' +| Coq_x34 -> '\052' +| Coq_x35 -> '\053' +| Coq_x36 -> '\054' +| Coq_x37 -> '\055' +| Coq_x38 -> '\056' +| Coq_x39 -> '\057' +| Coq_x3a -> '\058' +| Coq_x3b -> '\059' +| Coq_x3c -> '\060' +| Coq_x3d -> '\061' +| Coq_x3e -> '\062' +| Coq_x3f -> '\063' +| Coq_x40 -> '\064' +| Coq_x41 -> '\065' +| Coq_x42 -> '\066' +| Coq_x43 -> '\067' +| Coq_x44 -> '\068' +| Coq_x45 -> '\069' +| Coq_x46 -> '\070' +| Coq_x47 -> '\071' +| Coq_x48 -> '\072' +| Coq_x49 -> '\073' +| Coq_x4a -> '\074' +| Coq_x4b -> '\075' +| Coq_x4c -> '\076' +| Coq_x4d -> '\077' +| Coq_x4e -> '\078' +| Coq_x4f -> '\079' +| Coq_x50 -> '\080' +| Coq_x51 -> '\081' +| Coq_x52 -> '\082' +| Coq_x53 -> '\083' +| Coq_x54 -> '\084' +| Coq_x55 -> '\085' +| Coq_x56 -> '\086' +| Coq_x57 -> '\087' +| Coq_x58 -> '\088' +| Coq_x59 -> '\089' +| Coq_x5a -> '\090' +| Coq_x5b -> '\091' +| Coq_x5c -> '\092' +| Coq_x5d -> '\093' +| Coq_x5e -> '\094' +| Coq_x5f -> '\095' +| Coq_x60 -> '\096' +| Coq_x61 -> '\097' +| Coq_x62 -> '\098' +| Coq_x63 -> '\099' +| Coq_x64 -> '\100' +| Coq_x65 -> '\101' +| Coq_x66 -> '\102' +| Coq_x67 -> '\103' +| Coq_x68 -> '\104' +| Coq_x69 -> '\105' +| Coq_x6a -> '\106' +| Coq_x6b -> '\107' +| Coq_x6c -> '\108' +| Coq_x6d -> '\109' +| Coq_x6e -> '\110' +| Coq_x6f -> '\111' +| Coq_x70 -> '\112' +| Coq_x71 -> '\113' +| Coq_x72 -> '\114' +| Coq_x73 -> '\115' +| Coq_x74 -> '\116' +| Coq_x75 -> '\117' +| Coq_x76 -> '\118' +| Coq_x77 -> '\119' +| Coq_x78 -> '\120' +| Coq_x79 -> '\121' +| Coq_x7a -> '\122' +| Coq_x7b -> '\123' +| Coq_x7c -> '\124' +| Coq_x7d -> '\125' +| Coq_x7e -> '\126' +| Coq_x7f -> '\127' +| Coq_x80 -> '\128' +| Coq_x81 -> '\129' +| Coq_x82 -> '\130' +| Coq_x83 -> '\131' +| Coq_x84 -> '\132' +| Coq_x85 -> '\133' +| Coq_x86 -> '\134' +| Coq_x87 -> '\135' +| Coq_x88 -> '\136' +| Coq_x89 -> '\137' +| Coq_x8a -> '\138' +| Coq_x8b -> '\139' +| Coq_x8c -> '\140' +| Coq_x8d -> '\141' +| Coq_x8e -> '\142' +| Coq_x8f -> '\143' +| Coq_x90 -> '\144' +| Coq_x91 -> '\145' +| Coq_x92 -> '\146' +| Coq_x93 -> '\147' +| Coq_x94 -> '\148' +| Coq_x95 -> '\149' +| Coq_x96 -> '\150' +| Coq_x97 -> '\151' +| Coq_x98 -> '\152' +| Coq_x99 -> '\153' +| Coq_x9a -> '\154' +| Coq_x9b -> '\155' +| Coq_x9c -> '\156' +| Coq_x9d -> '\157' +| Coq_x9e -> '\158' +| Coq_x9f -> '\159' +| Coq_xa0 -> '\160' +| Coq_xa1 -> '\161' +| Coq_xa2 -> '\162' +| Coq_xa3 -> '\163' +| Coq_xa4 -> '\164' +| Coq_xa5 -> '\165' +| Coq_xa6 -> '\166' +| Coq_xa7 -> '\167' +| Coq_xa8 -> '\168' +| Coq_xa9 -> '\169' +| Coq_xaa -> '\170' +| Coq_xab -> '\171' +| Coq_xac -> '\172' +| Coq_xad -> '\173' +| Coq_xae -> '\174' +| Coq_xaf -> '\175' +| Coq_xb0 -> '\176' +| Coq_xb1 -> '\177' +| Coq_xb2 -> '\178' +| Coq_xb3 -> '\179' +| Coq_xb4 -> '\180' +| Coq_xb5 -> '\181' +| Coq_xb6 -> '\182' +| Coq_xb7 -> '\183' +| Coq_xb8 -> '\184' +| Coq_xb9 -> '\185' +| Coq_xba -> '\186' +| Coq_xbb -> '\187' +| Coq_xbc -> '\188' +| Coq_xbd -> '\189' +| Coq_xbe -> '\190' +| Coq_xbf -> '\191' +| Coq_xc0 -> '\192' +| Coq_xc1 -> '\193' +| Coq_xc2 -> '\194' +| Coq_xc3 -> '\195' +| Coq_xc4 -> '\196' +| Coq_xc5 -> '\197' +| Coq_xc6 -> '\198' +| Coq_xc7 -> '\199' +| Coq_xc8 -> '\200' +| Coq_xc9 -> '\201' +| Coq_xca -> '\202' +| Coq_xcb -> '\203' +| Coq_xcc -> '\204' +| Coq_xcd -> '\205' +| Coq_xce -> '\206' +| Coq_xcf -> '\207' +| Coq_xd0 -> '\208' +| Coq_xd1 -> '\209' +| Coq_xd2 -> '\210' +| Coq_xd3 -> '\211' +| Coq_xd4 -> '\212' +| Coq_xd5 -> '\213' +| Coq_xd6 -> '\214' +| Coq_xd7 -> '\215' +| Coq_xd8 -> '\216' +| Coq_xd9 -> '\217' +| Coq_xda -> '\218' +| Coq_xdb -> '\219' +| Coq_xdc -> '\220' +| Coq_xdd -> '\221' +| Coq_xde -> '\222' +| Coq_xdf -> '\223' +| Coq_xe0 -> '\224' +| Coq_xe1 -> '\225' +| Coq_xe2 -> '\226' +| Coq_xe3 -> '\227' +| Coq_xe4 -> '\228' +| Coq_xe5 -> '\229' +| Coq_xe6 -> '\230' +| Coq_xe7 -> '\231' +| Coq_xe8 -> '\232' +| Coq_xe9 -> '\233' +| Coq_xea -> '\234' +| Coq_xeb -> '\235' +| Coq_xec -> '\236' +| Coq_xed -> '\237' +| Coq_xee -> '\238' +| Coq_xef -> '\239' +| Coq_xf0 -> '\240' +| Coq_xf1 -> '\241' +| Coq_xf2 -> '\242' +| Coq_xf3 -> '\243' +| Coq_xf4 -> '\244' +| Coq_xf5 -> '\245' +| Coq_xf6 -> '\246' +| Coq_xf7 -> '\247' +| Coq_xf8 -> '\248' +| Coq_xf9 -> '\249' +| Coq_xfa -> '\250' +| Coq_xfb -> '\251' +| Coq_xfc -> '\252' +| Coq_xfd -> '\253' +| Coq_xfe -> '\254' +| Coq_xff -> '\255' +let byte_of_char = function +| '\000' -> Coq_x00 +| '\001' -> Coq_x01 +| '\002' -> Coq_x02 +| '\003' -> Coq_x03 +| '\004' -> Coq_x04 +| '\005' -> Coq_x05 +| '\006' -> Coq_x06 +| '\007' -> Coq_x07 +| '\008' -> Coq_x08 +| '\009' -> Coq_x09 +| '\010' -> Coq_x0a +| '\011' -> Coq_x0b +| '\012' -> Coq_x0c +| '\013' -> Coq_x0d +| '\014' -> Coq_x0e +| '\015' -> Coq_x0f +| '\016' -> Coq_x10 +| '\017' -> Coq_x11 +| '\018' -> Coq_x12 +| '\019' -> Coq_x13 +| '\020' -> Coq_x14 +| '\021' -> Coq_x15 +| '\022' -> Coq_x16 +| '\023' -> Coq_x17 +| '\024' -> Coq_x18 +| '\025' -> Coq_x19 +| '\026' -> Coq_x1a +| '\027' -> Coq_x1b +| '\028' -> Coq_x1c +| '\029' -> Coq_x1d +| '\030' -> Coq_x1e +| '\031' -> Coq_x1f +| '\032' -> Coq_x20 +| '\033' -> Coq_x21 +| '\034' -> Coq_x22 +| '\035' -> Coq_x23 +| '\036' -> Coq_x24 +| '\037' -> Coq_x25 +| '\038' -> Coq_x26 +| '\039' -> Coq_x27 +| '\040' -> Coq_x28 +| '\041' -> Coq_x29 +| '\042' -> Coq_x2a +| '\043' -> Coq_x2b +| '\044' -> Coq_x2c +| '\045' -> Coq_x2d +| '\046' -> Coq_x2e +| '\047' -> Coq_x2f +| '\048' -> Coq_x30 +| '\049' -> Coq_x31 +| '\050' -> Coq_x32 +| '\051' -> Coq_x33 +| '\052' -> Coq_x34 +| '\053' -> Coq_x35 +| '\054' -> Coq_x36 +| '\055' -> Coq_x37 +| '\056' -> Coq_x38 +| '\057' -> Coq_x39 +| '\058' -> Coq_x3a +| '\059' -> Coq_x3b +| '\060' -> Coq_x3c +| '\061' -> Coq_x3d +| '\062' -> Coq_x3e +| '\063' -> Coq_x3f +| '\064' -> Coq_x40 +| '\065' -> Coq_x41 +| '\066' -> Coq_x42 +| '\067' -> Coq_x43 +| '\068' -> Coq_x44 +| '\069' -> Coq_x45 +| '\070' -> Coq_x46 +| '\071' -> Coq_x47 +| '\072' -> Coq_x48 +| '\073' -> Coq_x49 +| '\074' -> Coq_x4a +| '\075' -> Coq_x4b +| '\076' -> Coq_x4c +| '\077' -> Coq_x4d +| '\078' -> Coq_x4e +| '\079' -> Coq_x4f +| '\080' -> Coq_x50 +| '\081' -> Coq_x51 +| '\082' -> Coq_x52 +| '\083' -> Coq_x53 +| '\084' -> Coq_x54 +| '\085' -> Coq_x55 +| '\086' -> Coq_x56 +| '\087' -> Coq_x57 +| '\088' -> Coq_x58 +| '\089' -> Coq_x59 +| '\090' -> Coq_x5a +| '\091' -> Coq_x5b +| '\092' -> Coq_x5c +| '\093' -> Coq_x5d +| '\094' -> Coq_x5e +| '\095' -> Coq_x5f +| '\096' -> Coq_x60 +| '\097' -> Coq_x61 +| '\098' -> Coq_x62 +| '\099' -> Coq_x63 +| '\100' -> Coq_x64 +| '\101' -> Coq_x65 +| '\102' -> Coq_x66 +| '\103' -> Coq_x67 +| '\104' -> Coq_x68 +| '\105' -> Coq_x69 +| '\106' -> Coq_x6a +| '\107' -> Coq_x6b +| '\108' -> Coq_x6c +| '\109' -> Coq_x6d +| '\110' -> Coq_x6e +| '\111' -> Coq_x6f +| '\112' -> Coq_x70 +| '\113' -> Coq_x71 +| '\114' -> Coq_x72 +| '\115' -> Coq_x73 +| '\116' -> Coq_x74 +| '\117' -> Coq_x75 +| '\118' -> Coq_x76 +| '\119' -> Coq_x77 +| '\120' -> Coq_x78 +| '\121' -> Coq_x79 +| '\122' -> Coq_x7a +| '\123' -> Coq_x7b +| '\124' -> Coq_x7c +| '\125' -> Coq_x7d +| '\126' -> Coq_x7e +| '\127' -> Coq_x7f +| '\128' -> Coq_x80 +| '\129' -> Coq_x81 +| '\130' -> Coq_x82 +| '\131' -> Coq_x83 +| '\132' -> Coq_x84 +| '\133' -> Coq_x85 +| '\134' -> Coq_x86 +| '\135' -> Coq_x87 +| '\136' -> Coq_x88 +| '\137' -> Coq_x89 +| '\138' -> Coq_x8a +| '\139' -> Coq_x8b +| '\140' -> Coq_x8c +| '\141' -> Coq_x8d +| '\142' -> Coq_x8e +| '\143' -> Coq_x8f +| '\144' -> Coq_x90 +| '\145' -> Coq_x91 +| '\146' -> Coq_x92 +| '\147' -> Coq_x93 +| '\148' -> Coq_x94 +| '\149' -> Coq_x95 +| '\150' -> Coq_x96 +| '\151' -> Coq_x97 +| '\152' -> Coq_x98 +| '\153' -> Coq_x99 +| '\154' -> Coq_x9a +| '\155' -> Coq_x9b +| '\156' -> Coq_x9c +| '\157' -> Coq_x9d +| '\158' -> Coq_x9e +| '\159' -> Coq_x9f +| '\160' -> Coq_xa0 +| '\161' -> Coq_xa1 +| '\162' -> Coq_xa2 +| '\163' -> Coq_xa3 +| '\164' -> Coq_xa4 +| '\165' -> Coq_xa5 +| '\166' -> Coq_xa6 +| '\167' -> Coq_xa7 +| '\168' -> Coq_xa8 +| '\169' -> Coq_xa9 +| '\170' -> Coq_xaa +| '\171' -> Coq_xab +| '\172' -> Coq_xac +| '\173' -> Coq_xad +| '\174' -> Coq_xae +| '\175' -> Coq_xaf +| '\176' -> Coq_xb0 +| '\177' -> Coq_xb1 +| '\178' -> Coq_xb2 +| '\179' -> Coq_xb3 +| '\180' -> Coq_xb4 +| '\181' -> Coq_xb5 +| '\182' -> Coq_xb6 +| '\183' -> Coq_xb7 +| '\184' -> Coq_xb8 +| '\185' -> Coq_xb9 +| '\186' -> Coq_xba +| '\187' -> Coq_xbb +| '\188' -> Coq_xbc +| '\189' -> Coq_xbd +| '\190' -> Coq_xbe +| '\191' -> Coq_xbf +| '\192' -> Coq_xc0 +| '\193' -> Coq_xc1 +| '\194' -> Coq_xc2 +| '\195' -> Coq_xc3 +| '\196' -> Coq_xc4 +| '\197' -> Coq_xc5 +| '\198' -> Coq_xc6 +| '\199' -> Coq_xc7 +| '\200' -> Coq_xc8 +| '\201' -> Coq_xc9 +| '\202' -> Coq_xca +| '\203' -> Coq_xcb +| '\204' -> Coq_xcc +| '\205' -> Coq_xcd +| '\206' -> Coq_xce +| '\207' -> Coq_xcf +| '\208' -> Coq_xd0 +| '\209' -> Coq_xd1 +| '\210' -> Coq_xd2 +| '\211' -> Coq_xd3 +| '\212' -> Coq_xd4 +| '\213' -> Coq_xd5 +| '\214' -> Coq_xd6 +| '\215' -> Coq_xd7 +| '\216' -> Coq_xd8 +| '\217' -> Coq_xd9 +| '\218' -> Coq_xda +| '\219' -> Coq_xdb +| '\220' -> Coq_xdc +| '\221' -> Coq_xdd +| '\222' -> Coq_xde +| '\223' -> Coq_xdf +| '\224' -> Coq_xe0 +| '\225' -> Coq_xe1 +| '\226' -> Coq_xe2 +| '\227' -> Coq_xe3 +| '\228' -> Coq_xe4 +| '\229' -> Coq_xe5 +| '\230' -> Coq_xe6 +| '\231' -> Coq_xe7 +| '\232' -> Coq_xe8 +| '\233' -> Coq_xe9 +| '\234' -> Coq_xea +| '\235' -> Coq_xeb +| '\236' -> Coq_xec +| '\237' -> Coq_xed +| '\238' -> Coq_xee +| '\239' -> Coq_xef +| '\240' -> Coq_xf0 +| '\241' -> Coq_xf1 +| '\242' -> Coq_xf2 +| '\243' -> Coq_xf3 +| '\244' -> Coq_xf4 +| '\245' -> Coq_xf5 +| '\246' -> Coq_xf6 +| '\247' -> Coq_xf7 +| '\248' -> Coq_xf8 +| '\249' -> Coq_xf9 +| '\250' -> Coq_xfa +| '\251' -> Coq_xfb +| '\252' -> Coq_xfc +| '\253' -> Coq_xfd +| '\254' -> Coq_xfe +| '\255' -> Coq_xff diff --git a/template-coq/src/caml_byte.mli b/template-coq/src/caml_byte.mli new file mode 100644 index 000000000..1dc00a1c5 --- /dev/null +++ b/template-coq/src/caml_byte.mli @@ -0,0 +1,4 @@ +open Byte + +val char_of_byte : byte -> char +val byte_of_char : char -> byte \ No newline at end of file diff --git a/template-coq/src/caml_bytestring.ml b/template-coq/src/caml_bytestring.ml new file mode 100644 index 000000000..3d733650f --- /dev/null +++ b/template-coq/src/caml_bytestring.ml @@ -0,0 +1,27 @@ +open Byte +open Caml_byte + +let rec caml_bytestring_length_aux s acc = + match s with + | Bytestring.String.EmptyString -> acc + | Bytestring.String.String (_, s) -> caml_bytestring_length_aux s (succ acc) + +let caml_bytestring_length s = caml_bytestring_length_aux s 0 + +let caml_string_of_bytestring (l : Bytestring.String.t) : string = + let len = caml_bytestring_length l in + let buf = Bytes.create len in + let rec aux i = function + | Bytestring.String.EmptyString -> () + | Bytestring.String.String (c, cs) -> + Bytes.set buf i (Caml_byte.char_of_byte c); aux (succ i) cs + in + aux 0 l; + Bytes.to_string buf + +let bytestring_of_caml_string (s : string) : Bytestring.String.t = + let rec aux acc i = + if i < 0 then acc + else aux (Bytestring.String.String (byte_of_char s.[i], acc)) (i - 1) + in aux Bytestring.String.EmptyString (String.length s - 1) + \ No newline at end of file diff --git a/template-coq/src/caml_nat.ml b/template-coq/src/caml_nat.ml new file mode 100644 index 000000000..5a78380a1 --- /dev/null +++ b/template-coq/src/caml_nat.ml @@ -0,0 +1,11 @@ +let nat_of_caml_int i = + let rec aux acc i = + if i < 0 then acc + else aux (Datatypes.S acc) (i - 1) + in aux Datatypes.O (i - 1) + +let rec caml_int_of_nat_aux n acc = + match n with + | Datatypes.O -> acc + | Datatypes.S x -> caml_int_of_nat_aux x (succ acc) +let caml_int_of_nat n = caml_int_of_nat_aux n 0 diff --git a/template-coq/src/caml_nat.mli b/template-coq/src/caml_nat.mli new file mode 100644 index 000000000..b11cab111 --- /dev/null +++ b/template-coq/src/caml_nat.mli @@ -0,0 +1,3 @@ + +val caml_int_of_nat : Datatypes.nat -> int +val nat_of_caml_int : int -> Datatypes.nat \ No newline at end of file diff --git a/template-coq/src/constr_denoter.ml b/template-coq/src/constr_denoter.ml index 2ae9a86a6..97fb9bdc0 100644 --- a/template-coq/src/constr_denoter.ml +++ b/template-coq/src/constr_denoter.ml @@ -81,15 +81,9 @@ struct let unquote_char trm = let (h,args) = app_full trm [] in - if constr_equall h tAscii then - match args with - a :: b :: c :: d :: e :: f :: g :: h :: [] -> - let bits = List.rev [a;b;c;d;e;f;g;h] in - let v = List.fold_left (fun a n -> (a lsl 1) lor if unquote_bool n then 1 else 0) 0 bits in - char_of_int v - | _ -> bad_term_verb trm "unquote_char" - else - not_supported trm + match Constr.kind h with + | Constr.Construct ((ind, idx), u) -> Char.chr (idx - 1) + | _ -> not_supported_verb trm "unquote_char" let unquote_string trm = let rec go n trm = @@ -324,8 +318,8 @@ struct not_supported_verb trm "unquote_inductive" - let unquote_int=unquote_nat - let print_term=Printer.pr_constr_env (Global.env ()) Evd.empty + let unquote_int = unquote_nat + let print_term = Printer.pr_constr_env (Global.env ()) Evd.empty let unquote_global_reference (trm : Constr.t) (* of type global_reference *) : GlobRef.t = diff --git a/template-coq/src/constr_quoter.ml b/template-coq/src/constr_quoter.ml index 26267f63d..73bb106f1 100644 --- a/template-coq/src/constr_quoter.ml +++ b/template-coq/src/constr_quoter.ml @@ -104,23 +104,21 @@ struct (* Quote OCaml int to Coq nat *) let quote_int = (* the cache is global but only accessible through quote_int *) - let cache = Hashtbl.create 10 in + let cache = Hashtbl.create 1023 in let rec recurse i = try Hashtbl.find cache i - with - Not_found -> - if i = 0 then - let result = Lazy.force tO in - let _ = Hashtbl.add cache i result in - result - else - let result = constr_mkApp (tS, [| recurse (i - 1) |]) in - let _ = Hashtbl.add cache i result in - result + with Not_found -> + if i = 0 then + let result = Lazy.force tO in + let _ = Hashtbl.add cache i result in + result + else + let result = constr_mkApp (tS, [| recurse (i - 1) |]) in + let _ = Hashtbl.add cache i result in + result in - fun i -> - if i >= 0 then recurse i else - CErrors.anomaly (str "Negative int can't be unquoted to nat.") + fun i -> if i >= 0 then recurse i else + CErrors.anomaly (str "Negative int can't be quoted to nat.") let quote_bool b = if b then Lazy.force ttrue else Lazy.force tfalse @@ -128,10 +126,17 @@ struct let quote_int63 i = constr_mkApp (tInt, [| Constr.mkInt i |]) let quote_float64 f = constr_mkApp (tFloat, [| Constr.mkFloat f |]) + let quote_inductive (kn, i) = + constr_mkApp (tmkInd, [| kn; i |]) + let byte_ind = + lazy (let ty = Lazy.force tByte in + match Constr.kind ty with + | Constr.Ind (ind, u) -> ind + | _ -> failwith "byte_ind : tByte is not bound to an inductive type") + let quote_char i = - constr_mkApp (tAscii, Array.of_list (List.map (fun m -> quote_bool ((i land m) = m)) - (List.rev [128;64;32;16;8;4;2;1]))) + Constr.mkConstruct (Lazy.force byte_ind, (i+1)) let chars = lazy (Array.init 255 quote_char) @@ -349,9 +354,6 @@ struct to_coq_list (prodl tident tTerm) (List.map (fun (a, b) -> pairl tident tTerm a b) d) - let quote_inductive (kn, i) = - constr_mkApp (tmkInd, [| kn; i |]) - let quote_dirpath dp = let l = DirPath.repr dp in to_coq_listl tident (List.map quote_ident l) diff --git a/template-coq/src/constr_reification.ml b/template-coq/src/constr_reification.ml index 15029a5a0..d4a4d83d5 100644 --- a/template-coq/src/constr_reification.ml +++ b/template-coq/src/constr_reification.ml @@ -91,7 +91,7 @@ struct let unit_tt = resolve "metacoq.unit.intro" - let tAscii = resolve "metacoq.ascii.intro" + let tByte = resolve "metacoq.byte.type" let tlist = resolve "metacoq.list.type" let c_nil = resolve "metacoq.list.nil" let c_cons = resolve "metacoq.list.cons" diff --git a/template-coq/src/genbytes.ml b/template-coq/src/genbytes.ml new file mode 100644 index 000000000..6a39256cb --- /dev/null +++ b/template-coq/src/genbytes.ml @@ -0,0 +1,60 @@ +(* +Printf.printf "let int_of_byte = \n";; + +for i = 0 to 255 do + Printf.printf "| Coq_x%02x -> %i\n" i i +done *) +(* +Printf.printf "open Byte\n";; + +Printf.printf "let char_of_byte = function\n";; + +for i = 0 to 255 do + Printf.printf "| Coq_x%02x -> '\\%03i'\n" i i +done;; + +Printf.printf "let byte_of_char = function\n";; +for i = 0 to 255 do + Printf.printf "| '\\%03i' -> Coq_x%02x\n" i i +done;; +*) +(* +Printf.printf "From Coq Require Import Byte.\n\n";; + +for i = 0 to 255 do + Printf.printf "Definition is_byte_x%02x (b : byte) := match b with | x%02x => true | _ => false end.\n" i i +done;; + +Printf.printf "Definition eqb (b b' : byte) :=\n match b with\n";; + +for i = 0 to 255 do + Printf.printf " | x%02x => is_byte_x%02x b'\n" i i +done;; +Printf.printf" end.\n";; + +for i = 0 to 255 do + Printf.printf "Definition compare_byte_x%02x (b : byte) :=\n" i; + Printf.printf " match b with\n"; + if i != 0 then Printf.printf " "; + for j = 0 to i - 1 do + Printf.printf "| x%02x " j + done; + if i != 0 then Printf.printf "=> Gt\n"; + Printf.printf " | x%02x => Eq\n" i; + if i != 255 then + Printf.printf " | _ => Lt\n"; + Printf.printf " end.\n" +done;; + +Printf.printf "Definition compare (b b' : byte) :=\n match b with\n";; + +for i = 0 to 255 do + Printf.printf " | x%02x => compare_byte_x%02x b'\n" i i +done;; +Printf.printf" end.\n";; +*) + + +for i = 0 to 255 do + Printf.printf "| x%02x -> <%%x%02x%%>\n" i i +done diff --git a/template-coq/src/quoter.ml b/template-coq/src/quoter.ml index 8502b06c7..2c9431486 100644 --- a/template-coq/src/quoter.ml +++ b/template-coq/src/quoter.ml @@ -466,7 +466,7 @@ struct univsty oib.mind_user_lc) univs mib.mind_packets - let quote_term_rec bypass env trm = + let quote_term_rec ~bypass ?(with_universes=true) env trm = let visited_terms = ref Names.KNset.empty in let visited_types = ref Mindset.empty in let universes = ref Univ.Level.Set.empty in @@ -553,10 +553,17 @@ struct (x,y) in let (tm, _) = quote_rem () env trm in - let univs = Univ.Level.Set.union (Vars.universes_of_constr trm) !universes in - let decls = List.fold_right (fun (kn, d) acc -> Q.add_global_decl kn d acc) !constants (Q.empty_global_declarations ()) in - let univs = Univ.Level.Set.filter (fun l -> Option.is_empty (Univ.Level.var_index l)) univs in - let univs = quote_ugraph ~kept:univs (Environ.universes env) in + let decls = List.fold_right (fun (kn, d) acc -> Q.add_global_decl kn d acc) !constants (Q.empty_global_declarations ()) in + let univs = + if with_universes then + let univs = Univ.Level.Set.union (Vars.universes_of_constr trm) !universes in + let univs = Univ.Level.Set.filter (fun l -> Option.is_empty (Univ.Level.var_index l)) univs in + quote_ugraph ~kept:univs (Environ.universes env) + else + (debug Pp.(fun () -> str"Skipping universes: "); + time (Pp.str"Quoting empty universe context") + (fun uctx -> Q.quote_univ_contextset uctx) Univ.ContextSet.empty) + in let env = Q.mk_global_env univs decls in Q.mk_program env tm diff --git a/template-coq/src/run_extractable.ml b/template-coq/src/run_extractable.ml index 136d31a3e..18c3e217b 100644 --- a/template-coq/src/run_extractable.ml +++ b/template-coq/src/run_extractable.ml @@ -6,8 +6,6 @@ open Tm_util open Ast_quoter open Ast_denoter -(* todo : replace string_to_list and converse by (un)quote_string *) - let to_reduction_strategy (s : Common0.reductionStrategy) : Plugin_core.reduction_strategy = match s with | Common0.Coq_cbv -> Plugin_core.rs_cbv @@ -17,11 +15,11 @@ let to_reduction_strategy (s : Common0.reductionStrategy) : Plugin_core.reductio | Common0.Coq_lazy -> Plugin_core.rs_lazy | Common0.Coq_unfold x -> failwith "not yet implemented: to_reduction_strategy" -let to_qualid (c : char list) : Libnames.qualid = - Libnames.qualid_of_string (list_to_string c) +let to_qualid c : Libnames.qualid = + Libnames.qualid_of_string (unquote_string c) -let of_qualid (q : Libnames.qualid) : char list = - string_to_list (Libnames.string_of_qualid q) +let of_qualid (q : Libnames.qualid) = + quote_string (Libnames.string_of_qualid q) (* todo(gmm): this definition adapted from quoter.ml *) let quote_rel_decl env = function @@ -209,8 +207,8 @@ let rec interp_tm (t : 'a coq_TM) : 'a tm = | Coq_tmReturn x -> tmReturn x | Coq_tmBind (c, k) -> tmBind (interp_tm c) (fun x -> interp_tm (k x)) | Coq_tmPrint t -> Obj.magic (tmPrint (to_constr t)) - | Coq_tmMsg msg -> Obj.magic (tmMsg (list_to_string msg)) - | Coq_tmFail err -> tmFailString (list_to_string err) + | Coq_tmMsg msg -> Obj.magic (tmMsg (unquote_string msg)) + | Coq_tmFail err -> tmFailString (unquote_string err) | Coq_tmEval (r,t) -> tmBind (tmEval (to_reduction_strategy r) (to_constr t)) (fun x -> Obj.magic (tmOfConstr x)) @@ -235,7 +233,7 @@ let rec interp_tm (t : 'a coq_TM) : 'a tm = tmMap (fun x -> Obj.magic (List.map quote_global_reference x)) (tmLocate (to_qualid id)) | Coq_tmCurrentModPath -> - tmMap (fun mp -> Obj.magic (string_to_list (Names.ModPath.to_string mp))) + tmMap (fun mp -> Obj.magic (quote_string (Names.ModPath.to_string mp))) tmCurrentModPath | Coq_tmQuoteInductive kn -> tmBind (tmQuoteInductive (unquote_kn kn)) diff --git a/template-coq/src/run_template_monad.ml b/template-coq/src/run_template_monad.ml index 39f2edd6c..7bde86623 100644 --- a/template-coq/src/run_template_monad.ml +++ b/template-coq/src/run_template_monad.ml @@ -407,7 +407,7 @@ let rec run_template_program_rec ~poly ?(intactic=false) (k : Constr.t Plugin_co in k ~st env evm qt | TmQuoteRecTransp (bypass, trm) -> let bypass = unquote_bool (reduce_all env evm bypass) in - let qt = quote_term_rec bypass env trm in + let qt = quote_term_rec ~bypass ~with_universes:true env trm in k ~st env evm qt | TmQuoteInd (name, strict) -> let kn = unquote_kn (reduce_all env evm name) in diff --git a/template-coq/src/tm_util.ml b/template-coq/src/tm_util.ml index af2731488..a09a536cd 100644 --- a/template-coq/src/tm_util.ml +++ b/template-coq/src/tm_util.ml @@ -58,23 +58,6 @@ let debug (m : unit ->Pp.t) = type ('a,'b) sum = Left of 'a | Right of 'b -(* todo(gmm): these are helper functions *) -let string_to_list (s : string) : char list = - let rec aux acc i = - if i < 0 then acc - else aux (s.[i] :: acc) (i - 1) - in aux [] (String.length s - 1) - -let list_to_string (l : char list) : string = - let buf = Bytes.create (List.length l) in - let rec aux i = function - | [] -> () - | c :: cs -> - Bytes.set buf i c; aux (succ i) cs - in - aux 0 l; - Bytes.to_string buf - let rec filter_map f l = match l with | [] -> [] diff --git a/template-coq/theories/AstUtils.v b/template-coq/theories/AstUtils.v index 4d4abfd43..9383a275a 100644 --- a/template-coq/theories/AstUtils.v +++ b/template-coq/theories/AstUtils.v @@ -14,7 +14,28 @@ Definition string_of_prim_int (i:Uint63.int) : string := Definition string_of_float (f : PrimFloat.float) := "". -Fixpoint string_of_term (t : term) := +Module string_of_term_tree. + Import bytestring.Tree. + Infix "^" := append. + + Definition string_of_predicate {term} (f : term -> t) (p : predicate term) := + "(" ^ "(" ^ concat "," (map f (pparams p)) ^ ")" + ^ "," ^ string_of_universe_instance (puinst p) + ^ ",(" ^ String.concat "," (map (string_of_name ∘ binder_name) (pcontext p)) ^ ")" + ^ "," ^ f (preturn p) ^ ")". + + Definition string_of_branch (f : term -> t) (b : branch term) := + "([" ^ String.concat "," (map (string_of_name ∘ binder_name) (bcontext b)) ^ "], " + ^ f (bbody b) ^ ")". + + Definition string_of_def {A} (f : A -> t) (def : def A) := + "(" ^ string_of_name (binder_name (dname def)) + ^ "," ^ string_of_relevance (binder_relevance (dname def)) + ^ "," ^ f (dtype def) + ^ "," ^ f (dbody def) + ^ "," ^ string_of_nat (rarg def) ^ ")". + + Fixpoint string_of_term (t : term) : Tree.t := match t with | tRel n => "Rel(" ^ string_of_nat n ^ ")" | tVar n => "Var(" ^ n ^ ")" @@ -53,7 +74,10 @@ Fixpoint string_of_term (t : term) := (* | tInt i => "Int(" ^ string_of_prim_int i ^ ")" | tFloat f => "Float(" ^ string_of_float f ^ ")" *) end. - +End string_of_term_tree. + +Definition string_of_term := Tree.to_string ∘ string_of_term_tree.string_of_term. + Fixpoint destArity Γ (t : term) := match t with | tProd na t b => destArity (Γ ,, vass na t) b @@ -184,7 +208,7 @@ Fixpoint lookup_mind_decl (id : kername) (decls : global_declarations) := match decls with | nil => None | (kn, InductiveDecl d) :: tl => - if eq_kername kn id then Some d else lookup_mind_decl id tl + if kn == id then Some d else lookup_mind_decl id tl | _ :: tl => lookup_mind_decl id tl end. diff --git a/template-coq/theories/BasicAst.v b/template-coq/theories/BasicAst.v index 48cb1b88e..c5dd0970b 100644 --- a/template-coq/theories/BasicAst.v +++ b/template-coq/theories/BasicAst.v @@ -1,75 +1,10 @@ (* Distributed under the terms of the MIT license. *) From Coq Require Import ssreflect Morphisms Orders Setoid. From MetaCoq.Template Require Import utils. +From MetaCoq.Template Require Export Kernames. From Coq Require Floats.SpecFloat. From Equations Require Import Equations. -(** ** Reification of names ** *) - -(** [Comment taken from Coq's code] - - Id.t is the type of identifiers, that is morally a subset of strings which - only contains Unicode characters of the Letter kind (and a few more). - => [ident] - - Name.t is an ad-hoc variant of Id.t option allowing to handle optionally - named objects. - => [name] - - DirPath.t represents generic paths as sequences of identifiers. - => [dirpath] - - Label.t is an equivalent of Id.t made distinct for semantical purposes. - => [ident] - - ModPath.t are module paths. - => [modpath] - - KerName.t are absolute names of objects in Coq. - => [kername] - - And also : - - Constant.t => [kername] - - variable => [ident] - - MutInd.t => [kername] - - inductive => [inductive] - - constructor => [inductive * nat] - - Projection.t => [projection] - - GlobRef.t => global_reference - - Finally, we define the models of primitive types (uint63 and floats64). -*) - -Definition ident := string. (* e.g. nat *) -Definition qualid := string. (* e.g. Datatypes.nat *) - -(** Type of directory paths. Essentially a list of module identifiers. The - order is reversed to improve sharing. E.g. A.B.C is ["C";"B";"A"] *) -Definition dirpath := list ident. -#[global] Instance dirpath_eqdec : Classes.EqDec dirpath := _. - -Definition string_of_dirpath : dirpath -> string - := String.concat "." ∘ rev. - -(** The module part of the kernel name. - - MPfile is for toplevel libraries, i.e. .vo files - - MPbound are parameters of functors - - MPdot is for submodules -*) -Inductive modpath := -| MPfile (dp : dirpath) -| MPbound (dp : dirpath) (id : ident) (i : nat) -| MPdot (mp : modpath) (id : ident). -Derive NoConfusion EqDec for modpath. - -Fixpoint string_of_modpath (mp : modpath) : string := - match mp with - | MPfile dp => string_of_dirpath dp - | MPbound dp id _ => string_of_dirpath dp ^ "." ^ id - | MPdot mp id => string_of_modpath mp ^ "." ^ id - end. - -(** The absolute names of objects seen by kernel *) -Definition kername := modpath × ident. -#[global] Instance kername_eqdec : Classes.EqDec kername := _. - -Definition string_of_kername (kn : kername) := - string_of_modpath kn.1 ^ "." ^ kn.2. - (** Identifiers that are allowed to be anonymous (i.e. "_" in concrete syntax). *) Inductive name : Set := | nAnon @@ -119,96 +54,6 @@ Definition string_of_relevance (r : relevance) := | Irrelevant => "Irrelevant" end. -(** Designation of a (particular) inductive type. *) -Record inductive : Set := mkInd { inductive_mind : kername ; - inductive_ind : nat }. -Arguments mkInd _%string _%nat. - -Derive NoConfusion EqDec for inductive. - -Definition string_of_inductive (i : inductive) := - string_of_kername (inductive_mind i) ^ "," ^ string_of_nat (inductive_ind i). - -Definition projection : Set := inductive * nat (* params *) * nat (* argument *). - -(** Kernel declaration references [global_reference] *) -Inductive global_reference := -| VarRef : ident -> global_reference -| ConstRef : kername -> global_reference -| IndRef : inductive -> global_reference -| ConstructRef : inductive -> nat -> global_reference. - -Derive NoConfusion EqDec for global_reference. - - -Definition string_of_gref gr : string := - match gr with - | VarRef v => v - | ConstRef s => string_of_kername s - | IndRef (mkInd s n) => - "Inductive " ^ string_of_kername s ^ " " ^ (string_of_nat n) - | ConstructRef (mkInd s n) k => - "Constructor " ^ string_of_kername s ^ " " ^ (string_of_nat n) ^ " " ^ (string_of_nat k) - end. - -Definition kername_eq_dec (k k0 : kername) : {k = k0} + {k <> k0} := Classes.eq_dec k k0. -#[global] Hint Resolve kername_eq_dec : eq_dec. - -Definition gref_eq_dec (gr gr' : global_reference) : {gr = gr'} + {~ gr = gr'} := Classes.eq_dec gr gr'. - -Definition ident_eq (x y : ident) := - match string_compare x y with - | Eq => true - | _ => false - end. - -Lemma ident_eq_spec x y : reflect (x = y) (ident_eq x y). -Proof. - unfold ident_eq. destruct (string_compare_eq x y). - destruct string_compare; constructor; auto. - - now apply H. - - intro Heq; specialize (H0 Heq). discriminate. - - intro Heq; specialize (H0 Heq). discriminate. -Qed. - -(* todo : better ? *) -Definition eq_kername (k k0 : kername) : bool := - match kername_eq_dec k k0 with - | left _ => true - | right _ => false - end. - -Lemma eq_kername_refl kn : eq_kername kn kn. -Proof. - unfold eq_kername. destruct (kername_eq_dec kn kn); cbnr. - contradiction. -Qed. - -Definition eq_inductive i i' := - let 'mkInd i n := i in - let 'mkInd i' n' := i' in - eq_kername i i' && Nat.eqb n n'. - -Definition eq_constant := eq_kername. - -Definition eq_projection p p' := - let '(ind, pars, arg) := p in - let '(ind', pars', arg') := p' in - eq_inductive ind ind' && Nat.eqb pars pars' && Nat.eqb arg arg'. - -Lemma eq_inductive_refl i : eq_inductive i i. -Proof. - destruct i as [mind k]. - unfold eq_inductive. now rewrite eq_kername_refl PeanoNat.Nat.eqb_refl. -Qed. - -Lemma eq_projection_refl i : eq_projection i i. -Proof. - destruct i as [[mind k] p]. - unfold eq_projection. - now rewrite eq_inductive_refl !PeanoNat.Nat.eqb_refl. -Qed. - (** The kind of a cast *) Inductive cast_kind : Set := | VmCast diff --git a/template-coq/theories/Checker.v b/template-coq/theories/Checker.v index 0e4f74d50..593cd815e 100644 --- a/template-coq/theories/Checker.v +++ b/template-coq/theories/Checker.v @@ -366,7 +366,7 @@ Fixpoint eq_term `{checker_flags} (φ : universes_graph) (t u : term) {struct t} match t, u with | tRel n, tRel n' => Nat.eqb n n' | tEvar ev args, tEvar ev' args' => Nat.eqb ev ev' && forallb2 (eq_term φ) args args' - | tVar id, tVar id' => eq_string id id' + | tVar id, tVar id' => eqb id id' | tSort s, tSort s' => check_eqb_universe φ s s' | tCast f k T, tCast f' k' T' => eq_term φ f f' && eq_term φ T T' | tApp f args, tApp f' args' => eq_term φ f f' && forallb2 (eq_term φ) args args' @@ -398,7 +398,7 @@ Fixpoint leq_term `{checker_flags} (φ : universes_graph) (t u : term) {struct t match t, u with | tRel n, tRel n' => Nat.eqb n n' | tEvar ev args, tEvar ev' args' => Nat.eqb ev ev' && forallb2 (eq_term φ) args args' - | tVar id, tVar id' => eq_string id id' + | tVar id, tVar id' => eqb id id' | tSort s, tSort s' => check_leqb_universe φ s s' | tApp f args, tApp f' args' => eq_term φ f f' && forallb2 (eq_term φ) args args' | tCast f k T, tCast f' k' T' => eq_term φ f f' && eq_term φ T T' diff --git a/template-coq/theories/Constants.v b/template-coq/theories/Constants.v index 5399e00bf..37b770f8a 100644 --- a/template-coq/theories/Constants.v +++ b/template-coq/theories/Constants.v @@ -1,15 +1,14 @@ (* Distributed under the terms of the MIT license. *) -From MetaCoq.Template Require Import BasicAst uGraph TemplateMonad +From MetaCoq.Template Require Import bytestring BasicAst uGraph TemplateMonad TemplateMonad.Extractable. (* Base types *) -Register Coq.Strings.String.string as metacoq.string.type. -Register Coq.Strings.String.EmptyString as metacoq.string.nil. -Register Coq.Strings.String.String as metacoq.string.cons. +Register bytestring.String.t as metacoq.string.type. +Register bytestring.String.EmptyString as metacoq.string.nil. +Register bytestring.String.String as metacoq.string.cons. -Register Coq.Strings.Ascii.ascii as metacoq.ascii.type. -Register Coq.Strings.Ascii.Ascii as metacoq.ascii.intro. +Register Coq.Init.Byte.byte as metacoq.byte.type. Register Coq.Init.Datatypes.nat as metacoq.nat.type. Register Coq.Init.Datatypes.O as metacoq.nat.zero. @@ -62,15 +61,21 @@ Register MetaCoq.Template.BasicAst.aname as metacoq.ast.aname. Register MetaCoq.Template.BasicAst.nAnon as metacoq.ast.nAnon. Register MetaCoq.Template.BasicAst.nNamed as metacoq.ast.nNamed. -Register MetaCoq.Template.BasicAst.ident as metacoq.ast.ident. -Register MetaCoq.Template.BasicAst.kername as metacoq.ast.kername. -Register MetaCoq.Template.BasicAst.modpath as metacoq.ast.modpath. -Register MetaCoq.Template.BasicAst.MPfile as metacoq.ast.MPfile. -Register MetaCoq.Template.BasicAst.MPbound as metacoq.ast.MPbound. -Register MetaCoq.Template.BasicAst.MPdot as metacoq.ast.MPdot. +Register MetaCoq.Template.Kernames.ident as metacoq.ast.ident. +Register MetaCoq.Template.Kernames.kername as metacoq.ast.kername. +Register MetaCoq.Template.Kernames.modpath as metacoq.ast.modpath. +Register MetaCoq.Template.Kernames.MPfile as metacoq.ast.MPfile. +Register MetaCoq.Template.Kernames.MPbound as metacoq.ast.MPbound. +Register MetaCoq.Template.Kernames.MPdot as metacoq.ast.MPdot. +Register MetaCoq.Template.Kernames.inductive as metacoq.ast.inductive. +Register MetaCoq.Template.Kernames.mkInd as metacoq.ast.mkInd. +Register MetaCoq.Template.Kernames.global_reference as metacoq.ast.global_reference. +Register MetaCoq.Template.Kernames.VarRef as metacoq.ast.VarRef. +Register MetaCoq.Template.Kernames.ConstRef as metacoq.ast.ConstRef. +Register MetaCoq.Template.Kernames.IndRef as metacoq.ast.IndRef. +Register MetaCoq.Template.Kernames.ConstructRef as metacoq.ast.ConstructRef. + Register MetaCoq.Template.BasicAst.name as metacoq.ast.name. -Register MetaCoq.Template.BasicAst.inductive as metacoq.ast.inductive. -Register MetaCoq.Template.BasicAst.mkInd as metacoq.ast.mkInd. Register MetaCoq.Template.BasicAst.def as metacoq.ast.def. Register MetaCoq.Template.BasicAst.mkdef as metacoq.ast.mkdef. Register MetaCoq.Template.BasicAst.cast_kind as metacoq.ast.cast_kind. @@ -83,11 +88,6 @@ Register MetaCoq.Template.BasicAst.recursivity_kind as metacoq.ast.recursivity_k Register MetaCoq.Template.BasicAst.Finite as metacoq.ast.Finite. Register MetaCoq.Template.BasicAst.CoFinite as metacoq.ast.CoFinite. Register MetaCoq.Template.BasicAst.BiFinite as metacoq.ast.BiFinite. -Register MetaCoq.Template.BasicAst.global_reference as metacoq.ast.global_reference. -Register MetaCoq.Template.BasicAst.VarRef as metacoq.ast.VarRef. -Register MetaCoq.Template.BasicAst.ConstRef as metacoq.ast.ConstRef. -Register MetaCoq.Template.BasicAst.IndRef as metacoq.ast.IndRef. -Register MetaCoq.Template.BasicAst.ConstructRef as metacoq.ast.ConstructRef. Register MetaCoq.Template.BasicAst.fresh_evar_id as metacoq.ast.fresh_evar_id. (* Universes *) diff --git a/template-coq/theories/EnvMap.v b/template-coq/theories/EnvMap.v index a046a6709..48251411c 100644 --- a/template-coq/theories/EnvMap.v +++ b/template-coq/theories/EnvMap.v @@ -9,9 +9,6 @@ Import String2pos. Implicit Types (cf:checker_flags). -Module KernameMap := FMapAVL.Make KernameOT.OT. -Module KernameMapFact := FMapFacts.WProperties KernameMap. - Module EnvMap. (* We keep the definition of EnvMap polymorphic over the data associated to a kername *) Section Poly. @@ -36,7 +33,7 @@ Module EnvMap. intros ne. unfold lookup, add. rewrite KernameMapFact.F.add_neq_o //. - intros eq. apply KernameOT.compare_eq in eq. congruence. + intros eq. apply Kername.compare_eq in eq. congruence. Qed. Lemma gss (e : t) kn kn' g : kn = kn' -> @@ -45,7 +42,7 @@ Module EnvMap. intros eq. unfold lookup, add. rewrite KernameMapFact.F.add_eq_o //. - now apply KernameOT.compare_eq. + now apply Kername.compare_eq. Qed. Definition equal (g g' : t) := KernameMap.Equal g g'. @@ -148,10 +145,10 @@ Module EnvMap. intros frΣ frk eq. intros k'. rewrite KernameMapFact.F.remove_o. - destruct KernameMap.E.eq_dec. eapply KernameOT.compare_eq in e0. subst k'. + destruct KernameMap.E.eq_dec. eapply Kername.compare_eq in e0. subst k'. - rewrite {}eq. induction frk. now cbn. rewrite of_global_env_cons //. depelim frΣ. simpl in H0 |- *. - rewrite KernameMapFact.F.add_neq_o //. intros c. eapply KernameOT.compare_eq in c. contradiction. + rewrite KernameMapFact.F.add_neq_o //. intros c. eapply Kername.compare_eq in c. contradiction. now apply IHfrk. - rewrite KernameMapFact.F.add_neq_o //. Qed. @@ -163,11 +160,11 @@ Module EnvMap. unfold repr, equal, remove, add. intros neq k''. rewrite KernameMapFact.F.remove_o. - destruct KernameMap.E.eq_dec. eapply KernameOT.compare_eq in e0. subst k'. - - rewrite KernameMapFact.F.add_neq_o //. intros c. eapply KernameOT.compare_eq in c. contradiction. + destruct KernameMap.E.eq_dec. eapply Kername.compare_eq in e0. subst k'. + - rewrite KernameMapFact.F.add_neq_o //. intros c. eapply Kername.compare_eq in c. contradiction. rewrite KernameMapFact.F.remove_o. destruct KernameMap.E.eq_dec => //. - elimtype False; apply n. now apply KernameOT.compare_eq. + elimtype False; apply n. now apply Kername.compare_eq. - rewrite !KernameMapFact.F.add_o //. destruct (KernameMap.E.eq_dec k k'') => //. rewrite KernameMapFact.F.remove_neq_o //. @@ -193,8 +190,7 @@ Module EnvMap. intros hin. red in eq. rewrite eq in hin. now eapply KernameMapFact.F.empty_in_iff in hin. - - cbn -[of_global_env]. - change (eq_kername k a.1) with (eqb k a.1). + - cbn -[of_global_env eqb]. destruct (eqb_spec k a.1). * subst. rewrite of_global_env_cons //. diff --git a/template-coq/theories/Environment.v b/template-coq/theories/Environment.v index befa9df8c..f43c316e6 100644 --- a/template-coq/theories/Environment.v +++ b/template-coq/theories/Environment.v @@ -504,7 +504,7 @@ Module Environment (T : Term). match Σ with | nil => None | d :: tl => - if eq_kername kn d.1 then Some d.2 + if kn == d.1 then Some d.2 else lookup_global tl kn end. diff --git a/template-coq/theories/EtaExpand.v b/template-coq/theories/EtaExpand.v index f796284d3..224de3870 100644 --- a/template-coq/theories/EtaExpand.v +++ b/template-coq/theories/EtaExpand.v @@ -5,9 +5,9 @@ Coq's conversion, the proof is essentially [eq_refl]. All dependencies are also expanded.*) -From Coq Require Import List PeanoNat Bool Ascii String Lia. +From Coq Require Import List PeanoNat Bool Lia. From MetaCoq.Template Require Export - utils.MCUtils (* Utility functions *) + utils (* Utility functions *) monad_utils (* Monadic notations *) uGraph (* The graph of universes *) BasicAst (* The basic AST structures *) @@ -18,8 +18,8 @@ From MetaCoq.Template Require Export UnivSubst (* Substitution of universe instances *) Typing (* Typing judgment *). -Open Scope string. Open Scope nat. +Open Scope bs. Import Template.Ast. Import ListNotations. @@ -307,7 +307,8 @@ Proof. * cbn in H. inversion H; subst. cbn. simpl_list. destruct l. -- cbn. econstructor; eauto. - -- cbn. eapply expanded_tConstruct_app; eauto. cbn. now rewrite H6. todo "lifting". + -- cbn. eapply expanded_tConstruct_app; eauto. cbn. now rewrite H6. + todo "lifting". * eapply expanded_mkApps_tConstruct. split. split. red. all: eauto. rewrite rev_map_spec. simpl_list. rewrite EE. lia. todo "induction on X0". + assert (Forall(fun t : term => expanded Σ0.1 (eta_expand Σ0.1.(declarations) t)) l). { diff --git a/template-coq/theories/Extraction.v b/template-coq/theories/Extraction.v index e068f4c12..0a005cd9f 100644 --- a/template-coq/theories/Extraction.v +++ b/template-coq/theories/Extraction.v @@ -7,14 +7,9 @@ From Coq Require Ascii Extraction ZArith NArith. From MetaCoq.Template Require Import utils Ast Reflect Induction. -From Coq Require Import FSets ExtrOcamlBasic ExtrOcamlString ExtrOCamlFloats +From Coq Require Import FSets ExtrOcamlBasic ExtrOCamlFloats ExtrOCamlInt63. From MetaCoq.Template Require Import MC_ExtrOCamlZPosInt. - -Extract Constant ascii_compare => - "fun x y -> match Char.compare x y with 0 -> Eq | x when x < 0 -> Lt | _ -> Gt". -Extract Constant Ascii.compare => - "fun x y -> match Char.compare x y with 0 -> Eq | x when x < 0 -> Lt | _ -> Gt". Extract Inductive Equations.Init.sigma => "( * )" ["(,)"]. Extract Constant Equations.Init.pr1 => "fst". diff --git a/template-coq/theories/Kernames.v b/template-coq/theories/Kernames.v index 6670cb881..5613bff39 100644 --- a/template-coq/theories/Kernames.v +++ b/template-coq/theories/Kernames.v @@ -1,15 +1,83 @@ (* Distributed under the terms of the MIT license. *) -From Coq Require Import Lia MSetList OrderedTypeAlt OrderedTypeEx Ascii String. -From MetaCoq.Template Require Import MCUtils BasicAst AstUtils. +From Coq Require Import Lia MSetList OrderedTypeAlt OrderedTypeEx FMapAVL FMapFacts MSetAVL MSetFacts MSetProperties. +From MetaCoq.Template Require Import utils. From Coq Require Import ssreflect. +From Equations Require Import Equations. +Local Open Scope string_scope2. Definition compare_ident := string_compare. +(** ** Reification of names ** *) + +(** [Comment taken from Coq's code] + - Id.t is the type of identifiers, that is morally a subset of strings which + only contains Unicode characters of the Letter kind (and a few more). + => [ident] + - Name.t is an ad-hoc variant of Id.t option allowing to handle optionally + named objects. + => [name] + - DirPath.t represents generic paths as sequences of identifiers. + => [dirpath] + - Label.t is an equivalent of Id.t made distinct for semantical purposes. + => [ident] + - ModPath.t are module paths. + => [modpath] + - KerName.t are absolute names of objects in Coq. + => [kername] + + And also : + - Constant.t => [kername] + - variable => [ident] + - MutInd.t => [kername] + - inductive => [inductive] + - constructor => [inductive * nat] + - Projection.t => [projection] + - GlobRef.t => global_reference + + Finally, we define the models of primitive types (uint63 and floats64). +*) + +Definition ident := string. (* e.g. nat *) +Definition qualid := string. (* e.g. Datatypes.nat *) + +(** Type of directory paths. Essentially a list of module identifiers. The + order is reversed to improve sharing. E.g. A.B.C is ["C";"B";"A"] *) +Definition dirpath := list ident. + Module IdentOT := StringOT. Module DirPathOT := ListOrderedType IdentOT. - + +#[global] Instance dirpath_eqdec : Classes.EqDec dirpath := _. + +Definition string_of_dirpath (dp : dirpath) : string := + String.concat "." (List.rev dp). + +(** The module part of the kernel name. + - MPfile is for toplevel libraries, i.e. .vo files + - MPbound are parameters of functors + - MPdot is for submodules +*) +Inductive modpath := +| MPfile (dp : dirpath) +| MPbound (dp : dirpath) (id : ident) (i : nat) +| MPdot (mp : modpath) (id : ident). +Derive NoConfusion for modpath. + +Fixpoint string_of_modpath (mp : modpath) : string := + match mp with + | MPfile dp => string_of_dirpath dp + | MPbound dp id n => string_of_dirpath dp ^ "." ^ id ^ "." ^ string_of_nat n + | MPdot mp id => string_of_modpath mp ^ "." ^ id + end. + +(** The absolute names of objects seen by kernel *) +Definition kername := modpath × ident. + +Definition string_of_kername (kn : kername) := + string_of_modpath kn.1 ^ "." ^ kn.2. + (* Eval compute in DirPathOT.compare ["foo"; "bar"] ["baz"]. *) @@ -119,6 +187,21 @@ End ModPathComp. Module ModPathOT := OrderedType_from_Alt ModPathComp. +Program Definition modpath_eq_dec (x y : modpath) : { x = y } + { x <> y } := + match ModPathComp.compare x y with + | Eq => left _ + | _ => right _ + end. +Next Obligation. + symmetry in Heq_anonymous. + now eapply ModPathComp.compare_eq in Heq_anonymous. +Qed. +Next Obligation. + rewrite ModPathOT.eq_refl in H. congruence. +Qed. + +#[global] Instance modpath_EqDec : Classes.EqDec modpath := { eq_dec := modpath_eq_dec }. + Module KernameComp. Definition t := kername. @@ -152,7 +235,7 @@ Module KernameComp. End KernameComp. -Module KernameOT. +Module Kername. Include KernameComp. Module OT := OrderedType_from_Alt KernameComp. @@ -191,19 +274,44 @@ Module KernameOT. now apply irreflexivity in H. Qed. - Definition eq_dec : forall x y, {eq x y} + {~ eq x y}. - Proof. - intros k k'. apply kername_eq_dec. + Definition eqb kn kn' := + match compare kn kn' with + | Eq => true + | _ => false + end. + + #[global, program] Instance reflect_kername : ReflectEq kername := { + eqb := eqb + }. + Next Obligation. + unfold eqb. destruct compare eqn:e; constructor. + - now apply compare_eq in e. + -intros e'; subst. now rewrite OT.eq_refl in e. + -intros e'; subst. now rewrite OT.eq_refl in e. Defined. + + Definition eq_dec : forall (x y : t), { x = y } + { x <> y } := Classes.eq_dec. + +End Kername. + +Module KernameMap := FMapAVL.Make Kername.OT. +Module KernameMapFact := FMapFacts.WProperties KernameMap. -End KernameOT. +Notation eq_kername := (eqb (A:=kername)) (only parsing). + +Lemma eq_kername_refl kn : eq_kername kn kn. +Proof. + eapply ReflectEq.eqb_refl. +Qed. + +Definition eq_constant := eq_kername. (* Local Open Scope string_scope.*) (* Eval compute in KernameOT.compare (MPfile ["fdejrkjl"], "A") (MPfile ["lfrk;k"], "B"). *) -Module KernameSet := MSetAVL.Make KernameOT. -Module KernameSetFact := MSetFacts.WFactsOn KernameOT KernameSet. -Module KernameSetProp := MSetProperties.WPropertiesOn KernameOT KernameSet. +Module KernameSet := MSetAVL.Make Kername. +Module KernameSetFact := MSetFacts.WFactsOn Kername KernameSet. +Module KernameSetProp := MSetProperties.WPropertiesOn Kername KernameSet. Lemma knset_in_fold_left {A} kn f (l : list A) acc : KernameSet.In kn (fold_left (fun acc x => KernameSet.union (f x) acc) l acc) <-> @@ -220,3 +328,96 @@ Proof. destruct ina as [<-|ina']; intuition auto. right. now exists a'. Qed. + +(** Designation of a (particular) inductive type. *) +Record inductive : Set := mkInd { inductive_mind : kername ; + inductive_ind : nat }. +Arguments mkInd _%bs _%nat. + +Derive NoConfusion for inductive. + +Definition string_of_inductive (i : inductive) := + string_of_kername (inductive_mind i) ^ "," ^ string_of_nat (inductive_ind i). + +Definition eq_inductive_def i i' := + let 'mkInd i n := i in + let 'mkInd i' n' := i' in + eqb (i, n) (i', n'). + +#[global, program] Instance reflect_eq_inductive : ReflectEq inductive := { + eqb := eq_inductive_def +}. +Next Obligation. + destruct x as [m n], y as [m' n']; cbn -[eqb]. + case: eqb_spec ; nodec. + cbn. constructor. noconf p; reflexivity. +Qed. + +Notation eq_inductive := (eqb (A:=inductive)). + +Lemma eq_inductive_refl i : eq_inductive i i. +Proof. + apply ReflectEq.eqb_refl. +Qed. + +Definition projection : Set := inductive * nat (* params *) * nat (* argument *). + +Definition eq_projection (p p' : projection) := eqb p p'. + +#[global, program] Instance reflect_eq_projection : ReflectEq projection := { + eqb := eq_projection +}. +Next Obligation. + unfold eq_projection. + case: eqb_spec ; nodec. + cbn. now constructor. +Qed. + +Lemma eq_projection_refl i : eq_projection i i. +Proof. + apply ReflectEq.eqb_refl. +Qed. + +(** Kernel declaration references [global_reference] *) +Inductive global_reference := +| VarRef : ident -> global_reference +| ConstRef : kername -> global_reference +| IndRef : inductive -> global_reference +| ConstructRef : inductive -> nat -> global_reference. + +Derive NoConfusion for global_reference. + +Definition string_of_gref gr : string := + match gr with + | VarRef v => v + | ConstRef s => string_of_kername s + | IndRef (mkInd s n) => + "Inductive " ^ string_of_kername s ^ " " ^ (string_of_nat n) + | ConstructRef (mkInd s n) k => + "Constructor " ^ string_of_kername s ^ " " ^ (string_of_nat n) ^ " " ^ (string_of_nat k) + end. + +Definition gref_eqb (x y : global_reference) : bool := + match x, y with + | VarRef i, VarRef i' => eqb i i' + | ConstRef c, ConstRef c' => eqb c c' + | IndRef i, IndRef i' => eqb i i' + | ConstructRef i k, ConstructRef i' k' => eqb i i' && eqb k k' + | _, _ => false + end. + +#[global, program] Instance grep_reflect_eq : ReflectEq global_reference := + {| eqb := gref_eqb |}. +Next Obligation. + destruct x, y; cbn; try constructor; try congruence. + - destruct (eqb_spec i i0); constructor; subst; auto; congruence. + - destruct (eqb_spec k k0); constructor; subst; auto; congruence. + - destruct (eqb_spec i i0); constructor; subst; auto; congruence. + - destruct (eqb_spec i i0); subst; cbn; auto; try constructor; try congruence. + destruct (eqb_spec n n0); constructor; subst; congruence. +Defined. + +Definition gref_eq_dec (gr gr' : global_reference) : {gr = gr'} + {~ gr = gr'} := Classes.eq_dec gr gr'. + +#[global] Hint Resolve Kername.eq_dec : eq_dec. + diff --git a/template-coq/theories/Pretty.v b/template-coq/theories/Pretty.v index 5ecda6ab0..7cda8f889 100644 --- a/template-coq/theories/Pretty.v +++ b/template-coq/theories/Pretty.v @@ -3,7 +3,6 @@ From MetaCoq Require Import utils Ast AstUtils Environment LiftSubst Universes. (** * Pretty printing *) - Section print_term. Context (Σ : global_env_ext). @@ -20,7 +19,7 @@ Section print_term. end. Definition is_fresh (Γ : list ident) (id : ident) := - List.forallb (fun id' => negb (ident_eq id id')) Γ. + List.forallb (fun id' => negb (eqb id id')) Γ. Definition lookup_ind_decl ind i := match lookup_env Σ ind with @@ -43,7 +42,7 @@ Section print_term. | tConst c u => "x" | tInd (mkInd i k) u => match lookup_ind_decl i k with - | Some body => substring 0 1 (body.(ind_name)) + | Some body => String.substring 0 1 (body.(ind_name)) | None => "X" end (* | tInt _ => "i" *) @@ -110,15 +109,28 @@ Section print_term. end in aux Γ (MCList.rev Γ'). - Definition print_defs (print_term : list ident -> bool -> term -> string) +End print_term. + +Module PrintTermTree. + Import bytestring.Tree. + Infix "^" := append. + + Section env. + Context (Σ : global_env_ext). + + Definition print_def {A} (f : A -> t) (g : A -> t) (def : def A) := + string_of_name (binder_name (dname def)) ^ " { struct " ^ string_of_nat (rarg def) ^ " }" ^ + " : " ^ f (dtype def) ^ " := " ^ nl ^ g (dbody def). + + Definition print_defs (print_term : list ident -> bool -> term -> t) Γ (defs : mfixpoint term) := let ctx' := fix_context defs in - print_list (print_def (print_term Γ true) (print_term (fresh_names Γ ctx') true)) + print_list (print_def (print_term Γ true) (print_term (fresh_names Σ Γ ctx') true)) (nl ^ " with ") defs. (* TODO: SPROP: we ignore relevance on printing, maybe add print config? *) - Fixpoint print_term (Γ : list ident) (top : bool) (t : term) {struct t} := + Fixpoint print_term (Γ : list ident) (top : bool) (t : term) {struct t} : Tree.t := match t with | tRel n => match nth_error Γ n with @@ -130,7 +142,7 @@ Section print_term. | tSort s => string_of_sort s | tCast c k t => parens top (print_term Γ true c ^ ":" ^ print_term Γ true t) | tProd na dom codom => - let na' := (fresh_name Γ na.(binder_name) (Some dom)) in + let na' := (fresh_name Σ Γ na.(binder_name) (Some dom)) in if (noccur_between 0 1 codom) then parens top (print_term Γ false dom ^ " → " ^ print_term (na' :: Γ) true codom) @@ -138,11 +150,11 @@ Section print_term. ("∀ " ^ na' ^ " : " ^ print_term Γ false dom ^ ", " ^ print_term (na' :: Γ) true codom) | tLambda na dom body => - let na' := (fresh_name Γ na.(binder_name) (Some dom)) in + let na' := (fresh_name Σ Γ na.(binder_name) (Some dom)) in parens top ("fun " ^ na' ^ " : " ^ print_term Γ true dom ^ " ⇒ " ^ print_term (na' :: Γ) true body) | tLetIn na def dom body => - let na' := (fresh_name Γ na.(binder_name) (Some dom)) in + let na' := (fresh_name Σ Γ na.(binder_name) (Some dom)) in parens top ("let " ^ na' ^ " : " ^ print_term Γ true dom ^ " := " ^ print_term Γ true def ^ " in " ^ nl ^ print_term (na' :: Γ) true body) @@ -150,13 +162,13 @@ Section print_term. parens top (print_term Γ false f ^ " " ^ print_list (print_term Γ false) " " l) | tConst c u => string_of_kername c ^ print_universe_instance u | tInd (mkInd i k) u => - match lookup_ind_decl i k with + match lookup_ind_decl Σ i k with | Some oib => oib.(ind_name) ^ print_universe_instance u | None => "UnboundInd(" ^ string_of_inductive (mkInd i k) ^ "," ^ string_of_universe_instance u ^ ")" end | tConstruct (mkInd i k as ind) l u => - match lookup_ind_decl i k with + match lookup_ind_decl Σ i k with | Some oib => match nth_error oib.(ind_ctors) l with | Some cb => cb.(cstr_name) ^ print_universe_instance u @@ -169,7 +181,7 @@ Section print_term. ^ string_of_universe_instance u ^ ")" end | tCase {| ci_ind := mkInd mind i as ind; ci_npar := pars |} p t brs => - match lookup_ind_decl mind i with + match lookup_ind_decl Σ mind i with | Some oib => match build_return_context ind oib p with | None => @@ -178,17 +190,17 @@ Section print_term. string_of_list (pretty_string_of_branch string_of_term) brs ^ ")" | Some Γret => - let Γret := fresh_names Γ Γret in + let Γret := fresh_names Σ Γ Γret in let ret_binders := firstn #|pcontext p| Γret in let (as_name, indices) := (hd "_" ret_binders, MCList.rev (tail ret_binders)) in - let in_args := repeat "_" #|pparams p| ++ indices in - let in_str := oib.(ind_name) ^ String.concat "" (map (fun a => " " ^ a) in_args) in + let in_args := (repeat "_" #|pparams p| ++ indices)%list in + let in_str := oib.(ind_name) ^ concat "" (map (fun a : bytestring.string => " " ^ a) in_args) in let fix print_branch Γ names prbr {struct names} := match names with | [] => "⇒ " ^ prbr Γ | na :: l => - let na' := (fresh_name Γ na.(binder_name) None) in + let na' := (fresh_name Σ Γ na.(binder_name) None) in na' ^ " " ^ print_branch (na' :: Γ) l prbr end in @@ -210,7 +222,7 @@ Section print_term. string_of_list (pretty_string_of_branch string_of_term) brs ^ ")" end | tProj (mkInd mind i as ind, pars, k) c => - match lookup_ind_decl mind i with + match lookup_ind_decl Σ mind i with | Some oib => match nth_error oib.(ind_projs) k with | Some (na, _) => print_term Γ false c ^ ".(" ^ na ^ ")" @@ -234,74 +246,83 @@ Section print_term. | tFloat f => "Float(" ^ string_of_float f ^ ")" *) end. -End print_term. - -Definition pr_context_decl (Σ : global_env_ext) Γ (c : context_decl) : ident * string := - match c with - | {| decl_name := na; decl_type := ty; decl_body := None |} => - let na' := (fresh_name Σ Γ na.(binder_name) (Some ty)) in - (na', ("(" ++ na' ++ " : " ++ print_term Σ Γ true ty ++ ")")%string) - | {| decl_name := na; decl_type := ty; decl_body := Some b |} => - let na' := (fresh_name Σ Γ na.(binder_name) (Some ty)) in - (na', ("(" ++ na' ++ " : " ++ print_term Σ Γ true ty ++ " := " ++ - print_term Σ Γ true b ++ ")")%string) - end. + Definition pr_context_decl Γ (c : context_decl) : ident * t := + match c with + | {| decl_name := na; decl_type := ty; decl_body := None |} => + let na' := (fresh_name Σ Γ na.(binder_name) (Some ty)) in + (na', ("(" ^ na' ^ " : " ^ print_term Γ true ty ^ ")")) + | {| decl_name := na; decl_type := ty; decl_body := Some b |} => + let na' := (fresh_name Σ Γ na.(binder_name) (Some ty)) in + (na', ("(" ^ na' ^ " : " ^ print_term Γ true ty ^ " := " ^ + print_term Γ true b ^ ")")) + end. -Fixpoint print_context Σ Γ Δ := - match Δ with - | [] => (Γ, ""%string) - | d :: decls => - let '(Γ, s) := print_context Σ Γ decls in - let '(na, s') := pr_context_decl Σ Γ d in - match decls with - | [] => (na :: Γ, (s ++ s')%string) - | _ => (na :: Γ, (s ++ " " ++ s')%string) - end - end. + Fixpoint print_context Γ Δ : list ident * t := + match Δ with + | [] => (Γ, "" : t) + | d :: decls => + let '(Γ, s) := print_context Γ decls in + let '(na, s') := pr_context_decl Γ d in + match decls with + | [] => (na :: Γ, s ^ s') + | _ => (na :: Γ, s ^ " " ^ s') + end + end. -Definition print_one_cstr Σ Γ (mib : mutual_inductive_body) (c : constructor_body) : string := - let '(Γargs, s) := print_context Σ Γ c.(cstr_args) in - c.(cstr_name) ++ " : " ++ s ++ "_" ++ print_list (print_term Σ Γargs true) " " c.(cstr_indices). + Definition print_one_cstr Γ (mib : mutual_inductive_body) (c : constructor_body) : t := + let '(Γargs, s) := print_context Γ c.(cstr_args) in + c.(cstr_name) ^ " : " ^ s ^ "_" ^ print_list (print_term Γargs true) " " c.(cstr_indices). -Definition print_one_ind (short : bool) Σ Γ (mib : mutual_inductive_body) (oib : one_inductive_body) : string := - let '(Γpars, spars) := print_context Σ Γ mib.(ind_params) in - let '(Γinds, sinds) := print_context Σ Γpars oib.(ind_indices) in - oib.(ind_name) ++ spars ++ sinds ++ print_term Σ Γinds true (tSort oib.(ind_sort)) ++ ":=" ++ nl ++ - if short then "..." - else print_list (print_one_cstr Σ Γpars mib) nl oib.(ind_ctors). + Definition print_one_ind (short : bool) Γ (mib : mutual_inductive_body) (oib : one_inductive_body) : t := + let '(Γpars, spars) := print_context Γ mib.(ind_params) in + let '(Γinds, sinds) := print_context Γpars oib.(ind_indices) in + oib.(ind_name) ^ spars ^ sinds ^ print_term Γinds true (tSort oib.(ind_sort)) ^ ":=" ^ nl ^ + if short then "..." + else print_list (print_one_cstr Γpars mib) nl oib.(ind_ctors). + End env. -Fixpoint print_env_aux (short : bool) (prefix : nat) (Σ : global_env) (acc : string) := - match prefix with - | 0 => match Σ.(declarations) with [] => acc | _ => ("..." ++ nl ++ acc)%string end - | S n => - let univs := Σ.(Env.universes) in - match Σ.(declarations) with - | [] => acc - | (kn, InductiveDecl mib) :: Σ => - let Σ' := ({| Env.universes := univs; declarations := Σ |}, mib.(ind_universes)) in - let names := fresh_names Σ' [] (arities_context mib.(ind_bodies)) in - print_env_aux short n Σ'.1 - ("Inductive " ++ - print_list (print_one_ind short Σ' names mib) nl mib.(ind_bodies) ++ "." ++ - nl ++ acc)%string - | (kn, ConstantDecl cb) :: Σ => - let Σ' := ({| Env.universes := univs; declarations := Σ |}, cb.(cst_universes)) in - print_env_aux short n Σ'.1 - ((match cb.(cst_body) with - | Some _ => "Definition " - | None => "Axiom " - end) ++ string_of_kername kn ++ " : " ++ print_term Σ' nil true cb.(cst_type) ++ - match cb.(cst_body) with - | Some b => - if short then ("..." ++ nl)%string - else (" := " ++ nl ++ print_term Σ' nil true b ++ "." ++ nl) - | None => "." - end ++ acc)%string - end - end. + Fixpoint print_env_aux (short : bool) (prefix : nat) (Σ : global_env) (acc : t) : t := + match prefix with + | 0 => match Σ.(declarations) with [] => acc | _ => ("..." ^ nl ^ acc) end + | S n => + let univs := Σ.(Env.universes) in + match Σ.(declarations) with + | [] => acc + | (kn, InductiveDecl mib) :: Σ => + let Σ' := ({| Env.universes := univs; declarations := Σ |}, mib.(ind_universes)) in + let names := fresh_names Σ' [] (arities_context mib.(ind_bodies)) in + print_env_aux short n Σ'.1 + ("Inductive " ^ + print_list (print_one_ind Σ' short names mib) nl mib.(ind_bodies) ^ "." ^ + nl ^ acc) + | (kn, ConstantDecl cb) :: Σ => + let Σ' := ({| Env.universes := univs; declarations := Σ |}, cb.(cst_universes)) in + print_env_aux short n Σ'.1 + ((match cb.(cst_body) with + | Some _ => "Definition " + | None => "Axiom " + end) ^ string_of_kername kn ^ " : " ^ print_term Σ' nil true cb.(cst_type) ^ + match cb.(cst_body) with + | Some b => + if short then ("..." ^ nl) + else (" := " ^ nl ^ print_term Σ' nil true b ^ "." ^ nl) + | None => "." + end ^ acc) + end + end. + + Definition print_env (short : bool) (prefix : nat) Σ := + print_env_aux short prefix Σ (Tree.string ""). + + Definition print_program (short : bool) (prefix : nat) (p : program) : t := + print_env short prefix (fst p) ^ nl ^ print_term (empty_ext (fst p)) nil true (snd p). + +End PrintTermTree. + +Definition print_term Σ Γ top := Tree.to_string ∘ PrintTermTree.print_term Σ Γ top. -Definition print_env (short : bool) (prefix : nat) Σ := print_env_aux short prefix Σ EmptyString. +Definition print_env (short : bool) (prefix : nat) Σ := + Tree.to_string (PrintTermTree.print_env short prefix Σ). Definition print_program (short : bool) (prefix : nat) (p : program) : string := - print_env short prefix (fst p) ++ nl ++ - print_term (empty_ext (fst p)) nil true (snd p). + Tree.to_string (PrintTermTree.print_program short prefix p). diff --git a/template-coq/theories/Reflect.v b/template-coq/theories/Reflect.v index 468a5ab1a..482784eaa 100644 --- a/template-coq/theories/Reflect.v +++ b/template-coq/theories/Reflect.v @@ -5,108 +5,8 @@ From MetaCoq.Template Require Import utils BasicAst Universes. Require Import ssreflect. From Equations Require Import Equations. -(* Some reflection / EqDec lemmata *) - -Class ReflectEq A := { - eqb : A -> A -> bool ; - eqb_spec : forall x y : A, reflect (x = y) (eqb x y) -}. - -Lemma eqb_eq {A} `{ReflectEq A} (x y : A) : eqb x y -> x = y. -Proof. - elim: eqb_spec; auto. - discriminate. -Qed. - -#[global] Instance ReflectEq_EqDec : - forall A, ReflectEq A -> EqDec A. -Proof. - intros A [eqb h] x y. - destruct (h x y). - - left. assumption. - - right. assumption. -Defined. - -Definition eq_dec_to_bool {A} `{EqDec A} x y := - match eq_dec x y with - | left _ => true - | right _ => false - end. - -#[global] (* Not an instance to avoid loops? *) -Lemma EqDec_ReflectEq : forall A `{EqDec A}, ReflectEq A. -Proof. - intros A h. - unshelve econstructor. - - eapply eq_dec_to_bool. - - unfold eq_dec_to_bool. - intros x y. destruct (eq_dec x y). - all: constructor ; assumption. -Defined. - -Ltac nodec := - let bot := fresh "bot" in - try solve [ constructor ; intro bot ; inversion bot ; subst ; tauto ]. - -Definition eq_option {A} (eqA : A -> A -> bool) (u v : option A) : bool := - match u, v with - | Some u, Some v => eqA u v - | None, None => true - | _, _ => false - end. - -#[global] Instance reflect_option : forall {A}, ReflectEq A -> ReflectEq (option A). -Proof. - intros A RA. refine {| eqb := eq_option eqb |}. - intros x y. destruct x, y. - all: cbn. - all: try solve [ constructor ; easy ]. - destruct (eqb_spec a a0) ; nodec. - constructor. f_equal. assumption. -Defined. - -Fixpoint eq_list {A} (eqA : A -> A -> bool) (l l' : list A) : bool := - match l, l' with - | a :: l, a' :: l' => - if eqA a a' then eq_list eqA l l' - else false - | [], [] => true - | _, _ => false - end. - -#[global] Instance reflect_list : forall {A}, ReflectEq A -> ReflectEq (list A). -Proof. - intros A RA. refine {| eqb := eq_list eqb |}. - intro x. induction x ; intro y ; destruct y. - - cbn. constructor. reflexivity. - - cbn. constructor. discriminate. - - cbn. constructor. discriminate. - - cbn. destruct (eqb_spec a a0) ; nodec. - destruct (IHx y) ; nodec. - subst. constructor. reflexivity. -Defined. - -#[global] Program Instance reflect_string : ReflectEq string := { - eqb := eq_string -}. -Next Obligation. - rename x into s, y into s'. - destruct (string_dec s s'). - - subst. rewrite eq_string_refl. constructor. reflexivity. - - assert (string_compare s s' <> Eq). - { intro bot. apply n. apply string_compare_eq. assumption. } - unfold eq_string. destruct (string_compare s s'). - + tauto. - + constructor. assumption. - + constructor. assumption. -Defined. - -#[global] Instance reflect_nat : ReflectEq nat := { - eqb_spec := Nat.eqb_spec -}. - -#[program,global] Instance reflect_prim_int : ReflectEq Numbers.Cyclic.Int63.Uint63.int := - { eqb := Numbers.Cyclic.Int63.Uint63.eqb }. +#[program,global] Instance reflect_prim_int : ReflectEq Numbers.Cyclic.Int63.Int63.int := + { eqb := Numbers.Cyclic.Int63.Int63.eqb }. Next Obligation. destruct (Uint63.eqb x y) eqn:eq; constructor. now apply (Numbers.Cyclic.Int63.Uint63.eqb_spec x y) in eq. @@ -122,34 +22,12 @@ Instance reflect_prim_float : ReflectEq PrimFloat.float := { eqb x y := eqb (ReflectEq := EqDec_ReflectEq SpecFloat.spec_float) (FloatOps.Prim2SF x) (FloatOps.Prim2SF y) }. Next Obligation. intros. cbn -[eqb]. - destruct (eqb_spec (ReflectEq := EqDec_ReflectEq SpecFloat.spec_float) (FloatOps.Prim2SF x) (FloatOps.Prim2SF y)); constructor. + destruct (eqb_spec (ReflectEq := EqDec_ReflectEq SpecFloat.spec_float) (FloatOps.Prim2SF x) (FloatOps.Prim2SF y)) as [H|H]; constructor. now apply FloatAxioms.Prim2SF_inj. - intros e; apply n. rewrite e. + intros e; apply H. rewrite e. reflexivity. Qed. -Definition eq_level l1 l2 := - match l1, l2 with - | Level.lzero, Level.lzero => true - | Level.Level s1, Level.Level s2 => eqb s1 s2 - | Level.Var n1, Level.Var n2 => eqb n1 n2 - | _, _ => false - end. - -#[global, program] Instance reflect_level : ReflectEq Level.t := { - eqb := eq_level -}. -Next Obligation. - destruct x, y. - all: unfold eq_level. - all: try solve [ constructor ; reflexivity ]. - all: try solve [ constructor ; discriminate ]. - - destruct (eqb_spec s s0) ; nodec. - constructor. f_equal. assumption. - - destruct (eqb_spec n n0) ; nodec. - constructor. subst. reflexivity. -Defined. - Definition eq_prop_level l1 l2 := match l1, l2 with | PropLevel.lProp, PropLevel.lProp => true @@ -185,47 +63,6 @@ Next Obligation. constructor; cong. Defined. -Definition eq_prod {A B} (eqA : A -> A -> bool) (eqB : B -> B -> bool) x y := - let '(a1, b1) := x in - let '(a2, b2) := y in - if eqA a1 a2 then eqB b1 b2 - else false. - -Local Obligation Tactic := idtac. -#[global, program] Instance reflect_prod : forall {A B}, ReflectEq A -> ReflectEq B -> ReflectEq (A * B) := { - eqb := eq_prod eqb eqb -}. -Next Obligation. - intros A B RA RB [x y] [u v]. - unfold eq_prod. - destruct (eqb_spec x u) ; nodec. - destruct (eqb_spec y v) ; nodec. - subst. constructor. reflexivity. -Defined. - -Lemma eq_prod_refl : - forall A B (eqA : A -> A -> bool) (eqB : B -> B -> bool), - (forall a, eqA a a) -> - (forall b, eqB b b) -> - forall p, eq_prod eqA eqB p p. -Proof. - intros A B eqA eqB eqA_refl eqB_refl [a b]. - simpl. rewrite eqA_refl. apply eqB_refl. -Qed. - -Definition eq_bool b1 b2 : bool := - if b1 then b2 else negb b2. - -#[global, program] Instance reflect_bool : ReflectEq bool := { - eqb := eq_bool -}. -Next Obligation. - intros x y. unfold eq_bool. - destruct x, y. - all: constructor. - all: try reflexivity. - all: discriminate. -Defined. Definition eq_name na nb := match na, nb with @@ -278,26 +115,6 @@ Next Obligation. constructor; destruct x, y; simpl in *; cong. Defined. -#[global, program] Instance reflect_kername : ReflectEq kername := { - eqb := eq_kername -}. -Next Obligation. - intros; unfold eq_kername; destruct kername_eq_dec; now constructor. -Qed. - - -#[global, program] Instance reflect_inductive : ReflectEq inductive := { - eqb := eq_inductive -}. -Next Obligation. - intros i i'. destruct i as [m n], i' as [m' n']; cbn. - change (eq_kername m m') with (eqb m m'). - change (n =? n') with (eqb n n'). - destruct (eqb_spec m m') ; nodec. - destruct (eqb_spec n n') ; nodec. - cbn. constructor. subst. reflexivity. -Defined. - Definition eq_def {A} `{ReflectEq A} (d1 d2 : def A) : bool := match d1, d2 with | mkdef n1 t1 b1 a1, mkdef n2 t2 b2 a2 => @@ -348,44 +165,8 @@ Proof. now inversion 1. intros ->. f_equal. apply uip. Qed. -(* move in Universes.v ?? *) -#[global] Instance eq_dec_UnivExpr : EqDec UnivExpr.t. -Proof. intros e e'. decide equality; apply eq_dec. Qed. - -#[global] Instance eq_dec_univ0 : EqDec Universe.nonEmptyUnivExprSet. -Proof. - intros u v. - assert (H : {UnivExprSet.elements u = UnivExprSet.elements v} - + {~ UnivExprSet.elements u = UnivExprSet.elements v}). { - decide equality. apply eq_dec. } - destruct H as [H|H]; [left; now apply eq_universe_iff' in H|right]. - intro X; apply H; now apply eq_universe_iff' in X. -Defined. - -#[global] Instance eq_dec_univ : EqDec Universe.t. -Proof. - red. decide equality. - apply eq_dec_univ0. -Defined. - -#[global] Instance reflect_eq_univ : ReflectEq Universe.t := EqDec_ReflectEq _. - #[global] Instance reflect_case_info : ReflectEq case_info := EqDec_ReflectEq case_info. -Definition eq_sig_true {A f} `{ReflectEq A} (x y : { z : A | f z = true }) : bool := - let '(exist x hx) := x in - let '(exist y hy) := y in - eqb x y. - -#[global, program] Instance reflect_sig_true {A f} `{ReflectEq A} : ReflectEq ({ z : A | f z = true }) := { - eqb := eq_sig_true -}. -Next Obligation. - intros A f RA. intros [x hx] [y hy]. simpl. - destruct (eqb_spec x y) ; nodec. subst. - constructor. pose proof (uip hx hy). subst. reflexivity. -Defined. - Derive NoConfusion NoConfusionHom for sig. Derive NoConfusion NoConfusionHom for prod. @@ -435,12 +216,11 @@ Proof. cong. Defined. -#[global] Instance Z_as_int : ReflectEq Int.Z_as_Int.t. -Proof. - refine {| eqb := Z.eqb |}. - apply Z.eqb_spec. -Defined. - +#[global, program] Instance Z_as_int : ReflectEq Int.Z_as_Int.t := + { eqb := Z.eqb }. +Next Obligation. + apply (reflect_reflectProp_2 Z.eqb_spec). +Qed. Scheme level_lt_ind_dep := Induction for Level.lt_ Sort Prop. Scheme constraint_type_lt_ind_dep := Induction for ConstraintType.lt_ Sort Prop. diff --git a/template-coq/theories/ReflectAst.v b/template-coq/theories/ReflectAst.v index e3f942703..0c36570de 100644 --- a/template-coq/theories/ReflectAst.v +++ b/template-coq/theories/ReflectAst.v @@ -30,9 +30,9 @@ Local Ltac term_dec_tac term_dec := | x : list aname, y : list aname |- _ => fcase (eq_dec x y) | n : nat, m : nat |- _ => fcase (Nat.eq_dec n m) - | i : ident, i' : ident |- _ => fcase (string_dec i i') - | i : kername, i' : kername |- _ => fcase (kername_eq_dec i i') - | i : string, i' : kername |- _ => fcase (string_dec i i') + | i : ident, i' : ident |- _ => fcase (eq_dec i i') + | i : kername, i' : kername |- _ => fcase (eq_dec i i') + | i : string, i' : kername |- _ => fcase (eq_dec i i') | n : name, n' : name |- _ => fcase (eq_dec n n') | n : aname, n' : aname |- _ => fcase (eq_dec n n') | i : inductive, i' : inductive |- _ => fcase (eq_dec i i') diff --git a/template-coq/theories/Typing.v b/template-coq/theories/Typing.v index be385100a..9c93b544b 100644 --- a/template-coq/theories/Typing.v +++ b/template-coq/theories/Typing.v @@ -682,7 +682,7 @@ Definition isCoFinite (r : recursivity_kind) := Definition check_recursivity_kind (Σ : global_env) ind r := match lookup_env Σ ind with - | Some (InductiveDecl mib) => Reflect.eqb mib.(ind_finite) r + | Some (InductiveDecl mib) => eqb mib.(ind_finite) r | _ => false end. @@ -708,7 +708,7 @@ Definition wf_fixpoint (Σ : global_env) mfix := | Some (ind :: inds) => (* Check that mutually recursive fixpoints are all on the same mututal inductive block *) - forallb (Reflect.eqb ind) inds && + forallb (eqb ind) inds && check_recursivity_kind Σ ind Finite | _ => false end. @@ -731,7 +731,7 @@ Definition wf_cofixpoint (Σ : global_env) mfix := | Some (ind :: inds) => (* Check that mutually recursive cofixpoints are all producing coinductives in the same mututal coinductive block *) - forallb (Reflect.eqb ind) inds && + forallb (eqb ind) inds && check_recursivity_kind Σ ind CoFinite | _ => false end. @@ -1539,12 +1539,11 @@ Lemma lookup_on_global_env `{checker_flags} {P} {Σ : global_env} {c decl} : Proof. unfold on_global_env. destruct Σ as [univs Σ]; cbn. intros [cu ond]. - induction ond; simpl in * => //. - unfold eq_kername. destruct kername_eq_dec; subst. - - intros [= ->]. - exists ({| universes := univs; declarations := Σ |}, udecl). + induction ond; cbn in * => //. + case: eqb_specT => [-> [= <-]| ne]. + - exists ({| universes := univs; declarations := Σ |}, udecl). split; try constructor; tas. - cbn. now split => //; exists [(kn, decl)]. + cbn. now split => //; exists [(kn, d)]. - intros hl. destruct (IHond hl) as [[Σ' udecl'] [ong [[equ ext] ond']]]. exists (Σ', udecl'). cbn in equ |- *. subst univs. repeat split; cbn; auto; try apply ong. diff --git a/template-coq/theories/TypingWf.v b/template-coq/theories/TypingWf.v index c556557d2..9188fab05 100644 --- a/template-coq/theories/TypingWf.v +++ b/template-coq/theories/TypingWf.v @@ -196,21 +196,17 @@ Proof. destruct Σ as [univs Σ]; cbn in *. rewrite /lookup_env /on_global_env /=. induction Σ in univs, Σ', k, d |- *; cbn => //. - destruct eq_kername eqn:eqk. + destruct (eqb_spec k a.1) as [e|e]. * move=> wfΣ' [=]. intros <- ext. destruct ext as [univeq [Σ'' eq]] => /=. cbn in *. subst univs. rewrite eq in wfΣ'. destruct Σ' as [univs' Σ']; cbn in *. subst Σ'. destruct wfΣ' as [cu wfΣ']. induction Σ''. - + cbn. now rewrite eqk. - + cbn. destruct (eq_kername k a0.1) eqn:eqk' => //. - change (eq_kername k a.1) with (Reflect.eqb k a.1) in eqk. - change (eq_kername k a0.1) with (Reflect.eqb k a0.1) in eqk'. - eapply eqb_eq in eqk. apply eqb_eq in eqk'. subst. - apply on_global_decls_extends_not_fresh in wfΣ'; eauto. - apply IHΣ''. - now depelim wfΣ'. + + cbn. now rewrite e eq_kername_refl. + + cbn. destruct (eqb_spec k a0.1) => //. subst. + { apply on_global_decls_extends_not_fresh in wfΣ'; eauto. } + subst. apply IHΣ''. now depelim wfΣ'. * intros HΣ' Hl [univeq [Σ'' eq]]; cbn in *. subst univs. rewrite eq in HΣ'. destruct HΣ'. eapply IHΣ; tea. split; eauto. now rewrite eq. diff --git a/template-coq/theories/Universes.v b/template-coq/theories/Universes.v index c6ecb6659..abfe621a3 100644 --- a/template-coq/theories/Universes.v +++ b/template-coq/theories/Universes.v @@ -1,6 +1,6 @@ From Coq Require Import MSetList MSetAVL MSetFacts MSetProperties MSetDecide. -From MetaCoq.Template Require Import utils BasicAst config. From Equations Require Import Equations. +From MetaCoq.Template Require Import utils BasicAst config. Require Import ssreflect. Local Open Scope nat_scope. @@ -40,13 +40,14 @@ Inductive allowed_eliminations : Set := | IntoAny. Derive NoConfusion EqDec for allowed_eliminations. + (** Levels are Set or Level or Var *) Module Level. Inductive t_ : Set := | lzero | Level (_ : string) | Var (_ : nat) (* these are debruijn indices *). - Derive NoConfusion EqDec for t_. + Derive NoConfusion for t_. Definition t := t_. @@ -84,11 +85,6 @@ Module Level. Definition eq : t -> t -> Prop := eq. Definition eq_equiv : Equivalence eq := _. - Definition eq_dec (l1 l2 : t) : {l1 = l2}+{l1 <> l2}. - Proof. - decide equality. apply string_dec. apply Peano_dec.eq_nat_dec. - Defined. - Inductive lt_ : t -> t -> Prop := | ltSetLevel s : lt_ lzero (Level s) | ltSetVar n : lt_ lzero (Var n) @@ -137,7 +133,7 @@ Module Level. Global Instance eqb_refl : Reflexive eqb. Proof. intros []; unfold eqb; cbnr. - - rewrite (ssreflect.iffRL (string_compare_eq s s)). all: auto. reflexivity. + - rewrite (ssreflect.iffRL (string_compare_eq _ _)). all: auto. reflexivity. - rewrite Nat.compare_refl. reflexivity. Qed. @@ -145,7 +141,7 @@ Module Level. Proof. destruct l, l'; cbn; try constructor; try reflexivity; try discriminate. - apply iff_reflect. unfold eqb; cbn. - destruct (CompareSpec_string s s0); split; intro HH; + destruct (CompareSpec_string t0 t1); split; intro HH; try reflexivity; try discriminate; try congruence. all: inversion HH; subst; now apply irreflexivity in H. - apply iff_reflect. unfold eqb; cbn. @@ -154,7 +150,31 @@ Module Level. all: inversion HH; subst; now apply Nat.lt_irrefl in H. Qed. + Definition eq_level l1 l2 := + match l1, l2 with + | Level.lzero, Level.lzero => true + | Level.Level s1, Level.Level s2 => ReflectEq.eqb s1 s2 + | Level.Var n1, Level.Var n2 => ReflectEq.eqb n1 n2 + | _, _ => false + end. + + #[global, program] Instance reflect_level : ReflectEq Level.t := { + eqb := eq_level + }. + Next Obligation. + destruct x, y. + all: unfold eq_level. + all: try solve [ constructor ; reflexivity ]. + all: try solve [ constructor ; discriminate ]. + - destruct (ReflectEq.eqb_spec t0 t1) ; nodec. + constructor. f_equal. assumption. + - destruct (ReflectEq.eqb_spec n n0) ; nodec. + constructor. subst. reflexivity. + Defined. + Definition eq_leibniz (x y : t) : eq x y -> x = y := id. + + Definition eq_dec : forall (l1 l2 : t), {l1 = l2}+{l1 <> l2} := Classes.eq_dec. End Level. @@ -311,10 +331,9 @@ Module UnivExpr. destruct (Nat.compare_spec n n0); repeat constructor; tas. congruence. Qed. - Definition eq_dec (l1 l2 : t) : {l1 = l2} + {l1 <> l2}. - Proof. - decide equality; apply eq_dec. - Defined. + Global Instance reflect_t : ReflectEq t := reflect_prod _ _ . + + Definition eq_dec : forall (l1 l2 : t), {l1 = l2} + {l1 <> l2} := Classes.eq_dec. Definition eq_leibniz (x y : t) : eq x y -> x = y := id. @@ -338,14 +357,20 @@ Module UnivExprSetProp := WPropertiesOn UnivExpr UnivExprSet. (* We have decidable equality w.r.t leibniz equality for sets of levels. This means universes also have a decidable equality. *) -#[global] Instance univexprset_eq_dec : Classes.EqDec UnivExprSet.t. -Proof. - intros p p'. - destruct (UnivExprSet.eq_dec p p'). - - now left; eapply UnivExprSet.eq_leibniz in e. - - right. intros ->. apply n. reflexivity. +#[global, program] Instance univexprset_reflect : ReflectEq UnivExprSet.t := + { eqb := UnivExprSet.equal }. +Next Obligation. + destruct (UnivExprSet.equal x y) eqn:e; constructor. + eapply UnivExprSet.equal_spec in e. + now eapply UnivExprSet.eq_leibniz. + intros e'. + subst y. + pose proof (@UnivExprSetFact.equal_1 x x). + forward H. reflexivity. congruence. Qed. +#[global] Instance univexprset_eq_dec : Classes.EqDec UnivExprSet.t := Classes.eq_dec. + Module Universe. (** A universe is a list of universe expressions which is: - sorted @@ -356,19 +381,24 @@ Module Universe. t_ne : UnivExprSet.is_empty t_set = false }. Derive NoConfusion for nonEmptyUnivExprSet. - - (** This needs a propositional UIP proof to show that [is_empty = false] is a set *) - Set Equations With UIP. - #[global] Instance nonEmptyUnivExprSet_eqdec : EqDec nonEmptyUnivExprSet. - Proof. eqdec_proof. Qed. + + (* We use uip on the is_empty condition *) + #[global, program] Instance univexprset_reflect : ReflectEq nonEmptyUnivExprSet := + { eqb x y := eqb x.(t_set) y.(t_set) }. + Next Obligation. + destruct (eqb_spec (t_set x) (t_set y)); constructor. + destruct x, y; cbn in *. subst. + now rewrite (uip t_ne0 t_ne1). + intros e; subst x; apply H. + reflexivity. + Qed. + + #[global] Instance eq_dec_univ0 : EqDec Universe.nonEmptyUnivExprSet := eq_dec. Inductive t_ := lProp | lSProp | lType (_ : nonEmptyUnivExprSet). Derive NoConfusion for t_. - #[global] Instance t_eqdec : EqDec t_. - Proof. eqdec_proof. Qed. - Definition t := t_. Coercion t_set : nonEmptyUnivExprSet >-> UnivExprSet.t. @@ -377,9 +407,19 @@ Module Universe. match u1, u2 with | lSProp, lSProp => true | lProp, lProp => true - | lType e1, lType e2 => UnivExprSet.equal e1.(t_set) e2.(t_set) + | lType e1, lType e2 => eqb e1 e2 | _,_ => false end. + + #[global, program] Instance reflect_eq_universe : ReflectEq t := + { eqb := eqb }. + Next Obligation. + destruct x, y; cbn; try constructor; auto; try congruence. + destruct (eqb_spec n n0); constructor. now f_equal. + congruence. + Qed. + + #[global] Instance eq_dec_univ : EqDec t := eq_dec. Definition make' (e : UnivExpr.t) : nonEmptyUnivExprSet := {| t_set := UnivExprSet.singleton e; diff --git a/template-coq/theories/common/uGraph.v b/template-coq/theories/common/uGraph.v index 97104334a..37c1ba248 100644 --- a/template-coq/theories/common/uGraph.v +++ b/template-coq/theories/common/uGraph.v @@ -75,7 +75,7 @@ Module VariableLevel. Definition eq_dec : forall x y : t, {x = y} + {x <> y}. intros [s|n] [s'|n']; try now constructor. - destruct (string_dec s s'); [left|right]; congruence. + destruct (Classes.eq_dec s s'); [left|right]; congruence. destruct (PeanoNat.Nat.eq_dec n n'); [left|right]; congruence. Defined. @@ -151,7 +151,7 @@ Module GoodConstraint. Definition eq_dec : forall x y : t, {eq x y} + {~ eq x y}. unfold eq. decide equality. all: try apply VariableLevel.eq_dec. - apply Z.eq_dec. all:apply string_dec || apply Peano_dec.eq_nat_dec. + apply Z.eq_dec. all:apply Classes.eq_dec || apply Peano_dec.eq_nat_dec. Defined. Reserved Notation "x Datatypes.Lt | gc_le _ _ _, _ => Gt | gc_lt_set_level n s, gc_lt_set_level n' s' => - compare_cont (Nat.compare n n') (string_compare s s')%string + compare_cont (Nat.compare n n') (string_compare s s') | _, gc_lt_set_level _ _ => Datatypes.Lt | gc_lt_set_level _ _, _ => Gt | gc_le_set_var n s, gc_le_set_var n' s' => @@ -229,11 +229,11 @@ Module GoodConstraint. apply Z.compare_eq in e'; subst. intros H; apply VariableLevel.compare_eq in H; subst. reflexivity. destruct (Nat.compare_spec n n0) => /= //; subst. - destruct (CompareSpec_string s s0) => /= //; red in H; subst => //. + rewrite StringOT.compare_eq => -> //. destruct (Nat.compare_spec n n1) => /= //; subst. destruct (Nat.compare_spec n0 n2) => /= //; subst => //. destruct (Nat.compare_spec n n0) => /= //; subst. - destruct (CompareSpec_string s s0) => /= //; red in H; subst => //. + rewrite (StringOT.compare_eq) => -> //. destruct (Nat.compare_spec n n1) => /= //; subst. destruct (Nat.compare_spec n0 n2) => /= //; subst => //. Qed. @@ -883,7 +883,7 @@ Proof. exists (VariableLevel.Level l'); intuition. exists (VariableLevel.Var l'); intuition. * intros [[l' [[H1|H1] H2]]|H]. right. subst a. destruct l'; apply EdgeSet.add_spec; left; tas. - destruct l'; left; [exists (VariableLevel.Level s)|exists (VariableLevel.Var n)]; intuition. + destruct l'; left; [exists (VariableLevel.Level t0)|exists (VariableLevel.Var n)]; intuition. right. destruct a; tas; apply EdgeSet.add_spec; right; tas. Qed. @@ -901,7 +901,7 @@ Proof. destruct gc; try apply (Hi.p2 _ Hgc). apply Hi. simpl. apply Hi. - apply Hi. - - cbn. intros l Hl. sq. destruct l. + - cbn. intros l Hl. sq. destruct l as [|s|n]. exists (pathOf_refl _ _). sq. simpl. reflexivity. assert (He: EdgeSet.In (edge_of_level (VariableLevel.Level s)) (wGraph.E (make_graph uctx))). { apply make_graph_E. left. exists (VariableLevel.Level s). intuition. } @@ -987,7 +987,7 @@ Section MakeGraph. : forall x, VSet.In x uctx.1 -> labelling_of_valuation (valuation_of_labelling l) x = l x. Proof. - destruct x; cbnr. + destruct x as [|s|n]; cbnr. - intros _. now apply proj1 in Hl; cbn in Hl. - intro Hs. apply Nat2Pos.id. assert (HH: EdgeSet.In (lzero, Z.of_nat 1, vtn (VariableLevel.Level s)) (wGraph.E G)). { @@ -2447,12 +2447,12 @@ Section AddLevelsCstrs. rewrite {}IHl. split. * intros [[c [eq inl]]|?]; firstorder auto. - destruct a; simpl in *; auto. + destruct a as [|s|n]; simpl in *; auto. rewrite -> EdgeSet.add_spec in H. intuition auto. subst e. left; exists (Level.Level s); intuition auto. rewrite -> EdgeSet.add_spec in H. intuition auto. subst e. left; eexists; intuition eauto. reflexivity. - * intros [[[] [[= <-] [->|inl]]]|?]; simpl; auto; + * intros [[[|s|n] [[= <-] [->|inl]]]|?]; simpl; auto; rewrite -> ?EdgeSet.add_spec; simpl; intuition auto. left. exists (Level.Level s); auto. left. exists (Level.Var n); auto. diff --git a/template-coq/theories/utils.v b/template-coq/theories/utils.v index 99fcaf2b1..16e27e7f3 100644 --- a/template-coq/theories/utils.v +++ b/template-coq/theories/utils.v @@ -4,13 +4,13 @@ (** If you don't want to have the following scopes opened you should *) (** not import this file or close them. *) -From Coq Require Export Bool ZArith Arith Lia String List. +From Coq Require Export Bool ZArith Arith Lia List. From MetaCoq.Template Require Export utils.MCUtils monad_utils. Global Set Asymmetric Patterns. -Global Open Scope string_scope. +Global Open Scope bs_scope. Global Open Scope list_scope. @@ -19,7 +19,7 @@ Global Open Scope nat_scope. (** We keep [++] for lists and use [^] for stings. *) Declare Scope string_scope2. -Notation "s1 ^ s2" := (String.append s1 s2) : string_scope2. +Notation "s1 ^ s2" := (bytestring.String.append s1 s2) : string_scope2. Open Scope string_scope2. (** We want [*] from type_scope but [+] from nat_scope. *) diff --git a/template-coq/theories/utils/ByteCompare.v b/template-coq/theories/utils/ByteCompare.v new file mode 100644 index 000000000..2a3582bd6 --- /dev/null +++ b/template-coq/theories/utils/ByteCompare.v @@ -0,0 +1,7 @@ +From Coq Require Import Byte NArith.BinNat. + +Definition eqb (x y : byte) := + N.eqb (Byte.to_N x) (Byte.to_N y). + +Definition compare (x y : byte) := + N.compare (Byte.to_N x) (Byte.to_N y). diff --git a/template-coq/theories/utils/ByteCompareSpec.v b/template-coq/theories/utils/ByteCompareSpec.v new file mode 100644 index 000000000..1159c5e63 --- /dev/null +++ b/template-coq/theories/utils/ByteCompareSpec.v @@ -0,0 +1,104 @@ +From Coq Require Import Byte NArith. +From MetaCoq.Template Require Import ReflectEq ByteCompare. +From Equations Require Import Equations. + +Derive NoConfusion for comparison. + +Section ByteNoConf. + Local Ltac solve_noconf ::= idtac. + Derive NoConfusion for byte. + Next Obligation. + destruct a; abstract (destruct b; try exact (False_ind _ H); exact eq_refl). + Defined. + Next Obligation. + destruct b; cbn; exact I. + Defined. + Next Obligation. + destruct a; abstract (destruct b; try (exact (False_ind _ e)); destruct e; exact eq_refl). + Defined. + Next Obligation. + destruct b; cbn; exact eq_refl. + Defined. +End ByteNoConf. + +Lemma reflect_equiv P Q b : P <-> Q -> Bool.reflect P b -> Bool.reflect Q b. +Proof. + intros eq r. destruct r. + - now constructor; apply eq. + - constructor. now rewrite <- eq. +Qed. + +Require Import MCCompare. + +Definition lt x y := compare x y = Lt. + +Lemma compare_equiv x y : compare x y = N.compare (Byte.to_N x) (Byte.to_N y). +Proof. + reflexivity. +Qed. +(* +Proof. + destruct x; abstract (destruct y; exact eq_refl). +Qed. *) + +Lemma compare_opp x y : compare x y = CompOpp (compare y x). +Proof. + rewrite !compare_equiv. + apply N.compare_antisym. +Qed. + +Lemma compare_eq x y : compare x y = Eq -> x = y. +Proof. + rewrite compare_equiv. + intros Hcomp. + eapply N.compare_eq in Hcomp. + eapply (f_equal of_N) in Hcomp. + rewrite !of_to_N in Hcomp. now noconf Hcomp. +Qed. + +Lemma compare_spec x y : CompareSpec (x = y) (lt x y) (lt y x) (compare x y). +Proof. + destruct (compare x y) eqn:comp. + - constructor. now apply compare_eq in comp. + - constructor. exact comp. + - constructor. red. now rewrite compare_opp, comp. +Qed. + +From Coq Require Import Lia. + +Lemma lt_trans x y z : lt x y -> lt y z -> lt x z. +Proof. + unfold lt. + rewrite !compare_equiv. + rewrite !N.compare_lt_iff. lia. +Qed. + +Lemma lt_not_eq x y : lt x y -> x <> y. +Proof. + unfold lt. + rewrite !compare_equiv. + rewrite !N.compare_lt_iff. intros Hlt Heq. subst. lia. +Qed. + +Lemma compare_eq_refl x : compare x x = Eq. +Proof. + rewrite compare_equiv. + eapply N.compare_refl. +Qed. + +Definition eqb_compare x y : eqb x y = match compare x y with Eq => true | _ => false end. +Proof. + unfold eqb. + apply N.eqb_compare. + (* destruct x; cbn; abstract (destruct y; cbn; exact eq_refl). *) +Qed. + +Global Program Instance byte_reflect_eq : ReflectEq byte := + {| ReflectEq.eqb := eqb |}. +Next Obligation. + rewrite eqb_compare. + destruct (compare_spec x y); constructor; auto. + all:apply lt_not_eq in H. + - assumption. + - now apply not_eq_sym. +Qed. \ No newline at end of file diff --git a/template-coq/theories/utils/ByteCompare_opt.v b/template-coq/theories/utils/ByteCompare_opt.v new file mode 100644 index 000000000..9502d44dd --- /dev/null +++ b/template-coq/theories/utils/ByteCompare_opt.v @@ -0,0 +1,2310 @@ +From Coq Require Import Byte. + +Definition is_byte_x00 (b : byte) := match b with | x00 => true | _ => false end. +Definition is_byte_x01 (b : byte) := match b with | x01 => true | _ => false end. +Definition is_byte_x02 (b : byte) := match b with | x02 => true | _ => false end. +Definition is_byte_x03 (b : byte) := match b with | x03 => true | _ => false end. +Definition is_byte_x04 (b : byte) := match b with | x04 => true | _ => false end. +Definition is_byte_x05 (b : byte) := match b with | x05 => true | _ => false end. +Definition is_byte_x06 (b : byte) := match b with | x06 => true | _ => false end. +Definition is_byte_x07 (b : byte) := match b with | x07 => true | _ => false end. +Definition is_byte_x08 (b : byte) := match b with | x08 => true | _ => false end. +Definition is_byte_x09 (b : byte) := match b with | x09 => true | _ => false end. +Definition is_byte_x0a (b : byte) := match b with | x0a => true | _ => false end. +Definition is_byte_x0b (b : byte) := match b with | x0b => true | _ => false end. +Definition is_byte_x0c (b : byte) := match b with | x0c => true | _ => false end. +Definition is_byte_x0d (b : byte) := match b with | x0d => true | _ => false end. +Definition is_byte_x0e (b : byte) := match b with | x0e => true | _ => false end. +Definition is_byte_x0f (b : byte) := match b with | x0f => true | _ => false end. +Definition is_byte_x10 (b : byte) := match b with | x10 => true | _ => false end. +Definition is_byte_x11 (b : byte) := match b with | x11 => true | _ => false end. +Definition is_byte_x12 (b : byte) := match b with | x12 => true | _ => false end. +Definition is_byte_x13 (b : byte) := match b with | x13 => true | _ => false end. +Definition is_byte_x14 (b : byte) := match b with | x14 => true | _ => false end. +Definition is_byte_x15 (b : byte) := match b with | x15 => true | _ => false end. +Definition is_byte_x16 (b : byte) := match b with | x16 => true | _ => false end. +Definition is_byte_x17 (b : byte) := match b with | x17 => true | _ => false end. +Definition is_byte_x18 (b : byte) := match b with | x18 => true | _ => false end. +Definition is_byte_x19 (b : byte) := match b with | x19 => true | _ => false end. +Definition is_byte_x1a (b : byte) := match b with | x1a => true | _ => false end. +Definition is_byte_x1b (b : byte) := match b with | x1b => true | _ => false end. +Definition is_byte_x1c (b : byte) := match b with | x1c => true | _ => false end. +Definition is_byte_x1d (b : byte) := match b with | x1d => true | _ => false end. +Definition is_byte_x1e (b : byte) := match b with | x1e => true | _ => false end. +Definition is_byte_x1f (b : byte) := match b with | x1f => true | _ => false end. +Definition is_byte_x20 (b : byte) := match b with | x20 => true | _ => false end. +Definition is_byte_x21 (b : byte) := match b with | x21 => true | _ => false end. +Definition is_byte_x22 (b : byte) := match b with | x22 => true | _ => false end. +Definition is_byte_x23 (b : byte) := match b with | x23 => true | _ => false end. +Definition is_byte_x24 (b : byte) := match b with | x24 => true | _ => false end. +Definition is_byte_x25 (b : byte) := match b with | x25 => true | _ => false end. +Definition is_byte_x26 (b : byte) := match b with | x26 => true | _ => false end. +Definition is_byte_x27 (b : byte) := match b with | x27 => true | _ => false end. +Definition is_byte_x28 (b : byte) := match b with | x28 => true | _ => false end. +Definition is_byte_x29 (b : byte) := match b with | x29 => true | _ => false end. +Definition is_byte_x2a (b : byte) := match b with | x2a => true | _ => false end. +Definition is_byte_x2b (b : byte) := match b with | x2b => true | _ => false end. +Definition is_byte_x2c (b : byte) := match b with | x2c => true | _ => false end. +Definition is_byte_x2d (b : byte) := match b with | x2d => true | _ => false end. +Definition is_byte_x2e (b : byte) := match b with | x2e => true | _ => false end. +Definition is_byte_x2f (b : byte) := match b with | x2f => true | _ => false end. +Definition is_byte_x30 (b : byte) := match b with | x30 => true | _ => false end. +Definition is_byte_x31 (b : byte) := match b with | x31 => true | _ => false end. +Definition is_byte_x32 (b : byte) := match b with | x32 => true | _ => false end. +Definition is_byte_x33 (b : byte) := match b with | x33 => true | _ => false end. +Definition is_byte_x34 (b : byte) := match b with | x34 => true | _ => false end. +Definition is_byte_x35 (b : byte) := match b with | x35 => true | _ => false end. +Definition is_byte_x36 (b : byte) := match b with | x36 => true | _ => false end. +Definition is_byte_x37 (b : byte) := match b with | x37 => true | _ => false end. +Definition is_byte_x38 (b : byte) := match b with | x38 => true | _ => false end. +Definition is_byte_x39 (b : byte) := match b with | x39 => true | _ => false end. +Definition is_byte_x3a (b : byte) := match b with | x3a => true | _ => false end. +Definition is_byte_x3b (b : byte) := match b with | x3b => true | _ => false end. +Definition is_byte_x3c (b : byte) := match b with | x3c => true | _ => false end. +Definition is_byte_x3d (b : byte) := match b with | x3d => true | _ => false end. +Definition is_byte_x3e (b : byte) := match b with | x3e => true | _ => false end. +Definition is_byte_x3f (b : byte) := match b with | x3f => true | _ => false end. +Definition is_byte_x40 (b : byte) := match b with | x40 => true | _ => false end. +Definition is_byte_x41 (b : byte) := match b with | x41 => true | _ => false end. +Definition is_byte_x42 (b : byte) := match b with | x42 => true | _ => false end. +Definition is_byte_x43 (b : byte) := match b with | x43 => true | _ => false end. +Definition is_byte_x44 (b : byte) := match b with | x44 => true | _ => false end. +Definition is_byte_x45 (b : byte) := match b with | x45 => true | _ => false end. +Definition is_byte_x46 (b : byte) := match b with | x46 => true | _ => false end. +Definition is_byte_x47 (b : byte) := match b with | x47 => true | _ => false end. +Definition is_byte_x48 (b : byte) := match b with | x48 => true | _ => false end. +Definition is_byte_x49 (b : byte) := match b with | x49 => true | _ => false end. +Definition is_byte_x4a (b : byte) := match b with | x4a => true | _ => false end. +Definition is_byte_x4b (b : byte) := match b with | x4b => true | _ => false end. +Definition is_byte_x4c (b : byte) := match b with | x4c => true | _ => false end. +Definition is_byte_x4d (b : byte) := match b with | x4d => true | _ => false end. +Definition is_byte_x4e (b : byte) := match b with | x4e => true | _ => false end. +Definition is_byte_x4f (b : byte) := match b with | x4f => true | _ => false end. +Definition is_byte_x50 (b : byte) := match b with | x50 => true | _ => false end. +Definition is_byte_x51 (b : byte) := match b with | x51 => true | _ => false end. +Definition is_byte_x52 (b : byte) := match b with | x52 => true | _ => false end. +Definition is_byte_x53 (b : byte) := match b with | x53 => true | _ => false end. +Definition is_byte_x54 (b : byte) := match b with | x54 => true | _ => false end. +Definition is_byte_x55 (b : byte) := match b with | x55 => true | _ => false end. +Definition is_byte_x56 (b : byte) := match b with | x56 => true | _ => false end. +Definition is_byte_x57 (b : byte) := match b with | x57 => true | _ => false end. +Definition is_byte_x58 (b : byte) := match b with | x58 => true | _ => false end. +Definition is_byte_x59 (b : byte) := match b with | x59 => true | _ => false end. +Definition is_byte_x5a (b : byte) := match b with | x5a => true | _ => false end. +Definition is_byte_x5b (b : byte) := match b with | x5b => true | _ => false end. +Definition is_byte_x5c (b : byte) := match b with | x5c => true | _ => false end. +Definition is_byte_x5d (b : byte) := match b with | x5d => true | _ => false end. +Definition is_byte_x5e (b : byte) := match b with | x5e => true | _ => false end. +Definition is_byte_x5f (b : byte) := match b with | x5f => true | _ => false end. +Definition is_byte_x60 (b : byte) := match b with | x60 => true | _ => false end. +Definition is_byte_x61 (b : byte) := match b with | x61 => true | _ => false end. +Definition is_byte_x62 (b : byte) := match b with | x62 => true | _ => false end. +Definition is_byte_x63 (b : byte) := match b with | x63 => true | _ => false end. +Definition is_byte_x64 (b : byte) := match b with | x64 => true | _ => false end. +Definition is_byte_x65 (b : byte) := match b with | x65 => true | _ => false end. +Definition is_byte_x66 (b : byte) := match b with | x66 => true | _ => false end. +Definition is_byte_x67 (b : byte) := match b with | x67 => true | _ => false end. +Definition is_byte_x68 (b : byte) := match b with | x68 => true | _ => false end. +Definition is_byte_x69 (b : byte) := match b with | x69 => true | _ => false end. +Definition is_byte_x6a (b : byte) := match b with | x6a => true | _ => false end. +Definition is_byte_x6b (b : byte) := match b with | x6b => true | _ => false end. +Definition is_byte_x6c (b : byte) := match b with | x6c => true | _ => false end. +Definition is_byte_x6d (b : byte) := match b with | x6d => true | _ => false end. +Definition is_byte_x6e (b : byte) := match b with | x6e => true | _ => false end. +Definition is_byte_x6f (b : byte) := match b with | x6f => true | _ => false end. +Definition is_byte_x70 (b : byte) := match b with | x70 => true | _ => false end. +Definition is_byte_x71 (b : byte) := match b with | x71 => true | _ => false end. +Definition is_byte_x72 (b : byte) := match b with | x72 => true | _ => false end. +Definition is_byte_x73 (b : byte) := match b with | x73 => true | _ => false end. +Definition is_byte_x74 (b : byte) := match b with | x74 => true | _ => false end. +Definition is_byte_x75 (b : byte) := match b with | x75 => true | _ => false end. +Definition is_byte_x76 (b : byte) := match b with | x76 => true | _ => false end. +Definition is_byte_x77 (b : byte) := match b with | x77 => true | _ => false end. +Definition is_byte_x78 (b : byte) := match b with | x78 => true | _ => false end. +Definition is_byte_x79 (b : byte) := match b with | x79 => true | _ => false end. +Definition is_byte_x7a (b : byte) := match b with | x7a => true | _ => false end. +Definition is_byte_x7b (b : byte) := match b with | x7b => true | _ => false end. +Definition is_byte_x7c (b : byte) := match b with | x7c => true | _ => false end. +Definition is_byte_x7d (b : byte) := match b with | x7d => true | _ => false end. +Definition is_byte_x7e (b : byte) := match b with | x7e => true | _ => false end. +Definition is_byte_x7f (b : byte) := match b with | x7f => true | _ => false end. +Definition is_byte_x80 (b : byte) := match b with | x80 => true | _ => false end. +Definition is_byte_x81 (b : byte) := match b with | x81 => true | _ => false end. +Definition is_byte_x82 (b : byte) := match b with | x82 => true | _ => false end. +Definition is_byte_x83 (b : byte) := match b with | x83 => true | _ => false end. +Definition is_byte_x84 (b : byte) := match b with | x84 => true | _ => false end. +Definition is_byte_x85 (b : byte) := match b with | x85 => true | _ => false end. +Definition is_byte_x86 (b : byte) := match b with | x86 => true | _ => false end. +Definition is_byte_x87 (b : byte) := match b with | x87 => true | _ => false end. +Definition is_byte_x88 (b : byte) := match b with | x88 => true | _ => false end. +Definition is_byte_x89 (b : byte) := match b with | x89 => true | _ => false end. +Definition is_byte_x8a (b : byte) := match b with | x8a => true | _ => false end. +Definition is_byte_x8b (b : byte) := match b with | x8b => true | _ => false end. +Definition is_byte_x8c (b : byte) := match b with | x8c => true | _ => false end. +Definition is_byte_x8d (b : byte) := match b with | x8d => true | _ => false end. +Definition is_byte_x8e (b : byte) := match b with | x8e => true | _ => false end. +Definition is_byte_x8f (b : byte) := match b with | x8f => true | _ => false end. +Definition is_byte_x90 (b : byte) := match b with | x90 => true | _ => false end. +Definition is_byte_x91 (b : byte) := match b with | x91 => true | _ => false end. +Definition is_byte_x92 (b : byte) := match b with | x92 => true | _ => false end. +Definition is_byte_x93 (b : byte) := match b with | x93 => true | _ => false end. +Definition is_byte_x94 (b : byte) := match b with | x94 => true | _ => false end. +Definition is_byte_x95 (b : byte) := match b with | x95 => true | _ => false end. +Definition is_byte_x96 (b : byte) := match b with | x96 => true | _ => false end. +Definition is_byte_x97 (b : byte) := match b with | x97 => true | _ => false end. +Definition is_byte_x98 (b : byte) := match b with | x98 => true | _ => false end. +Definition is_byte_x99 (b : byte) := match b with | x99 => true | _ => false end. +Definition is_byte_x9a (b : byte) := match b with | x9a => true | _ => false end. +Definition is_byte_x9b (b : byte) := match b with | x9b => true | _ => false end. +Definition is_byte_x9c (b : byte) := match b with | x9c => true | _ => false end. +Definition is_byte_x9d (b : byte) := match b with | x9d => true | _ => false end. +Definition is_byte_x9e (b : byte) := match b with | x9e => true | _ => false end. +Definition is_byte_x9f (b : byte) := match b with | x9f => true | _ => false end. +Definition is_byte_xa0 (b : byte) := match b with | xa0 => true | _ => false end. +Definition is_byte_xa1 (b : byte) := match b with | xa1 => true | _ => false end. +Definition is_byte_xa2 (b : byte) := match b with | xa2 => true | _ => false end. +Definition is_byte_xa3 (b : byte) := match b with | xa3 => true | _ => false end. +Definition is_byte_xa4 (b : byte) := match b with | xa4 => true | _ => false end. +Definition is_byte_xa5 (b : byte) := match b with | xa5 => true | _ => false end. +Definition is_byte_xa6 (b : byte) := match b with | xa6 => true | _ => false end. +Definition is_byte_xa7 (b : byte) := match b with | xa7 => true | _ => false end. +Definition is_byte_xa8 (b : byte) := match b with | xa8 => true | _ => false end. +Definition is_byte_xa9 (b : byte) := match b with | xa9 => true | _ => false end. +Definition is_byte_xaa (b : byte) := match b with | xaa => true | _ => false end. +Definition is_byte_xab (b : byte) := match b with | xab => true | _ => false end. +Definition is_byte_xac (b : byte) := match b with | xac => true | _ => false end. +Definition is_byte_xad (b : byte) := match b with | xad => true | _ => false end. +Definition is_byte_xae (b : byte) := match b with | xae => true | _ => false end. +Definition is_byte_xaf (b : byte) := match b with | xaf => true | _ => false end. +Definition is_byte_xb0 (b : byte) := match b with | xb0 => true | _ => false end. +Definition is_byte_xb1 (b : byte) := match b with | xb1 => true | _ => false end. +Definition is_byte_xb2 (b : byte) := match b with | xb2 => true | _ => false end. +Definition is_byte_xb3 (b : byte) := match b with | xb3 => true | _ => false end. +Definition is_byte_xb4 (b : byte) := match b with | xb4 => true | _ => false end. +Definition is_byte_xb5 (b : byte) := match b with | xb5 => true | _ => false end. +Definition is_byte_xb6 (b : byte) := match b with | xb6 => true | _ => false end. +Definition is_byte_xb7 (b : byte) := match b with | xb7 => true | _ => false end. +Definition is_byte_xb8 (b : byte) := match b with | xb8 => true | _ => false end. +Definition is_byte_xb9 (b : byte) := match b with | xb9 => true | _ => false end. +Definition is_byte_xba (b : byte) := match b with | xba => true | _ => false end. +Definition is_byte_xbb (b : byte) := match b with | xbb => true | _ => false end. +Definition is_byte_xbc (b : byte) := match b with | xbc => true | _ => false end. +Definition is_byte_xbd (b : byte) := match b with | xbd => true | _ => false end. +Definition is_byte_xbe (b : byte) := match b with | xbe => true | _ => false end. +Definition is_byte_xbf (b : byte) := match b with | xbf => true | _ => false end. +Definition is_byte_xc0 (b : byte) := match b with | xc0 => true | _ => false end. +Definition is_byte_xc1 (b : byte) := match b with | xc1 => true | _ => false end. +Definition is_byte_xc2 (b : byte) := match b with | xc2 => true | _ => false end. +Definition is_byte_xc3 (b : byte) := match b with | xc3 => true | _ => false end. +Definition is_byte_xc4 (b : byte) := match b with | xc4 => true | _ => false end. +Definition is_byte_xc5 (b : byte) := match b with | xc5 => true | _ => false end. +Definition is_byte_xc6 (b : byte) := match b with | xc6 => true | _ => false end. +Definition is_byte_xc7 (b : byte) := match b with | xc7 => true | _ => false end. +Definition is_byte_xc8 (b : byte) := match b with | xc8 => true | _ => false end. +Definition is_byte_xc9 (b : byte) := match b with | xc9 => true | _ => false end. +Definition is_byte_xca (b : byte) := match b with | xca => true | _ => false end. +Definition is_byte_xcb (b : byte) := match b with | xcb => true | _ => false end. +Definition is_byte_xcc (b : byte) := match b with | xcc => true | _ => false end. +Definition is_byte_xcd (b : byte) := match b with | xcd => true | _ => false end. +Definition is_byte_xce (b : byte) := match b with | xce => true | _ => false end. +Definition is_byte_xcf (b : byte) := match b with | xcf => true | _ => false end. +Definition is_byte_xd0 (b : byte) := match b with | xd0 => true | _ => false end. +Definition is_byte_xd1 (b : byte) := match b with | xd1 => true | _ => false end. +Definition is_byte_xd2 (b : byte) := match b with | xd2 => true | _ => false end. +Definition is_byte_xd3 (b : byte) := match b with | xd3 => true | _ => false end. +Definition is_byte_xd4 (b : byte) := match b with | xd4 => true | _ => false end. +Definition is_byte_xd5 (b : byte) := match b with | xd5 => true | _ => false end. +Definition is_byte_xd6 (b : byte) := match b with | xd6 => true | _ => false end. +Definition is_byte_xd7 (b : byte) := match b with | xd7 => true | _ => false end. +Definition is_byte_xd8 (b : byte) := match b with | xd8 => true | _ => false end. +Definition is_byte_xd9 (b : byte) := match b with | xd9 => true | _ => false end. +Definition is_byte_xda (b : byte) := match b with | xda => true | _ => false end. +Definition is_byte_xdb (b : byte) := match b with | xdb => true | _ => false end. +Definition is_byte_xdc (b : byte) := match b with | xdc => true | _ => false end. +Definition is_byte_xdd (b : byte) := match b with | xdd => true | _ => false end. +Definition is_byte_xde (b : byte) := match b with | xde => true | _ => false end. +Definition is_byte_xdf (b : byte) := match b with | xdf => true | _ => false end. +Definition is_byte_xe0 (b : byte) := match b with | xe0 => true | _ => false end. +Definition is_byte_xe1 (b : byte) := match b with | xe1 => true | _ => false end. +Definition is_byte_xe2 (b : byte) := match b with | xe2 => true | _ => false end. +Definition is_byte_xe3 (b : byte) := match b with | xe3 => true | _ => false end. +Definition is_byte_xe4 (b : byte) := match b with | xe4 => true | _ => false end. +Definition is_byte_xe5 (b : byte) := match b with | xe5 => true | _ => false end. +Definition is_byte_xe6 (b : byte) := match b with | xe6 => true | _ => false end. +Definition is_byte_xe7 (b : byte) := match b with | xe7 => true | _ => false end. +Definition is_byte_xe8 (b : byte) := match b with | xe8 => true | _ => false end. +Definition is_byte_xe9 (b : byte) := match b with | xe9 => true | _ => false end. +Definition is_byte_xea (b : byte) := match b with | xea => true | _ => false end. +Definition is_byte_xeb (b : byte) := match b with | xeb => true | _ => false end. +Definition is_byte_xec (b : byte) := match b with | xec => true | _ => false end. +Definition is_byte_xed (b : byte) := match b with | xed => true | _ => false end. +Definition is_byte_xee (b : byte) := match b with | xee => true | _ => false end. +Definition is_byte_xef (b : byte) := match b with | xef => true | _ => false end. +Definition is_byte_xf0 (b : byte) := match b with | xf0 => true | _ => false end. +Definition is_byte_xf1 (b : byte) := match b with | xf1 => true | _ => false end. +Definition is_byte_xf2 (b : byte) := match b with | xf2 => true | _ => false end. +Definition is_byte_xf3 (b : byte) := match b with | xf3 => true | _ => false end. +Definition is_byte_xf4 (b : byte) := match b with | xf4 => true | _ => false end. +Definition is_byte_xf5 (b : byte) := match b with | xf5 => true | _ => false end. +Definition is_byte_xf6 (b : byte) := match b with | xf6 => true | _ => false end. +Definition is_byte_xf7 (b : byte) := match b with | xf7 => true | _ => false end. +Definition is_byte_xf8 (b : byte) := match b with | xf8 => true | _ => false end. +Definition is_byte_xf9 (b : byte) := match b with | xf9 => true | _ => false end. +Definition is_byte_xfa (b : byte) := match b with | xfa => true | _ => false end. +Definition is_byte_xfb (b : byte) := match b with | xfb => true | _ => false end. +Definition is_byte_xfc (b : byte) := match b with | xfc => true | _ => false end. +Definition is_byte_xfd (b : byte) := match b with | xfd => true | _ => false end. +Definition is_byte_xfe (b : byte) := match b with | xfe => true | _ => false end. +Definition is_byte_xff (b : byte) := match b with | xff => true | _ => false end. +Definition eqb (b b' : byte) := + match b with + | x00 => is_byte_x00 b' + | x01 => is_byte_x01 b' + | x02 => is_byte_x02 b' + | x03 => is_byte_x03 b' + | x04 => is_byte_x04 b' + | x05 => is_byte_x05 b' + | x06 => is_byte_x06 b' + | x07 => is_byte_x07 b' + | x08 => is_byte_x08 b' + | x09 => is_byte_x09 b' + | x0a => is_byte_x0a b' + | x0b => is_byte_x0b b' + | x0c => is_byte_x0c b' + | x0d => is_byte_x0d b' + | x0e => is_byte_x0e b' + | x0f => is_byte_x0f b' + | x10 => is_byte_x10 b' + | x11 => is_byte_x11 b' + | x12 => is_byte_x12 b' + | x13 => is_byte_x13 b' + | x14 => is_byte_x14 b' + | x15 => is_byte_x15 b' + | x16 => is_byte_x16 b' + | x17 => is_byte_x17 b' + | x18 => is_byte_x18 b' + | x19 => is_byte_x19 b' + | x1a => is_byte_x1a b' + | x1b => is_byte_x1b b' + | x1c => is_byte_x1c b' + | x1d => is_byte_x1d b' + | x1e => is_byte_x1e b' + | x1f => is_byte_x1f b' + | x20 => is_byte_x20 b' + | x21 => is_byte_x21 b' + | x22 => is_byte_x22 b' + | x23 => is_byte_x23 b' + | x24 => is_byte_x24 b' + | x25 => is_byte_x25 b' + | x26 => is_byte_x26 b' + | x27 => is_byte_x27 b' + | x28 => is_byte_x28 b' + | x29 => is_byte_x29 b' + | x2a => is_byte_x2a b' + | x2b => is_byte_x2b b' + | x2c => is_byte_x2c b' + | x2d => is_byte_x2d b' + | x2e => is_byte_x2e b' + | x2f => is_byte_x2f b' + | x30 => is_byte_x30 b' + | x31 => is_byte_x31 b' + | x32 => is_byte_x32 b' + | x33 => is_byte_x33 b' + | x34 => is_byte_x34 b' + | x35 => is_byte_x35 b' + | x36 => is_byte_x36 b' + | x37 => is_byte_x37 b' + | x38 => is_byte_x38 b' + | x39 => is_byte_x39 b' + | x3a => is_byte_x3a b' + | x3b => is_byte_x3b b' + | x3c => is_byte_x3c b' + | x3d => is_byte_x3d b' + | x3e => is_byte_x3e b' + | x3f => is_byte_x3f b' + | x40 => is_byte_x40 b' + | x41 => is_byte_x41 b' + | x42 => is_byte_x42 b' + | x43 => is_byte_x43 b' + | x44 => is_byte_x44 b' + | x45 => is_byte_x45 b' + | x46 => is_byte_x46 b' + | x47 => is_byte_x47 b' + | x48 => is_byte_x48 b' + | x49 => is_byte_x49 b' + | x4a => is_byte_x4a b' + | x4b => is_byte_x4b b' + | x4c => is_byte_x4c b' + | x4d => is_byte_x4d b' + | x4e => is_byte_x4e b' + | x4f => is_byte_x4f b' + | x50 => is_byte_x50 b' + | x51 => is_byte_x51 b' + | x52 => is_byte_x52 b' + | x53 => is_byte_x53 b' + | x54 => is_byte_x54 b' + | x55 => is_byte_x55 b' + | x56 => is_byte_x56 b' + | x57 => is_byte_x57 b' + | x58 => is_byte_x58 b' + | x59 => is_byte_x59 b' + | x5a => is_byte_x5a b' + | x5b => is_byte_x5b b' + | x5c => is_byte_x5c b' + | x5d => is_byte_x5d b' + | x5e => is_byte_x5e b' + | x5f => is_byte_x5f b' + | x60 => is_byte_x60 b' + | x61 => is_byte_x61 b' + | x62 => is_byte_x62 b' + | x63 => is_byte_x63 b' + | x64 => is_byte_x64 b' + | x65 => is_byte_x65 b' + | x66 => is_byte_x66 b' + | x67 => is_byte_x67 b' + | x68 => is_byte_x68 b' + | x69 => is_byte_x69 b' + | x6a => is_byte_x6a b' + | x6b => is_byte_x6b b' + | x6c => is_byte_x6c b' + | x6d => is_byte_x6d b' + | x6e => is_byte_x6e b' + | x6f => is_byte_x6f b' + | x70 => is_byte_x70 b' + | x71 => is_byte_x71 b' + | x72 => is_byte_x72 b' + | x73 => is_byte_x73 b' + | x74 => is_byte_x74 b' + | x75 => is_byte_x75 b' + | x76 => is_byte_x76 b' + | x77 => is_byte_x77 b' + | x78 => is_byte_x78 b' + | x79 => is_byte_x79 b' + | x7a => is_byte_x7a b' + | x7b => is_byte_x7b b' + | x7c => is_byte_x7c b' + | x7d => is_byte_x7d b' + | x7e => is_byte_x7e b' + | x7f => is_byte_x7f b' + | x80 => is_byte_x80 b' + | x81 => is_byte_x81 b' + | x82 => is_byte_x82 b' + | x83 => is_byte_x83 b' + | x84 => is_byte_x84 b' + | x85 => is_byte_x85 b' + | x86 => is_byte_x86 b' + | x87 => is_byte_x87 b' + | x88 => is_byte_x88 b' + | x89 => is_byte_x89 b' + | x8a => is_byte_x8a b' + | x8b => is_byte_x8b b' + | x8c => is_byte_x8c b' + | x8d => is_byte_x8d b' + | x8e => is_byte_x8e b' + | x8f => is_byte_x8f b' + | x90 => is_byte_x90 b' + | x91 => is_byte_x91 b' + | x92 => is_byte_x92 b' + | x93 => is_byte_x93 b' + | x94 => is_byte_x94 b' + | x95 => is_byte_x95 b' + | x96 => is_byte_x96 b' + | x97 => is_byte_x97 b' + | x98 => is_byte_x98 b' + | x99 => is_byte_x99 b' + | x9a => is_byte_x9a b' + | x9b => is_byte_x9b b' + | x9c => is_byte_x9c b' + | x9d => is_byte_x9d b' + | x9e => is_byte_x9e b' + | x9f => is_byte_x9f b' + | xa0 => is_byte_xa0 b' + | xa1 => is_byte_xa1 b' + | xa2 => is_byte_xa2 b' + | xa3 => is_byte_xa3 b' + | xa4 => is_byte_xa4 b' + | xa5 => is_byte_xa5 b' + | xa6 => is_byte_xa6 b' + | xa7 => is_byte_xa7 b' + | xa8 => is_byte_xa8 b' + | xa9 => is_byte_xa9 b' + | xaa => is_byte_xaa b' + | xab => is_byte_xab b' + | xac => is_byte_xac b' + | xad => is_byte_xad b' + | xae => is_byte_xae b' + | xaf => is_byte_xaf b' + | xb0 => is_byte_xb0 b' + | xb1 => is_byte_xb1 b' + | xb2 => is_byte_xb2 b' + | xb3 => is_byte_xb3 b' + | xb4 => is_byte_xb4 b' + | xb5 => is_byte_xb5 b' + | xb6 => is_byte_xb6 b' + | xb7 => is_byte_xb7 b' + | xb8 => is_byte_xb8 b' + | xb9 => is_byte_xb9 b' + | xba => is_byte_xba b' + | xbb => is_byte_xbb b' + | xbc => is_byte_xbc b' + | xbd => is_byte_xbd b' + | xbe => is_byte_xbe b' + | xbf => is_byte_xbf b' + | xc0 => is_byte_xc0 b' + | xc1 => is_byte_xc1 b' + | xc2 => is_byte_xc2 b' + | xc3 => is_byte_xc3 b' + | xc4 => is_byte_xc4 b' + | xc5 => is_byte_xc5 b' + | xc6 => is_byte_xc6 b' + | xc7 => is_byte_xc7 b' + | xc8 => is_byte_xc8 b' + | xc9 => is_byte_xc9 b' + | xca => is_byte_xca b' + | xcb => is_byte_xcb b' + | xcc => is_byte_xcc b' + | xcd => is_byte_xcd b' + | xce => is_byte_xce b' + | xcf => is_byte_xcf b' + | xd0 => is_byte_xd0 b' + | xd1 => is_byte_xd1 b' + | xd2 => is_byte_xd2 b' + | xd3 => is_byte_xd3 b' + | xd4 => is_byte_xd4 b' + | xd5 => is_byte_xd5 b' + | xd6 => is_byte_xd6 b' + | xd7 => is_byte_xd7 b' + | xd8 => is_byte_xd8 b' + | xd9 => is_byte_xd9 b' + | xda => is_byte_xda b' + | xdb => is_byte_xdb b' + | xdc => is_byte_xdc b' + | xdd => is_byte_xdd b' + | xde => is_byte_xde b' + | xdf => is_byte_xdf b' + | xe0 => is_byte_xe0 b' + | xe1 => is_byte_xe1 b' + | xe2 => is_byte_xe2 b' + | xe3 => is_byte_xe3 b' + | xe4 => is_byte_xe4 b' + | xe5 => is_byte_xe5 b' + | xe6 => is_byte_xe6 b' + | xe7 => is_byte_xe7 b' + | xe8 => is_byte_xe8 b' + | xe9 => is_byte_xe9 b' + | xea => is_byte_xea b' + | xeb => is_byte_xeb b' + | xec => is_byte_xec b' + | xed => is_byte_xed b' + | xee => is_byte_xee b' + | xef => is_byte_xef b' + | xf0 => is_byte_xf0 b' + | xf1 => is_byte_xf1 b' + | xf2 => is_byte_xf2 b' + | xf3 => is_byte_xf3 b' + | xf4 => is_byte_xf4 b' + | xf5 => is_byte_xf5 b' + | xf6 => is_byte_xf6 b' + | xf7 => is_byte_xf7 b' + | xf8 => is_byte_xf8 b' + | xf9 => is_byte_xf9 b' + | xfa => is_byte_xfa b' + | xfb => is_byte_xfb b' + | xfc => is_byte_xfc b' + | xfd => is_byte_xfd b' + | xfe => is_byte_xfe b' + | xff => is_byte_xff b' + end. +Definition compare_byte_x00 (b : byte) := + match b with + | x00 => Eq + | _ => Lt + end. +Definition compare_byte_x01 (b : byte) := + match b with + | x00 => Gt + | x01 => Eq + | _ => Lt + end. +Definition compare_byte_x02 (b : byte) := + match b with + | x00 | x01 => Gt + | x02 => Eq + | _ => Lt + end. +Definition compare_byte_x03 (b : byte) := + match b with + | x00 | x01 | x02 => Gt + | x03 => Eq + | _ => Lt + end. +Definition compare_byte_x04 (b : byte) := + match b with + | x00 | x01 | x02 | x03 => Gt + | x04 => Eq + | _ => Lt + end. +Definition compare_byte_x05 (b : byte) := + match b with + | x00 | x01 | x02 | x03 | x04 => Gt + | x05 => Eq + | _ => Lt + end. +Definition compare_byte_x06 (b : byte) := + match b with + | x00 | x01 | x02 | x03 | x04 | x05 => Gt + | x06 => Eq + | _ => Lt + end. +Definition compare_byte_x07 (b : byte) := + match b with + | x00 | x01 | x02 | x03 | x04 | x05 | x06 => Gt + | x07 => Eq + | _ => Lt + end. +Definition compare_byte_x08 (b : byte) := + match b with + | x00 | x01 | x02 | x03 | x04 | x05 | x06 | x07 => Gt + | x08 => Eq + | _ => Lt + end. +Definition compare_byte_x09 (b : byte) := + match b with + | x00 | x01 | x02 | x03 | x04 | x05 | x06 | x07 | x08 => Gt + | x09 => Eq + | _ => Lt + end. +Definition compare_byte_x0a (b : byte) := + match b with + | x00 | x01 | x02 | x03 | x04 | x05 | x06 | x07 | x08 | x09 => Gt + | x0a => Eq + | _ => Lt + end. +Definition compare_byte_x0b (b : byte) := + match b with + | x00 | x01 | x02 | x03 | x04 | x05 | x06 | x07 | x08 | x09 | x0a => Gt + | x0b => Eq + | _ => Lt + end. +Definition compare_byte_x0c (b : byte) := + match b with + | x00 | x01 | x02 | x03 | x04 | x05 | x06 | x07 | x08 | x09 | x0a | x0b => Gt + | x0c => Eq + | _ => Lt + end. +Definition compare_byte_x0d (b : byte) := + match b with + | x00 | x01 | x02 | x03 | x04 | x05 | x06 | x07 | x08 | x09 | x0a | x0b | x0c => Gt + | x0d => Eq + | _ => Lt + end. +Definition compare_byte_x0e (b : byte) := + match b with + | x00 | x01 | x02 | x03 | x04 | x05 | x06 | x07 | x08 | x09 | x0a | x0b | x0c | x0d => Gt + | x0e => Eq + | _ => Lt + end. +Definition compare_byte_x0f (b : byte) := + match b with + | x00 | x01 | x02 | x03 | x04 | x05 | x06 | x07 | x08 | x09 | x0a | x0b | x0c | x0d | x0e => Gt + | x0f => Eq + | _ => Lt + end. +Definition compare_byte_x10 (b : byte) := + match b with + | x00 | x01 | x02 | x03 | x04 | x05 | x06 | x07 | x08 | x09 | x0a | x0b | x0c | x0d | x0e | x0f => Gt + | x10 => Eq + | _ => Lt + end. +Definition compare_byte_x11 (b : byte) := + match b with + | x00 | x01 | x02 | x03 | x04 | x05 | x06 | x07 | x08 | x09 | x0a | x0b | x0c | x0d | x0e | x0f | x10 => Gt + | x11 => Eq + | _ => Lt + end. +Definition compare_byte_x12 (b : byte) := + match b with + | x00 | x01 | x02 | x03 | x04 | x05 | x06 | x07 | x08 | x09 | x0a | x0b | x0c | x0d | x0e | x0f | x10 | x11 => Gt + | x12 => Eq + | _ => Lt + end. +Definition compare_byte_x13 (b : byte) := + match b with + | x00 | x01 | x02 | x03 | x04 | x05 | x06 | x07 | x08 | x09 | x0a | x0b | x0c | x0d | x0e | x0f | x10 | x11 | x12 => Gt + | x13 => Eq + | _ => Lt + end. +Definition compare_byte_x14 (b : byte) := + match b with + | x00 | x01 | x02 | x03 | x04 | x05 | x06 | x07 | x08 | x09 | x0a | x0b | x0c | x0d | x0e | x0f | x10 | x11 | x12 | x13 => Gt + | x14 => Eq + | _ => Lt + end. +Definition compare_byte_x15 (b : byte) := + match b with + | x00 | x01 | x02 | x03 | x04 | x05 | x06 | x07 | x08 | x09 | x0a | x0b | x0c | x0d | x0e | x0f | x10 | x11 | x12 | x13 | x14 => Gt + | x15 => Eq + | _ => Lt + end. +Definition compare_byte_x16 (b : byte) := + match b with + | x00 | x01 | x02 | x03 | x04 | x05 | x06 | x07 | x08 | x09 | x0a | x0b | x0c | x0d | x0e | x0f | x10 | x11 | x12 | x13 | x14 | x15 => Gt + | x16 => Eq + | _ => Lt + end. +Definition compare_byte_x17 (b : byte) := + match b with + | x00 | x01 | x02 | x03 | x04 | x05 | x06 | x07 | x08 | x09 | x0a | x0b | x0c | x0d | x0e | x0f | x10 | x11 | x12 | x13 | x14 | x15 | x16 => Gt + | x17 => Eq + | _ => Lt + end. +Definition compare_byte_x18 (b : byte) := + match b with + | x00 | x01 | x02 | x03 | x04 | x05 | x06 | x07 | x08 | x09 | x0a | x0b | x0c | x0d | x0e | x0f | x10 | x11 | x12 | x13 | x14 | x15 | x16 | x17 => Gt + | x18 => Eq + | _ => Lt + end. +Definition compare_byte_x19 (b : byte) := + match b with + | 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 => Gt + | x19 => Eq + | _ => Lt + end. +Definition compare_byte_x1a (b : byte) := + match b with + | 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 => Gt + | x1a => Eq + | _ => Lt + end. +Definition compare_byte_x1b (b : byte) := + match b with + | 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 => Gt + | x1b => Eq + | _ => Lt + end. +Definition compare_byte_x1c (b : byte) := + match b with + | 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 => Gt + | x1c => Eq + | _ => Lt + end. +Definition compare_byte_x1d (b : byte) := + match b with + | 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 => Gt + | x1d => Eq + | _ => Lt + end. +Definition compare_byte_x1e (b : byte) := + match b with + | 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 => Gt + | x1e => Eq + | _ => Lt + end. +Definition compare_byte_x1f (b : byte) := + match b with + | 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 => Gt + | x1f => Eq + | _ => Lt + end. +Definition compare_byte_x20 (b : byte) := + match b with + | 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 => Gt + | x20 => Eq + | _ => Lt + end. +Definition compare_byte_x21 (b : byte) := + match b with + | 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 => Gt + | x21 => Eq + | _ => Lt + end. +Definition compare_byte_x22 (b : byte) := + match b with + | 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 => Gt + | x22 => Eq + | _ => Lt + end. +Definition compare_byte_x23 (b : byte) := + match b with + | 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 => Gt + | x23 => Eq + | _ => Lt + end. +Definition compare_byte_x24 (b : byte) := + match b with + | 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 => Gt + | x24 => Eq + | _ => Lt + end. +Definition compare_byte_x25 (b : byte) := + match b with + | 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 => Gt + | x25 => Eq + | _ => Lt + end. +Definition compare_byte_x26 (b : byte) := + match b with + | 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 => Gt + | x26 => Eq + | _ => Lt + end. +Definition compare_byte_x27 (b : byte) := + match b with + | 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 => Gt + | x27 => Eq + | _ => Lt + end. +Definition compare_byte_x28 (b : byte) := + match b with + | 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 => Gt + | x28 => Eq + | _ => Lt + end. +Definition compare_byte_x29 (b : byte) := + match b with + | 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 => Gt + | x29 => Eq + | _ => Lt + end. +Definition compare_byte_x2a (b : byte) := + match b with + | 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 => Gt + | x2a => Eq + | _ => Lt + end. +Definition compare_byte_x2b (b : byte) := + match b with + | 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 => Gt + | x2b => Eq + | _ => Lt + end. +Definition compare_byte_x2c (b : byte) := + match b with + | 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 => Gt + | x2c => Eq + | _ => Lt + end. +Definition compare_byte_x2d (b : byte) := + match b with + | 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 => Gt + | x2d => Eq + | _ => Lt + end. +Definition compare_byte_x2e (b : byte) := + match b with + | 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 => Gt + | x2e => Eq + | _ => Lt + end. +Definition compare_byte_x2f (b : byte) := + match b with + | 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 => Gt + | x2f => Eq + | _ => Lt + end. +Definition compare_byte_x30 (b : byte) := + match b with + | 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 => Gt + | x30 => Eq + | _ => Lt + end. +Definition compare_byte_x31 (b : byte) := + match b with + | 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 => Gt + | x31 => Eq + | _ => Lt + end. +Definition compare_byte_x32 (b : byte) := + match b with + | 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 => Gt + | x32 => Eq + | _ => Lt + end. +Definition compare_byte_x33 (b : byte) := + match b with + | 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 => Gt + | x33 => Eq + | _ => Lt + end. +Definition compare_byte_x34 (b : byte) := + match b with + | 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 => Gt + | x34 => Eq + | _ => Lt + end. +Definition compare_byte_x35 (b : byte) := + match b with + | 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 => Gt + | x35 => Eq + | _ => Lt + end. +Definition compare_byte_x36 (b : byte) := + match b with + | 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 => Gt + | x36 => Eq + | _ => Lt + end. +Definition compare_byte_x37 (b : byte) := + match b with + | 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 => Gt + | x37 => Eq + | _ => Lt + end. +Definition compare_byte_x38 (b : byte) := + match b with + | 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 => Gt + | x38 => Eq + | _ => Lt + end. +Definition compare_byte_x39 (b : byte) := + match b with + | 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 => Gt + | x39 => Eq + | _ => Lt + end. +Definition compare_byte_x3a (b : byte) := + match b with + | 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 => Gt + | x3a => Eq + | _ => Lt + end. +Definition compare_byte_x3b (b : byte) := + match b with + | 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 => Gt + | x3b => Eq + | _ => Lt + end. +Definition compare_byte_x3c (b : byte) := + match b with + | 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 => Gt + | x3c => Eq + | _ => Lt + end. +Definition compare_byte_x3d (b : byte) := + match b with + | 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 => Gt + | x3d => Eq + | _ => Lt + end. +Definition compare_byte_x3e (b : byte) := + match b with + | 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 => Gt + | x3e => Eq + | _ => Lt + end. +Definition compare_byte_x3f (b : byte) := + match b with + | 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 => Gt + | x3f => Eq + | _ => Lt + end. +Definition compare_byte_x40 (b : byte) := + match b with + | 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 => Gt + | x40 => Eq + | _ => Lt + end. +Definition compare_byte_x41 (b : byte) := + match b with + | 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 => Gt + | x41 => Eq + | _ => Lt + end. +Definition compare_byte_x42 (b : byte) := + match b with + | 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 => Gt + | x42 => Eq + | _ => Lt + end. +Definition compare_byte_x43 (b : byte) := + match b with + | 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 => Gt + | x43 => Eq + | _ => Lt + end. +Definition compare_byte_x44 (b : byte) := + match b with + | 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 => Gt + | x44 => Eq + | _ => Lt + end. +Definition compare_byte_x45 (b : byte) := + match b with + | 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 => Gt + | x45 => Eq + | _ => Lt + end. +Definition compare_byte_x46 (b : byte) := + match b with + | 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 => Gt + | x46 => Eq + | _ => Lt + end. +Definition compare_byte_x47 (b : byte) := + match b with + | 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 => Gt + | x47 => Eq + | _ => Lt + end. +Definition compare_byte_x48 (b : byte) := + match b with + | 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 => Gt + | x48 => Eq + | _ => Lt + end. +Definition compare_byte_x49 (b : byte) := + match b with + | 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 => Gt + | x49 => Eq + | _ => Lt + end. +Definition compare_byte_x4a (b : byte) := + match b with + | 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 => Gt + | x4a => Eq + | _ => Lt + end. +Definition compare_byte_x4b (b : byte) := + match b with + | 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 => Gt + | x4b => Eq + | _ => Lt + end. +Definition compare_byte_x4c (b : byte) := + match b with + | 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 => Gt + | x4c => Eq + | _ => Lt + end. +Definition compare_byte_x4d (b : byte) := + match b with + | 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 => Gt + | x4d => Eq + | _ => Lt + end. +Definition compare_byte_x4e (b : byte) := + match b with + | 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 => Gt + | x4e => Eq + | _ => Lt + end. +Definition compare_byte_x4f (b : byte) := + match b with + | 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 => Gt + | x4f => Eq + | _ => Lt + end. +Definition compare_byte_x50 (b : byte) := + match b with + | 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 => Gt + | x50 => Eq + | _ => Lt + end. +Definition compare_byte_x51 (b : byte) := + match b with + | 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 => Gt + | x51 => Eq + | _ => Lt + end. +Definition compare_byte_x52 (b : byte) := + match b with + | 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 => Gt + | x52 => Eq + | _ => Lt + end. +Definition compare_byte_x53 (b : byte) := + match b with + | 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 => Gt + | x53 => Eq + | _ => Lt + end. +Definition compare_byte_x54 (b : byte) := + match b with + | 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 => Gt + | x54 => Eq + | _ => Lt + end. +Definition compare_byte_x55 (b : byte) := + match b with + | 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 => Gt + | x55 => Eq + | _ => Lt + end. +Definition compare_byte_x56 (b : byte) := + match b with + | 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 => Gt + | x56 => Eq + | _ => Lt + end. +Definition compare_byte_x57 (b : byte) := + match b with + | 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 => Gt + | x57 => Eq + | _ => Lt + end. +Definition compare_byte_x58 (b : byte) := + match b with + | 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 => Gt + | x58 => Eq + | _ => Lt + end. +Definition compare_byte_x59 (b : byte) := + match b with + | 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 => Gt + | x59 => Eq + | _ => Lt + end. +Definition compare_byte_x5a (b : byte) := + match b with + | 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 => Gt + | x5a => Eq + | _ => Lt + end. +Definition compare_byte_x5b (b : byte) := + match b with + | 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 => Gt + | x5b => Eq + | _ => Lt + end. +Definition compare_byte_x5c (b : byte) := + match b with + | 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 => Gt + | x5c => Eq + | _ => Lt + end. +Definition compare_byte_x5d (b : byte) := + match b with + | 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 => Gt + | x5d => Eq + | _ => Lt + end. +Definition compare_byte_x5e (b : byte) := + match b with + | 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 => Gt + | x5e => Eq + | _ => Lt + end. +Definition compare_byte_x5f (b : byte) := + match b with + | 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 => Gt + | x5f => Eq + | _ => Lt + end. +Definition compare_byte_x60 (b : byte) := + match b with + | 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 => Gt + | x60 => Eq + | _ => Lt + end. +Definition compare_byte_x61 (b : byte) := + match b with + | 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 => Gt + | x61 => Eq + | _ => Lt + end. +Definition compare_byte_x62 (b : byte) := + match b with + | 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 => Gt + | x62 => Eq + | _ => Lt + end. +Definition compare_byte_x63 (b : byte) := + match b with + | 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 => Gt + | x63 => Eq + | _ => Lt + end. +Definition compare_byte_x64 (b : byte) := + match b with + | 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 => Gt + | x64 => Eq + | _ => Lt + end. +Definition compare_byte_x65 (b : byte) := + match b with + | 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 => Gt + | x65 => Eq + | _ => Lt + end. +Definition compare_byte_x66 (b : byte) := + match b with + | 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 => Gt + | x66 => Eq + | _ => Lt + end. +Definition compare_byte_x67 (b : byte) := + match b with + | 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 => Gt + | x67 => Eq + | _ => Lt + end. +Definition compare_byte_x68 (b : byte) := + match b with + | 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 => Gt + | x68 => Eq + | _ => Lt + end. +Definition compare_byte_x69 (b : byte) := + match b with + | 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 => Gt + | x69 => Eq + | _ => Lt + end. +Definition compare_byte_x6a (b : byte) := + match b with + | 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 => Gt + | x6a => Eq + | _ => Lt + end. +Definition compare_byte_x6b (b : byte) := + match b with + | 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 => Gt + | x6b => Eq + | _ => Lt + end. +Definition compare_byte_x6c (b : byte) := + match b with + | 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 => Gt + | x6c => Eq + | _ => Lt + end. +Definition compare_byte_x6d (b : byte) := + match b with + | 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 => Gt + | x6d => Eq + | _ => Lt + end. +Definition compare_byte_x6e (b : byte) := + match b with + | 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 => Gt + | x6e => Eq + | _ => Lt + end. +Definition compare_byte_x6f (b : byte) := + match b with + | 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 => Gt + | x6f => Eq + | _ => Lt + end. +Definition compare_byte_x70 (b : byte) := + match b with + | 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 => Gt + | x70 => Eq + | _ => Lt + end. +Definition compare_byte_x71 (b : byte) := + match b with + | 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 => Gt + | x71 => Eq + | _ => Lt + end. +Definition compare_byte_x72 (b : byte) := + match b with + | 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 => Gt + | x72 => Eq + | _ => Lt + end. +Definition compare_byte_x73 (b : byte) := + match b with + | 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 => Gt + | x73 => Eq + | _ => Lt + end. +Definition compare_byte_x74 (b : byte) := + match b with + | 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 => Gt + | x74 => Eq + | _ => Lt + end. +Definition compare_byte_x75 (b : byte) := + match b with + | 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 => Gt + | x75 => Eq + | _ => Lt + end. +Definition compare_byte_x76 (b : byte) := + match b with + | 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 => Gt + | x76 => Eq + | _ => Lt + end. +Definition compare_byte_x77 (b : byte) := + match b with + | 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 => Gt + | x77 => Eq + | _ => Lt + end. +Definition compare_byte_x78 (b : byte) := + match b with + | 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 => Gt + | x78 => Eq + | _ => Lt + end. +Definition compare_byte_x79 (b : byte) := + match b with + | 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 => Gt + | x79 => Eq + | _ => Lt + end. +Definition compare_byte_x7a (b : byte) := + match b with + | 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 => Gt + | x7a => Eq + | _ => Lt + end. +Definition compare_byte_x7b (b : byte) := + match b with + | 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 => Gt + | x7b => Eq + | _ => Lt + end. +Definition compare_byte_x7c (b : byte) := + match b with + | 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 => Gt + | x7c => Eq + | _ => Lt + end. +Definition compare_byte_x7d (b : byte) := + match b with + | 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 => Gt + | x7d => Eq + | _ => Lt + end. +Definition compare_byte_x7e (b : byte) := + match b with + | 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 => Gt + | x7e => Eq + | _ => Lt + end. +Definition compare_byte_x7f (b : byte) := + match b with + | 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 => Gt + | x7f => Eq + | _ => Lt + end. +Definition compare_byte_x80 (b : byte) := + match b with + | 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 => Gt + | x80 => Eq + | _ => Lt + end. +Definition compare_byte_x81 (b : byte) := + match b with + | 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 => Gt + | x81 => Eq + | _ => Lt + end. +Definition compare_byte_x82 (b : byte) := + match b with + | 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 => Gt + | x82 => Eq + | _ => Lt + end. +Definition compare_byte_x83 (b : byte) := + match b with + | 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 => Gt + | x83 => Eq + | _ => Lt + end. +Definition compare_byte_x84 (b : byte) := + match b with + | 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 => Gt + | x84 => Eq + | _ => Lt + end. +Definition compare_byte_x85 (b : byte) := + match b with + | 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 => Gt + | x85 => Eq + | _ => Lt + end. +Definition compare_byte_x86 (b : byte) := + match b with + | 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 => Gt + | x86 => Eq + | _ => Lt + end. +Definition compare_byte_x87 (b : byte) := + match b with + | 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 => Gt + | x87 => Eq + | _ => Lt + end. +Definition compare_byte_x88 (b : byte) := + match b with + | 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 => Gt + | x88 => Eq + | _ => Lt + end. +Definition compare_byte_x89 (b : byte) := + match b with + | 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 => Gt + | x89 => Eq + | _ => Lt + end. +Definition compare_byte_x8a (b : byte) := + match b with + | 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 => Gt + | x8a => Eq + | _ => Lt + end. +Definition compare_byte_x8b (b : byte) := + match b with + | 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 => Gt + | x8b => Eq + | _ => Lt + end. +Definition compare_byte_x8c (b : byte) := + match b with + | 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 => Gt + | x8c => Eq + | _ => Lt + end. +Definition compare_byte_x8d (b : byte) := + match b with + | 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 => Gt + | x8d => Eq + | _ => Lt + end. +Definition compare_byte_x8e (b : byte) := + match b with + | 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 => Gt + | x8e => Eq + | _ => Lt + end. +Definition compare_byte_x8f (b : byte) := + match b with + | 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 => Gt + | x8f => Eq + | _ => Lt + end. +Definition compare_byte_x90 (b : byte) := + match b with + | 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 => Gt + | x90 => Eq + | _ => Lt + end. +Definition compare_byte_x91 (b : byte) := + match b with + | 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 => Gt + | x91 => Eq + | _ => Lt + end. +Definition compare_byte_x92 (b : byte) := + match b with + | 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 => Gt + | x92 => Eq + | _ => Lt + end. +Definition compare_byte_x93 (b : byte) := + match b with + | 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 => Gt + | x93 => Eq + | _ => Lt + end. +Definition compare_byte_x94 (b : byte) := + match b with + | 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 => Gt + | x94 => Eq + | _ => Lt + end. +Definition compare_byte_x95 (b : byte) := + match b with + | 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 => Gt + | x95 => Eq + | _ => Lt + end. +Definition compare_byte_x96 (b : byte) := + match b with + | 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 => Gt + | x96 => Eq + | _ => Lt + end. +Definition compare_byte_x97 (b : byte) := + match b with + | 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 => Gt + | x97 => Eq + | _ => Lt + end. +Definition compare_byte_x98 (b : byte) := + match b with + | 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 => Gt + | x98 => Eq + | _ => Lt + end. +Definition compare_byte_x99 (b : byte) := + match b with + | 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 => Gt + | x99 => Eq + | _ => Lt + end. +Definition compare_byte_x9a (b : byte) := + match b with + | 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 => Gt + | x9a => Eq + | _ => Lt + end. +Definition compare_byte_x9b (b : byte) := + match b with + | 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 => Gt + | x9b => Eq + | _ => Lt + end. +Definition compare_byte_x9c (b : byte) := + match b with + | 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 => Gt + | x9c => Eq + | _ => Lt + end. +Definition compare_byte_x9d (b : byte) := + match b with + | 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 => Gt + | x9d => Eq + | _ => Lt + end. +Definition compare_byte_x9e (b : byte) := + match b with + | 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 => Gt + | x9e => Eq + | _ => Lt + end. +Definition compare_byte_x9f (b : byte) := + match b with + | 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 => Gt + | x9f => Eq + | _ => Lt + end. +Definition compare_byte_xa0 (b : byte) := + match b with + | 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 => Gt + | xa0 => Eq + | _ => Lt + end. +Definition compare_byte_xa1 (b : byte) := + match b with + | 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 => Gt + | xa1 => Eq + | _ => Lt + end. +Definition compare_byte_xa2 (b : byte) := + match b with + | 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 => Gt + | xa2 => Eq + | _ => Lt + end. +Definition compare_byte_xa3 (b : byte) := + match b with + | 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 => Gt + | xa3 => Eq + | _ => Lt + end. +Definition compare_byte_xa4 (b : byte) := + match b with + | 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 => Gt + | xa4 => Eq + | _ => Lt + end. +Definition compare_byte_xa5 (b : byte) := + match b with + | 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 => Gt + | xa5 => Eq + | _ => Lt + end. +Definition compare_byte_xa6 (b : byte) := + match b with + | 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 => Gt + | xa6 => Eq + | _ => Lt + end. +Definition compare_byte_xa7 (b : byte) := + match b with + | 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 => Gt + | xa7 => Eq + | _ => Lt + end. +Definition compare_byte_xa8 (b : byte) := + match b with + | 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 => Gt + | xa8 => Eq + | _ => Lt + end. +Definition compare_byte_xa9 (b : byte) := + match b with + | 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 => Gt + | xa9 => Eq + | _ => Lt + end. +Definition compare_byte_xaa (b : byte) := + match b with + | 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 => Gt + | xaa => Eq + | _ => Lt + end. +Definition compare_byte_xab (b : byte) := + match b with + | 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 => Gt + | xab => Eq + | _ => Lt + end. +Definition compare_byte_xac (b : byte) := + match b with + | 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 => Gt + | xac => Eq + | _ => Lt + end. +Definition compare_byte_xad (b : byte) := + match b with + | 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 => Gt + | xad => Eq + | _ => Lt + end. +Definition compare_byte_xae (b : byte) := + match b with + | 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 => Gt + | xae => Eq + | _ => Lt + end. +Definition compare_byte_xaf (b : byte) := + match b with + | 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 => Gt + | xaf => Eq + | _ => Lt + end. +Definition compare_byte_xb0 (b : byte) := + match b with + | 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 => Gt + | xb0 => Eq + | _ => Lt + end. +Definition compare_byte_xb1 (b : byte) := + match b with + | 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 => Gt + | xb1 => Eq + | _ => Lt + end. +Definition compare_byte_xb2 (b : byte) := + match b with + | 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 => Gt + | xb2 => Eq + | _ => Lt + end. +Definition compare_byte_xb3 (b : byte) := + match b with + | 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 => Gt + | xb3 => Eq + | _ => Lt + end. +Definition compare_byte_xb4 (b : byte) := + match b with + | 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 => Gt + | xb4 => Eq + | _ => Lt + end. +Definition compare_byte_xb5 (b : byte) := + match b with + | 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 => Gt + | xb5 => Eq + | _ => Lt + end. +Definition compare_byte_xb6 (b : byte) := + match b with + | 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 => Gt + | xb6 => Eq + | _ => Lt + end. +Definition compare_byte_xb7 (b : byte) := + match b with + | 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 => Gt + | xb7 => Eq + | _ => Lt + end. +Definition compare_byte_xb8 (b : byte) := + match b with + | 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 => Gt + | xb8 => Eq + | _ => Lt + end. +Definition compare_byte_xb9 (b : byte) := + match b with + | 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 => Gt + | xb9 => Eq + | _ => Lt + end. +Definition compare_byte_xba (b : byte) := + match b with + | 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 => Gt + | xba => Eq + | _ => Lt + end. +Definition compare_byte_xbb (b : byte) := + match b with + | 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 => Gt + | xbb => Eq + | _ => Lt + end. +Definition compare_byte_xbc (b : byte) := + match b with + | 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 => Gt + | xbc => Eq + | _ => Lt + end. +Definition compare_byte_xbd (b : byte) := + match b with + | 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 => Gt + | xbd => Eq + | _ => Lt + end. +Definition compare_byte_xbe (b : byte) := + match b with + | 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 => Gt + | xbe => Eq + | _ => Lt + end. +Definition compare_byte_xbf (b : byte) := + match b with + | 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 => Gt + | xbf => Eq + | _ => Lt + end. +Definition compare_byte_xc0 (b : byte) := + match b with + | 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 => Gt + | xc0 => Eq + | _ => Lt + end. +Definition compare_byte_xc1 (b : byte) := + match b with + | 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 => Gt + | xc1 => Eq + | _ => Lt + end. +Definition compare_byte_xc2 (b : byte) := + match b with + | 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 => Gt + | xc2 => Eq + | _ => Lt + end. +Definition compare_byte_xc3 (b : byte) := + match b with + | 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 => Gt + | xc3 => Eq + | _ => Lt + end. +Definition compare_byte_xc4 (b : byte) := + match b with + | 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 => Gt + | xc4 => Eq + | _ => Lt + end. +Definition compare_byte_xc5 (b : byte) := + match b with + | 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 => Gt + | xc5 => Eq + | _ => Lt + end. +Definition compare_byte_xc6 (b : byte) := + match b with + | 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 => Gt + | xc6 => Eq + | _ => Lt + end. +Definition compare_byte_xc7 (b : byte) := + match b with + | 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 => Gt + | xc7 => Eq + | _ => Lt + end. +Definition compare_byte_xc8 (b : byte) := + match b with + | 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 => Gt + | xc8 => Eq + | _ => Lt + end. +Definition compare_byte_xc9 (b : byte) := + match b with + | 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 => Gt + | xc9 => Eq + | _ => Lt + end. +Definition compare_byte_xca (b : byte) := + match b with + | 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 => Gt + | xca => Eq + | _ => Lt + end. +Definition compare_byte_xcb (b : byte) := + match b with + | 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 => Gt + | xcb => Eq + | _ => Lt + end. +Definition compare_byte_xcc (b : byte) := + match b with + | 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 => Gt + | xcc => Eq + | _ => Lt + end. +Definition compare_byte_xcd (b : byte) := + match b with + | 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 => Gt + | xcd => Eq + | _ => Lt + end. +Definition compare_byte_xce (b : byte) := + match b with + | 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 => Gt + | xce => Eq + | _ => Lt + end. +Definition compare_byte_xcf (b : byte) := + match b with + | 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 => Gt + | xcf => Eq + | _ => Lt + end. +Definition compare_byte_xd0 (b : byte) := + match b with + | 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 => Gt + | xd0 => Eq + | _ => Lt + end. +Definition compare_byte_xd1 (b : byte) := + match b with + | 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 => Gt + | xd1 => Eq + | _ => Lt + end. +Definition compare_byte_xd2 (b : byte) := + match b with + | 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 => Gt + | xd2 => Eq + | _ => Lt + end. +Definition compare_byte_xd3 (b : byte) := + match b with + | 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 => Gt + | xd3 => Eq + | _ => Lt + end. +Definition compare_byte_xd4 (b : byte) := + match b with + | 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 => Gt + | xd4 => Eq + | _ => Lt + end. +Definition compare_byte_xd5 (b : byte) := + match b with + | 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 => Gt + | xd5 => Eq + | _ => Lt + end. +Definition compare_byte_xd6 (b : byte) := + match b with + | 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 => Gt + | xd6 => Eq + | _ => Lt + end. +Definition compare_byte_xd7 (b : byte) := + match b with + | 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 => Gt + | xd7 => Eq + | _ => Lt + end. +Definition compare_byte_xd8 (b : byte) := + match b with + | 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 => Gt + | xd8 => Eq + | _ => Lt + end. +Definition compare_byte_xd9 (b : byte) := + match b with + | 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 => Gt + | xd9 => Eq + | _ => Lt + end. +Definition compare_byte_xda (b : byte) := + match b with + | 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 => Gt + | xda => Eq + | _ => Lt + end. +Definition compare_byte_xdb (b : byte) := + match b with + | 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 => Gt + | xdb => Eq + | _ => Lt + end. +Definition compare_byte_xdc (b : byte) := + match b with + | 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 => Gt + | xdc => Eq + | _ => Lt + end. +Definition compare_byte_xdd (b : byte) := + match b with + | 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 => Gt + | xdd => Eq + | _ => Lt + end. +Definition compare_byte_xde (b : byte) := + match b with + | 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 => Gt + | xde => Eq + | _ => Lt + end. +Definition compare_byte_xdf (b : byte) := + match b with + | 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 => Gt + | xdf => Eq + | _ => Lt + end. +Definition compare_byte_xe0 (b : byte) := + match b with + | 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 => Gt + | xe0 => Eq + | _ => Lt + end. +Definition compare_byte_xe1 (b : byte) := + match b with + | 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 => Gt + | xe1 => Eq + | _ => Lt + end. +Definition compare_byte_xe2 (b : byte) := + match b with + | 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 => Gt + | xe2 => Eq + | _ => Lt + end. +Definition compare_byte_xe3 (b : byte) := + match b with + | 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 => Gt + | xe3 => Eq + | _ => Lt + end. +Definition compare_byte_xe4 (b : byte) := + match b with + | 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 => Gt + | xe4 => Eq + | _ => Lt + end. +Definition compare_byte_xe5 (b : byte) := + match b with + | 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 => Gt + | xe5 => Eq + | _ => Lt + end. +Definition compare_byte_xe6 (b : byte) := + match b with + | 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 => Gt + | xe6 => Eq + | _ => Lt + end. +Definition compare_byte_xe7 (b : byte) := + match b with + | 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 => Gt + | xe7 => Eq + | _ => Lt + end. +Definition compare_byte_xe8 (b : byte) := + match b with + | 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 => Gt + | xe8 => Eq + | _ => Lt + end. +Definition compare_byte_xe9 (b : byte) := + match b with + | 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 => Gt + | xe9 => Eq + | _ => Lt + end. +Definition compare_byte_xea (b : byte) := + match b with + | 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 => Gt + | xea => Eq + | _ => Lt + end. +Definition compare_byte_xeb (b : byte) := + match b with + | 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 => Gt + | xeb => Eq + | _ => Lt + end. +Definition compare_byte_xec (b : byte) := + match b with + | 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 => Gt + | xec => Eq + | _ => Lt + end. +Definition compare_byte_xed (b : byte) := + match b with + | 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 => Gt + | xed => Eq + | _ => Lt + end. +Definition compare_byte_xee (b : byte) := + match b with + | 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 => Gt + | xee => Eq + | _ => Lt + end. +Definition compare_byte_xef (b : byte) := + match b with + | 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 => Gt + | xef => Eq + | _ => Lt + end. +Definition compare_byte_xf0 (b : byte) := + match b with + | 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 => Gt + | xf0 => Eq + | _ => Lt + end. +Definition compare_byte_xf1 (b : byte) := + match b with + | 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 => Gt + | xf1 => Eq + | _ => Lt + end. +Definition compare_byte_xf2 (b : byte) := + match b with + | 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 => Gt + | xf2 => Eq + | _ => Lt + end. +Definition compare_byte_xf3 (b : byte) := + match b with + | 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 => Gt + | xf3 => Eq + | _ => Lt + end. +Definition compare_byte_xf4 (b : byte) := + match b with + | 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 => Gt + | xf4 => Eq + | _ => Lt + end. +Definition compare_byte_xf5 (b : byte) := + match b with + | 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 => Gt + | xf5 => Eq + | _ => Lt + end. +Definition compare_byte_xf6 (b : byte) := + match b with + | 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 => Gt + | xf6 => Eq + | _ => Lt + end. +Definition compare_byte_xf7 (b : byte) := + match b with + | 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 => Gt + | xf7 => Eq + | _ => Lt + end. +Definition compare_byte_xf8 (b : byte) := + match b with + | 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 => Gt + | xf8 => Eq + | _ => Lt + end. +Definition compare_byte_xf9 (b : byte) := + match b with + | 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 => Gt + | xf9 => Eq + | _ => Lt + end. +Definition compare_byte_xfa (b : byte) := + match b with + | 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 => Gt + | xfa => Eq + | _ => Lt + end. +Definition compare_byte_xfb (b : byte) := + match b with + | 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 => Gt + | xfb => Eq + | _ => Lt + end. +Definition compare_byte_xfc (b : byte) := + match b with + | 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 => Gt + | xfc => Eq + | _ => Lt + end. +Definition compare_byte_xfd (b : byte) := + match b with + | 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 => Gt + | xfd => Eq + | _ => Lt + end. +Definition compare_byte_xfe (b : byte) := + match b with + | 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 => Gt + | xfe => Eq + | _ => Lt + end. +Definition compare_byte_xff (b : byte) := + match b with + | 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 => Gt + | xff => Eq + end. +Definition compare (b b' : byte) := + match b with + | x00 => compare_byte_x00 b' + | x01 => compare_byte_x01 b' + | x02 => compare_byte_x02 b' + | x03 => compare_byte_x03 b' + | x04 => compare_byte_x04 b' + | x05 => compare_byte_x05 b' + | x06 => compare_byte_x06 b' + | x07 => compare_byte_x07 b' + | x08 => compare_byte_x08 b' + | x09 => compare_byte_x09 b' + | x0a => compare_byte_x0a b' + | x0b => compare_byte_x0b b' + | x0c => compare_byte_x0c b' + | x0d => compare_byte_x0d b' + | x0e => compare_byte_x0e b' + | x0f => compare_byte_x0f b' + | x10 => compare_byte_x10 b' + | x11 => compare_byte_x11 b' + | x12 => compare_byte_x12 b' + | x13 => compare_byte_x13 b' + | x14 => compare_byte_x14 b' + | x15 => compare_byte_x15 b' + | x16 => compare_byte_x16 b' + | x17 => compare_byte_x17 b' + | x18 => compare_byte_x18 b' + | x19 => compare_byte_x19 b' + | x1a => compare_byte_x1a b' + | x1b => compare_byte_x1b b' + | x1c => compare_byte_x1c b' + | x1d => compare_byte_x1d b' + | x1e => compare_byte_x1e b' + | x1f => compare_byte_x1f b' + | x20 => compare_byte_x20 b' + | x21 => compare_byte_x21 b' + | x22 => compare_byte_x22 b' + | x23 => compare_byte_x23 b' + | x24 => compare_byte_x24 b' + | x25 => compare_byte_x25 b' + | x26 => compare_byte_x26 b' + | x27 => compare_byte_x27 b' + | x28 => compare_byte_x28 b' + | x29 => compare_byte_x29 b' + | x2a => compare_byte_x2a b' + | x2b => compare_byte_x2b b' + | x2c => compare_byte_x2c b' + | x2d => compare_byte_x2d b' + | x2e => compare_byte_x2e b' + | x2f => compare_byte_x2f b' + | x30 => compare_byte_x30 b' + | x31 => compare_byte_x31 b' + | x32 => compare_byte_x32 b' + | x33 => compare_byte_x33 b' + | x34 => compare_byte_x34 b' + | x35 => compare_byte_x35 b' + | x36 => compare_byte_x36 b' + | x37 => compare_byte_x37 b' + | x38 => compare_byte_x38 b' + | x39 => compare_byte_x39 b' + | x3a => compare_byte_x3a b' + | x3b => compare_byte_x3b b' + | x3c => compare_byte_x3c b' + | x3d => compare_byte_x3d b' + | x3e => compare_byte_x3e b' + | x3f => compare_byte_x3f b' + | x40 => compare_byte_x40 b' + | x41 => compare_byte_x41 b' + | x42 => compare_byte_x42 b' + | x43 => compare_byte_x43 b' + | x44 => compare_byte_x44 b' + | x45 => compare_byte_x45 b' + | x46 => compare_byte_x46 b' + | x47 => compare_byte_x47 b' + | x48 => compare_byte_x48 b' + | x49 => compare_byte_x49 b' + | x4a => compare_byte_x4a b' + | x4b => compare_byte_x4b b' + | x4c => compare_byte_x4c b' + | x4d => compare_byte_x4d b' + | x4e => compare_byte_x4e b' + | x4f => compare_byte_x4f b' + | x50 => compare_byte_x50 b' + | x51 => compare_byte_x51 b' + | x52 => compare_byte_x52 b' + | x53 => compare_byte_x53 b' + | x54 => compare_byte_x54 b' + | x55 => compare_byte_x55 b' + | x56 => compare_byte_x56 b' + | x57 => compare_byte_x57 b' + | x58 => compare_byte_x58 b' + | x59 => compare_byte_x59 b' + | x5a => compare_byte_x5a b' + | x5b => compare_byte_x5b b' + | x5c => compare_byte_x5c b' + | x5d => compare_byte_x5d b' + | x5e => compare_byte_x5e b' + | x5f => compare_byte_x5f b' + | x60 => compare_byte_x60 b' + | x61 => compare_byte_x61 b' + | x62 => compare_byte_x62 b' + | x63 => compare_byte_x63 b' + | x64 => compare_byte_x64 b' + | x65 => compare_byte_x65 b' + | x66 => compare_byte_x66 b' + | x67 => compare_byte_x67 b' + | x68 => compare_byte_x68 b' + | x69 => compare_byte_x69 b' + | x6a => compare_byte_x6a b' + | x6b => compare_byte_x6b b' + | x6c => compare_byte_x6c b' + | x6d => compare_byte_x6d b' + | x6e => compare_byte_x6e b' + | x6f => compare_byte_x6f b' + | x70 => compare_byte_x70 b' + | x71 => compare_byte_x71 b' + | x72 => compare_byte_x72 b' + | x73 => compare_byte_x73 b' + | x74 => compare_byte_x74 b' + | x75 => compare_byte_x75 b' + | x76 => compare_byte_x76 b' + | x77 => compare_byte_x77 b' + | x78 => compare_byte_x78 b' + | x79 => compare_byte_x79 b' + | x7a => compare_byte_x7a b' + | x7b => compare_byte_x7b b' + | x7c => compare_byte_x7c b' + | x7d => compare_byte_x7d b' + | x7e => compare_byte_x7e b' + | x7f => compare_byte_x7f b' + | x80 => compare_byte_x80 b' + | x81 => compare_byte_x81 b' + | x82 => compare_byte_x82 b' + | x83 => compare_byte_x83 b' + | x84 => compare_byte_x84 b' + | x85 => compare_byte_x85 b' + | x86 => compare_byte_x86 b' + | x87 => compare_byte_x87 b' + | x88 => compare_byte_x88 b' + | x89 => compare_byte_x89 b' + | x8a => compare_byte_x8a b' + | x8b => compare_byte_x8b b' + | x8c => compare_byte_x8c b' + | x8d => compare_byte_x8d b' + | x8e => compare_byte_x8e b' + | x8f => compare_byte_x8f b' + | x90 => compare_byte_x90 b' + | x91 => compare_byte_x91 b' + | x92 => compare_byte_x92 b' + | x93 => compare_byte_x93 b' + | x94 => compare_byte_x94 b' + | x95 => compare_byte_x95 b' + | x96 => compare_byte_x96 b' + | x97 => compare_byte_x97 b' + | x98 => compare_byte_x98 b' + | x99 => compare_byte_x99 b' + | x9a => compare_byte_x9a b' + | x9b => compare_byte_x9b b' + | x9c => compare_byte_x9c b' + | x9d => compare_byte_x9d b' + | x9e => compare_byte_x9e b' + | x9f => compare_byte_x9f b' + | xa0 => compare_byte_xa0 b' + | xa1 => compare_byte_xa1 b' + | xa2 => compare_byte_xa2 b' + | xa3 => compare_byte_xa3 b' + | xa4 => compare_byte_xa4 b' + | xa5 => compare_byte_xa5 b' + | xa6 => compare_byte_xa6 b' + | xa7 => compare_byte_xa7 b' + | xa8 => compare_byte_xa8 b' + | xa9 => compare_byte_xa9 b' + | xaa => compare_byte_xaa b' + | xab => compare_byte_xab b' + | xac => compare_byte_xac b' + | xad => compare_byte_xad b' + | xae => compare_byte_xae b' + | xaf => compare_byte_xaf b' + | xb0 => compare_byte_xb0 b' + | xb1 => compare_byte_xb1 b' + | xb2 => compare_byte_xb2 b' + | xb3 => compare_byte_xb3 b' + | xb4 => compare_byte_xb4 b' + | xb5 => compare_byte_xb5 b' + | xb6 => compare_byte_xb6 b' + | xb7 => compare_byte_xb7 b' + | xb8 => compare_byte_xb8 b' + | xb9 => compare_byte_xb9 b' + | xba => compare_byte_xba b' + | xbb => compare_byte_xbb b' + | xbc => compare_byte_xbc b' + | xbd => compare_byte_xbd b' + | xbe => compare_byte_xbe b' + | xbf => compare_byte_xbf b' + | xc0 => compare_byte_xc0 b' + | xc1 => compare_byte_xc1 b' + | xc2 => compare_byte_xc2 b' + | xc3 => compare_byte_xc3 b' + | xc4 => compare_byte_xc4 b' + | xc5 => compare_byte_xc5 b' + | xc6 => compare_byte_xc6 b' + | xc7 => compare_byte_xc7 b' + | xc8 => compare_byte_xc8 b' + | xc9 => compare_byte_xc9 b' + | xca => compare_byte_xca b' + | xcb => compare_byte_xcb b' + | xcc => compare_byte_xcc b' + | xcd => compare_byte_xcd b' + | xce => compare_byte_xce b' + | xcf => compare_byte_xcf b' + | xd0 => compare_byte_xd0 b' + | xd1 => compare_byte_xd1 b' + | xd2 => compare_byte_xd2 b' + | xd3 => compare_byte_xd3 b' + | xd4 => compare_byte_xd4 b' + | xd5 => compare_byte_xd5 b' + | xd6 => compare_byte_xd6 b' + | xd7 => compare_byte_xd7 b' + | xd8 => compare_byte_xd8 b' + | xd9 => compare_byte_xd9 b' + | xda => compare_byte_xda b' + | xdb => compare_byte_xdb b' + | xdc => compare_byte_xdc b' + | xdd => compare_byte_xdd b' + | xde => compare_byte_xde b' + | xdf => compare_byte_xdf b' + | xe0 => compare_byte_xe0 b' + | xe1 => compare_byte_xe1 b' + | xe2 => compare_byte_xe2 b' + | xe3 => compare_byte_xe3 b' + | xe4 => compare_byte_xe4 b' + | xe5 => compare_byte_xe5 b' + | xe6 => compare_byte_xe6 b' + | xe7 => compare_byte_xe7 b' + | xe8 => compare_byte_xe8 b' + | xe9 => compare_byte_xe9 b' + | xea => compare_byte_xea b' + | xeb => compare_byte_xeb b' + | xec => compare_byte_xec b' + | xed => compare_byte_xed b' + | xee => compare_byte_xee b' + | xef => compare_byte_xef b' + | xf0 => compare_byte_xf0 b' + | xf1 => compare_byte_xf1 b' + | xf2 => compare_byte_xf2 b' + | xf3 => compare_byte_xf3 b' + | xf4 => compare_byte_xf4 b' + | xf5 => compare_byte_xf5 b' + | xf6 => compare_byte_xf6 b' + | xf7 => compare_byte_xf7 b' + | xf8 => compare_byte_xf8 b' + | xf9 => compare_byte_xf9 b' + | xfa => compare_byte_xfa b' + | xfb => compare_byte_xfb b' + | xfc => compare_byte_xfc b' + | xfd => compare_byte_xfd b' + | xfe => compare_byte_xfe b' + | xff => compare_byte_xff b' + end. diff --git a/template-coq/theories/utils/MCCompare.v b/template-coq/theories/utils/MCCompare.v index 250387de3..99c8eea34 100644 --- a/template-coq/theories/utils/MCCompare.v +++ b/template-coq/theories/utils/MCCompare.v @@ -1,4 +1,4 @@ -From Coq Require Import ssreflect Extraction OrderedType Orders Ascii String. +From Coq Require Import ssreflect Extraction OrderedType Orders. From Equations Require Import Equations. Lemma CompareSpec_Proper : Proper (iff ==> iff ==> iff ==> Logic.eq ==> iff) CompareSpec. @@ -109,296 +109,6 @@ Notation bool_compare := BoolOT.compare. Local Ltac trd := cbn in *; try reflexivity; try discriminate. -Module AsciiOT <: UsualOrderedType. - Definition t := ascii. - - Definition eq : t -> t -> Prop := eq. - Definition eq_equiv : Equivalence eq := _. - - Definition compare x y := - let 'Ascii a b c d e f g h := x in - let 'Ascii a' b' c' d' e' f' g' h' := y in - bool_compare a a' - ?? bool_compare b b' - ?? bool_compare c c' - ?? bool_compare d d' - ?? bool_compare e e' - ?? bool_compare f f' - ?? bool_compare g g' - ?? bool_compare h h'. - - Ltac tryone a a' H := - destruct a, a'; simpl in H; try (reflexivity || discriminate). - - Lemma compare_eq : forall x y, compare x y = Eq <-> x = y. - Proof. - destruct x as [a b c d e f g h]. - destruct y as [a' b' c' d' e' f' g' h']. - split. intros H. - tryone a a' H; - tryone b b' H; - tryone c c' H; - tryone d d' H; - tryone e e' H; - tryone f f' H; - tryone g g' H; - tryone h h' H; reflexivity. - intros H; injection H. intros; subst. - destruct a', b', c', d', e', f', g', h'; reflexivity. - Qed. - - Lemma compare_refl x : compare x x = Eq. - Proof. - now apply compare_eq. - Qed. - - Lemma compare_Lt x y : compare x y = Gt <-> compare y x = Lt. - Proof. - destruct x as [a b c d e f g h]. - destruct y as [a' b' c' d' e' f' g' h']. - split. - intros H. - tryone a a' H; - tryone b b' H; - tryone c c' H; - tryone d d' H; - tryone e e' H; - tryone f f' H; - tryone g g' H; - tryone h h' H; try reflexivity. - intros H. - tryone a a' H; - tryone b b' H; - tryone c c' H; - tryone d d' H; - tryone e e' H; - tryone f f' H; - tryone g g' H; - tryone h h' H; try reflexivity. - Qed. - - Definition lt x y := compare x y = Lt. - - Definition compare_spec (x y : t) : CompareSpec (eq x y) (lt x y) (lt y x) (compare x y). - Proof. - case_eq (compare x y). - intros. - - constructor 1. - now apply compare_eq. - - intros. apply CompLt. - destruct x as [a b c d e f g h]. - destruct y as [a' b' c' d' e' f' g' h']. - unfold lt. apply H. - - intros. apply CompGt. red. now apply compare_Lt. - Qed. - - Lemma compare_sym (x y : t) : compare x y = CompOpp (compare y x). - Proof. - destruct (compare_spec x y). - - red in H; subst. now rewrite compare_refl. - - symmetry. rewrite (proj2 (compare_Lt _ _)); auto. - - symmetry. rewrite (proj1 (compare_Lt _ _)); auto. - now apply compare_Lt. - Qed. - - Lemma lt_irreflexive : Irreflexive lt. - Proof. - intro x. destruct x. unfold complement, lt; cbn. - destruct b, b0, b1, b2, b3, b4, b5, b6; cbn; discriminate. - Qed. - - Lemma lt_transitive : Transitive lt. - Proof. - intros [a b c d e f g h] [a' b' c' d' e' f' g' h'] - [a'' b'' c'' d'' e'' f'' g'' h'']. - unfold lt, compare. - intros H1 H2. - destruct a, a', a''; trd; - destruct b, b', b''; trd; - destruct c, c', c''; trd; - destruct d, d', d''; trd; - destruct e, e', e''; trd; - destruct f, f', f''; trd; - destruct g, g', g''; trd; - destruct h, h', h''; trd; - eapply transitive_bool_lt; eassumption. - Qed. - - Global Instance lt_strorder : StrictOrder lt. - Proof. - split. - - apply lt_irreflexive. - - apply lt_transitive. - Qed. - - Definition lt_compat : Proper (eq ==> eq ==> iff) lt. - Proof. - unfold eq. intros x y e z t e'. subst; reflexivity. - Qed. - - (* Bonus *) - Definition eqb (l1 l2 : t) : bool - := match compare l1 l2 with Eq => true | _ => false end. - - Definition eq_leibniz (x y : t) : eq x y -> x = y := id. - - Definition eq_dec (l1 l2 : t) : {l1 = l2}+{l1 <> l2}. - Proof. - destruct (compare l1 l2) eqn:e. - apply compare_eq in e. now left. - right; intros eq. destruct (compare_spec l1 l2); try discriminate. subst. now apply irreflexivity in H. - right; intros eq. destruct (compare_spec l1 l2); try discriminate. subst. now apply irreflexivity in H. - Defined. - -End AsciiOT. - -Notation ascii_compare := AsciiOT.compare. - -Module StringOT <: UsualOrderedType. - Definition t := string. - - Definition eq : t -> t -> Prop := eq. - Definition eq_equiv : Equivalence eq := _. - - Fixpoint compare x y := - match x, y with - | EmptyString, EmptyString => Eq - | String a s, String b s' => - match ascii_compare a b with - | Eq => compare s s' - | Lt => Lt - | Gt => Gt - end - | EmptyString, String _ _ => Lt - | String _ _, EmptyString => Gt - end. - - Definition lt x y : Prop := compare x y = Lt. - - Lemma compare_eq : forall x y : string, compare x y = Eq <-> eq x y. - Proof. - split. - induction x in y |- *. - + destruct y. reflexivity. - discriminate. - + destruct y. discriminate. - simpl. destruct (AsciiOT.compare_spec a a0); red in H; subst; try discriminate. - intros eq. red. f_equal; auto. now apply IHx. - + intros ->. - induction y. reflexivity. - simpl. now rewrite AsciiOT.compare_refl. - Qed. - - Lemma compare_refl x : compare x x = Eq. - Proof. - now apply compare_eq. - Qed. - - Lemma compare_lt : forall x y : string, compare x y = Lt <-> compare y x = Gt. - Proof. - split. - - induction x in y |- *. - + destruct y. discriminate. - reflexivity. - + destruct y. discriminate. - simpl. destruct (AsciiOT.compare_spec a a0); intros H'; red in H; subst; try discriminate. - rewrite AsciiOT.compare_refl //; auto. - apply AsciiOT.compare_Lt in H. now rewrite H. - - induction x in y |- *. - + destruct y. discriminate. - reflexivity. - + destruct y. discriminate. - simpl. rewrite (AsciiOT.compare_sym a a0). - destruct (AsciiOT.compare_spec a0 a); cbn; try congruence. - auto. - Qed. - - Definition compare_spec (x y : t) : CompareSpec (eq x y) (lt x y) (lt y x) (compare x y). - Proof. - case_eq (compare x y); intros H. - - apply CompEq. now apply compare_eq. - - apply CompLt; assumption. - - apply CompGt. red. now apply compare_lt. - Qed. - - #[local] Instance lt_transitive : Transitive lt. - Proof. - red. unfold lt. - intro s; induction s. - - induction y; cbn. - + intuition. - + intros [|]. discriminate. reflexivity. - - intros [|y1 y2]. discriminate. - intros [|z1 z2]. discriminate. - cbn. case_eq (ascii_compare a y1); try discriminate. - + intro Ha; apply AsciiOT.compare_eq in Ha; subst. - destruct (ascii_compare y1 z1); try discriminate. - intros; eauto. reflexivity. - + intros Ha _. case_eq (ascii_compare y1 z1); try discriminate. - intros Hy1; apply AsciiOT.compare_eq in Hy1; subst. now rewrite Ha. - intro Hy1. eapply AsciiOT.lt_transitive in Ha. - specialize (Ha Hy1). now rewrite Ha. - Qed. - - Lemma compare_sym (x y : string) : compare x y = CompOpp (compare y x). - Proof. - destruct (compare_spec x y). - - red in H; subst. - rewrite (proj2 (compare_eq _ _)); auto. reflexivity. - - rewrite (proj1 (compare_lt _ _)); auto. - - rewrite (proj2 (compare_lt _ _)); auto. - red in H. - now apply compare_lt. - Qed. - - Lemma compare_trans (x y z : string) c : compare x y = c -> compare y z = c -> compare x z = c. - Proof. - destruct (compare_spec x y); subst; intros <-; - destruct (compare_spec y z); subst; try congruence. - eapply transitivity in H0. 2:eassumption. red in H0. subst. now rewrite compare_refl. - eapply transitivity in H0. 2:eassumption. now red in H0. - eapply transitivity in H. 2:eassumption. - now apply compare_lt in H. - Qed. - - Definition lt_irreflexive : Irreflexive lt. - Proof. - intro x. red; unfold lt. - now rewrite compare_refl. - Qed. - - Global Instance lt_strorder : StrictOrder lt. - Proof. - split. - - apply lt_irreflexive. - - apply lt_transitive. - Qed. - - Definition lt_compat : Proper (eq ==> eq ==> iff) lt. - Proof. - unfold eq. intros x y e z t e'. subst; reflexivity. - Qed. - - (* Bonus *) - Definition eqb (l1 l2 : t) : bool := - match compare l1 l2 with Eq => true | _ => false end. - - Definition eq_leibniz (x y : t) : eq x y -> x = y := id. - - Definition eq_dec (l1 l2 : t) : {l1 = l2}+{l1 <> l2}. - Proof. - destruct (compare l1 l2) eqn:e. - apply compare_eq in e. now left. - right; intros eq. destruct (compare_spec l1 l2); try discriminate. subst. now apply irreflexivity in H. - right; intros eq. destruct (compare_spec l1 l2); try discriminate. subst. now apply irreflexivity in H. - Defined. - -End StringOT. - -Notation string_compare := StringOT.compare. -Notation string_compare_eq := StringOT.compare_eq. -Notation CompareSpec_string := StringOT.compare_spec. - Module ListOrderedType (A : UsualOrderedType) <: UsualOrderedType. Definition t := list A.t. Import List. Import ListNotations. @@ -413,12 +123,7 @@ Module ListOrderedType (A : UsualOrderedType) <: UsualOrderedType. Definition eq : t -> t -> Prop := eq. Definition eq_equiv : Equivalence eq := _. - - Definition eq_dec (l1 l2 : t) : {l1 = l2}+{l1 <> l2}. - Proof. - decide equality. apply A.eq_dec. - Defined. - + Inductive lt_ : t -> t -> Prop := | lt_nil_cons hd tl : lt_ [] (hd :: tl) | lt_cons_cons_hd hd tl hd' tl' : A.lt hd hd' -> lt_ (hd :: tl) (hd' :: tl') @@ -511,5 +216,22 @@ Module ListOrderedType (A : UsualOrderedType) <: UsualOrderedType. Definition eq_leibniz (x y : t) : eq x y -> x = y := id. + Program Definition eqb_dec (x y : t) : { x = y } + { x <> y } := + match eqb x y with + | true => left _ + | false => right _ + end. + Next Obligation. + unfold eqb in Heq_anonymous. + destruct (compare x y) eqn:hc; try congruence. + now eapply compare_eq in hc. + Qed. + Next Obligation. + unfold eqb in Heq_anonymous. + rewrite compare_refl in Heq_anonymous. now discriminate. + Qed. + + Global Instance eq_dec : EqDec t := { eq_dec := eqb_dec }. + End ListOrderedType. diff --git a/template-coq/theories/utils/MCString.v b/template-coq/theories/utils/MCString.v index 1ec07e68b..c5d6ef402 100644 --- a/template-coq/theories/utils/MCString.v +++ b/template-coq/theories/utils/MCString.v @@ -1,11 +1,12 @@ -From Coq Require Import String Decimal DecimalString ZArith. -From MetaCoq.Template Require Import MCCompare. +From Coq Require Import ssreflect ssrbool Decimal DecimalString ZArith. +From MetaCoq.Template Require Import MCCompare bytestring ReflectEq. -Local Open Scope string_scope. +Local Open Scope bs. +Notation string := String.t. (** When defining [Show] instance for your own datatypes, you sometimes need to start a new line for better printing. [nl] is a shorthand for it. *) -Definition nl : string := String (Ascii.ascii_of_nat 10) EmptyString. +Definition nl : string := String.String "010"%byte String.EmptyString. Definition string_of_list_aux {A} (f : A -> string) (sep : string) (l : list A) : string := let fix aux l := @@ -25,8 +26,23 @@ Definition print_list {A} (f : A -> string) (sep : string) (l : list A) : string Definition parens (top : bool) (s : string) := if top then s else "(" ++ s ++ ")". +Fixpoint string_of_uint n := + match n with + | Nil => "" + | D0 n => "0" ++ string_of_uint n + | D1 n => "1" ++ string_of_uint n + | D2 n => "2" ++ string_of_uint n + | D3 n => "3" ++ string_of_uint n + | D4 n => "4" ++ string_of_uint n + | D5 n => "5" ++ string_of_uint n + | D6 n => "6" ++ string_of_uint n + | D7 n => "7" ++ string_of_uint n + | D8 n => "8" ++ string_of_uint n + | D9 n => "9" ++ string_of_uint n + end. + Definition string_of_nat n : string := - DecimalString.NilEmpty.string_of_uint (Nat.to_uint n). + string_of_uint (Nat.to_uint n). #[global] Hint Resolve String.string_dec : eq_dec. @@ -40,15 +56,3 @@ Definition string_of_Z (z : Z) : string := | Zpos p => string_of_positive p | Zneg p => "-" ++ string_of_positive p end. - -Definition eq_string s s' := - match string_compare s s' with - | Eq => true - | _ => false - end. - -Lemma eq_string_refl x : is_true (eq_string x x). -Proof. - unfold eq_string. - now rewrite (proj2 (string_compare_eq x x) eq_refl). -Qed. diff --git a/template-coq/theories/utils/MCUtils.v b/template-coq/theories/utils/MCUtils.v index c36249ab3..95e65de6e 100644 --- a/template-coq/theories/utils/MCUtils.v +++ b/template-coq/theories/utils/MCUtils.v @@ -12,6 +12,8 @@ Require Export MCPrelude MCSquash MCRelations MCString + ReflectEq + bytestring . Tactic Notation "destruct" "?" := @@ -186,8 +188,8 @@ Proof. destruct q. apply Eqdep_dec.UIP_refl_bool. Qed. -Axiom todo : String.string -> forall {A}, A. +Axiom todo : string -> forall {A}, A. Ltac todo s := exact (todo s). From Coq Require Import Extraction. -Extract Constant todo => "fun s -> failwith (String.concat """" (List.map (String.make 1) s))". +Extract Constant todo => "fun s -> failwith (Caml_bytestring.caml_string_of_bytestring s)". diff --git a/template-coq/theories/utils/ReflectEq.v b/template-coq/theories/utils/ReflectEq.v new file mode 100644 index 000000000..5dc72dfbe --- /dev/null +++ b/template-coq/theories/utils/ReflectEq.v @@ -0,0 +1,205 @@ +From Coq Require Import ssreflect ssrbool List. +From Equations Require Import Equations. +Import ListNotations. + +(* Some reflection / EqDec lemmata *) + +Inductive reflectProp (P : Prop) : bool -> Prop := + | reflectP : P -> reflectProp P true + | reflectF : ~ P -> reflectProp P false. +Derive Signature for reflectProp. + +Lemma elimP {T} {b} : reflectProp T b -> b -> T. +Proof. intros [] => //. Qed. +Coercion elimP : reflectProp >-> Funclass. + +Lemma introP {T} {b} : reflectProp T b -> T -> b. +Proof. intros [] => //. Qed. + +Hint View for move/ introP|2. +Hint View for apply/ introP|2. + +Lemma reflect_reflectProp {P b} : reflect P b -> reflectProp P b. +Proof. + intros []; constructor; auto. +Qed. + +(** If one really wants a computational version. *) +Remark reflectProp_reflect {P b} : reflectProp P b -> reflect P b. +Proof. + now destruct b; intros H; constructor; depelim H. +Defined. + +Lemma reflect_reflectProp_1 {A} {P : A -> Prop} {b} : (forall x, reflect (P x) (b x)) -> (forall x, reflectProp (P x) (b x)). +Proof. + intros f x. now apply reflect_reflectProp. +Qed. + +Lemma reflect_reflectProp_2 {A B} {P : A -> B -> Prop} {b} : (forall x y, reflect (P x y) (b x y)) -> (forall x y, reflectProp (P x y) (b x y)). +Proof. + intros f x y. now apply reflect_reflectProp. +Qed. + +Class ReflectEq A := { + eqb : A -> A -> bool ; + eqb_spec : forall x y : A, reflectProp (x = y) (eqb x y) (* Prevent using reflect in computational content *) +}. +Arguments eqb : simpl never. +Infix "==" := eqb (at level 70). + +(* If one needs to eliminate a decidable equality in Type, e.g. when working with PCUIC derivations. *) +Lemma eqb_specT {A} {HR : ReflectEq A} (x y : A) : reflect (x = y) (x == y). +Proof. + eapply reflectProp_reflect. apply eqb_spec. +Qed. + +Lemma eqb_eq {A} `{ReflectEq A} (x y : A) : x == y -> x = y. +Proof. + elim: eqb_spec; auto. + discriminate. +Qed. + +Lemma eqb_refl {A} {R : ReflectEq A} (x : A) : x == x. +Proof. + destruct (eqb_spec x x); auto. +Qed. + +Lemma neqb {A} {R : ReflectEq A} (x y : A) : ~~ (x == y) <-> x <> y. +Proof. + destruct (eqb_spec x y); auto; subst; intuition auto. +Qed. + +#[global, program] Instance ReflectEq_EqDec {A} (R : ReflectEq A) : EqDec A := { + eq_dec := fun x y => + match eqb x y with + | true => left _ + | false => right _ + end }. +Next Obligation. + now apply eqb_eq. +Qed. +Next Obligation. + rewrite eqb_refl in Heq_anonymous. + discriminate. +Qed. + +Definition eq_dec_to_bool {A} `{EqDec A} x y := + match eq_dec x y with + | left _ => true + | right _ => false + end. + +(* Not an instance to avoid loops and making boolean definitions depend on sumbool ones *) +#[global, program] +Definition EqDec_ReflectEq A {E : EqDec A} : ReflectEq A := + {| eqb := eq_dec_to_bool |}. +Next Obligation. +Proof. + unfold eq_dec_to_bool. + destruct (eq_dec x y). + all: constructor ; assumption. +Qed. + +Ltac nodec := + let bot := fresh "bot" in + try solve [ constructor ; intro bot ; inversion bot ; subst ; tauto ]. + +Definition eq_option {A} (eqA : A -> A -> bool) (u v : option A) : bool := + match u, v with + | Some u, Some v => eqA u v + | None, None => true + | _, _ => false + end. + +#[global] Instance reflect_option : forall {A}, ReflectEq A -> ReflectEq (option A). +Proof. + intros A RA. refine {| eqb := eq_option eqb |}. + intros x y. destruct x, y. + all: cbn. + all: try solve [ constructor ; easy ]. + destruct (eqb_spec a a0) ; nodec. + constructor. f_equal. assumption. +Defined. + +Fixpoint eq_list {A} (eqA : A -> A -> bool) (l l' : list A) : bool := + match l, l' with + | a :: l, a' :: l' => + if eqA a a' then eq_list eqA l l' + else false + | [], [] => true + | _, _ => false + end. + +#[global] Instance reflect_list : forall {A}, ReflectEq A -> ReflectEq (list A). +Proof. + intros A RA. refine {| eqb := eq_list eqb |}. + intro x. induction x ; intro y ; destruct y. + - cbn. constructor. reflexivity. + - cbn. constructor. discriminate. + - cbn. constructor. discriminate. + - cbn. destruct (eqb_spec a a0) ; nodec. + destruct (IHx y) ; nodec. + subst. constructor. reflexivity. +Defined. + +#[global] Instance reflect_nat : ReflectEq nat := { + eqb_spec := reflect_reflectProp_2 PeanoNat.Nat.eqb_spec +}. + + +Definition eq_bool b1 b2 : bool := + if b1 then b2 else negb b2. + +Local Obligation Tactic := idtac. + +#[global, program] Instance reflect_bool : ReflectEq bool := { + eqb := eq_bool +}. +Next Obligation. + intros x y. unfold eq_bool. + destruct x, y. + all: constructor. + all: try reflexivity. + all: discriminate. +Defined. + +Definition eq_sig_true {A f} `{ReflectEq A} (x y : { z : A | f z = true }) : bool := + proj1_sig x == proj1_sig y. + +#[global, program] Instance reflect_sig_true {A f} `{ReflectEq A} : ReflectEq ({ z : A | f z = true }) := { + eqb := eq_sig_true +}. +Next Obligation. + intros A f RA. intros [x hx] [y hy]. unfold eq_sig_true; cbn. + destruct (eqb_spec x y) ; nodec. subst. + constructor. pose proof (uip hx hy). subst. reflexivity. +Defined. + + +Definition eq_prod {A B} (eqA : A -> A -> bool) (eqB : B -> B -> bool) x y := + let '(a1, b1) := x in + let '(a2, b2) := y in + if eqA a1 a2 then eqB b1 b2 + else false. + +#[global, program] Instance reflect_prod : forall {A B}, ReflectEq A -> ReflectEq B -> ReflectEq (A * B) := { + eqb := eq_prod eqb eqb +}. +Next Obligation. + intros A B RA RB [x y] [u v]. + unfold eq_prod. + destruct (eqb_spec x u) ; nodec. + destruct (eqb_spec y v) ; nodec. + subst. constructor. reflexivity. +Defined. + +Lemma eq_prod_refl : + forall A B (eqA : A -> A -> bool) (eqB : B -> B -> bool), + (forall a, eqA a a) -> + (forall b, eqB b b) -> + forall p, eq_prod eqA eqB p p. +Proof. + intros A B eqA eqB eqA_refl eqB_refl [a b]. + simpl. rewrite eqA_refl. apply eqB_refl. +Qed. + diff --git a/template-coq/theories/utils/bytestring.v b/template-coq/theories/utils/bytestring.v new file mode 100644 index 000000000..d8f0cad4b --- /dev/null +++ b/template-coq/theories/utils/bytestring.v @@ -0,0 +1,499 @@ +(** + * Copyright (C) 2020 BedRock Systems, Inc. + * All rights reserved. + * + * SPDX-License-Identifier: LGPL-2.1 WITH BedRock Exception for use over network, + * see repository root for details. + *) + +Require Coq.Strings.String ssrbool. +Require Import ssreflect. +Require Import Coq.NArith.NArith. +Require Import Coq.micromega.Lia. +From Equations Require Import Equations. +Set Primitive Projections. +Set Default Proof Using "Type". +From MetaCoq.Template Require Import MCCompare ReflectEq. +From MetaCoq.Template Require ByteCompare ByteCompareSpec. +(** bytes *) + +Definition byte_parse (b : Byte.byte) : Byte.byte := b. +Definition byte_print (b : Byte.byte) : Byte.byte := b. + +Delimit Scope byte_scope with byte. +String Notation Byte.byte byte_parse byte_print : byte_scope. +Bind Scope byte_scope with Byte.byte. + +Declare Scope bs_scope. +Delimit Scope bs_scope with bs. + +(** bytestrings *) +Module String. + Inductive t : Set := + | EmptyString + | String (_ : Byte.byte) (_ : t). + Bind Scope bs_scope with t. + + Fixpoint print (b : t) : list Byte.byte := + match b with + | EmptyString => nil + | String b bs => b :: print bs + end. + + Fixpoint parse (b : list Byte.byte) : t := + match b with + | nil => EmptyString + | List.cons b bs => String b (parse bs) + end. + + Lemma print_parse_inv: + forall x : t, parse (print x) = x. + Proof. + induction x; simpl; intros; auto. + f_equal; auto. + Qed. + + Fixpoint append (x y : t) : t := + match x with + | EmptyString => y + | String x xs => String x (append xs y) + end. + + Notation "x ++ y" := (append x y) : bs_scope. + + Fixpoint to_string (b : t) : String.string := + match b with + | EmptyString => Strings.String.EmptyString + | String x xs => Strings.String.String (Ascii.ascii_of_byte x) (to_string xs) + end. + + Fixpoint of_string (b : String.string) : t := + match b with + | Strings.String.EmptyString => EmptyString + | Strings.String.String x xs => String (Ascii.byte_of_ascii x) (of_string xs) + end%bs. + + Fixpoint rev (acc s : t) : t := + match s with + | EmptyString => acc + | String s ss => rev (String s acc) ss + end. + + (** *** Substrings *) + + (** [substring n m s] returns the substring of [s] that starts + at position [n] and of length [m]; + if this does not make sense it returns [""] *) + + Fixpoint substring (n m : nat) (s : t) : t := + match n, m, s with + | 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. + + Fixpoint prefix (s1 s2 : t) {struct s1} : bool := + match s1 with + | EmptyString => true + | String x xs => + match s2 with + | EmptyString => false + | String y ys => + if Byte.eqb x y then prefix xs ys + else false + end + end%bs. + + Fixpoint index (n : nat) (s1 s2 : t) {struct s2} : option nat := + match s2 with + | EmptyString => + match n with + | 0 => match s1 with + | EmptyString => Some 0 + | String _ _ => None + end + | S _ => None + end + | String _ s2' => + match n with + | 0 => + if prefix s1 s2 + then Some 0 + else match index 0 s1 s2' with + | Some n0 => Some (S n0) + | None => None + end + | S n' => match index n' s1 s2' with + | Some n0 => Some (S n0) + | None => None + end + end + end%bs. + + Fixpoint length (l : t) : nat := + match l with + | EmptyString => 0 + | String _ l => S (length l) + end. + + Local Fixpoint contains (start: nat) (keys: list t) (fullname: t) :bool := + match keys with + | List.cons kh ktl => + match index start kh fullname with + | Some n => contains (n + length kh) ktl fullname + | None => false + end + | List.nil => true + end. + + Fixpoint eqb (a b : t) : bool := + match a , b with + | EmptyString , EmptyString => true + | String x xs , String y ys => + if ByteCompare.eqb x y then eqb xs ys else false + | _ , _ => false + end. + + Fixpoint compare (xs ys : t) : comparison := + match xs , ys with + | EmptyString , EmptyString => Eq + | EmptyString , _ => Lt + | _ , EmptyString => Gt + | String x xs , String y ys => + match ByteCompare.compare x y with + | Eq => compare xs ys + | x => x + end + end. + + Lemma eqb_compare xs ys : eqb xs ys = match compare xs ys with Eq => true | _ => false end. + Proof. + induction xs in ys |- *; destruct ys => /= //. + rewrite ByteCompareSpec.eqb_compare. + destruct ByteCompare.compare => //. + Qed. + + Fixpoint concat (sep : t) (s : list t) : t := + match s with + | nil => EmptyString + | cons s nil => s + | cons s xs => s ++ sep ++ concat sep xs + end. + +End String. + +Definition bs := String.t. +Notation string := String.t. + +Bind Scope bs_scope with bs. + +String Notation String.t String.parse String.print : bs_scope. + +Notation "x ++ y" := (String.append x y) : bs_scope. + +Import String. + +(** comparison *) +Require Import Orders Coq.Structures.OrderedType. + +Lemma to_N_inj : forall x y, Byte.to_N x = Byte.to_N y <-> x = y. +Proof. + split. + 2: destruct 1; reflexivity. + intros. + assert (Some x = Some y). + { do 2 rewrite <- Byte.of_to_N. + destruct H. reflexivity. } + injection H0. auto. +Qed. + +Module OT_byte <: OrderedType.OrderedType with Definition t := Byte.byte. + Definition t := Byte.byte. + Definition eq := @Logic.eq t. + Definition lt := fun l r => ByteCompare.compare l r = Lt. + Theorem eq_refl : forall x : t, eq x x. + Proof. + intros; apply eq_refl. + Qed. + Theorem eq_sym : forall x y : t, eq x y -> eq y x. + Proof. + apply eq_sym. + Qed. + Theorem eq_trans : forall x y z : t, eq x y -> eq y z -> eq x z. + Proof. + apply eq_trans. + Qed. + Theorem lt_trans : forall x y z : t, lt x y -> lt y z -> lt x z. + Proof. + exact ByteCompareSpec.lt_trans. + Qed. + Theorem lt_not_eq : forall x y : t, lt x y -> not (eq x y). + Proof. + apply ByteCompareSpec.lt_not_eq. + Qed. + Definition compare (x y : t) : OrderedType.Compare lt eq x y. + refine ( + match ByteCompare.compare x y as X return ByteCompare.compare x y = X -> OrderedType.Compare lt eq x y with + | Eq => fun pf => OrderedType.EQ _ + | Lt => fun pf => OrderedType.LT pf + | Gt => fun pf => OrderedType.GT _ + end (Logic.eq_refl)). + now apply ByteCompareSpec.compare_eq in pf. + rewrite ByteCompareSpec.compare_opp in pf. + apply CompOpp_iff in pf. apply pf. + Defined. + + Definition eq_dec : forall x y : t, {eq x y} + {not (eq x y)} := Classes.eq_dec. +End OT_byte. + +Global Instance byte_eqdec : Classes.EqDec Byte.byte := _. + +Module StringOT <: UsualOrderedType. + Definition t := string. + + Definition eq : t -> t -> Prop := eq. + Definition eq_equiv : Equivalence eq := _. + + Definition compare := String.compare. + Definition lt x y : Prop := compare x y = Lt. + + Theorem compare_spec : forall x y, CompareSpec (x = y) (lt x y) (lt y x) (compare x y). + Proof. + induction x; destruct y; simpl. + - constructor; reflexivity. + - constructor. reflexivity. + - constructor. reflexivity. + - unfold lt; simpl. destruct (ByteCompareSpec.compare_spec b b0); simpl. + + subst. destruct (IHx y); constructor; eauto. + congruence. now rewrite ByteCompareSpec.compare_eq_refl. + + constructor; auto. + + red in H. rewrite H. constructor; auto. + Qed. + + Theorem eq_refl : forall x : t, eq x x. + Proof. + reflexivity. + Qed. + Theorem eq_sym : forall x y : t, eq x y -> eq y x. + Proof. + eapply eq_sym. + Qed. + Theorem eq_trans : forall x y z : t, eq x y -> eq y z -> eq x z. + Proof. + eapply eq_trans. + Qed. + Theorem lt_trans : forall x y z : t, lt x y -> lt y z -> lt x z. + Proof. + unfold lt. + induction x; destruct y; simpl; try congruence. + - destruct z; congruence. + - destruct (ByteCompareSpec.compare_spec b b0); subst. + + destruct z; auto. + destruct (ByteCompare.compare b0 b); auto. + eauto. + + destruct z; auto. + red in H. + destruct (ByteCompareSpec.compare_spec b0 b1). + * subst. intros _. now rewrite H. + * generalize (OT_byte.lt_trans _ _ _ H H0). unfold OT_byte.lt. + intro X; rewrite X. auto. + * congruence. + + congruence. + Qed. + Theorem lt_not_eq : forall x y : t, lt x y -> not (eq x y). + Proof. + unfold lt, eq. + intros. intro. subst. + induction y; simpl in *; try congruence. + rewrite ByteCompareSpec.compare_eq_refl in H. auto. + Qed. + + #[global] Program Instance reflect_eq_string : ReflectEq t := { + eqb := eqb + }. + Next Obligation. + rename x into s, y into s'. + destruct (eqb s s') eqn:e; constructor. + - rewrite String.eqb_compare in e. fold (compare s s') in e. + now destruct (compare_spec s s'). + - rewrite String.eqb_compare in e. + fold (compare s s') in e. + destruct (compare_spec s s') => //. + now apply lt_not_eq. now apply not_eq_sym, lt_not_eq. + Qed. + + Definition eq_dec : forall x y : t, {eq x y} + {not (eq x y)} := Classes.eq_dec. + + Lemma compare_eq : forall x y : string, compare x y = Eq <-> eq x y. + Proof. + intros. + destruct (compare_spec x y); intuition auto; try congruence. + - apply lt_not_eq in H; contradiction. + - apply lt_not_eq in H. symmetry in H0. contradiction. + Qed. + + Lemma compare_refl x : compare x x = Eq. + Proof. + now apply compare_eq. + Qed. + + Lemma compare_lt : forall x y : string, compare x y = Lt <-> compare y x = Gt. + Proof. + intros x y. + destruct (compare_spec x y). + all:split; try congruence; subst. + - intros hc. + fold (compare y y) in hc. now rewrite compare_refl in hc. + - intros _. + destruct (compare_spec y x); subst. + * eapply lt_not_eq in H. elim H; reflexivity. + * eapply lt_trans in H; try eassumption. + eapply lt_not_eq in H. elim H; reflexivity. + * reflexivity. + Qed. + + #[local] Instance lt_transitive : Transitive lt. + Proof. + red. eapply lt_trans. + Qed. + + Lemma compare_sym (x y : string) : compare x y = CompOpp (compare y x). + Proof. + destruct (compare_spec x y). + - subst. + rewrite (proj2 (compare_eq _ _)); auto. reflexivity. + - rewrite (proj1 (compare_lt _ _)); auto. + - rewrite (proj2 (compare_lt _ _)); auto. + red in H. + now apply compare_lt. + Qed. + + Lemma compare_trans (x y z : string) c : compare x y = c -> compare y z = c -> compare x z = c. + Proof. + destruct (compare_spec x y); subst; intros <-; + destruct (compare_spec y z); subst; try congruence. + eapply transitivity in H0. 2:eassumption. intros; apply H0. + eapply transitivity in H. 2:eassumption. + now apply compare_lt in H. + Qed. + + Definition lt_irreflexive : Irreflexive lt. + Proof. + intro x. red; unfold lt. + now rewrite compare_refl. + Qed. + + Global Instance lt_strorder : StrictOrder lt. + Proof. + split. + - apply lt_irreflexive. + - apply lt_transitive. + Qed. + + Definition lt_compat : Proper (eq ==> eq ==> iff) lt. + Proof. + unfold eq. intros x y e z t e'. subst; reflexivity. + Qed. + + Definition eq_leibniz (x y : t) : eq x y -> x = y := id. + +End StringOT. + +Notation string_compare := StringOT.compare. +Notation string_compare_eq := StringOT.compare_eq. +Notation CompareSpec_string := StringOT.compare_spec. + +(** To perform efficient pretty printing, one needs to use a tree structure + to avoid quadratic overhead of appending strings. *) +Module Tree. + Local Open Scope bs_scope. + Inductive t := + | string : String.t -> t + | append : t -> t -> t. + + Coercion string : String.t >-> t. + + Infix "++" := append. + + Fixpoint to_rev_list_aux t acc := + match t with + | string s => cons s acc + | append s s' => to_rev_list_aux s' (to_rev_list_aux s acc) + end. + + Fixpoint to_string_acc acc l := + match l with + | nil => acc + | cons s xs => to_string_acc (String.append s acc) xs + end. + + Definition to_string t := + let l := to_rev_list_aux t nil in + to_string_acc "" l. + + (* Definition test := "a" ++ "b" ++ "v" ++ "c". + Eval compute in to_string test. *) + + (* Fixpoint to_string_acc t acc := + match t with + | string s => String.append s acc + | append s s' => to_string_acc s (to_string_acc s' acc) + end. + + Definition to_string t := to_string_acc t "". *) + + Definition string_of_list_aux {A} (f : A -> t) (sep : t) (l : list A) := + let fix aux l := + match l return t with + | nil => "" + | cons a nil => f a + | cons a l => f a ++ sep ++ aux l + end + in aux l. + + Definition string_of_list {A} (f : A -> t) l := + "[" ++ string_of_list_aux f "," l ++ "]". + + Definition print_list {A} (f : A -> t) (sep : t) (l : list A) : t := + string_of_list_aux f sep l. + + Fixpoint concat (sep : t) (s : list t) : t := + match s with + | nil => EmptyString + | cons s nil => s + | cons s xs => s ++ sep ++ concat sep xs + end. + + Definition parens (top : bool) (s : t) := + if top then s else "(" ++ s ++ ")". + +End Tree. + +(* Tests *) +(* + +Local Open Scope bs_scope. +Import String. +Definition x := Eval compute in Tree.to_string (Tree.concat (Tree.string "") (List.repeat (Tree.string "frliqhgrnvcrlkejflqrjfljkl") 10000)). +Definition y := Eval compute in String.concat "" (List.repeat "frliqhgrnvcrlkejflqrjfljkl" 10000) ++ "'". + +Time Eval compute in String.compare x y. (* 0.02s *) + +Import String. +Fixpoint String.compare' (xs ys : bs) : comparison := + match xs , ys with + | EmptyString, EmptyString => Eq + | EmptyString , _ => Lt + | _ , EmptyString => Gt + | String x xs , String y ys => + match compare x y with + | Eq => String.compare' xs ys + | x => x + end + end. + +Time Eval vm_compute in String.compare x y. (* 0.14s *) +Time Eval vm_compute in String.compare' x y. (* 0.03s~0.04s, i.e. more than 3 times faster *) +*) \ No newline at end of file diff --git a/template-coq/update_plugin.sh b/template-coq/update_plugin.sh index ee9c2afba..171799dd3 100755 --- a/template-coq/update_plugin.sh +++ b/template-coq/update_plugin.sh @@ -1,6 +1,8 @@ #!/usr/bin/env bash SED=`which gsed || which sed` -TOCOPY="ast_denoter.ml ast_quoter.ml denoter.ml plugin_core.ml plugin_core.mli reification.ml quoter.ml run_extractable.ml run_extractable.mli tm_util.ml" +TOCOPY="ast_denoter.ml ast_quoter.ml denoter.ml plugin_core.ml plugin_core.mli reification.ml quoter.ml run_extractable.ml run_extractable.mli tm_util.ml \ + caml_nat.mli caml_nat.ml caml_byte.ml caml_byte.mli caml_bytestring.ml" +SED=`which gsed || which sed` # Test if gen-src is older than src if [[ ! -f "gen-src/denoter.ml" || diff --git a/test-suite/bug6.v b/test-suite/bug6.v index 59070a02e..45e724e4e 100644 --- a/test-suite/bug6.v +++ b/test-suite/bug6.v @@ -1,4 +1,5 @@ Require Import MetaCoq.Template.Loader. +Import bytestring. Definition I (t:Type) (x:t) : t := x. Definition II := I (forall t:Type, t -> t) I. diff --git a/test-suite/bugkncst.v b/test-suite/bugkncst.v index ef0767b26..e00e73606 100644 --- a/test-suite/bugkncst.v +++ b/test-suite/bugkncst.v @@ -13,8 +13,6 @@ Function Plus1 (n: nat) {measure id n} : nat := Defined. Require Import Coq.Lists.List. -Require Import Coq.Strings.String. -Require Import Coq.Strings.Ascii. Require Import Coq.Bool.Bool. Require Import MetaCoq.Template.Loader. Require Import MetaCoq.Template.All. @@ -41,7 +39,7 @@ Fixpoint pocc_term (n:nat) (t:term): bool := | tLambda _ ty t => pocc_term n t || pocc_term n ty | tLetIn _ dfn ty t => pocc_term n dfn || pocc_term n t || pocc_term n ty | tApp fn args => pocc_term n fn || fold_left orb (map (pocc_term n) args) false - | tConst nm _ => if kername_eq_dec str nm then true else false + | tConst nm _ => if eqb str nm then true else false | tCase _ ty mch brs => existsb (pocc_term n) (pparams ty) || pocc_term n (preturn ty) || pocc_term n mch || @@ -55,7 +53,7 @@ Fixpoint pocc_term (n:nat) (t:term): bool := (** does [tConst str] occur anywhere in a program? **) Definition bound_global_decl (d : kername * global_decl) : bool := - if kername_eq_dec str (fst d) then true else false. + if eq_kername str (fst d) then true else false. Definition bound_program (p : program) := List.existsb bound_global_decl (fst p).(declarations). @@ -70,7 +68,7 @@ Definition pocc_program p := pocc_term 2000 (snd p) || List.existsb pocc_global_ End occ_term_Sec. -Require Import List String. Import ListNotations. Open Scope string_scope. +Require Import List. Import ListNotations. Open Scope bs_scope. (* MetaCoq Test Unquote (tConst (MPfile ["Nat"; "PeanoNat"; "Arith"; "Coq"], "pred") []) . *) Eval vm_compute in (eq_refl : pocc_program (MPfile ["Nat"; "PeanoNat"; "Arith"; "Coq"], "pred") p_Plus1 = false). diff --git a/test-suite/extractable.v b/test-suite/extractable.v index c101ac890..458c4f6bd 100644 --- a/test-suite/extractable.v +++ b/test-suite/extractable.v @@ -5,7 +5,7 @@ From MetaCoq.Template Require Import From MetaCoq.Template.TemplateMonad Require Import Common Extractable. -Local Open Scope string_scope. +Local Open Scope bs_scope. Import MCMonadNotation. Notation "<% x %>" := (ltac:(let p y := exact y in quote_term x p)) @@ -24,7 +24,7 @@ MetaCoq Run (fun t => tmPrint t)). MetaCoq Run - (tmBind (tmDefinition "two" None <% 1 + 1 %>) + (tmBind (tmDefinition "two"%bs None <% 1 + 1 %>) (fun kn => tmPrint (Ast.tConst kn nil))). MetaCoq Run diff --git a/test-suite/issue27.v b/test-suite/issue27.v index e95fb3c1c..8414b0fa4 100644 --- a/test-suite/issue27.v +++ b/test-suite/issue27.v @@ -1,4 +1,4 @@ Require Import MetaCoq.Template.All. -Require Export String List. -Open Scope string. +Require Export List. +Open Scope bs_scope. MetaCoq Run (tmLemma "test" (@nil nat = @nil nat)). diff --git a/test-suite/issue28.v b/test-suite/issue28.v index a488c0160..efa3baea4 100644 --- a/test-suite/issue28.v +++ b/test-suite/issue28.v @@ -1,6 +1,6 @@ -Require Import MetaCoq.Template.All MetaCoq.Template.Pretty. -Require Export String List. -Open Scope string. +Require Import MetaCoq.Template.All MetaCoq.Template.utils.bytestring MetaCoq.Template.Pretty. +Require Export List. +Open Scope bs_scope. Import ListNotations. Import MCMonadNotation. diff --git a/test-suite/opaque.v b/test-suite/opaque.v index 001f4a692..5c76a50ef 100644 --- a/test-suite/opaque.v +++ b/test-suite/opaque.v @@ -1,11 +1,11 @@ -From Coq Require Import String List PeanoNat. +From Coq Require Import List Nat. From MetaCoq.Template Require Import All. -Import MCMonadNotation String ListNotations. +Import MCMonadNotation ListNotations. Definition foo : nat. exact 0. Qed. -Local Open Scope string_scope. +Local Open Scope bs_scope. MetaCoq Quote Recursively Definition foo_syn := foo. MetaCoq Quote Recursively Definition comm_syn := PeanoNat.Nat.add_comm. diff --git a/test-suite/plugin-demo/src/g_demo_plugin.mlg b/test-suite/plugin-demo/src/g_demo_plugin.mlg index 7f9797c32..1cc2ca482 100644 --- a/test-suite/plugin-demo/src/g_demo_plugin.mlg +++ b/test-suite/plugin-demo/src/g_demo_plugin.mlg @@ -13,12 +13,12 @@ END VERNAC COMMAND EXTEND Make_lens_vernac CLASSIFIED AS SIDEFF STATE program | [ "Make" "Lens" ident(name) ] -> - { Run_extractable.run_vernac (genLensN (Tm_util.string_to_list (Names.Id.to_string name))) } + { Run_extractable.run_vernac (genLensN (Caml_bytestring.bytestring_of_caml_string (Names.Id.to_string name))) } END VERNAC COMMAND EXTEND Unquote_vernac CLASSIFIED AS SIDEFF STATE program | [ "LookupPrint" ident(name) ] -> { let nstr = Names.Id.to_string name in - Run_extractable.run_vernac (lookupPrint (Tm_util.string_to_list nstr)) } + Run_extractable.run_vernac (lookupPrint (Caml_bytestring.bytestring_of_caml_string nstr)) } END diff --git a/test-suite/plugin-demo/theories/MyPlugin.v b/test-suite/plugin-demo/theories/MyPlugin.v index 6d455f49c..07b73fe52 100644 --- a/test-suite/plugin-demo/theories/MyPlugin.v +++ b/test-suite/plugin-demo/theories/MyPlugin.v @@ -1,11 +1,12 @@ Require Import Coq.Lists.List. From MetaCoq.Template Require Import - Ast + bytestring Ast Loader TemplateMonad.Extractable. Import TemplateMonad.Extractable. From MetaCoq Require Import Template.BasicAst Template.AstUtils Ast. +Open Scope bs_scope. Notation TemplateMonad := TM. Fixpoint mconcat (ls : list (TemplateMonad unit)) : TM unit := @@ -45,7 +46,7 @@ Notation "<% x %>" := (ltac:(let p y := exact y in quote_term x p)) (* | ConstructRef ind _ => ind.(inductive_mind) *) (* end. *) -Definition tmResolve (nm : String.string) : TM (option kername) := +Definition tmResolve (nm : String.t) : TM (option kername) := tmBind (tmLocate nm) (fun gr => match gr with @@ -76,9 +77,7 @@ Fixpoint countTo (n : nat) : list nat := | S m => countTo m ++ (m :: nil) end. -Require Import String. -Open Scope string_scope. -Definition prepend (ls : string) (i : ident) : ident := +Definition prepend (ls : String.t) (i : ident) : ident := ls ++ i. Definition cBuild_Lens := <% Build_Lens %>. @@ -155,14 +154,14 @@ Definition opBind {A B} (a: option A) (f: A -> option B) : option B := | None => None end. -Definition genLensN (baseName : String.string) : TM unit := +Definition genLensN (baseName : String.t) : TM unit := tmBind (tmLocate baseName) (fun gr => match gr with | (IndRef kn) :: _ => let name := kn.(inductive_mind) in let ty := Ast.tInd - {| BasicAst.inductive_mind := name - ; BasicAst.inductive_ind := 0 (* TODO: fix for mutual records *) |} List.nil in + {| Kernames.inductive_mind := name + ; Kernames.inductive_ind := 0 (* TODO: fix for mutual records *) |} List.nil in tmBind (tmQuoteInductive name) (fun ind => match getFields ind with | Some info => @@ -182,7 +181,7 @@ Definition genLensN (baseName : String.string) : TM unit := end). -Definition tmQuoteConstantR (nm : String.string) (bypass : bool) : TM _ := +Definition tmQuoteConstantR (nm : String.t) (bypass : bool) : TM _ := tmBind (tmLocate nm) (fun gr => match gr with @@ -192,7 +191,7 @@ Definition tmQuoteConstantR (nm : String.string) (bypass : bool) : TM _ := | _ => tmReturn None end). -Definition lookupPrint (baseName : String.string) : TM unit := +Definition lookupPrint (baseName : String.t) : TM unit := tmBind (tmQuoteConstantR baseName true) (fun b => match b with diff --git a/test-suite/proj.v b/test-suite/proj.v index fc04ec8d5..a8d08ab60 100644 --- a/test-suite/proj.v +++ b/test-suite/proj.v @@ -23,7 +23,7 @@ Proof. |- ?T => quote_term T (fun x => pose (qgoal:=x)) end. match goal with - H:= context [Ast.tProj (BasicAst.mkInd _ 0, 1, 0) _] |- _ => idtac + H:= context [Ast.tProj (Kernames.mkInd _ 0, 1, 0) _] |- _ => idtac end. reflexivity. Qed. diff --git a/test-suite/tmFreshName.v b/test-suite/tmFreshName.v index f5ce13c68..bf76fb57e 100644 --- a/test-suite/tmFreshName.v +++ b/test-suite/tmFreshName.v @@ -1,15 +1,15 @@ -Require Import List Arith String. +Require Import List Arith. Require Import MetaCoq.Template.All. Import ListNotations MCMonadNotation. -MetaCoq Run (x <- tmFreshName ("x" ++ "y")%string ;; +MetaCoq Run (x <- tmFreshName ("x" ++ "y")%bs ;; tmDefinition x 0). Check (eq_refl : xy = 0). -MetaCoq Run (x <- tmFreshName ("x" ++ "y")%string ;; +MetaCoq Run (x <- tmFreshName ("x" ++ "y")%bs ;; tmDefinition x 1). Check (eq_refl : xy0 = 1). diff --git a/test-suite/unfold.v b/test-suite/unfold.v index 898588aa3..f5eb695c5 100644 --- a/test-suite/unfold.v +++ b/test-suite/unfold.v @@ -1,4 +1,5 @@ From MetaCoq.Template Require Import utils All. +Import MCMonadNotation. MetaCoq Test Quote negb. MetaCoq Run (tmBind (tmEval (unfold (MPfile ["Datatypes"; "Init"; "Coq"], "negb")) negb) tmPrint). diff --git a/test-suite/univ.v b/test-suite/univ.v index e55ed11cf..e3dc6ec46 100644 --- a/test-suite/univ.v +++ b/test-suite/univ.v @@ -1,7 +1,7 @@ From MetaCoq.Template Require Import All. -Require Import String List Arith. +Require Import List Arith. Import ListNotations MCMonadNotation. -Open Scope string. +Open Scope bs_scope. Notation "'unfolded' d" := ltac:(let y := eval unfold d in d in exact y) (at level 100, only parsing). diff --git a/translations/translation_utils.v b/translations/translation_utils.v index 2756bfe92..5d2dee521 100644 --- a/translations/translation_utils.v +++ b/translations/translation_utils.v @@ -18,11 +18,10 @@ Fixpoint lookup_tsl_table (E : tsl_table) (gr : global_reference) match E with | nil => None | hd :: tl => - if gref_eq_dec gr (fst hd) then Some (snd hd) + if gref_eqb gr (fst hd) then Some (snd hd) else lookup_tsl_table tl gr end. - Definition tsl_context := (global_env_ext * tsl_table)%type. Definition emptyTC : tsl_context := (empty_ext empty_global_env, []). @@ -91,9 +90,9 @@ Definition nNamed n := {| binder_name := nNamed n; binder_relevance := Relevant Definition tsl_name f := map_binder_annot (tsl_name0 f). -Definition tmDebug {A} : A -> TemplateMonad unit - (* := tmPrint. *) - := fun _ => ret tt. +Definition tmDebug {A} : A -> TemplateMonad unit := + fun _ => ret tt. + (* fun x => tmPrint x ;; ret tt. *) Definition add_global_decl d (Σ : global_env_ext) : global_env_ext