From ba0c595774fba96088ff4fd29d2384eac714903f Mon Sep 17 00:00:00 2001 From: SimonBoulier Date: Fri, 7 Feb 2020 23:27:35 +0100 Subject: [PATCH] Big change in universes: - change the representation of UnivExpr - change the representation of universes (sorted, non empty list w/o duplicate) - prove completness of the comparison algorithm --- .gitignore | 4 + checker/src/metacoq_checker_plugin.mlpack | 4 +- checker/theories/Checker.v | 2 +- checker/theories/Closed.v | 27 +- checker/theories/Reflect.v | 46 +- checker/theories/Retyping.v | 2 +- checker/theories/Substitution.v | 34 +- checker/theories/WcbvEval.v | 25 +- checker/theories/Weakening.v | 48 +- erasure/_PluginProject.in | 8 +- erasure/src/metacoq_erasure_plugin.mlpack | 4 +- erasure/theories/EArities.v | 30 +- erasure/theories/ELiftSubst.v | 4 +- erasure/theories/ESubstitution.v | 26 +- erasure/theories/ErasureCorrectness.v | 174 +- erasure/theories/ErasureFunction.v | 24 +- erasure/theories/Extract.v | 6 +- erasure/theories/SafeErasureFunction.v | 4 +- examples/demo.v | 4 +- pcuic/theories/PCUICAlpha.v | 26 +- pcuic/theories/PCUICAst.v | 14 +- pcuic/theories/PCUICAstUtils.v | 30 +- pcuic/theories/PCUICChecker.v | 2 +- pcuic/theories/PCUICClosed.v | 32 +- pcuic/theories/PCUICConfluence.v | 54 +- pcuic/theories/PCUICContextConversion.v | 2 +- pcuic/theories/PCUICElimination.v | 10 +- pcuic/theories/PCUICEquality.v | 294 +-- pcuic/theories/PCUICLiftSubst.v | 8 +- pcuic/theories/PCUICNameless.v | 133 +- pcuic/theories/PCUICParallelReduction.v | 16 +- .../PCUICParallelReductionConfluence.v | 28 +- pcuic/theories/PCUICPosition.v | 14 +- pcuic/theories/PCUICPrincipality.v | 6 +- pcuic/theories/PCUICReduction.v | 12 +- pcuic/theories/PCUICReflect.v | 58 +- pcuic/theories/PCUICRetyping.v | 6 +- pcuic/theories/PCUICSN.v | 4 +- pcuic/theories/PCUICSafeLemmata.v | 28 +- pcuic/theories/PCUICSigmaCalculus.v | 2 +- pcuic/theories/PCUICSubstitution.v | 6 +- pcuic/theories/PCUICTyping.v | 14 +- pcuic/theories/PCUICUnivSubst.v | 171 +- pcuic/theories/PCUICUnivSubstitution.v | 342 +++- pcuic/theories/PCUICWcbvEval.v | 24 +- pcuic/theories/TemplateToPCUICCorrectness.v | 8 +- safechecker/_PluginProject.in | 8 +- .../src/metacoq_safechecker_plugin.mlpack | 4 +- safechecker/theories/PCUICSafeChecker.v | 86 +- safechecker/theories/PCUICSafeConversion.v | 20 +- safechecker/theories/PCUICSafeReduce.v | 6 +- safechecker/theories/PCUICSafeRetyping.v | 2 +- template-coq/_PluginProject | 10 +- .../gen-src/metacoq_template_plugin.mlpack | 3 +- template-coq/src/ast_denoter.ml | 8 +- template-coq/src/ast_quoter.ml | 32 +- template-coq/src/constr_denoter.ml | 76 +- template-coq/src/constr_quoted.ml | 14 +- template-coq/src/constr_quoter.ml | 2 +- template-coq/src/monad_extraction.mlpack | 1 + template-coq/theories/Ast.v | 12 +- template-coq/theories/AstUtils.v | 43 +- template-coq/theories/BasicAst.v | 4 +- template-coq/theories/Constants.v | 17 +- template-coq/theories/Environment.v | 8 +- template-coq/theories/EnvironmentTyping.v | 20 +- template-coq/theories/Induction.v | 4 +- template-coq/theories/LiftSubst.v | 8 +- template-coq/theories/Pretty.v | 2 +- template-coq/theories/TemplateMonad/Common.v | 2 +- template-coq/theories/TemplateMonad/Core.v | 2 +- .../theories/TemplateMonad/Extractable.v | 2 +- template-coq/theories/Typing.v | 20 +- template-coq/theories/TypingWf.v | 2 +- template-coq/theories/UnivSubst.v | 215 +- template-coq/theories/Universes.v | 1782 +++++++++++++---- template-coq/theories/common/uGraph.v | 1630 ++++++++------- template-coq/theories/utils.v | 234 +-- template-coq/theories/utils/All_Forall.v | 62 +- template-coq/theories/utils/MCList.v | 35 +- template-coq/theories/utils/MCOption.v | 24 + template-coq/theories/utils/MCProd.v | 16 + template-coq/theories/utils/wGraph.v | 323 +-- test-suite/erasure_test.v | 2 +- test-suite/univ.v | 13 +- 85 files changed, 3775 insertions(+), 2799 deletions(-) diff --git a/.gitignore b/.gitignore index ae5c31ce2..45b99f663 100644 --- a/.gitignore +++ b/.gitignore @@ -359,6 +359,10 @@ pcuic/src/pCUICUtils.mli /erasure/src/pCUICSafeRetyping.mli /erasure/src/safeErasureFunction.ml /erasure/src/safeErasureFunction.mli +erasure/src/mSetWeakList.ml +erasure/src/mSetWeakList.mli +erasure/src/utils.ml +erasure/src/utils.mli template-coq/gen-src/cRelationClasses.mli.rej .DS_Store diff --git a/checker/src/metacoq_checker_plugin.mlpack b/checker/src/metacoq_checker_plugin.mlpack index 7fd7661f2..34f9bd2fb 100644 --- a/checker/src/metacoq_checker_plugin.mlpack +++ b/checker/src/metacoq_checker_plugin.mlpack @@ -1,9 +1,9 @@ -Compare_dec -MSetList EqdepFacts Wf WGraph Monad_utils +Utils +MSetWeakList UGraph0 EnvironmentTyping Typing0 diff --git a/checker/theories/Checker.v b/checker/theories/Checker.v index 385c94dec..070e2b36c 100644 --- a/checker/theories/Checker.v +++ b/checker/theories/Checker.v @@ -546,7 +546,7 @@ Section Typecheck. | None => raise (NotEnoughFuel fuel) end. - Definition reduce_to_sort Γ (t : term) : typing_result universe := + Definition reduce_to_sort Γ (t : term) : typing_result Universe.t := match t with | tSort s => ret s | _ => diff --git a/checker/theories/Closed.v b/checker/theories/Closed.v index fcf1cf6e5..637c1f043 100644 --- a/checker/theories/Closed.v +++ b/checker/theories/Closed.v @@ -26,7 +26,7 @@ Proof. rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length, ?Nat.add_assoc; simpl closed in *; solve_all; unfold compose, test_def, test_snd in *; - try solve [simpl lift; simpl closed; f_equal; auto; repeat (toProp; solve_all)]; try easy. + try solve [simpl lift; simpl closed; f_equal; auto; repeat (rtoProp; solve_all)]; try easy. - elim (Nat.leb_spec k' n0); intros. simpl. elim (Nat.ltb_spec); auto. apply Nat.ltb_lt in H. lia. @@ -41,7 +41,7 @@ Proof. induction t in n, k, k' |- * using term_forall_list_ind; intros; simpl in *; rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length, ?Nat.add_assoc in *; - simpl closed in *; repeat (toProp; solve_all); try change_Sk; + simpl closed in *; repeat (rtoProp; solve_all); try change_Sk; unfold compose, test_def, on_snd, test_snd in *; simpl in *; eauto with all. - revert H0. @@ -51,8 +51,8 @@ Proof. - specialize (IHt2 n (S k) (S k')). eauto with all. - specialize (IHt2 n (S k) (S k')). eauto with all. - specialize (IHt3 n (S k) (S k')). eauto with all. - - toProp. solve_all. specialize (H1 n (#|m| + k) (#|m| + k')). eauto with all. - - toProp. solve_all. specialize (H1 n (#|m| + k) (#|m| + k')). eauto with all. + - rtoProp. solve_all. specialize (H1 n (#|m| + k) (#|m| + k')). eauto with all. + - rtoProp. solve_all. specialize (H1 n (#|m| + k) (#|m| + k')). eauto with all. Qed. Lemma closedn_mkApps k f u: @@ -85,7 +85,7 @@ Proof. induction t in k' |- * using term_forall_list_ind; intros; simpl in *; rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; - simpl closed in *; try change_Sk; repeat (toProp; solve_all); + simpl closed in *; try change_Sk; repeat (rtoProp; solve_all); unfold compose, test_def, on_snd, test_snd in *; simpl in *; eauto with all. - elim (Nat.leb_spec k' n); intros. simpl. @@ -105,10 +105,10 @@ Proof. - specialize (IHt3 (S k')). rewrite <- Nat.add_succ_comm in IHt3. eauto. - rewrite closedn_mkApps; solve_all. - - toProp; solve_all. rewrite -> !Nat.add_assoc in *. + - rtoProp; solve_all. rewrite -> !Nat.add_assoc in *. specialize (H0 (#|m| + k')). unfold is_true. rewrite <- H0. f_equal. lia. unfold is_true. rewrite <- H2. f_equal. lia. - - toProp; solve_all. rewrite -> !Nat.add_assoc in *. + - rtoProp; solve_all. rewrite -> !Nat.add_assoc in *. specialize (H0 (#|m| + k')). unfold is_true. rewrite <- H0. f_equal. lia. unfold is_true. rewrite <- H2. f_equal. lia. Qed. @@ -249,6 +249,7 @@ Proof. now rewrite app_length // plus_n_Sm. Qed. + Lemma typecheck_closed `{cf : checker_flags} : env_prop (fun Σ Γ t T => closedn #|Γ| t && closedn #|Γ| T). @@ -317,7 +318,7 @@ Proof. eauto using closed_upwards with arith. - intuition auto. solve_all. unfold test_snd. simpl in *. - toProp; eauto. + rtoProp; eauto. apply closedn_mkApps; auto. rewrite forallb_app. simpl. rewrite H1. rewrite forallb_skipn; auto. @@ -342,26 +343,26 @@ Proof. - split. solve_all. destruct x; simpl in *. - unfold test_def. simpl. toProp. + unfold test_def. simpl. rtoProp. split. rewrite -> app_context_length in *. rewrite -> Nat.add_comm in *. eapply closedn_lift_inv in H2; eauto. lia. subst types. now rewrite app_context_length fix_context_length in H1. - eapply nth_error_all in H0; eauto. simpl in H0. intuition. toProp. + eapply nth_error_all in H0; eauto. simpl in H0. intuition. rtoProp. subst types. rewrite app_context_length in H1. rewrite Nat.add_comm in H1. now eapply closedn_lift_inv in H1. - split. solve_all. destruct x; simpl in *. - unfold test_def. simpl. toProp. + unfold test_def. simpl. rtoProp. split. rewrite -> app_context_length in *. rewrite -> Nat.add_comm in *. eapply closedn_lift_inv in H2; eauto. lia. subst types. now rewrite -> app_context_length, fix_context_length in H1. eapply (nth_error_all) in H; eauto. simpl in *. - intuition. toProp. + intuition. rtoProp. subst types. rewrite app_context_length in H1. rewrite Nat.add_comm in H1. now eapply closedn_lift_inv in H1. @@ -435,7 +436,7 @@ Proof. eapply on_global_env_impl; cycle 1. apply (env_prop_sigma _ typecheck_closed _ X). red; intros. unfold lift_typing in *. destruct T; intuition auto with wf. - destruct X1 as [s0 Hs0]. simpl. toProp; intuition. + destruct X1 as [s0 Hs0]. simpl. rtoProp; intuition. Qed. Lemma closed_decl_upwards k d : closed_decl k d -> forall k', k <= k' -> closed_decl k' d. diff --git a/checker/theories/Reflect.v b/checker/theories/Reflect.v index d4ea8f2b8..995ba3ae2 100644 --- a/checker/theories/Reflect.v +++ b/checker/theories/Reflect.v @@ -226,13 +226,13 @@ Proof. unfold eq_inductive, eqb; cbn. now rewrite eq_string_refl Nat.eqb_refl. Defined. -Definition eq_def {A : Set} `{ReflectEq A} (d1 d2 : def A) : bool := +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 => eqb n1 n2 && eqb t1 t2 && eqb b1 b2 && eqb a1 a2 end. -#[program] Instance reflect_def : forall {A : Set} `{ReflectEq A}, ReflectEq (def A) := { +#[program] Instance reflect_def : forall {A} `{ReflectEq A}, ReflectEq (def A) := { eqb := eq_def }. Next Obligation. @@ -246,25 +246,6 @@ Next Obligation. cbn. constructor. subst. reflexivity. Defined. -Fixpoint eq_non_empty_list {A : Set} (eqA : A -> A -> bool) (l l' : non_empty_list A) : bool := - match l, l' with - | NEL.sing a, NEL.sing a' => eqA a a' - | NEL.cons a l, NEL.cons a' l' => - eqA a a' && eq_non_empty_list eqA l l' - | _, _ => false - end. - -#[program] Instance reflect_non_empty_list : - forall {A : Set} `{ReflectEq A}, ReflectEq (non_empty_list A) := - { eqb := eq_non_empty_list eqb }. -Next Obligation. - induction x, y; cbn. - destruct (eqb_spec a a0); constructor; congruence. - constructor; congruence. - constructor; congruence. - destruct (eqb_spec a a0), (IHx y); constructor; congruence. -Defined. - Fixpoint eq_cast_kind (c c' : cast_kind) : bool := match c, c' with | VmCast, VmCast @@ -293,13 +274,16 @@ Local Ltac fcase c := let e := fresh "e" in case c ; intro e ; [ subst ; try (left ; reflexivity) | finish ]. +Instance eq_dec_univ : EqDec Universe.t. +Admitted. + Local Ltac term_dec_tac term_dec := repeat match goal with | t : term, u : term |- _ => fcase (term_dec t u) - | u : universe, u' : universe |- _ => fcase (eq_dec u u') - | x : universe_instance, y : universe_instance |- _ => + | u : Universe.t, u' : Universe.t |- _ => fcase (eq_dec u u') + | x : Instance.t, y : Instance.t |- _ => fcase (eq_dec x y) - | x : list Level.t, y : universe_instance |- _ => + | x : list Level.t, y : Instance.t |- _ => fcase (eq_dec x y) | n : nat, m : nat |- _ => fcase (Nat.eq_dec n m) | i : ident, i' : ident |- _ => fcase (string_dec i i') @@ -315,18 +299,18 @@ Local Ltac term_dec_tac term_dec := Derive NoConfusion NoConfusionHom for term. -Derive EqDec for term. -Next Obligation. - induction x using term_forall_list_rect ; intro t ; +Instance EqDec_term : EqDec term. +Proof. + intro x; induction x using term_forall_list_rect ; intro t ; destruct t ; try (right ; discriminate). all: term_dec_tac term_dec. - - induction H in args |- *. + - induction X in args |- *. + destruct args. * left. reflexivity. * right. discriminate. + destruct args. * right. discriminate. - * destruct (IHAll args) ; nodec. + * destruct (IHX args) ; nodec. destruct (p t) ; nodec. subst. left. inversion e. reflexivity. - destruct (IHx1 t1) ; nodec. @@ -343,11 +327,11 @@ Next Obligation. destruct (IHx3 t3) ; nodec. subst. left. reflexivity. - destruct (IHx t) ; nodec. - subst. induction H in args |- *. + subst. induction X in args |- *. + destruct args. all: nodec. left. reflexivity. + destruct args. all: nodec. - destruct (IHAll args). all: nodec. + destruct (IHX args). all: nodec. destruct (p t0). all: nodec. subst. inversion e. subst. left. reflexivity. diff --git a/checker/theories/Retyping.v b/checker/theories/Retyping.v index 42e0b9f14..baa60c0d4 100644 --- a/checker/theories/Retyping.v +++ b/checker/theories/Retyping.v @@ -109,7 +109,7 @@ Section TypeOf. end end. - Definition sort_of (Γ : context) (t : term) : typing_result universe := + Definition sort_of (Γ : context) (t : term) : typing_result Universe.t := ty <- type_of Γ t;; type_of_as_sort type_of Γ ty. diff --git a/checker/theories/Substitution.v b/checker/theories/Substitution.v index a5cdc5479..bacf8c9af 100644 --- a/checker/theories/Substitution.v +++ b/checker/theories/Substitution.v @@ -34,7 +34,7 @@ Inductive subs `{cf : checker_flags} Σ (Γ : context) : list term -> context -> (** Linking a context (with let-ins), an instance (reversed substitution) for its assumptions and a well-formed substitution for it. *) -Inductive context_subst : context -> list term -> list term -> Set := +Inductive context_subst : context -> list term -> list term -> Type := | context_subst_nil : context_subst [] [] [] | context_subst_ass Γ args s na t a : context_subst Γ args s -> @@ -920,7 +920,7 @@ Proof. f_equal. rewrite (lift_to_extended_list_k _ _ 1) map_map_compose. pose proof (to_extended_list_k_spec Γ k). - solve_all. destruct H3 as [n [-> Hn]]. + solve_all. destruct H2 as [n [-> Hn]]. rewrite /compose /lift (subst_app_decomp [a] s k); auto with wf. rewrite subst_rel_gt. simpl; lia. repeat (f_equal; simpl; try lia). @@ -929,7 +929,7 @@ Proof. rewrite -IHcontext_subst // to_extended_list_k_cons /=. rewrite (lift_to_extended_list_k _ _ 1) map_map_compose. pose proof (to_extended_list_k_spec Γ k). - solve_all. destruct H3 as [n [-> Hn]]. + solve_all. destruct H2 as [n [-> Hn]]. rewrite /compose /lift (subst_app_decomp [subst0 s b] s k); auto with wf. rewrite subst_rel_gt. simpl; lia. repeat (f_equal; simpl; try lia). @@ -943,7 +943,7 @@ Proof. induction 1. constructor. eapply All_app in wfargs as [wfargs wfa]. inv wfa. inv wfctx. constructor; intuition auto. - inv wfctx. red in H0. constructor; solve_all. apply wf_subst. solve_all. intuition auto with wf. + inv wfctx. red in H. constructor; solve_all. apply wf_subst. solve_all. intuition auto with wf. Qed. Lemma wf_make_context_subst ctx args s' : @@ -1235,7 +1235,7 @@ Lemma red1_mkApps_l Σ Γ M1 N1 M2 : Proof. induction M2 in M1, N1 |- *. simpl; auto. intros. specialize (IHM2 (mkApp M1 a) (mkApp N1 a)). - inv H0. + inv X. forward IHM2. apply wf_mkApp; auto. forward IHM2. auto. rewrite <- !mkApps_mkApp; auto. @@ -1247,17 +1247,17 @@ Lemma red1_mkApps_r Σ Γ M1 M2 N2 : Ast.wf M1 -> All Ast.wf M2 -> OnOne2 (red1 Σ Γ) M2 N2 -> red1 Σ Γ (mkApps M1 M2) (mkApps M1 N2). Proof. - intros. induction X in M1, H, H0 |- *. - inv H0. + intros. induction X0 in M1, H, X |- *. + inv X. destruct (isApp M1) eqn:Heq. destruct M1; try discriminate. simpl. constructor. apply OnOne2_app. constructor. auto. rewrite mkApps_tApp; try congruence. rewrite mkApps_tApp; try congruence. constructor. constructor. auto. - inv H0. - specialize (IHX (mkApp M1 hd)). forward IHX. - apply wf_mkApp; auto. forward IHX; auto. - now rewrite !mkApps_mkApp in IHX. + inv X. + specialize (IHX0 (mkApp M1 hd)). forward IHX0. + apply wf_mkApp; auto. forward IHX0; auto. + now rewrite !mkApps_mkApp in IHX0. Qed. Lemma wf_fix : @@ -1402,11 +1402,11 @@ Proof. now apply All_map. assert(Ast.wf M1). now inv wfM. assert(All Ast.wf M2). eapply Forall_All. now inv wfM. - clear -X H1 H2 wfΓ Hs. + clear -X H0 X1 wfΓ Hs. induction X; constructor; auto. intuition. - eapply b; eauto. now inv H2. - apply IHX. now inv H2. + eapply b; eauto. now inv X1. + apply IHX. now inv X1. - forward IHred1. now inv wfM. constructor. specialize (IHred1 _ _ (Γ'' ,, vass na M1) eq_refl). @@ -1592,13 +1592,13 @@ Proof. eapply All2_map. eapply All2_impl'. eassumption. eapply All_impl. eassumption. - cbn. apply All2_length in H3 as e. rewrite e. + cbn. apply All2_length in X0 as e. rewrite e. intros x [] y []; now split. - constructor. eapply All2_map. eapply All2_impl'. eassumption. eapply All_impl. eassumption. - cbn. apply All2_length in H3 as e. rewrite e. + cbn. apply All2_length in X0 as e. rewrite e. intros x [] y []; now split. Qed. @@ -1638,7 +1638,7 @@ Lemma subst_eq_context `{checker_flags} φ l l' n k : Proof. induction l in l', n, k |- *; inversion 1. constructor. rewrite !subst_context_snoc. constructor. - apply All2_length in H5 as e. rewrite e. + apply All2_length in X1 as e. rewrite e. now apply subst_eq_decl. now apply IHl. Qed. diff --git a/checker/theories/WcbvEval.v b/checker/theories/WcbvEval.v index 0d9bc059b..aa896e0f7 100644 --- a/checker/theories/WcbvEval.v +++ b/checker/theories/WcbvEval.v @@ -237,17 +237,17 @@ Section Wcbv. (forall i : nat, option_map decl_body (nth_error Γ i) = Some None -> P (tRel i) (tRel i)) -> (forall (c : ident) (decl : constant_body) (body : term), declared_constant Σ c decl -> - forall (u : universe_instance) (res : term), + forall (u : Instance.t) (res : term), cst_body decl = Some body -> eval (subst_instance_constr u body) res -> P (subst_instance_constr u body) res -> P (tConst c u) res) -> (forall (c : ident) (decl : constant_body), - declared_constant Σ c decl -> forall u : universe_instance, cst_body decl = None -> P (tConst c u) (tConst c u)) -> - (forall (ind : inductive) (pars : nat) (discr : term) (c : nat) (u : universe_instance) + declared_constant Σ c decl -> forall u : Instance.t, cst_body decl = None -> P (tConst c u) (tConst c u)) -> + (forall (ind : inductive) (pars : nat) (discr : term) (c : nat) (u : Instance.t) (args : list term) (p : term) (brs : list (nat × term)) (res : term), eval discr (mkApps (tConstruct ind c u) args) -> P discr (mkApps (tConstruct ind c u) args) -> eval (iota_red pars c args brs) res -> P (iota_red pars c args brs) res -> P (tCase (ind, pars) p discr brs) res) -> - (forall (i : inductive) (pars arg : nat) (discr : term) (args : list term) (k : nat) (u : universe_instance) + (forall (i : inductive) (pars arg : nat) (discr : term) (args : list term) (k : nat) (u : Instance.t) (a res : term), eval discr (mkApps (tConstruct i k u) args) -> P discr (mkApps (tConstruct i k u) args) -> @@ -368,7 +368,7 @@ Section Wcbv. - now eapply atom_mkApps in H. - intros * isapp appeq. move: (value_head_nApp H) => Ht. apply mkApps_eq_inj in appeq; intuition subst; auto. - - intros * isapp appeq. move: (isStuckfix_nApp H1) => Hf. + - intros * isapp appeq. move: (isStuckfix_nApp H) => Hf. apply mkApps_eq_inj in appeq; intuition subst; auto. Qed. @@ -394,8 +394,8 @@ Section Wcbv. destruct decl as [? [b|] ?]; try discriminate. constructor. constructor. - eapply value_stuck_fix => //. - now eapply All2_right in H. - - eapply All2_right in H2. + now eapply All2_right in X0. + - eapply All2_right in X0. depelim IHeve. destruct t; simpl in * |- *; try congruence. eapply value_app; auto. @@ -443,22 +443,21 @@ Section Wcbv. induction 1 using value_values_ind; simpl; auto using value. - now constructor. - assert (All2 eval l l). - { induction X; constructor; auto. eapply IHX. now inversion H0. } + { induction X0; constructor; auto. eapply IHX0. now inversion X. } move/implyP: (value_head_spec t). move/(_ H) => Ht. induction l using rev_ind. simpl. now eapply value_head_final. - eapply All_app in H0 as [Hl Hx]. inv Hx. - eapply All_app in X as [Hl' Hx']. inv Hx'. - eapply All2_app_r in X0 as [Hl'' Hx'']. + eapply All_app in X as [Hl Hx]. inv Hx. + eapply All_app in X0 as [Hl' Hx']. inv Hx'. + eapply All2_app_r in X1 as [Hl'' Hx'']. pose proof (value_head_nApp H). rewrite -{1}mkApps_tApp => //. rewrite is_empty_app /= // andb_false_r //. eapply eval_app_cong; auto. rewrite is_empty_app /= andb_false_r //. now eapply value_head_final. eapply All2_app; auto. - destruct f; try discriminate. - assert (All2 eval args args). - { clear H0. induction X; constructor; auto. eapply IHX. now inversion H. } + apply All_All2_refl in X0. eapply eval_fix_value => //. constructor; auto. Qed. diff --git a/checker/theories/Weakening.v b/checker/theories/Weakening.v index 3c46b9e99..22b80c6ba 100644 --- a/checker/theories/Weakening.v +++ b/checker/theories/Weakening.v @@ -709,23 +709,23 @@ Proof. induction T in n, k, U, Rl |- * using term_forall_list_rect; inversion 1; simpl; try (now constructor). - destruct (k <=? n0); constructor. - - constructor. subst. clear - X H3. - induction l in X, args', H3 |- *. - + inversion H3; constructor. - + inversion H3. inversion X. subst. + - constructor. subst. clear - X X1. + induction l in X, args', X1 |- *. + + inversion X1; constructor. + + inversion X1. inversion X. subst. simpl. constructor. all: easy. - - constructor. easy. clear - X H4. - induction l in X, args', H4 |- *. - + inversion H4; constructor. - + inversion H4; inversion X; subst. + - constructor. easy. clear - X X2. + induction l in X, args', X2 |- *. + + inversion X2; constructor. + + inversion X2; inversion X; subst. now constructor. - - constructor; try easy. clear - X H7. - induction l in X, brs', H7 |- *. - + inversion H7; constructor. - + inversion H7; inversion X; subst. + - constructor; try easy. clear - X X3. + induction l in X, brs', X3 |- *. + + inversion X3; constructor. + + inversion X3; inversion X; subst. constructor. cbn; easy. easy. - - constructor; try easy. clear - X H3. + - constructor; try easy. clear - X X1. assert (XX:forall k k', All2 (fun x y => eq_term_upto_univ Re Re (dtype x) (dtype y) × eq_term_upto_univ Re Re (dbody x) (dbody y) × @@ -733,15 +733,15 @@ Proof. (map (map_def (lift n k) (lift n (#|m| + k'))) m) (map (map_def (lift n k) (lift n (#|mfix'| + k'))) mfix')); [|now apply XX]. clear k. - induction m in X, mfix', H3 |- *. - + inversion H3; constructor. - + inversion H3; inversion X; subst. + induction m in X, mfix', X1 |- *. + + inversion X1; constructor. + + inversion X1; inversion X; subst. simpl. constructor. split. cbn; easy. cbn; erewrite All2_length by eassumption. easy. unfold tFixProp in IHm. cbn. rewrite !plus_n_Sm. now apply IHm. - - constructor; try easy. clear - X H3. + - constructor; try easy. clear - X X1. assert (XX:forall k k', All2 (fun x y => eq_term_upto_univ Re Re (dtype x) (dtype y) × eq_term_upto_univ Re Re (dbody x) (dbody y) × @@ -749,9 +749,9 @@ Proof. (map (map_def (lift n k) (lift n (#|m| + k'))) m) (map (map_def (lift n k) (lift n (#|mfix'| + k'))) mfix')); [|now apply XX]. clear k. - induction m in X, mfix', H3 |- *. - + inversion H3; constructor. - + inversion H3; inversion X; subst. + induction m in X, mfix', X1 |- *. + + inversion X1; constructor. + + inversion X1; inversion X; subst. simpl. constructor. split. cbn; easy. cbn; erewrite All2_length by eassumption. easy. @@ -799,10 +799,10 @@ Lemma lift_eq_context `{checker_flags} φ l l' n k : Proof. induction l in l', n, k |- *; intros; destruct l'; rewrite -> ?lift_context_snoc0. - constructor. - - inversion H0. - - inversion H0. - - inversion H0; subst. constructor. - + apply All2_length in H6. rewrite H6. + - inversion X. + - inversion X. + - inversion X; subst. constructor. + + apply All2_length in X1. rewrite X1. now apply lift_eq_decl. + now apply IHl. Qed. diff --git a/erasure/_PluginProject.in b/erasure/_PluginProject.in index db76665ba..36db00a9a 100644 --- a/erasure/_PluginProject.in +++ b/erasure/_PluginProject.in @@ -8,13 +8,13 @@ src/classes0.mli src/classes0.ml # From Coq's stdlib -# src/compare_dec.mli -# src/compare_dec.ml -src/mSetList.mli -src/mSetList.ml +src/mSetWeakList.mli +src/mSetWeakList.ml src/eqdepFacts.mli src/eqdepFacts.ml +src/utils.ml +src/utils.mli src/monad_utils.ml src/monad_utils.mli src/uGraph0.ml diff --git a/erasure/src/metacoq_erasure_plugin.mlpack b/erasure/src/metacoq_erasure_plugin.mlpack index df026e3c0..5e7ee33b4 100644 --- a/erasure/src/metacoq_erasure_plugin.mlpack +++ b/erasure/src/metacoq_erasure_plugin.mlpack @@ -1,9 +1,9 @@ Monad_utils -MSetList +MSetWeakList EqdepFacts +Utils WGraph -Monad_utils UGraph0 Typing0 diff --git a/erasure/theories/EArities.v b/erasure/theories/EArities.v index 15274a978..0d0db9eb5 100644 --- a/erasure/theories/EArities.v +++ b/erasure/theories/EArities.v @@ -46,7 +46,7 @@ Qed. Lemma isWfArity_prod_inv: forall (Σ : global_env_ext) (Γ : context) (x : name) (x0 x1 : term), - isWfArity typing Σ Γ (tProd x x0 x1) -> (∑ s : universe, Σ;;; Γ |- x0 : tSort s) × isWfArity typing Σ (Γ,, vass x x0) x1 + isWfArity typing Σ Γ (tProd x x0 x1) -> (∑ s : Universe.t, Σ;;; Γ |- x0 : tSort s) × isWfArity typing Σ (Γ,, vass x x0) x1 . intros. destruct X as (? & ? & ? & ?). cbn in e. eapply destArity_app_Some in e as (? & ? & ?); subst. @@ -199,7 +199,7 @@ Proof. Qed. Lemma invert_it_Ind_eq_prod: - forall (u : universe_instance) (i : inductive) (x : name) (x0 x1 : term) (x2 : context) (x3 : list term), + forall (u : Instance.t) (i : inductive) (x : name) (x0 x1 : term) (x2 : context) (x3 : list term), tProd x x0 x1 = it_mkProd_or_LetIn x2 (mkApps (tInd i u) x3) -> exists (L' : context) (l' : list term), x1 = it_mkProd_or_LetIn L' (mkApps (tInd i u) l'). Proof. intros u i x x0 x1 x2 x3 H0. @@ -423,26 +423,11 @@ Proof. - right. eexists. eapply subject_reduction ; eauto. Qed. -Lemma is_prop_sort_sup: - forall x1 x2 : universe, is_prop_sort (Universe.sup x1 x2) -> is_prop_sort x2. -Proof. - induction x1; cbn; intros. - - inv H. - - inv H. -Qed. - -Lemma is_prop_sort_prod x2 x3 : - is_prop_sort (Universe.sort_of_product x2 x3) -> is_prop_sort x3. -Proof. - intros. unfold Universe.sort_of_product in *. destruct ?; eauto. - eapply is_prop_sort_sup in H. eauto. -Qed. - Lemma sort_typing_spine: - forall (Σ : global_env_ext) (Γ : context) (L : list term) (u : universe) (x x0 : term), + forall (Σ : global_env_ext) (Γ : context) (L : list term) (u : Universe.t) (x x0 : term), wf Σ -> - is_prop_sort u -> - typing_spine Σ Γ x L x0 -> Σ;;; Γ |- x : tSort u -> ∑ u', Σ;;; Γ |- x0 : tSort u' × is_prop_sort u'. + Universe.is_prop u -> + typing_spine Σ Γ x L x0 -> Σ;;; Γ |- x : tSort u -> ∑ u', Σ;;; Γ |- x0 : tSort u' × Universe.is_prop u'. Proof. intros Σ Γ L u x x0 ? ? t1 c0. revert u H c0. @@ -532,10 +517,7 @@ Proof. eapply cumul_Sort_inv in c. eapply leq_universe_prop in c as []; cbn; eauto. eexists. split. eassumption. right. eexists. split. eassumption. - unfold Universe.sort_of_product in H. - destruct ?; eauto. - - eapply is_prop_sort_sup; eauto. + eapply is_prop_sort_prod; eassumption. - auto. Qed. diff --git a/erasure/theories/ELiftSubst.v b/erasure/theories/ELiftSubst.v index ef948b4fa..fdd6a9a6b 100644 --- a/erasure/theories/ELiftSubst.v +++ b/erasure/theories/ELiftSubst.v @@ -463,7 +463,7 @@ Proof. elim t using term_forall_list_ind; intros; try easy; rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; unfold test_def in *; - simpl closed in *; try solve [simpl lift; simpl closed; f_equal; auto; toProp; solve_all]; try easy. + simpl closed in *; try solve [simpl lift; simpl closed; f_equal; auto; rtoProp; solve_all]; try easy. - rewrite lift_rel_lt; auto. revert H. elim (Nat.ltb_spec n0 k); intros; try easy. Qed. @@ -474,7 +474,7 @@ Proof. elim t using term_forall_list_ind; intros; try lia; rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; simpl closed in *; unfold test_snd, test_def in *; - try solve [(try f_equal; simpl; repeat (toProp; solve_all); eauto)]. + try solve [(try f_equal; simpl; repeat (rtoProp; solve_all); eauto)]. - elim (ltb_spec n k'); auto. intros. apply ltb_lt in H. lia. diff --git a/erasure/theories/ESubstitution.v b/erasure/theories/ESubstitution.v index 479cc870d..cda2b7049 100644 --- a/erasure/theories/ESubstitution.v +++ b/erasure/theories/ESubstitution.v @@ -56,7 +56,7 @@ Lemma Informative_extends: (mdecl : PCUICAst.mutual_inductive_body) (idecl : PCUICAst.one_inductive_body), PCUICTyping.declared_inductive (fst Σ) mdecl ind idecl -> - forall (Σ' : global_env) (u0 : universe_instance), + forall (Σ' : global_env) (u0 : Instance.t), wf Σ' -> extends Σ Σ' -> Informative Σ ind -> Informative (Σ', Σ.2) ind. @@ -96,8 +96,8 @@ Proof. - econstructor. destruct isdecl. 2:eauto. eapply Informative_extends; eauto. - econstructor. - eapply All2_All_mix_left in H4; eauto. - eapply All2_impl. exact H4. + eapply All2_All_mix_left in X1; eauto. + eapply All2_impl. exact X1. intros ? ? [[[]] [? []]]. split; eauto. - eauto. @@ -205,13 +205,13 @@ Proof. eassumption. } rewrite app_length in *. subst types. rewrite fix_context_length in *. - rewrite (All2_length _ _ H5) in *. + rewrite (All2_length _ _ X2) in *. eapply erases_ctx_ext. eapply e4. rewrite lift_context_app. unfold app_context. rewrite !app_assoc. repeat f_equal. rewrite <- lift_fix_context. rewrite <- plus_n_O. - now rewrite (All2_length _ _ H5). + now rewrite (All2_length _ _ X2). - eauto. Qed. @@ -306,15 +306,15 @@ Proof. intros Σ wfΣ Γ0 wfΓ0; intros; simpl in * |-; subst Γ0. - inv H0. + cbn. destruct ?. destruct ?. - * eapply All2_nth_error_Some in H2; eauto. - destruct H2 as (? & ? & ?). + * eapply All2_nth_error_Some in X2; eauto. + destruct X2 as (? & ? & ?). rewrite e. erewrite <- subst_context_length. eapply substlet_typable in X1 as []. 2:exact E0. eapply erases_weakening; eauto. * erewrite All2_length; eauto. - eapply All2_nth_error_None in H2; eauto. - rewrite H2. econstructor. + eapply All2_nth_error_None in X2; eauto. + rewrite X2. econstructor. * econstructor. + econstructor. eapply is_type_subst; eauto. @@ -365,11 +365,11 @@ Proof. * eapply H3; eauto. * eapply All2_map. eapply All2_impl_In; eauto. - intros. destruct H11, x, y. cbn in *. subst. split; eauto. + intros. destruct H9, x, y. cbn in *. subst. split; eauto. eapply All2_All_left in X3. 2:{ intros ? ? [[[[? ?] e1] ?] ?]. exact e1. } - eapply In_nth_error in H9 as []. + eapply In_nth_error in H7 as []. eapply nth_error_all in X3; eauto. eapply X3; eauto. @@ -386,7 +386,7 @@ Proof. eapply All2_map. eapply All2_impl_In. eassumption. - intros. destruct H4 as (? & ? & ?). + intros. destruct H3 as (? & ? & ?). repeat split; eauto. eapply In_nth_error in H1 as []. eapply nth_error_all in X0; eauto. @@ -413,7 +413,7 @@ Proof. eapply All2_map. eapply All2_impl_In. eassumption. - intros. destruct H4 as (? & ? & ?). + intros. destruct H3 as (? & ? & ?). repeat split; eauto. eapply In_nth_error in H1 as []. eapply nth_error_all in X0; eauto. diff --git a/erasure/theories/ErasureCorrectness.v b/erasure/theories/ErasureCorrectness.v index 47920e62a..7ee23c9a4 100644 --- a/erasure/theories/ErasureCorrectness.v +++ b/erasure/theories/ErasureCorrectness.v @@ -35,16 +35,6 @@ Proof. induction T; cbn; intros; tauto. Qed. -Lemma is_prop_subst_instance: - forall (u : universe_instance) (x0 : universe), is_prop_sort x0 -> is_prop_sort (UnivSubst.subst_instance_univ u x0). -Proof. - intros u x0 i. destruct x0; cbn in *; try congruence. - unfold is_prop_sort in *. unfold Universe.level in *. - destruct t. destruct t; cbn in *; try congruence. - destruct b; cbn in *; try congruence. -Qed. - - Lemma wf_ext_wk_wf {cf:checker_flags} Σ : wf_ext_wk Σ -> wf Σ. Proof. intro H; apply H. Qed. @@ -156,12 +146,12 @@ Proof. decompose [prod] X2. intuition auto. eapply b0. subst types. - eapply conv_context_app; auto. now eapply typing_wf_local in a4. - eapply typing_wf_local in a4. subst types. + eapply conv_context_app; auto. eapply typing_wf_local; eassumption. + eapply typing_wf_local in a1. subst types. 2:eauto. eapply All_local_env_app_inv. - eapply All_local_env_app in a4. intuition auto. + eapply All_local_env_app in a1. intuition auto. (* clear -wfΣ X2 a2 b4 X1. *) eapply All_local_env_impl; eauto. simpl; intros. @@ -169,14 +159,14 @@ Proof. eapply PCUICContextConversion.context_conversion with (Γ ,,, Γ0); eauto. now eapply typing_wf_local; eauto. eapply conv_context_app; auto. eapply typing_wf_local; eauto. - eapply typing_wf_local in X2. + eapply typing_wf_local in X3. eapply PCUICSafeChecker.wf_local_app_inv. - eauto. eapply wf_local_rel_local in X2. - eapply wf_local_rel_app in X2 as []. rewrite app_context_nil_l in w0. + eauto. eapply wf_local_rel_local in X3. + eapply wf_local_rel_app in X3 as []. rewrite app_context_nil_l in w0. eapply wf_local_rel_conv; eauto. - destruct X2. exists x0. + destruct X3. exists x0. eapply PCUICContextConversion.context_conversion with (Γ ,,, Γ0); eauto. now eapply typing_wf_local; eauto. eapply conv_context_app; auto. eapply typing_wf_local; eauto. @@ -192,7 +182,7 @@ Qed. (** ** Erasure is stable under substituting universe constraints *) Lemma fix_context_subst_instance: - forall (mfix : list (BasicAst.def term)) (u : universe_instance), + forall (mfix : list (BasicAst.def term)) (u : Instance.t), map (map_decl (PCUICUnivSubst.subst_instance_constr u)) (PCUICLiftSubst.fix_context mfix) = PCUICLiftSubst.fix_context @@ -239,7 +229,7 @@ Proof. eapply All2_map_left. eapply All2_impl. eapply All2_All_mix_left. eapply All2_All_left. exact X3. intros ? ? [[[? e] ?] ?]. - exact e. exact H16. + exact e. exact X6. intros; cbn in *. destruct H. destruct p0. split; eauto. - assert (Hw : wf_local (Σ.1, univs) (subst_instance_context u (Γ ,,, types))). { (* rewrite subst_instance_context_app. *) @@ -266,8 +256,8 @@ Proof. cbn. econstructor; eauto. eapply All2_map_left. eapply All2_impl. eapply All2_All_mix_left. eassumption. - exact H7. - intros; cbn in *. destruct X3. destruct p0. destruct p0. + exact X3. + intros; cbn in *. destruct X4. destruct p0. destruct p0. destruct p. destruct p. repeat split; eauto. eapply e2 in e1. unfold PCUICUnivSubst.subst_instance_context in *. @@ -809,36 +799,36 @@ Proof. unfold unfold_fix in H0. rewrite e in H0. inv H0. eapply erases_mkApps_inv in He as [(? & ? & ? & ? & [] & ? & ? & ?) | (? & ? & ? & ? & ?)]; eauto. - + subst. assert (X100 := X1). eapply Is_type_app in X100 as[]. + + subst. assert (X100 := X2). eapply Is_type_app in X100 as[]. exists tBox. split. 2:{ assert (exists x5, Forall2 (EWcbvEval.eval Σ') x4 x5) as [x5]. { - eapply All2_app_inv in H3 as ([] & ? & ?). destruct p. + eapply All2_app_inv in X0 as ([] & ? & ?). destruct p. clear a0. subst. assert (forall x n, nth_error x3 n = Some x -> ∑ T, Σ;;; [] |- x : T). { intros. eapply typing_spine_inv with (arg := n + #|x2|) in t0 as []. 2:{ rewrite nth_error_app2. 2:lia. rewrite Nat.add_sub. eassumption. } eauto. } - clear - X3 a1 H7. revert X3 x4 H7; induction a1; intros. - ** inv H7. exists []; eauto. - ** inv H7. destruct (X3 x 0 eq_refl). + clear - X0 a1 H6. revert X0 x4 H6; induction a1; intros. + ** inv H6. exists []; eauto. + ** inv H6. destruct (X0 x 0 eq_refl). eapply r in t as (? & ? & ?); eauto. eapply IHa1 in H3 as (? & ?); eauto. - intros. eapply (X3 x2 (S n)). eassumption. + intros. eapply (X0 x2 (S n)). eassumption. } eapply eval_box_apps. eassumption. econstructor; eauto. } econstructor. eapply Is_type_eval. eauto. eassumption. rewrite <- mkApps_nested. - eapply Is_type_red. eauto. 2:exact X2. repeat eapply PCUICReduction.red_mkApps_f. + eapply Is_type_red. eauto. 2:exact X3. repeat eapply PCUICReduction.red_mkApps_f. eapply wcbeval_red; eauto. eauto. eauto. rewrite mkApps_nested; eauto. + subst. - destruct ?; inv H6. - eapply IHeval1 in H5 as (? & ? & ?); eauto. + destruct ?; inv H5. + eapply IHeval1 in H4 as (? & ? & ?); eauto. inv H0. - * assert (Hmfix' := H10). - eapply All2_nth_error_Some in H10 as (? & ? & ?); eauto. + * assert (Hmfix' := X2). + eapply All2_nth_error_Some in X2 as (? & ? & ?); eauto. destruct x1. cbn in *. subst. enough(Σ;;; [] ,,, subst_context (fix_subst mfix) 0 [] |- PCUICLiftSubst.subst (fix_subst mfix) 0 dbody ⇝ℇ subst (ETyping.fix_subst mfix') 0 (Extract.E.dbody x4)). @@ -863,15 +853,15 @@ Proof. eapply Forall2_nth_error_Some in H as (? & ? & ?); eauto. rewrite H. - unfold isConstruct_app in H10. + unfold isConstruct_app in H9. destruct (decompose_app t) eqn:EE. assert (E2 : fst (decompose_app t) = t1) by now rewrite EE. destruct t1. - all:inv H10. + all:inv H9. (* erewrite <- PCUICConfluence.fst_decompose_app_rec in E2. *) pose proof (decompose_app_rec_inv EE). - cbn in H9. subst. + cbn in H8. subst. assert (∑ T, Σ ;;; [] |- mkApps (tConstruct ind n ui) l : T) as [T' HT']. { eapply typing_spine_eval in t0; eauto. eapply typing_spine_inv in t0; eauto. @@ -882,7 +872,7 @@ Proof. --- subst. eapply nth_error_forall in Hv; eauto. eapply value_app_inv in Hv. subst. reflexivity. - --- subst. inv H9. + --- subst. inv H8. +++ eapply nth_error_forall in Hv; eauto. destruct x7 using rev_ind; cbn - [EAstUtils.decompose_app]. reflexivity. rewrite emkApps_snoc at 1. @@ -894,17 +884,17 @@ Proof. --- eauto. --- eauto. ** subst. rewrite H2. - now eapply Forall2_length in H7. + now eapply Forall2_length in H6. ** eapply subject_reduction. eauto. exact Hty. etransitivity. eapply PCUICReduction.red_mkApps. now eapply wcbeval_red. eapply All2_impl. exact X. intros. now eapply wcbeval_red. econstructor 2. econstructor. econstructor. unfold unfold_fix. now rewrite e, E. exact Hcon. - ++ clear - t0 H3 H7. revert x t0 x3 H7; induction H3; intros. - ** inv H7. exists []; eauto. - ** inv H7. inv t0. eapply r in H1 as (? & ? & ?); eauto. - eapply IHAll2 in X2 as (? & ? & ?); eauto. + ++ clear - t0 X0 H6. revert x t0 x3 H6; induction X0; intros. + ** inv H6. exists []; eauto. + ** inv H6. inv t0. eapply r in H1 as (? & ? & ?); eauto. + eapply IHX0 in X3 as (? & ? & ?); eauto. -- cbn. destruct p. destruct p. eapply (erases_subst Σ [] (PCUICLiftSubst.fix_context mfix) [] dbody (fix_subst mfix)) in e3; cbn; eauto. ++ eapply subslet_fix_subst. now eapply wf_ext_wf. all: eassumption. @@ -913,12 +903,12 @@ Proof. ++ eapply All2_from_nth_error. erewrite fix_subst_length, ETyping.fix_subst_length, All2_length; eauto. intros. - rewrite fix_subst_nth in H6. 2:{ now rewrite fix_subst_length in H0. } - rewrite efix_subst_nth in H8. 2:{ rewrite fix_subst_length in H0. + rewrite fix_subst_nth in H5. 2:{ now rewrite fix_subst_length in H0. } + rewrite efix_subst_nth in H7. 2:{ rewrite fix_subst_length in H0. erewrite <- All2_length; eauto. } - inv H6; inv H8. + inv H5; inv H7. erewrite All2_length; eauto. - * eapply Is_type_app in X1 as []. + * eapply Is_type_app in X2 as []. exists tBox. split. econstructor. eapply Is_type_eval. eauto. eassumption. @@ -929,12 +919,12 @@ Proof. { intros. eapply typing_spine_inv with (arg := n) in t0 as []. 2:{ eassumption. } eauto. } - clear - X2 H3 H7. revert X2 x3 H7; induction H3; intros. - ** inv H7. exists []; eauto. - ** inv H7. destruct (X2 x 0 eq_refl). + clear - X3 X0 H6. revert X3 x3 H6; induction X0; intros. + ** inv H6. exists []; eauto. + ** inv H6. destruct (X3 x 0 eq_refl). eapply r in t as (? & ? & ?); eauto. - eapply IHAll2 in H4 as (? & ?); eauto. - intros. eapply (X2 x2 (S n)). eassumption. + eapply IHX0 in H3 as (? & ?); eauto. + intros. eapply (X3 x2 (S n)). eassumption. } eapply eval_box_apps. eauto. eauto. eapply wf_ext_wf. eauto. eauto. @@ -945,7 +935,7 @@ Proof. + auto. - assert (Hty' := Hty). assert (Hunf := H). - assert (Hcon := H1). + assert (Hcon := H0). assert (Σ ;;; [] |- mkApps (tFix mfix idx) args ▷ mkApps (tFix mfix idx) args') by eauto. eapply type_mkApps_inv in Hty' as (? & ? & [] & ?); eauto. assert (Ht := t). @@ -954,36 +944,36 @@ Proof. eapply inversion_Fix in t as (? & ? & ? & ? & ? & ?). eapply erases_mkApps_inv in He as [(? & ? & ? & ? & [] & ? & ? & ?) | (? & ? & ? & ? & ?)]; eauto. - + subst. assert (X100 := X1). eapply Is_type_app in X100 as[]. + + subst. assert (X100 := X2). eapply Is_type_app in X100 as[]. exists tBox. split. 2:{ assert (exists x5, Forall2 (EWcbvEval.eval Σ') x4 x5) as [x5]. { - eapply All2_app_inv in H0 as ([] & ? & ?). destruct p. + eapply All2_app_inv in X0 as ([] & ? & ?). destruct p. clear a0. subst. assert (forall x n, nth_error x3 n = Some x -> ∑ T, Σ;;; [] |- x : T). { intros. eapply typing_spine_inv with (arg := n + #|x2|) in t0 as []. 2:{ rewrite nth_error_app2. 2:lia. rewrite Nat.add_sub. eassumption. } eauto. } - clear - X3 a1 H4. revert X3 x4 H4; induction a1; intros. - ** inv H4. exists []; eauto. - ** inv H4. destruct (X3 x 0 eq_refl). + clear - X0 a1 H3. revert X0 x4 H3; induction a1; intros. + ** inv H3. exists []; eauto. + ** inv H3. destruct (X0 x 0 eq_refl). eapply r in t as (? & ? & ?); eauto. - eapply IHa1 in H3 as (? & ?); eauto. - intros. eapply (X3 x2 (S n)). eassumption. + eapply IHa1 in H4 as (? & ?); eauto. + intros. eapply (X0 x2 (S n)). eassumption. } eapply eval_box_apps. eassumption. econstructor; eauto. } econstructor. eapply Is_type_eval. eauto. eassumption. rewrite <- mkApps_nested. - eapply Is_type_red. eauto. 2:exact X2. repeat eapply PCUICReduction.red_mkApps_f. + eapply Is_type_red. eauto. 2:exact X3. repeat eapply PCUICReduction.red_mkApps_f. eapply wcbeval_red; eauto. eauto. eauto. rewrite mkApps_nested; eauto. + subst. - eapply IHeval in H3 as (? & ? & ?); eauto. - assert (H2' := H2). - inv H2. - * assert (Hmfix' := H8). - eapply All2_nth_error_Some in H8 as (? & ? & ?); eauto. + eapply IHeval in H2 as (? & ? & ?); eauto. + assert (H2' := H1). + inv H1. + * assert (Hmfix' := X2). + eapply All2_nth_error_Some in X2 as (? & ? & ?); eauto. destruct x1. cbn in *. subst. (* enough(Σ;;; [] ,,, subst_context (fix_subst mfix) 0 [] |- PCUICLiftSubst.subst (fix_subst mfix) 0 dbody ⇝ℇ subst (ETyping.fix_subst mfix') 0 (Extract.E.dbody x4)). *) @@ -1004,47 +994,47 @@ Proof. eapply All2_nth_error_Some in Hmfix' as (? & ? & ?); eauto. destruct p. destruct p. cbn in *. subst. unfold ETyping.unfold_fix, unfold_fix in *. rewrite e4. rewrite e in Hcon. - clear H1. cbn in *. + clear H0. cbn in *. destruct dbody; cbn in *; try now inv Hcon. rewrite e6 in *. eapply erases_mkApps_inv in H as [ (? & ? & ? & ? & ? & ? & ? & ?) | (? & ? & ? & ? & ?) ]. - ** assert (EAstUtils.decompose_app (EAst.mkApps (E.tFix mfix' idx) L) = EAstUtils.decompose_app (EAst.mkApps tBox x7)) by now rewrite H7. - rewrite !EAstUtils.decompose_app_mkApps in H8; try reflexivity. - inv H8. + ** assert (EAstUtils.decompose_app (EAst.mkApps (E.tFix mfix' idx) L) = EAstUtils.decompose_app (EAst.mkApps tBox x7)) by now rewrite H6. + rewrite !EAstUtils.decompose_app_mkApps in H7; try reflexivity. + inv H7. ** unfold is_constructor, ETyping.is_constructor in *. destruct nth_error eqn:En; try now inv Hcon. - --- eapply Forall2_nth_error_Some in H5 as (? & ? & ?); eauto. - inv H1. + --- eapply Forall2_nth_error_Some in H4 as (? & ? & ?); eauto. + inv H0. +++ assert (EAstUtils.decompose_app (EAst.mkApps (E.tFix mfix' idx) L) = EAstUtils.decompose_app (EAst.mkApps (E.tFix mfix'0 idx) x6)) by now rewrite H. - rewrite !EAstUtils.decompose_app_mkApps in H1; try reflexivity. - inv H1. - rewrite H5. + rewrite !EAstUtils.decompose_app_mkApps in H0; try reflexivity. + inv H0. + rewrite H4. eapply is_construct_erases in Hcon; eauto. unfold EisConstruct_app in *. destruct ?; eauto. +++ assert (EAstUtils.decompose_app (EAst.mkApps (E.tFix mfix' idx) L) = EAstUtils.decompose_app (EAst.mkApps tBox x6)) by now rewrite H. - rewrite !EAstUtils.decompose_app_mkApps in H1; try reflexivity. - inv H1. - --- eapply Forall2_nth_error_None_l in H5; eauto. - inv H1. + rewrite !EAstUtils.decompose_app_mkApps in H0; try reflexivity. + inv H0. + --- eapply Forall2_nth_error_None_l in H4; eauto. + inv H0. +++ assert (EAstUtils.decompose_app (EAst.mkApps (E.tFix mfix' idx) L) = EAstUtils.decompose_app (EAst.mkApps (E.tFix mfix'0 idx) x6)) by now rewrite H. - rewrite !EAstUtils.decompose_app_mkApps in H1; try reflexivity. - inv H1. - rewrite H5. eauto. + rewrite !EAstUtils.decompose_app_mkApps in H0; try reflexivity. + inv H0. + rewrite H4. eauto. +++ assert (EAstUtils.decompose_app (EAst.mkApps (E.tFix mfix' idx) L) = EAstUtils.decompose_app (EAst.mkApps tBox x6)) by now rewrite H. - rewrite !EAstUtils.decompose_app_mkApps in H1; try reflexivity. - inv H1. + rewrite !EAstUtils.decompose_app_mkApps in H0; try reflexivity. + inv H0. ** auto. ** eapply subject_reduction. eauto. exact Hty. etransitivity. eapply PCUICReduction.red_mkApps. now eapply wcbeval_red. eapply All2_impl. exact X. intros. now eapply wcbeval_red. econstructor. - ++ clear - t0 H4 H0. revert x t0 x3 H4; induction H0; intros. - ** inv H4. exists []; eauto. - ** inv H4. inv t0. eapply r in X1 as (? & ? & ?); eauto. - eapply IHAll2 in X2 as (? & ? & ?); eauto. - * eapply Is_type_app in X1 as []. + ++ clear - t0 H3 X0. revert x t0 x3 H3; induction X0; intros. + ** inv H3. exists []; eauto. + ** inv H3. inv t0. eapply r in X2 as (? & ? & ?); eauto. + eapply IHX0 in X3 as (? & ? & ?); eauto. + * eapply Is_type_app in X2 as []. exists tBox. split. econstructor. eapply Is_type_eval. eauto. eassumption. @@ -1055,12 +1045,12 @@ Proof. { intros. eapply typing_spine_inv with (arg := n) in t0 as []. 2:{ eassumption. } eauto. } - clear - X2 H4 H0. revert X2 x3 H4; induction H0; intros. - ** inv H4. exists []; eauto. - ** inv H4. destruct (X2 x 0 eq_refl). + clear - X3 H3 X0. revert X3 x3 H3; induction X0; intros. + ** inv H3. exists []; eauto. + ** inv H3. destruct (X3 x 0 eq_refl). eapply r in t as (? & ? & ?); eauto. - eapply IHAll2 in H5 as (? & ?); eauto. - intros. eapply (X2 x2 (S n)). eassumption. + eapply IHX0 in H4 as (? & ?); eauto. + intros. eapply (X3 x2 (S n)). eassumption. } eapply eval_box_apps. eauto. eauto. eapply wf_ext_wf. eauto. eauto. diff --git a/erasure/theories/ErasureFunction.v b/erasure/theories/ErasureFunction.v index 5e6efa5fa..0d867c323 100644 --- a/erasure/theories/ErasureFunction.v +++ b/erasure/theories/ErasureFunction.v @@ -245,7 +245,7 @@ Next Obligation. pose proof (hnf_sound HΣ (h := h)). repeat match goal with [H : squash (red _ _ _ _ ) |- _ ] => destruct H end. destruct HΣ. - eapply PCUICConfluence.red_confluence in X as [t'' []]. 3:exact X0. 2:eauto. + eapply PCUICConfluence.red_confluence in X0 as [t'' []]. 3:exact X1. 2:eauto. eapply red_red' in r. inversion r; subst. - eapply invert_red_sort in r0; eauto. @@ -271,7 +271,7 @@ Next Obligation. pose proof (hnf_sound HΣ (h := h)). repeat match goal with [H : squash (red _ _ _ _ ) |- _ ] => destruct H end. destruct HΣ. - eapply PCUICConfluence.red_confluence in X1 as [t'' []]. 3:exact X2. 2:eauto. + eapply PCUICConfluence.red_confluence in X2 as [t'' []]. 3:exact X3. 2:eauto. eapply red_red' in r. inversion r; subst. - eapply invert_red_prod in r0 as (? & ? & [] & ?); eauto. @@ -388,7 +388,7 @@ Program Definition is_erasable (Sigma : PCUICAst.global_env_ext) (HΣ : ∥wf_ex ret (left _) else mlet (K; _) <- @make_graph_and_infer _ _ HΣ Gamma HΓ T ;; mlet (u;_) <- @reduce_to_sort _ Sigma _ Gamma K _ ;; - match is_prop_sort u with true => ret (left _) | false => ret (right _) end + match Universe.is_prop u with true => ret (left _) | false => ret (right _) end . Next Obligation. sq; eauto. Qed. Next Obligation. @@ -615,11 +615,11 @@ Proof. pose proof (Prelim.monad_map_All2 _ _ _ brs a2 E2). - eapply All2_All_left in X3. 2:{ intros. destruct X4. destruct p0. destruct p0. exact e0. } + eapply All2_All_left in X3. 2:{ intros. destruct X5. destruct p0. destruct p0. exact e0. } eapply All2_impl. eapply All2_All_mix_left. eassumption. eassumption. - intros. destruct H5. + intros. destruct H. destruct ?; inv e0. cbn. eauto. - econstructor. clear E. @@ -635,7 +635,7 @@ Proof. pose proof (Prelim.monad_map_All2 _ _ _ mfix a1 E1). eapply All2_impl. eapply All2_All_mix_left. exact X0. eassumption. - intros. destruct X1. cbn in *. unfold bind in e. cbn in e. + intros. destruct X2. cbn in *. unfold bind in e. cbn in e. repeat destruct ?; try congruence; inv e. cbn. repeat split; eauto. @@ -773,10 +773,10 @@ Proof. eapply All2_Forall2. eapply All2_impl. eassumption. - intros. cbn in H0. - unfold erase_one_inductive_body, bind in H0. cbn in H0. + intros. cbn in H. + unfold erase_one_inductive_body, bind in H. cbn in H. repeat destruct ?; try congruence. - inv H0. + inv H. unfold erases_one_inductive_body. cbn. destruct ?; cbn. (* unfold lift_opt_typing in E. *) (* destruct decompose_prod_n_assum eqn:E6; inv E. cbn. *) @@ -786,7 +786,7 @@ Proof. eapply All2_impl_In. eassumption. intros. destruct x0, p, y, p. cbn in *. destruct ?; try congruence. - inv H4. split; eauto. + inv H1. split; eauto. (* pose (t' := t). inv t'. cbn in *. *) destruct (erase_Some_typed E6) as [? []]. @@ -794,9 +794,9 @@ Proof. eapply erases_erase. 2:eauto. eauto. -- eapply All2_Forall2. eapply All2_impl_In. eassumption. - intros. cbn in H2. destruct x0, y. + intros. cbn in H1. destruct x0, y. destruct ?; try congruence; - inv H4. split; eauto. + inv H1. split; eauto. (* pose (t' := t). inv t'. cbn in *. *) destruct (erase_Some_typed E6) as [? []]. diff --git a/erasure/theories/Extract.v b/erasure/theories/Extract.v index a0256d575..d56738e28 100644 --- a/erasure/theories/Extract.v +++ b/erasure/theories/Extract.v @@ -11,7 +11,7 @@ Import MonadNotation. Local Existing Instance extraction_checker_flags. -Definition isErasable Σ Γ t := ∑ T, Σ ;;; Γ |- t : T × (isArity T + (∑ u, (Σ ;;; Γ |- T : tSort u) * is_prop_sort u))%type. +Definition isErasable Σ Γ t := ∑ T, Σ ;;; Γ |- t : T × (isArity T + (∑ u, (Σ ;;; Γ |- T : tSort u) * Universe.is_prop u))%type. Module E := EAst. Local Notation Ret t := t. @@ -47,9 +47,9 @@ Inductive erases (Σ : global_env_ext) (Γ : context) : term -> E.term -> Prop : | erases_tApp : forall (f u : term) (f' u' : E.term), Σ;;; Γ |- f ⇝ℇ f' -> Σ;;; Γ |- u ⇝ℇ u' -> Σ;;; Γ |- tApp f u ⇝ℇ E.tApp f' u' - | erases_tConst : forall (kn : kername) (u : universe_instance), + | erases_tConst : forall (kn : kername) (u : Instance.t), Σ;;; Γ |- tConst kn u ⇝ℇ E.tConst kn - | erases_tConstruct : forall (kn : inductive) (k : nat) (n : universe_instance), + | erases_tConstruct : forall (kn : inductive) (k : nat) (n : Instance.t), Σ;;; Γ |- tConstruct kn k n ⇝ℇ E.tConstruct kn k | erases_tCase1 : forall (ind : inductive) (npar : nat) (T c : term) (brs : list (nat × term)) (c' : E.term) diff --git a/erasure/theories/SafeErasureFunction.v b/erasure/theories/SafeErasureFunction.v index 09a0c65e6..eba17d367 100644 --- a/erasure/theories/SafeErasureFunction.v +++ b/erasure/theories/SafeErasureFunction.v @@ -216,7 +216,7 @@ Program Definition is_erasable (Sigma : PCUICAst.global_env_ext) (HΣ : ∥wf_ex ret (left _) else mlet (K; _) <- @type_of extraction_checker_flags Sigma _ _ Gamma T _ ;; mlet (u;_) <- @reduce_to_sort _ Sigma _ Gamma K _ ;; - match is_prop_sort u with true => ret (left _) | false => ret (right _) end + match Universe.is_prop u with true => ret (left _) | false => ret (right _) end . Next Obligation. sq; eauto. Qed. Next Obligation. @@ -298,7 +298,7 @@ Qed. (* ret (left _) *) (* else mlet (K; _) <- @make_graph_and_infer _ _ HΣ Gamma HΓ T ;; *) (* mlet (u;_) <- @reduce_to_sort _ Sigma _ Gamma K _ ;; *) -(* match is_prop_sort u with true => ret (left _) | false => ret (right _) end *) +(* match Universe.is_prop u with true => ret (left _) | false => ret (right _) end *) (* . *) (* Next Obligation. sq; eauto. Qed. *) (* Next Obligation. *) diff --git a/examples/demo.v b/examples/demo.v index 36561d13b..3a3f94c90 100644 --- a/examples/demo.v +++ b/examples/demo.v @@ -356,9 +356,9 @@ Inductive T : Type := Quote Recursively Definition TT := T. Unset Strict Unquote Universe Mode. -Make Definition t := (tSort (NEL.sing (Level.Level "Top.20000", false))). +Make Definition t := (tSort (Universe.make (Level.Level "Top.20000"))). Make Definition t' := (tSort fresh_universe). -Make Definition myProp := (tSort (Universe.make' (Level.lProp, false))). +Make Definition myProp := (tSort (Universe.make Level.lProp)). Make Definition myProp' := (tSort Universe.type0m). Make Definition mySet := (tSort (Universe.make Level.lSet)). diff --git a/pcuic/theories/PCUICAlpha.v b/pcuic/theories/PCUICAlpha.v index 0f68408f8..6edfd90fa 100644 --- a/pcuic/theories/PCUICAlpha.v +++ b/pcuic/theories/PCUICAlpha.v @@ -274,19 +274,19 @@ Qed. - intros cst u decl ? ? hdecl hcons v e. dependent destruction e. apply Forall2_eq in r. apply map_inj in r ; revgoals. - { intros x y h. inversion h. reflexivity. } + { apply Universe.make_inj. } subst. constructor ; auto. - intros ind u mdecl idecl isdecl ? ? hcons v e. dependent destruction e. apply Forall2_eq in r. apply map_inj in r ; revgoals. - { intros x y h. inversion h. reflexivity. } + { apply Universe.make_inj. } subst. econstructor ; eauto. - intros ind i u mdecl idecl cdecl isdecl ? ? ? v e. dependent destruction e. apply Forall2_eq in r. apply map_inj in r ; revgoals. - { intros x y h. inversion h. reflexivity. } + { apply Universe.make_inj. } subst. econstructor ; eauto. - intros ind u npar p c brs args mdecl idecl isdecl X X0 H pars ps pty @@ -698,38 +698,38 @@ Qed. revert t u Rle. fix aux 4. destruct 1; cbn; intros t'' u'' H' H0'; inv H'; inv H0'; try econstructor; eauto. - - revert args'0 args'1 H2 H3. + - revert args'0 args'1 X X0. induction a; simpl; intros args0 args'0 H1 H2. + inv H1; inv H2; constructor; eauto. + inv H1; inv H2. constructor; eauto. - apply Forall2_eq, map_inj in H2. apply Forall2_eq, map_inj in H3. congruence. - all: intros x y H; now inv H. + all: apply Universe.make_inj. - apply Forall2_eq, map_inj in H2. apply Forall2_eq, map_inj in H3. congruence. - all: intros x y H; now inv H. + all: apply Universe.make_inj. - apply Forall2_eq, map_inj in H3. apply Forall2_eq, map_inj in H4. congruence. - all: intros x y H; now inv H. - - revert brs'0 brs'1 H8 H11. + all: apply Universe.make_inj. + - revert brs'0 brs'1 X3 X6. induction a; simpl; intros args0 args'0 H1 H2. + inv H1; inv H2; constructor; eauto. + inv H1; inv H2. constructor; eauto. - destruct H5, H4, r. split; eauto. congruence. - - revert mfix'0 mfix'1 H2 H3. + destruct X3, X7, r. split; eauto. congruence. + - revert mfix'0 mfix'1 X X0. induction a; simpl; intros args0 args'0 H1 H2. + inv H1; inv H2; constructor; eauto. + inv H1; inv H2. constructor; eauto. - destruct H3 as [[? ?] ?], H1 as [[? ?] ?], r as [[? ?] ?]. + destruct X as [[? ?] ?], X1 as [[? ?] ?], r as [[? ?] ?]. repeat split; eauto. congruence. - - revert mfix'0 mfix'1 H2 H3. + - revert mfix'0 mfix'1 X X0. induction a; simpl; intros args0 args'0 H1 H2. + inv H1; inv H2; constructor; eauto. + inv H1; inv H2. constructor; eauto. - destruct H3 as [[? ?] ?], H1 as [[? ?] ?], r as [[? ?] ?]. + destruct X as [[? ?] ?], X1 as [[? ?] ?], r as [[? ?] ?]. repeat split; eauto. congruence. Qed. diff --git a/pcuic/theories/PCUICAst.v b/pcuic/theories/PCUICAst.v index cd49d4426..b85c1ad7b 100644 --- a/pcuic/theories/PCUICAst.v +++ b/pcuic/theories/PCUICAst.v @@ -15,18 +15,18 @@ Open Scope pcuic. In particular, it has binary applications and all terms are well-formed. Casts are absent as well. *) -Inductive term : Set := +Inductive term := | tRel (n : nat) | tVar (i : ident) (* For free variables (e.g. in a goal) *) | tEvar (n : nat) (l : list term) -| tSort (u : universe) +| tSort (u : Universe.t) | tProd (na : name) (A B : term) | tLambda (na : name) (A t : term) | tLetIn (na : name) (b B t : term) (* let na := b : B in t *) | tApp (u v : term) -| tConst (k : kername) (ui : universe_instance) -| tInd (ind : inductive) (ui : universe_instance) -| tConstruct (ind : inductive) (n : nat) (ui : universe_instance) +| tConst (k : kername) (ui : Instance.t) +| tInd (ind : inductive) (ui : Instance.t) +| tConstruct (ind : inductive) (n : nat) (ui : Instance.t) | tCase (indn : inductive * nat) (p c : term) (brs : list (nat * term)) (* # of parameters/type info/discriminee/branches *) | tProj (p : projection) (c : term) | tFix (mfix : mfixpoint term) (idx : nat) @@ -93,11 +93,11 @@ Inductive constant_entry := [x1:X1;...;xn:Xn]. *) -Inductive local_entry : Set := +Inductive local_entry := | LocalDef : term -> local_entry (* local let binding *) | LocalAssum : term -> local_entry. -Record one_inductive_entry : Set := { +Record one_inductive_entry := { mind_entry_typename : ident; mind_entry_arity : term; mind_entry_template : bool; (* template polymorphism *) diff --git a/pcuic/theories/PCUICAstUtils.v b/pcuic/theories/PCUICAstUtils.v index a73fbf8c3..fc13be592 100644 --- a/pcuic/theories/PCUICAstUtils.v +++ b/pcuic/theories/PCUICAstUtils.v @@ -100,14 +100,14 @@ Proof. destruct decl; reflexivity. Qed. Lemma map_cst_body f decl : option_map f (cst_body decl) = cst_body (map_constant_body f decl). Proof. destruct decl; reflexivity. Qed. -Definition map_def {A B : Set} (tyf bodyf : A -> B) (d : def A) := +Definition map_def {A B} (tyf bodyf : A -> B) (d : def A) := {| dname := d.(dname); dtype := tyf d.(dtype); dbody := bodyf d.(dbody); rarg := d.(rarg) |}. -Lemma map_dtype {A B : Set} (f : A -> B) (g : A -> B) (d : def A) : +Lemma map_dtype {A B} (f : A -> B) (g : A -> B) (d : def A) : f (dtype d) = dtype (map_def f g d). Proof. destruct d; reflexivity. Qed. -Lemma map_dbody {A B : Set} (f : A -> B) (g : A -> B) (d : def A) : +Lemma map_dbody {A B} (f : A -> B) (g : A -> B) (d : def A) : g (dbody d) = dbody (map_def f g d). Proof. destruct d; reflexivity. Qed. @@ -513,13 +513,13 @@ Lemma ind_projs_map f npars_ass arities n oib : map (on_snd (f (S npars_ass))) (ind_projs oib). Proof. destruct oib; simpl. reflexivity. Qed. -Definition test_def {A : Set} (tyf bodyf : A -> bool) (d : def A) := +Definition test_def {A} (tyf bodyf : A -> bool) (d : def A) := tyf d.(dtype) && bodyf d.(dbody). Definition tCaseBrsProp {A} (P : A -> Type) (l : list (nat * A)) := All (fun x => P (snd x)) l. -Definition tFixProp {A : Set} (P P' : A -> Type) (m : mfixpoint A) := +Definition tFixProp {A} (P P' : A -> Type) (m : mfixpoint A) := All (fun x : def A => P x.(dtype) * P' x.(dbody))%type m. @@ -528,21 +528,21 @@ Ltac merge_All := repeat toAll. -Lemma map_def_map_def {A B C : Set} (f f' : B -> C) (g g' : A -> B) (d : def A) : +Lemma map_def_map_def {A B C} (f f' : B -> C) (g g' : A -> B) (d : def A) : map_def f f' (map_def g g' d) = map_def (fun x => f (g x)) (fun x => f' (g' x)) d. Proof. destruct d; reflexivity. Qed. -Lemma compose_map_def {A B C : Set} (f f' : B -> C) (g g' : A -> B) : +Lemma compose_map_def {A B C} (f f' : B -> C) (g g' : A -> B) : compose (A:=def A) (map_def f f') (map_def g g') = map_def (compose f g) (compose f' g'). Proof. reflexivity. Qed. -Lemma map_def_id {t : Set} x : map_def (@id t) (@id t) x = x. +Lemma map_def_id {t} x : map_def (@id t) (@id t) x = x. Proof. now destruct x. Qed. Hint Rewrite @map_def_id @map_id : map. -Lemma map_def_spec {A B : Set} (P P' : A -> Type) (f f' g g' : A -> B) (x : def A) : +Lemma map_def_spec {A B} (P P' : A -> Type) (f f' g g' : A -> B) (x : def A) : P' x.(dbody) -> P x.(dtype) -> (forall x, P x -> f x = g x) -> (forall x, P' x -> f' x = g' x) -> map_def f f' x = map_def g g' x. @@ -551,7 +551,7 @@ Proof. rewrite !H // !H0 //. Qed. -Lemma case_brs_map_spec {A B : Set} {P : A -> Type} {l} {f g : A -> B} : +Lemma case_brs_map_spec {A B} {P : A -> Type} {l} {f g : A -> B} : tCaseBrsProp P l -> (forall x, P x -> f x = g x) -> map (on_snd f) l = map (on_snd g) l. Proof. @@ -560,7 +560,7 @@ Proof. apply on_snd_eq_spec; eauto. Qed. -Lemma tfix_map_spec {A B : Set} {P P' : A -> Type} {l} {f f' g g' : A -> B} : +Lemma tfix_map_spec {A B} {P P' : A -> Type} {l} {f f' g g' : A -> B} : tFixProp P P' l -> (forall x, P x -> f x = g x) -> (forall x, P' x -> f' x = g' x) -> map (map_def f f') l = map (map_def g g') l. @@ -572,7 +572,7 @@ Proof. Qed. -Lemma map_def_test_spec {A B : Set} +Lemma map_def_test_spec {A B} (P P' : A -> Type) (p p' : pred A) (f f' g g' : A -> B) (x : def A) : P x.(dtype) -> P' x.(dbody) -> (forall x, P x -> p x -> f x = g x) -> (forall x, P' x -> p' x -> f' x = g' x) -> @@ -584,7 +584,7 @@ Proof. rewrite !H // !H0 //; intuition auto. Qed. -Lemma case_brs_forallb_map_spec {A B : Set} {P : A -> Type} {p : A -> bool} +Lemma case_brs_forallb_map_spec {A B} {P : A -> Type} {p : A -> bool} {l} {f g : A -> B} : tCaseBrsProp P l -> forallb (test_snd p) l -> @@ -597,7 +597,7 @@ Proof. eapply on_snd_test_spec; eauto. Qed. -Lemma tfix_forallb_map_spec {A B : Set} {P P' : A -> Prop} {p p'} {l} {f f' g g' : A -> B} : +Lemma tfix_forallb_map_spec {A B} {P P' : A -> Prop} {p p'} {l} {f f' g g' : A -> B} : tFixProp P P' l -> forallb (test_def p p') l -> (forall x, P x -> p x -> f x = g x) -> @@ -1015,4 +1015,4 @@ Definition fst_ctx : global_env_ext -> global_env := fst. Coercion fst_ctx : global_env_ext >-> global_env. Definition empty_ext (Σ : global_env) : global_env_ext - := (Σ, Monomorphic_ctx ContextSet.empty). \ No newline at end of file + := (Σ, Monomorphic_ctx ContextSet.empty). diff --git a/pcuic/theories/PCUICChecker.v b/pcuic/theories/PCUICChecker.v index 0d6364e00..a275f4146 100644 --- a/pcuic/theories/PCUICChecker.v +++ b/pcuic/theories/PCUICChecker.v @@ -38,7 +38,7 @@ Inductive type_error := | NotAnInductive (t : term) | IllFormedFix (m : mfixpoint term) (i : nat) | UnsatisfiedConstraints (c : ConstraintSet.t) -| CannotTakeSuccessor (u : universe) +| CannotTakeSuccessor (u : Universe.t) | NotEnoughFuel (n : nat). diff --git a/pcuic/theories/PCUICClosed.v b/pcuic/theories/PCUICClosed.v index c314b57a0..108e50a09 100644 --- a/pcuic/theories/PCUICClosed.v +++ b/pcuic/theories/PCUICClosed.v @@ -26,7 +26,7 @@ Proof. rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length, ?Nat.add_assoc; simpl closed in *; solve_all; unfold compose, test_def, test_snd in *; - try solve [simpl lift; simpl closed; f_equal; auto; repeat (toProp; solve_all)]; try easy. + try solve [simpl lift; simpl closed; f_equal; auto; repeat (rtoProp; solve_all)]; try easy. - elim (Nat.leb_spec k' n0); intros. simpl. elim (Nat.ltb_spec); auto. apply Nat.ltb_lt in H. lia. @@ -41,7 +41,7 @@ Proof. induction t in n, k, k' |- * using term_forall_list_ind; intros; simpl in *; rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length, ?Nat.add_assoc in *; - simpl closed in *; repeat (toProp; solve_all); try change_Sk; + simpl closed in *; repeat (rtoProp; solve_all); try change_Sk; unfold compose, test_def, on_snd, test_snd in *; simpl in *; eauto with all. - revert H0. @@ -51,8 +51,8 @@ Proof. - specialize (IHt2 n (S k) (S k')). eauto with all. - specialize (IHt2 n (S k) (S k')). eauto with all. - specialize (IHt3 n (S k) (S k')). eauto with all. - - toProp. solve_all. specialize (b0 n (#|m| + k) (#|m| + k')). eauto with all. - - toProp. solve_all. specialize (b0 n (#|m| + k) (#|m| + k')). eauto with all. + - rtoProp. solve_all. specialize (b0 n (#|m| + k) (#|m| + k')). eauto with all. + - rtoProp. solve_all. specialize (b0 n (#|m| + k) (#|m| + k')). eauto with all. Qed. Lemma closedn_mkApps k f u: @@ -81,7 +81,7 @@ Proof. induction t in k' |- * using term_forall_list_ind; intros; simpl in *; rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; - simpl closed in *; try change_Sk; repeat (toProp; solve_all); + simpl closed in *; try change_Sk; repeat (rtoProp; solve_all); unfold compose, test_def, on_snd, test_snd in *; simpl in *; eauto with all. - elim (Nat.leb_spec k' n); intros. simpl. @@ -100,10 +100,10 @@ Proof. rewrite <- Nat.add_succ_comm in IHt2. eauto. - specialize (IHt3 (S k')). rewrite <- Nat.add_succ_comm in IHt3. eauto. - - toProp; solve_all. rewrite -> !Nat.add_assoc in *. + - rtoProp; solve_all. rewrite -> !Nat.add_assoc in *. specialize (b0 (#|m| + k')). unfold is_true. rewrite <- b0. f_equal. lia. unfold is_true. rewrite <- H0. f_equal. lia. - - toProp; solve_all. rewrite -> !Nat.add_assoc in *. + - rtoProp; solve_all. rewrite -> !Nat.add_assoc in *. specialize (b0 (#|m| + k')). unfold is_true. rewrite <- b0. f_equal. lia. unfold is_true. rewrite <- H0. f_equal. lia. Qed. @@ -165,7 +165,7 @@ Proof. induction t in σ, k, Hs |- * using term_forall_list_ind; intros; sigma; simpl in *; rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length, ?Nat.add_assoc in *; - simpl closed in *; repeat (toProp; f_equal; solve_all); try change_Sk; + simpl closed in *; repeat (rtoProp; f_equal; solve_all); try change_Sk; unfold compose, test_def, on_snd, test_snd in *; simpl in *; eauto with all. - revert Hs. @@ -178,9 +178,9 @@ Proof. - specialize (IHt2 σ (S k) H0). rewrite -{2}IHt2. now sigma. - specialize (IHt2 σ (S k) H0). rewrite -{2}IHt2. now sigma. - specialize (IHt3 σ (S k) H0). rewrite -{2}IHt3. now sigma. - - toProp. specialize (b0 σ (#|m| + k) H0). eapply map_def_id_spec; auto. + - rtoProp. specialize (b0 σ (#|m| + k) H0). eapply map_def_id_spec; auto. revert b0. now sigma. - - toProp. specialize (b0 σ (#|m| + k) H0). eapply map_def_id_spec; auto. + - rtoProp. specialize (b0 σ (#|m| + k) H0). eapply map_def_id_spec; auto. revert b0. now sigma. Qed. @@ -392,7 +392,7 @@ Proof. - intuition auto. + solve_all. unfold test_snd. simpl in *. - toProp; eauto. + rtoProp; eauto. + apply closedn_mkApps; auto. rewrite forallb_app. simpl. rewrite H1. rewrite forallb_skipn; auto. @@ -420,26 +420,26 @@ Proof. - clear H0. split. solve_all. destruct x; simpl in *. - unfold test_def. simpl. toProp. + unfold test_def. simpl. rtoProp. split. rewrite -> app_context_length in *. rewrite -> Nat.add_comm in *. eapply closedn_lift_inv in H1; eauto. lia. subst types. now rewrite app_context_length fix_context_length in H0. - eapply nth_error_all in H; eauto. simpl in H. intuition auto. toProp. + eapply nth_error_all in H; eauto. simpl in H. intuition auto. rtoProp. subst types. rewrite app_context_length in H0. rewrite Nat.add_comm in H0. now eapply closedn_lift_inv in H0. - split. solve_all. destruct x; simpl in *. - unfold test_def. simpl. toProp. + unfold test_def. simpl. rtoProp. split. rewrite -> app_context_length in *. rewrite -> Nat.add_comm in *. eapply closedn_lift_inv in H2; eauto. lia. subst types. now rewrite -> app_context_length, fix_context_length in H1. eapply (nth_error_all) in H; eauto. simpl in *. - intuition auto. toProp. + intuition auto. rtoProp. subst types. rewrite app_context_length in H1. rewrite Nat.add_comm in H1. now eapply closedn_lift_inv in H1. @@ -512,7 +512,7 @@ Proof. eapply on_global_env_impl; cycle 1. apply (env_prop_sigma _ typecheck_closed _ X). red; intros. unfold lift_typing in *. destruct T; intuition auto with wf. - destruct X1 as [s0 Hs0]. simpl. toProp; intuition. + destruct X1 as [s0 Hs0]. simpl. rtoProp; intuition. Qed. Lemma closed_decl_upwards k d : closed_decl k d -> forall k', k <= k' -> closed_decl k' d. diff --git a/pcuic/theories/PCUICConfluence.v b/pcuic/theories/PCUICConfluence.v index e4f1eaae9..9464770c6 100644 --- a/pcuic/theories/PCUICConfluence.v +++ b/pcuic/theories/PCUICConfluence.v @@ -9,7 +9,7 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICInduction PCUICLiftSubst PCUICUnivSubst PCUICTyping PCUICReduction PCUICWeakening PCUICSubstitution PCUICEquality PCUICReflect PCUICClosed PCUICParallelReduction PCUICParallelReductionConfluence - PCUICCumulativity. + PCUICCumulativity PCUICUnivSubstitution. (* Type-valued relations. *) Require Import CRelationClasses. @@ -668,7 +668,7 @@ Proof. × (forall Rle (u' : term), RelationClasses.Reflexive Rle -> RelationClasses.Transitive Rle -> - (forall u u'0 : universe, Re u u'0 -> Rle u u'0) -> + (forall u u'0 : Universe.t, Re u u'0 -> Rle u u'0) -> eq_term_upto_univ Re Rle (dbody x) u' -> ∑ v' : term, red1 Σ (Γ ,,, fix_context L) u' v' @@ -812,7 +812,7 @@ Proof. × (forall Rle (u' : term), RelationClasses.Reflexive Rle -> RelationClasses.Transitive Rle -> - (forall u u'0 : universe, Re u u'0 -> Rle u u'0) -> + (forall u u'0 : Universe.t, Re u u'0 -> Rle u u'0) -> eq_term_upto_univ Re Rle (dbody x) u' -> ∑ v' : term, red1 Σ (Γ ,,, fix_context L) u' v' @@ -946,7 +946,7 @@ Qed. Section RedEq. Context (Σ : global_env_ext). - Context {Re Rle : universe -> universe -> Prop} + Context {Re Rle : Universe.t -> Universe.t -> Prop} {refl : RelationClasses.Reflexive Re} {refl': RelationClasses.Reflexive Rle} {pres : SubstUnivPreserving Re} @@ -1009,23 +1009,23 @@ Instance clos_refl_trans_Signature (A : Type) (R : Relation.relation A) (x H : A clos_refl_trans_sig_pack A R x H. Unset Universe Polymorphism. -(* Derive Signature for red1. *) +Derive Signature for red1. -Definition red1_sig@{} (Σ : global_env) (index : sigma (fun _ : context => sigma (fun _ : term => term))) := - let Γ := pr1 index in let H := pr1 (pr2 index) in let H0 := pr2 (pr2 index) in red1 Σ Γ H H0. +(* Definition red1_sig@{} (Σ : global_env) (index : sigma (fun _ : context => sigma (fun _ : term => term))) := *) +(* let Γ := pr1 index in let H := pr1 (pr2 index) in let H0 := pr2 (pr2 index) in red1 Σ Γ H H0. *) -Universe red1u. +(* Universe red1u. *) -Definition red1_sig_pack@{} (Σ : global_env) (Γ : context) (H H0 : term) (red1_var : red1 Σ Γ H H0) - : sigma@{red1u} (fun index : sigma (fun _ : context => sigma (fun _ : term => term)) => - red1 Σ (pr1 index) (pr1 (pr2 index)) (pr2 (pr2 index))) - := - {| pr1 := {| pr1 := Γ; pr2 := {| pr1 := H; pr2 := H0 |} |}; pr2 := red1_var |}. +(* Definition red1_sig_pack@{} (Σ : global_env) (Γ : context) (H H0 : term) (red1_var : red1 Σ Γ H H0) *) +(* : sigma@{red1u} (fun index : sigma (fun _ : context => sigma (fun _ : term => term)) => *) +(* red1 Σ (pr1 index) (pr1 (pr2 index)) (pr2 (pr2 index))) *) +(* := *) +(* {| pr1 := {| pr1 := Γ; pr2 := {| pr1 := H; pr2 := H0 |} |}; pr2 := red1_var |}. *) -Instance red1_Signature@{} (Σ : global_env) (Γ : context) (H H0 : term) - : Signature@{red1u} (red1 Σ Γ H H0) (sigma (fun _ : context => sigma (fun _ : term => term))) (red1_sig Σ) := - red1_sig_pack Σ Γ H H0. -(** FIXME Equations *) +(* Instance red1_Signature@{} (Σ : global_env) (Γ : context) (H H0 : term) *) +(* : Signature@{red1u} (red1 Σ Γ H H0) (sigma (fun _ : context => sigma (fun _ : term => term))) (red1_sig Σ) := *) +(* red1_sig_pack Σ Γ H H0. *) +(* (** FIXME Equations *) ==> seems fixed now that term is not in Set anymore *) Lemma local_env_telescope P Γ Γ' Δ Δ' : All2_telescope (on_decl P) Γ Γ' Δ Δ' -> @@ -2394,8 +2394,8 @@ Section RedConfluence. option_map decl_body (nth_error Γ' n) = Some d. Proof. move: Γ' n d; induction Γ; destruct n; simpl; intros; try congruence. - noconf H. depelim H0. simpl. now rewrite -e. - depelim H0. simpl. apply IHΓ; auto. + noconf H. depelim X. simpl. now rewrite -e. + depelim X. simpl. apply IHΓ; auto. Qed. Lemma decl_type_eq_context_upto_names Γ Γ' n d : @@ -2404,8 +2404,8 @@ Section RedConfluence. option_map decl_type (nth_error Γ' n) = Some d. Proof. move: Γ' n d; induction Γ; destruct n; simpl; intros; try congruence. - noconf H. depelim H0. simpl. now rewrite -e0. - depelim H0. simpl. apply IHΓ; auto. + noconf H. depelim X. simpl. now rewrite -e0. + depelim X. simpl. apply IHΓ; auto. Qed. Lemma eq_context_upto_names_app Γ Γ' Δ : @@ -2466,17 +2466,17 @@ Section RedConfluence. red_ctx Γ' Δ'. Proof. intros. - induction H in H0, Δ, Δ', X |- *. depelim X. depelim H0. constructor. + induction X in X0, Δ, Δ', X1 |- *. depelim X1. depelim X0. constructor. destruct x as [na b ty], y as [na' b' ty']; simpl in *. subst. - depelim X. depelim H0. hnf in H1. noconf H1. + depelim X1. depelim X0. hnf in H. noconf H. red in o. simpl in *. subst. destruct y as [? [b'|] ?]; noconf e. - constructor; auto. eapply IHeq_context_upto_names; eauto. + constructor; auto. eapply IHX; eauto. red. eapply red_eq_context_upto_names; eauto. - hnf in H1. noconf H1. destruct o. depelim H0. simpl in *. + hnf in H. noconf H. destruct o. depelim X0. simpl in *. destruct y as [? [b'|] ?]; noconf e. subst; simpl in *. - constructor. eapply IHeq_context_upto_names; eauto. + constructor. eapply IHX; eauto. red. split; eauto using red_eq_context_upto_names. Qed. @@ -2487,7 +2487,7 @@ Section RedConfluence. reduce_goal. split. intros. eapply eq_context_upto_names_red_ctx; eauto. - intros. symmetry in H, H0. eapply eq_context_upto_names_red_ctx; eauto. + intros. symmetry in X, X0. eapply eq_context_upto_names_red_ctx; eauto. Qed. Lemma clos_rt_red1_alpha_out x y : diff --git a/pcuic/theories/PCUICContextConversion.v b/pcuic/theories/PCUICContextConversion.v index 8db10914d..e3f4568f5 100644 --- a/pcuic/theories/PCUICContextConversion.v +++ b/pcuic/theories/PCUICContextConversion.v @@ -11,7 +11,7 @@ From MetaCoq Require Import LibHypsNaming. From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICInduction PCUICLiftSubst PCUICUnivSubst PCUICTyping PCUICWeakeningEnv PCUICWeakening PCUICSubstitution PCUICClosed PCUICCumulativity PCUICGeneration PCUICReduction - PCUICParallelReduction PCUICEquality + PCUICParallelReduction PCUICEquality PCUICUnivSubstitution PCUICParallelReductionConfluence PCUICConfluence. From Equations Require Import Equations. diff --git a/pcuic/theories/PCUICElimination.v b/pcuic/theories/PCUICElimination.v index ef0cc9016..846919551 100644 --- a/pcuic/theories/PCUICElimination.v +++ b/pcuic/theories/PCUICElimination.v @@ -9,7 +9,7 @@ From MetaCoq.PCUIC Require Import PCUICTyping PCUICAst PCUICAstUtils PCUICInduct From Equations Require Import Equations. Require Import Equations.Prop.DepElim. -Definition Is_proof `{cf : checker_flags} Σ Γ t := ∑ T u, Σ ;;; Γ |- t : T × Σ ;;; Γ |- T : tSort u × is_prop_sort u. +Definition Is_proof `{cf : checker_flags} Σ Γ t := ∑ T u, Σ ;;; Γ |- t : T × Σ ;;; Γ |- T : tSort u × Universe.is_prop u. Lemma declared_inductive_inj `{cf : checker_flags} {Σ mdecl mdecl' ind idecl idecl'} : declared_inductive Σ mdecl' ind idecl' -> @@ -213,7 +213,7 @@ Variable Hcf : prop_sub_type = false. Lemma cumul_prop1 (Σ : global_env_ext) Γ A B u : wf Σ -> - is_prop_sort u -> Σ ;;; Γ |- B : tSort u -> + Universe.is_prop u -> Σ ;;; Γ |- B : tSort u -> Σ ;;; Γ |- A <= B -> Σ ;;; Γ |- A : tSort u. Proof. intros. induction X1. @@ -224,7 +224,7 @@ Admitted. (* cumul_prop1 *) Lemma cumul_prop2 (Σ : global_env_ext) Γ A B u : wf Σ -> - is_prop_sort u -> Σ ;;; Γ |- A <= B -> + Universe.is_prop u -> Σ ;;; Γ |- A <= B -> Σ ;;; Γ |- A : tSort u -> Σ ;;; Γ |- B : tSort u. Proof. Admitted. (* cumul_prop2 *) @@ -234,8 +234,8 @@ Lemma leq_universe_prop (Σ : global_env_ext) u1 u2 : (* @prop_sub_type cf = false -> *) wf Σ -> @leq_universe cf (global_ext_constraints Σ) u1 u2 -> - (is_prop_sort u1 \/ is_prop_sort u2) -> - (is_prop_sort u1 /\ is_prop_sort u2). + (Universe.is_prop u1 \/ Universe.is_prop u2) -> + (Universe.is_prop u1 /\ Universe.is_prop u2). Proof. intros. unfold leq_universe in *. (* rewrite H in H1. *) (* unfold leq_universe0 in H1. *) diff --git a/pcuic/theories/PCUICEquality.v b/pcuic/theories/PCUICEquality.v index 90af29ef8..af27e5df3 100644 --- a/pcuic/theories/PCUICEquality.v +++ b/pcuic/theories/PCUICEquality.v @@ -55,7 +55,7 @@ Qed. (** ** Syntactic equality up-to universes We don't look at printing annotations *) -Inductive eq_term_upto_univ (Re Rle : universe -> universe -> Prop) : term -> term -> Type := +Inductive eq_term_upto_univ (Re Rle : Universe.t -> Universe.t -> Prop) : term -> term -> Type := | eq_Rel n : eq_term_upto_univ Re Rle (tRel n) (tRel n) @@ -143,20 +143,6 @@ Definition eq_term `{checker_flags} φ := Definition leq_term `{checker_flags} φ := eq_term_upto_univ (eq_universe φ) (leq_universe φ). -(* TODO MOVE *) -Lemma Forall2_same : - forall A (P : A -> A -> Prop) l, - (forall x, P x x) -> - Forall2 P l l. -Proof. - intros A P l h. - induction l. - - constructor. - - constructor. - + eapply h. - + assumption. -Qed. - Instance eq_term_upto_univ_refl Re Rle : RelationClasses.Reflexive Re -> RelationClasses.Reflexive Rle -> @@ -319,18 +305,6 @@ Instance leq_term_preorder {cf:checker_flags} φ : PreOrder (leq_term φ) := {| PreOrder_Reflexive := leq_term_refl _; PreOrder_Transitive := leq_term_trans _ |}. -(* TODO MOVE *) -Lemma Forall2_sym : - forall A (P : A -> A -> Prop) l l', - Forall2 P l l' -> - Forall2 (fun x y => P y x) l' l. -Proof. - intros A P l l' h. - induction h. - - constructor. - - constructor. all: auto. -Qed. - Instance R_universe_instance_equiv R (hR : RelationClasses.Equivalence R) : RelationClasses.Equivalence (R_universe_instance R). Proof. @@ -529,7 +503,7 @@ Qed. (** ** Boolean version ** *) -Fixpoint eqb_term_upto_univ (equ lequ : universe -> universe -> bool) (u v : term) : bool := +Fixpoint eqb_term_upto_univ (equ lequ : Universe.t -> Universe.t -> bool) (u v : term) : bool := match u, v with | tRel n, tRel m => eqb n m @@ -629,21 +603,6 @@ Local Ltac ih := destruct (ih lequ Rle hle t') ; nodec ; subst end. -(* TODO MOVE *) -Lemma forallb2_Forall2 : - forall A (p : A -> A -> bool) l l', - forallb2 p l l' -> - Forall2 (fun x y => p x y) l l'. -Proof. - intros A p l l' h. - induction l in l', h |- *. - - destruct l'. 2: discriminate. - constructor. - - destruct l'. 1: discriminate. - simpl in h. apply andP in h as [? ?]. - constructor. all: auto. -Qed. - Lemma eqb_term_upto_univ_impl (equ lequ : _ -> _ -> bool) Re Rle : RelationClasses.subrelation equ Re -> RelationClasses.subrelation lequ Rle -> @@ -659,24 +618,24 @@ Proof. eapply All2_impl'; tea. eapply All_impl; tea. eauto. - constructor; eauto. - - intro. toProp. constructor; eauto. - - intro. toProp. constructor; eauto. - - intro. toProp. constructor; eauto. - - intro. toProp. constructor; eauto. + - intro. rtoProp. constructor; eauto. + - intro. rtoProp. constructor; eauto. + - intro. rtoProp. constructor; eauto. + - intro. rtoProp. constructor; eauto. - unfold kername in *. eqspec; [|discriminate]. - intro. toProp. constructor; eauto. + intro. rtoProp. constructor; eauto. apply forallb2_Forall2 in H0. eapply Forall2_impl; tea; eauto. - unfold kername in *. eqspec; [|discriminate]. - intro. toProp. constructor; eauto. + intro. rtoProp. constructor; eauto. apply forallb2_Forall2 in H0. eapply Forall2_impl; tea; eauto. - unfold kername in *. eqspec; [|discriminate]. eqspec; [|discriminate]. - intro. toProp. constructor; eauto. + intro. rtoProp. constructor; eauto. apply forallb2_Forall2 in H0. eapply Forall2_impl; tea; eauto. - - eqspec; [|discriminate]. intro. toProp. + - eqspec; [|discriminate]. intro. rtoProp. destruct indn. econstructor; eauto. apply forallb2_All2 in H0. eapply All2_impl'; tea. @@ -690,18 +649,18 @@ Proof. eapply All2_impl'; tea. red in X. eapply All_impl; tea. cbn -[eqb]. intros x X0 y. eqspec; [|rewrite andb_false_r; discriminate]. - intro. toProp. split; tas. split; eapply X0; tea. + intro. rtoProp. split; tas. split; eapply X0; tea. - eqspec; [|discriminate]. econstructor; eauto. cbn -[eqb] in H; apply forallb2_All2 in H. eapply All2_impl'; tea. red in X. eapply All_impl; tea. cbn -[eqb]. intros x X0 y. eqspec; [|rewrite andb_false_r; discriminate]. - intro. toProp. split; tas. split; eapply X0; tea. + intro. rtoProp. split; tas. split; eapply X0; tea. Qed. -Lemma reflect_eq_term_upto_univ equ lequ (Re Rle : universe -> universe -> Prop) : +Lemma reflect_eq_term_upto_univ equ lequ (Re Rle : Universe.t -> Universe.t -> Prop) : (forall u u', reflectT (Re u u') (equ u u')) -> (forall u u', reflectT (Rle u u') (lequ u u')) -> forall t t', reflectT (eq_term_upto_univ Re Rle t t') @@ -723,20 +682,18 @@ Proof. induction X in l0 |- *. + destruct l0. * constructor. constructor. constructor. - * constructor. intro bot. inversion bot. subst. - inversion H0. + * constructor. intro bot. inversion bot. inversion X. + destruct l0. - * constructor. intro bot. inversion bot. subst. - inversion H0. + * constructor. intro bot. inversion bot. subst. inversion X0. * cbn. destruct (p _ _ he t). -- destruct (IHX l0). ++ constructor. constructor. constructor ; try assumption. inversion e0. subst. assumption. ++ constructor. intro bot. inversion bot. subst. - inversion H0. subst. + inversion X0. subst. apply f. constructor. assumption. -- constructor. intro bot. apply f. - inversion bot. subst. inversion H0. subst. assumption. + inversion bot. subst. inversion X0. subst. assumption. - cbn - [eqb]. eqspecs. equspec equ he. equspec lequ hle. ih. constructor. constructor. assumption. - cbn - [eqb]. eqspecs. equspec equ he. equspec lequ hle. ih. @@ -817,9 +774,9 @@ Proof. + destruct brs. * constructor. constructor ; try assumption. constructor. - * constructor. intro bot. inversion bot. subst. inversion H8. + * constructor. intro bot. inversion bot. subst. inversion X2. + destruct brs. - * constructor. intro bot. inversion bot. subst. inversion H8. + * constructor. intro bot. inversion bot. subst. inversion X2. * cbn - [eqb]. inversion X. subst. destruct a, p. cbn - [eqb]. eqspecs. -- cbn - [eqb]. pose proof (X0 equ Re he t0) as hh. cbn in hh. @@ -831,11 +788,11 @@ Proof. inversion e2. subst. assumption. ** constructor. intro bot. apply f. inversion bot. subst. constructor ; try assumption. - inversion H8. subst. assumption. + inversion X4. subst. assumption. ++ constructor. intro bot. apply f. inversion bot. subst. - inversion H8. subst. destruct H3. assumption. + inversion X4. subst. destruct X5. assumption. -- constructor. intro bot. inversion bot. subst. - inversion H8. subst. destruct H3. cbn in e1. subst. + inversion X4. subst. destruct X5. cbn in e1. subst. apply n2. reflexivity. - cbn - [eqb]. eqspecs. equspec equ he. equspec lequ hle. ih. constructor. constructor ; assumption. @@ -843,9 +800,9 @@ Proof. cbn - [eqb]. induction m in X, mfix |- *. + destruct mfix. * constructor. constructor. constructor. - * constructor. intro bot. inversion bot. subst. inversion H0. + * constructor. intro bot. inversion bot. subst. inversion X0. + destruct mfix. - * constructor. intro bot. inversion bot. subst. inversion H0. + * constructor. intro bot. inversion bot. subst. inversion X0. * cbn - [eqb]. inversion X. subst. destruct X0 as [h1 h2]. destruct (h1 equ Re he (dtype d)). @@ -856,22 +813,22 @@ Proof. inversion e2. assumption. --- constructor. intro bot. apply f. inversion bot. subst. constructor. - inversion H0. subst. assumption. + inversion X0. subst. assumption. ** constructor. intro bot. inversion bot. subst. - apply n. inversion H0. subst. destruct H3 as [[? ?] ?]. + apply n. inversion X0. subst. destruct X2 as [[? ?] ?]. assumption. ++ constructor. intro bot. apply f. - inversion bot. subst. inversion H0. subst. - apply H3. + inversion bot. subst. inversion X0. subst. + apply X2. -- constructor. intro bot. apply f. - inversion bot. subst. inversion H0. subst. apply H3. + inversion bot. subst. inversion X0. subst. apply X2. - cbn - [eqb]. eqspecs. equspec equ he. equspec lequ hle. ih. cbn - [eqb]. induction m in X, mfix |- *. + destruct mfix. * constructor. constructor. constructor. - * constructor. intro bot. inversion bot. subst. inversion H0. + * constructor. intro bot. inversion bot. subst. inversion X0. + destruct mfix. - * constructor. intro bot. inversion bot. subst. inversion H0. + * constructor. intro bot. inversion bot. subst. inversion X0. * cbn - [eqb]. inversion X. subst. destruct X0 as [h1 h2]. destruct (h1 equ Re he (dtype d)). @@ -882,19 +839,19 @@ Proof. inversion e2. assumption. --- constructor. intro bot. apply f. inversion bot. subst. constructor. - inversion H0. subst. assumption. + inversion X0. subst. assumption. ** constructor. intro bot. inversion bot. subst. - apply n. inversion H0. subst. destruct H3 as [[? ?] ?]. + apply n. inversion X0. subst. destruct X2 as [[? ?] ?]. assumption. ++ constructor. intro bot. apply f. - inversion bot. subst. inversion H0. subst. - apply H3. + inversion bot. subst. inversion X0. subst. + apply X2. -- constructor. intro bot. apply f. - inversion bot. subst. inversion H0. subst. apply H3. + inversion bot. subst. inversion X0. subst. apply X2. Qed. Lemma eqb_term_upto_univ_refl : - forall (eqb leqb : universe -> universe -> bool) t, + forall (eqb leqb : Universe.t -> Universe.t -> bool) t, (forall u, eqb u u) -> (forall u, leqb u u) -> eqb_term_upto_univ eqb leqb t t. @@ -911,7 +868,7 @@ Proof. end. all: simpl. all: try solve [ eauto ]. - - induction H. + - induction X. + reflexivity. + simpl. rewrite -> p by assumption. auto. - eapply forallb2_map. eapply forallb2_refl. @@ -1071,15 +1028,15 @@ Proof. Qed. (* todo: rename *) -Definition nleq_term t t' := - eqb_term_upto_univ eqb eqb t t'. +(* Definition nleq_term t t' := *) +(* eqb_term_upto_univ eqb eqb t t'. *) -Corollary reflect_upto_names : - forall t t', reflectT (upto_names t t') (nleq_term t t'). -Proof. - intros t t'. eapply reflect_eq_term_upto_univ. - all: intros u u'; eapply reflect_reflectT, eqb_spec. -Qed. +(* Corollary reflect_upto_names : *) +(* forall t t', reflectT (upto_names t t') (nleq_term t t'). *) +(* Proof. *) +(* intros t t'. eapply reflect_eq_term_upto_univ. *) +(* all: intros u u'; eapply reflect_reflectT, eqb_spec. *) +(* Qed. *) Lemma upto_names_impl Re Rle : RelationClasses.Reflexive Re -> @@ -1115,7 +1072,7 @@ Qed. (** ** Equality on contexts ** *) -Inductive eq_context_upto (Re : universe -> universe -> Prop) : context -> context -> Type := +Inductive eq_context_upto (Re : Universe.t -> Universe.t -> Prop) : context -> context -> Type := | eq_context_nil : eq_context_upto Re [] [] | eq_context_vass na A Γ nb B Δ : eq_term_upto_univ Re Re A B -> @@ -1241,7 +1198,7 @@ Qed. Section ContextUpTo. - Context (Re : universe -> universe -> Prop). + Context (Re : Universe.t -> Universe.t -> Prop). Context (ReR : RelationClasses.Reflexive Re). Context (ReS : RelationClasses.Symmetric Re). Context (ReT : RelationClasses.Transitive Re). @@ -1297,148 +1254,13 @@ Lemma lift_eq_context `{checker_flags} φ l l' n k : Proof. induction l in l', n, k |- *; intros; destruct l'; rewrite -> ?lift_context_snoc0. constructor. - all: inversion H0; subst. constructor. - - apply All2_length in H6. rewrite H6. + all: inversion X; subst. constructor. + - apply All2_length in X1. rewrite X1. now apply lift_eq_decl. - now apply IHl. Qed. - -Lemma subst_instance_level_val u l v v' - (H1 : forall s, valuation_mono v s = valuation_mono v' s) - (H2 : forall n, val0 v (nth n u Level.lSet) = Z.of_nat (valuation_poly v' n)) - : val0 v (subst_instance_level u l) = val0 v' l. -Proof. - destruct l; cbn; try congruence. -Qed. - -Lemma eq_val v v' - (H1 : forall s, valuation_mono v s = valuation_mono v' s) - (H2 : forall n, valuation_poly v n = valuation_poly v' n) - : forall u, val v u = val v' u. -Proof. - assert (He : forall e, val1 v e = val1 v' e). { - intros [? []]; unfold val1; cbn. - all: destruct t; cbn; rewrite ?H1 ?H2; reflexivity. } - destruct u; cbn; rewrite He; auto. - generalize (val1 v' t). induction u; cbn; intro; now rewrite He. -Qed. - -Lemma is_prop_subst_instance_level u l - (Hu : forallb (negb ∘ Level.is_prop) u) - : Level.is_prop (subst_instance_level u l) = Level.is_prop l. -Proof. - destruct l; cbn; try reflexivity. - destruct (le_lt_dec #|u| n) as [HH|HH]. - + now rewrite nth_overflow. - + eapply (forallb_nth _ _ _ Level.lSet Hu) in HH. - destruct HH as [l [HH1 HH2]]. rewrite HH1. now apply ssrbool.negbTE. -Qed. - - -Lemma subst_instance_univ_val u l v v' - (Hu : forallb (negb ∘ Level.is_prop) u) - (H1 : forall s, valuation_mono v s = valuation_mono v' s) - (H2 : forall n, val0 v (nth n u Level.lSet) = Z.of_nat (valuation_poly v' n)) - : val v (subst_instance_univ u l) = val v' l. -Proof. - assert (He: forall e, val1 v (subst_instance_level_expr u e) = val1 v' e). { - clear l. intros [l []]; unfold val1; simpl. - - erewrite subst_instance_level_val; tea. - now rewrite is_prop_subst_instance_level. - - now apply subst_instance_level_val. } - destruct l; simpl. - - apply He. - - rewrite He. generalize (val1 v' t). induction l; simpl. - now rewrite He. - intro. rewrite !He. now apply IHl. -Qed. - - -Definition subst_instance_valuation (u : universe_instance) (v : valuation) := - {| valuation_mono := valuation_mono v ; - valuation_poly := fun i => Z.to_nat (val0 v (nth i u Level.lSet)) |}. - - -Lemma subst_instance_univ_val' u l v - (Hu : forallb (negb ∘ Level.is_prop) u) - : val v (subst_instance_univ u l) = val (subst_instance_valuation u v) l. -Proof. - eapply subst_instance_univ_val; auto. - cbn. intro; rewrite Z2Nat.id; auto. - destruct (le_lt_dec #|u| n) as [HH|HH]. - + now rewrite nth_overflow. - + eapply (forallb_nth _ _ _ Level.lSet Hu) in HH. - destruct HH as [?l [HH1 HH2]]. rewrite HH1. - destruct l0; try discriminate; cbn. - apply Zle_0_nat. -Qed. - - -Class SubstUnivPreserving Re := Build_SubstUnivPreserving : - forall s u1 u2, R_universe_instance Re u1 u2 -> - Re (subst_instance_univ u1 s) (subst_instance_univ u2 s). - -Lemma subst_equal_inst_inst Re : - SubstUnivPreserving Re -> - forall u u1 u2, R_universe_instance Re u1 u2 -> - R_universe_instance Re (subst_instance_instance u1 u) - (subst_instance_instance u2 u). -Proof. - intros hRe u. induction u; cbnr. constructor. - intros u1 u2; unfold R_universe_instance; cbn; constructor. - - now apply (hRe (Universe.make a) u1 u2). - - exact (IHu u1 u2 H). -Qed. - -Lemma eq_term_upto_univ_subst_instance_constr Re : - RelationClasses.Reflexive Re -> - SubstUnivPreserving Re -> - forall t u1 u2, - R_universe_instance Re u1 u2 -> - eq_term_upto_univ Re Re (subst_instance_constr u1 t) - (subst_instance_constr u2 t). -Proof. - intros ref hRe t. - induction t using term_forall_list_ind; intros u1 u2 hu. - all: cbn; try constructor; eauto using subst_equal_inst_inst. - all: eapply All2_map, All_All2; tea; cbn; intros; rdest; eauto. -Qed. - -Instance leq_term_SubstUnivPreserving {cf:checker_flags} φ : - SubstUnivPreserving (eq_universe φ). -Proof. - intros s u1 u2 hu. - unfold eq_universe in *; destruct check_univs; [|trivial]. - intros v Hv; cbn. - assert (Hl: forall l, (val0 v (subst_instance_level u1 l) - = val0 v (subst_instance_level u2 l))%Z). { - destruct l; cbnr. - apply Forall2_map_inv in hu. - induction n in u1, u2, hu |- *; cbnr. - - destruct hu; cbnr. now apply H. - - destruct hu; cbnr. now apply IHn. } - assert (He : forall e, (val1 v (subst_instance_level_expr u1 e) - = val1 v (subst_instance_level_expr u2 e))%Z). { - destruct e as [l []]; specialize (Hl l); unfold val1; cbnr; tas. - replace (Level.is_prop (subst_instance_level u2 l)) - with (Level.is_prop (subst_instance_level u1 l)). - + destruct ?; lia. - + destruct l; cbnr. clear Hl. - apply Forall2_map_inv in hu. - induction n in u1, u2, hu |- *. - * destruct hu; cbnr. - specialize (H _ Hv). - destruct x, y; cbn in *; try reflexivity; lia. - * destruct hu; cbnr. eauto. } - induction s. eapply He. - cbn -[val]. rewrite !val_cons. - specialize (He a). unfold subst_instance_univ in IHs. lia. -Qed. - - - Lemma eq_term_upto_univ_mkApps_inv Re u l u' l' : isApp u = false -> isApp u' = false -> @@ -1524,7 +1346,7 @@ Proof. intros A R x y h. assumption. Qed. -Lemma eq_term_upto_univ_flip (Re Rle Rle' : universe -> universe -> Prop) u v : +Lemma eq_term_upto_univ_flip (Re Rle Rle' : Universe.t -> Universe.t -> Prop) u v : RelationClasses.Reflexive Re -> RelationClasses.Reflexive Rle -> RelationClasses.Symmetric Re -> @@ -1558,17 +1380,13 @@ Proof. now eapply eq_term_upto_univ_sym. Qed. + Lemma eq_univ_make : forall u u', Forall2 eq (map Universe.make u) (map Universe.make u') -> u = u'. Proof. - intros u u' h. - revert u' h. - induction u ; intros u' h. - - destruct u' ; inversion h. reflexivity. - - destruct u' ; inversion h. subst. - f_equal. - + inversion H2. reflexivity. - + eapply IHu. assumption. -Qed. \ No newline at end of file + intros u u' H. eapply Forall2_map' in H. + eapply Forall2_eq. eapply Forall2_impl; tea. + clear. intros [] [] H; now inversion H. +Qed. diff --git a/pcuic/theories/PCUICLiftSubst.v b/pcuic/theories/PCUICLiftSubst.v index 694a20fc5..6b748a0f4 100644 --- a/pcuic/theories/PCUICLiftSubst.v +++ b/pcuic/theories/PCUICLiftSubst.v @@ -221,7 +221,7 @@ Qed. Hint Resolve -> on_snd_eq_id_spec : all. Hint Resolve -> on_snd_eq_spec : all. -Lemma map_def_eq_spec {A B : Set} (f f' g g' : A -> B) (x : def A) : +Lemma map_def_eq_spec {A B} (f f' g g' : A -> B) (x : def A) : f (dtype x) = g (dtype x) -> f' (dbody x) = g' (dbody x) -> map_def f f' x = map_def g g' x. @@ -230,7 +230,7 @@ Proof. Qed. Hint Resolve map_def_eq_spec : all. -Lemma map_def_id_spec {A : Set} (f f' : A -> A) (x : def A) : +Lemma map_def_id_spec {A} (f f' : A -> A) (x : def A) : f (dtype x) = (dtype x) -> f' (dbody x) = (dbody x) -> map_def f f' x = x. @@ -494,7 +494,7 @@ Proof. elim t using term_forall_list_ind; intros; try easy; rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; unfold test_def in *; - simpl closed in *; try solve [simpl lift; simpl closed; f_equal; auto; toProp; solve_all]; try easy. + simpl closed in *; try solve [simpl lift; simpl closed; f_equal; auto; rtoProp; solve_all]; try easy. - rewrite lift_rel_lt; auto. revert H. elim (Nat.ltb_spec n0 k); intros; try easy. - simpl lift. f_equal. solve_all. unfold test_def in b. toProp. solve_all. @@ -507,7 +507,7 @@ Proof. elim t using term_forall_list_ind; intros; try lia; rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; simpl closed in *; unfold test_snd, test_def in *; - try solve [(try f_equal; simpl; repeat (toProp; solve_all); eauto)]. + try solve [(try f_equal; simpl; repeat (rtoProp; solve_all); eauto)]. - elim (ltb_spec n k'); auto. intros. apply ltb_lt in H. lia. diff --git a/pcuic/theories/PCUICNameless.v b/pcuic/theories/PCUICNameless.v index 41059a422..03d01706c 100644 --- a/pcuic/theories/PCUICNameless.v +++ b/pcuic/theories/PCUICNameless.v @@ -43,7 +43,7 @@ Fixpoint nameless (t : term) : bool := forallb (test_def nameless nameless) mfix end. -Definition map_def_anon {A B : Set} (tyf bodyf : A -> B) (d : def A) := {| +Definition map_def_anon {A B} (tyf bodyf : A -> B) (d : def A) := {| dname := nAnon ; dtype := tyf d.(dtype) ; dbody := bodyf d.(dbody) ; @@ -111,10 +111,10 @@ Proof. reflexivity. + destruct args' ; try solve [ inversion h ]. inversion h. subst. - inversion H. subst. + inversion X. subst. cbn in hu, hv. destruct_andb. f_equal. - * eapply H2 ; assumption. + * eapply H0 ; assumption. * eapply IHl ; assumption. - f_equal ; try solve [ ih ]. eapply eq_univ_make. assumption. @@ -130,8 +130,8 @@ Proof. cbn in h1, h2. destruct_andb. inversion X. subst. f_equal. - * destruct a, p0. cbn in *. destruct H6. subst. - f_equal. eapply H11 ; assumption. + * destruct a, p0. cbn in *. destruct X0. subst. + f_equal. eapply H8; assumption. * eapply IHl ; assumption. - f_equal ; try solve [ ih ]. revert mfix' H1 H2 H H0 a. @@ -141,9 +141,9 @@ Proof. inversion X. subst. cbn in h1, h2, h3, h4. destruct_andb. f_equal. - * destruct a, d. cbn in *. destruct H2 as [[? ?] ?]. - destruct H1 as [Hty Hbod]. - unfold test_def in H7, H. cbn in H7, H. + * destruct a, d. cbn in *. destruct X0 as [[? ?] ?]. + destruct H0 as [Hty Hbod]. + unfold test_def in H4, H. cbn in H4, H. destruct_andb. anonify. f_equal. -- eapply Hty; assumption. @@ -158,9 +158,9 @@ Proof. inversion X. subst. cbn in h1, h2, h3, h4. destruct_andb. f_equal. - * destruct a, d. cbn in *. destruct H2 as [[? ?] ?]. - destruct H1 as [Hty Hbod]. - unfold test_def in H7, H. cbn in H7, H. + * destruct a, d. cbn in *. destruct X0 as [[? ?] ?]. + destruct H0 as [Hty Hbod]. + unfold test_def in H4, H. cbn in H4, H. destruct_andb. anonify. f_equal. -- eapply Hty; assumption. @@ -188,7 +188,7 @@ Proof. * cbn. eapply IHm. inversion X. subst. assumption. + induction m. * reflexivity. - * cbn. inversion X. subst. destruct H1. + * cbn. inversion X. subst. destruct H0. repeat (eapply andb_true_intro ; split). all: try assumption. eapply IHm. assumption. @@ -198,7 +198,7 @@ Proof. * cbn. eapply IHm. inversion X. subst. assumption. + induction m. * reflexivity. - * cbn. inversion X. subst. destruct H1. + * cbn. inversion X. subst. destruct H0. repeat (eapply andb_true_intro ; split). all: try assumption. eapply IHm. assumption. @@ -519,7 +519,7 @@ Proof. intros u b. induction b using term_forall_list_ind. all: try (simpl ; rewrite ?IHb, ?IHb1, ?IHb2, ?IHb3 ; reflexivity). - - simpl. f_equal. induction H. + - simpl. f_equal. rename X into H; induction H. + reflexivity. + simpl. rewrite p, IHAll. reflexivity. - simpl. rewrite IHb1, IHb2. f_equal. @@ -678,7 +678,7 @@ Proof. all: simpl. all: try congruence. - destruct (_ <=? _). all: reflexivity. - - f_equal. induction H. + - f_equal. rename X into H; induction H. + reflexivity. + simpl. f_equal. * eapply p. @@ -758,7 +758,7 @@ Proof. rewrite nth_error_map. destruct (nth_error _ _). + simpl. apply nl_lift. + rewrite map_length. reflexivity. - - f_equal. induction H. + - f_equal. rename X into H; induction H. + reflexivity. + simpl. f_equal. * eapply p. @@ -1261,50 +1261,50 @@ Proof. + now apply nl_cumul. Qed. -Corollary reflect_nleq_term : - forall t t', - reflect (nl t = nl t') (nleq_term t t'). -Proof. - intros t t'. - destruct (reflect_upto_names t t'). - - constructor. eapply eq_term_nl_eq. assumption. - - constructor. intro bot. apply f. - apply eq_term_upto_univ_nl_inv. - rewrite bot. - apply eq_term_upto_univ_refl. - all: auto. -Qed. +(* Corollary reflect_nleq_term : *) +(* forall t t', *) +(* reflect (nl t = nl t') (nleq_term t t'). *) +(* Proof. *) +(* intros t t'. *) +(* destruct (reflect_upto_names t t'). *) +(* - constructor. eapply eq_term_nl_eq. assumption. *) +(* - constructor. intro bot. apply f. *) +(* apply eq_term_upto_univ_nl_inv. *) +(* rewrite bot. *) +(* apply eq_term_upto_univ_refl. *) +(* all: auto. *) +(* Qed. *) -Lemma nleq_term_it_mkLambda_or_LetIn : - forall Γ u v, - nleq_term u v -> - nleq_term (it_mkLambda_or_LetIn Γ u) (it_mkLambda_or_LetIn Γ v). -Proof. - intros Γ u v h. - induction Γ as [| [na [b|] A] Γ ih ] in u, v, h |- *. - - assumption. - - simpl. cbn. apply ih. - eapply ssrbool.introT. - + eapply reflect_nleq_term. - + cbn. f_equal. - eapply ssrbool.elimT. - * eapply reflect_nleq_term. - * assumption. - - simpl. cbn. apply ih. - eapply ssrbool.introT. - + eapply reflect_nleq_term. - + cbn. f_equal. - eapply ssrbool.elimT. - * eapply reflect_nleq_term. - * assumption. -Qed. +(* Lemma nleq_term_it_mkLambda_or_LetIn : *) +(* forall Γ u v, *) +(* nleq_term u v -> *) +(* nleq_term (it_mkLambda_or_LetIn Γ u) (it_mkLambda_or_LetIn Γ v). *) +(* Proof. *) +(* intros Γ u v h. *) +(* induction Γ as [| [na [b|] A] Γ ih ] in u, v, h |- *. *) +(* - assumption. *) +(* - simpl. cbn. apply ih. *) +(* eapply ssrbool.introT. *) +(* + eapply reflect_nleq_term. *) +(* + cbn. f_equal. *) +(* eapply ssrbool.elimT. *) +(* * eapply reflect_nleq_term. *) +(* * assumption. *) +(* - simpl. cbn. apply ih. *) +(* eapply ssrbool.introT. *) +(* + eapply reflect_nleq_term. *) +(* + cbn. f_equal. *) +(* eapply ssrbool.elimT. *) +(* * eapply reflect_nleq_term. *) +(* * assumption. *) +(* Qed. *) Lemma nl_two M : nl (nl M) = nl M. Proof. induction M using term_forall_list_ind; cbnr. all: rewrite ?IHM1, ?IHM2, ?IHM3, ?IHM; cbnr. - - f_equal. induction H; cbnr. congruence. + - f_equal. induction X; cbnr. congruence. - f_equal. induction X; cbnr. f_equal; tas. destruct x; unfold on_snd; simpl in *. congruence. - f_equal. induction X; cbnr. f_equal; tas. @@ -1515,3 +1515,30 @@ Proof. (* + cbn. congruence. *) (* Qed. *) Admitted. + + + (* Lemma nleq_term_zipc : *) + (* forall u v π, *) + (* nleq_term u v -> *) + (* nleq_term (zipc u π) (zipc v π). *) + (* Proof. *) + (* intros u v π h. *) + (* eapply ssrbool.introT. *) + (* - eapply reflect_nleq_term. *) + (* - cbn. rewrite 2!nl_zipc. f_equal. *) + (* eapply ssrbool.elimT. *) + (* + eapply reflect_nleq_term. *) + (* + assumption. *) + (* Qed. *) + + (* Lemma nleq_term_zipx : *) + (* forall Γ u v π, *) + (* nleq_term u v -> *) + (* nleq_term (zipx Γ u π) (zipx Γ v π). *) + (* Proof. *) + (* intros Γ u v π h. *) + (* unfold zipx. *) + (* eapply nleq_term_it_mkLambda_or_LetIn. *) + (* eapply nleq_term_zipc. *) + (* assumption. *) + (* Qed. *) diff --git a/pcuic/theories/PCUICParallelReduction.v b/pcuic/theories/PCUICParallelReduction.v index fe2b149db..67a69af2e 100644 --- a/pcuic/theories/PCUICParallelReduction.v +++ b/pcuic/theories/PCUICParallelReduction.v @@ -592,7 +592,7 @@ Section ParallelReduction. (* All2_local_env (on_decl pred1) Γ Γ' -> *) (* All2_local_env (on_decl P) Γ Γ' -> *) (* P Γ Γ' (tRel i) (tRel i)) -> *) - (* (forall (Γ Γ' : context) (ind : inductive) (pars c : nat) (u : universe_instance) (args0 args1 : list term) *) + (* (forall (Γ Γ' : context) (ind : inductive) (pars c : nat) (u : Instance.t) (args0 args1 : list term) *) (* (p : term) (brs0 brs1 : list (nat * term)), *) (* All2_local_env (on_decl pred1) Γ Γ' -> *) (* All2_local_env (on_decl P) Γ Γ' -> *) @@ -624,13 +624,13 @@ Section ParallelReduction. (* All2_local_env (on_decl pred1) Γ Γ' -> *) (* All2_local_env (on_decl P) Γ Γ' -> *) (* declared_constant Σ c decl -> *) - (* forall u : universe_instance, cst_body decl = Some body -> *) + (* forall u : Instance.t, cst_body decl = Some body -> *) (* P Γ Γ' (tConst c u) (subst_instance_constr u body)) -> *) - (* (forall (Γ Γ' : context) (c : ident) (u : universe_instance), *) + (* (forall (Γ Γ' : context) (c : ident) (u : Instance.t), *) (* All2_local_env (on_decl pred1) Γ Γ' -> *) (* All2_local_env (on_decl P) Γ Γ' -> *) (* P Γ Γ' (tConst c u) (tConst c u)) -> *) - (* (forall (Γ Γ' : context) (i : inductive) (pars narg : nat) (k : nat) (u : universe_instance) *) + (* (forall (Γ Γ' : context) (i : inductive) (pars narg : nat) (k : nat) (u : Instance.t) *) (* (args0 args1 : list term) (arg1 : term), *) (* All2_local_env (on_decl pred1) Γ Γ' -> *) (* All2_local_env (on_decl P) Γ Γ' -> *) @@ -770,7 +770,7 @@ Section ParallelReduction. All2_local_env (on_decl pred1) Γ Γ' -> Pctx Γ Γ' -> P Γ Γ' (tRel i) (tRel i)) -> - (forall (Γ Γ' : context) (ind : inductive) (pars c : nat) (u : universe_instance) (args0 args1 : list term) + (forall (Γ Γ' : context) (ind : inductive) (pars c : nat) (u : Instance.t) (args0 args1 : list term) (p : term) (brs0 brs1 : list (nat * term)), All2_local_env (on_decl pred1) Γ Γ' -> Pctx Γ Γ' -> @@ -818,13 +818,13 @@ Section ParallelReduction. All2_local_env (on_decl pred1) Γ Γ' -> Pctx Γ Γ' -> declared_constant Σ c decl -> - forall u : universe_instance, cst_body decl = Some body -> + forall u : Instance.t, cst_body decl = Some body -> P Γ Γ' (tConst c u) (subst_instance_constr u body)) -> - (forall (Γ Γ' : context) (c : ident) (u : universe_instance), + (forall (Γ Γ' : context) (c : ident) (u : Instance.t), All2_local_env (on_decl pred1) Γ Γ' -> Pctx Γ Γ' -> P Γ Γ' (tConst c u) (tConst c u)) -> - (forall (Γ Γ' : context) (i : inductive) (pars narg : nat) (k : nat) (u : universe_instance) + (forall (Γ Γ' : context) (i : inductive) (pars narg : nat) (k : nat) (u : Instance.t) (args0 args1 : list term) (arg1 : term), All2_local_env (on_decl pred1) Γ Γ' -> Pctx Γ Γ' -> diff --git a/pcuic/theories/PCUICParallelReductionConfluence.v b/pcuic/theories/PCUICParallelReductionConfluence.v index c242a806b..7ae5188ed 100644 --- a/pcuic/theories/PCUICParallelReductionConfluence.v +++ b/pcuic/theories/PCUICParallelReductionConfluence.v @@ -670,7 +670,7 @@ Section Confluence. | _ => True }. Transparent discr_construct_cofix. - Inductive construct_cofix_view : term -> Set := + Inductive construct_cofix_view : term -> Type := | construct_cofix_construct c u i : construct_cofix_view (tConstruct c u i) | construct_cofix_cofix mfix idx : construct_cofix_view (tCoFix mfix idx) | construct_cofix_other t : discr_construct_cofix t -> construct_cofix_view t. @@ -685,7 +685,7 @@ Section Confluence. isConstruct _ => false. Transparent isConstruct. - Inductive construct_view : term -> Set := + Inductive construct_view : term -> Type := | construct_construct c u i : construct_view (tConstruct c u i) | construct_other t : ~~ isConstruct t -> construct_view t. @@ -700,7 +700,7 @@ Section Confluence. | _ => false end. - Inductive fix_app_view : term -> Set := + Inductive fix_app_view : term -> Type := | fix_app_fix mfix i l : l <> [] -> fix_app_view (mkApps (tFix mfix i) l) | fix_app_other t : ~~ isFix_app t -> fix_app_view t. @@ -708,8 +708,8 @@ Section Confluence. Proof. induction t; try solve [apply fix_app_other; simpl; auto]. destruct IHt1. pose proof (fix_app_fix mfix i (l ++ [t2])). - forward H. intros eq. destruct l; discriminate. - now rewrite -mkApps_nested in H. + forward X. intros eq. destruct l; discriminate. + now rewrite -mkApps_nested in X. destruct t; try solve [apply fix_app_other; simpl; auto]. apply (fix_app_fix mfix idx [t2]). congruence. Qed. @@ -721,7 +721,7 @@ Section Confluence. | _ => false end. - Inductive fix_lambda_view : term -> Set := + Inductive fix_lambda_view : term -> Type := | fix_lambda_lambda na b t : fix_lambda_view (tLambda na b t) | fix_lambda_fix mfix i : fix_lambda_view (tFix mfix i) | fix_lambda_other t : ~~ isFixLambda t -> fix_lambda_view t. @@ -1248,9 +1248,9 @@ Section Confluence. rewrite rho_ctx_over_app. simpl. unfold app_context. unfold app_context in IHm. erewrite <- IHm. simpl. cbn. - depelim H. rewrite -e. + depelim X. rewrite -e. rewrite - app_assoc. - f_equal. now depelim H. + f_equal. now depelim X. Qed. Lemma fold_fix_context_over' Γ m : @@ -1838,7 +1838,7 @@ Section Confluence. - simpl. rewrite inst_closed0. rewrite closedn_subst_instance_constr; auto. eapply declared_decl_closed in H; auto. hnf in H. rewrite H0 in H. - toProp; auto. + rtoProp; auto. econstructor; eauto with pcuic. - (* Proj-Construct *) @@ -2036,7 +2036,7 @@ Section Confluence. | _ => True end. - Inductive fix_view : term -> Set := + Inductive fix_view : term -> Type := | fix_fix mfix i : fix_view (tFix mfix i) | fix_other t : discr_fix t -> fix_view t. @@ -2064,7 +2064,7 @@ Section Confluence. | _ => false end. - Inductive lambda_app_view : term -> Set := + Inductive lambda_app_view : term -> Type := | lambda_app_fix na t b a : lambda_app_view (tApp (tLambda na t b) a) | lambda_app_other t : ~~ lambda_app_discr t -> lambda_app_view t. @@ -2777,7 +2777,7 @@ Section Confluence. Proof. apply decompose_app_rec_rename. Qed. (* TODO rename isConstruct to isConstruct *) - Lemma nisConstruct_elim {A} {t} {a : inductive -> nat -> universe_instance -> A} {b : A} : + Lemma nisConstruct_elim {A} {t} {a : inductive -> nat -> Instance.t -> A} {b : A} : ~~ isConstruct t -> match t with | tConstruct ind n u => a ind n u @@ -2873,7 +2873,7 @@ Section Confluence. destruct (nth_error _ (rarg d)) eqn:Hisc. simpl option_map. cbv beta iota. assert (rho Δ (rename r t0) = rename r (rho Γ t0)) as ->. - eapply nth_error_all in H3; eauto. now simpl in H3. + eapply nth_error_all in X; eauto. now simpl in X. simpl. assert (Hbod: ∀ (Γ Δ : list context_decl) (r : nat → nat), renaming Γ Δ r → rename r (rho Γ (dbody d)) = rho Δ (rename r (dbody d))). @@ -2903,7 +2903,7 @@ Section Confluence. rewrite map_fix_subst // !map_map_compose. assert (0 < list_size size l). { destruct l; simpl; auto. congruence. lia. } - clear -H H2 H4. + clear -H H2 H3. unfold fix_subst. generalize #|mfix|. induction n; simpl; auto. rewrite IHn. f_equal. diff --git a/pcuic/theories/PCUICPosition.v b/pcuic/theories/PCUICPosition.v index 08e3f5068..aa2d52c62 100644 --- a/pcuic/theories/PCUICPosition.v +++ b/pcuic/theories/PCUICPosition.v @@ -679,7 +679,17 @@ Inductive stack : Type := Notation "'ε'" := (Empty). -Derive NoConfusion NoConfusionHom EqDec for stack. +Derive NoConfusion NoConfusionHom for stack. + +Instance EqDec_def {A} : EqDec A -> EqDec (def A). +Proof. + intros X x y. decide equality; apply eq_dec. +Defined. + +Instance EqDec_stack : EqDec stack. +Proof. + intros x y. decide equality; apply eq_dec. +Defined. Instance reflect_stack : ReflectEq stack := let h := EqDec_ReflectEq stack in _. @@ -1418,4 +1428,4 @@ Definition context_clos (R : term -> term -> Type) u v := Definition context_env_clos (R : context -> term -> term -> Type) Γ u v := ∑ u' v' π, R (Γ ,,, stack_context π) u' v' × - (u = zipc u' π /\ v = zipc v' π). \ No newline at end of file + (u = zipc u' π /\ v = zipc v' π). diff --git a/pcuic/theories/PCUICPrincipality.v b/pcuic/theories/PCUICPrincipality.v index 356758ae8..1e27b851f 100644 --- a/pcuic/theories/PCUICPrincipality.v +++ b/pcuic/theories/PCUICPrincipality.v @@ -605,9 +605,9 @@ Section Principality. - apply inversion_Sort in hA as iA. 2: auto. apply inversion_Sort in hB as iB. 2: auto. repeat outsum. repeat outtimes. subst. - inversion e. subst. - repeat insum. repeat intimes. - all: try eassumption. + assert (x0 = x) as ee. { + clear -e. destruct x, x0; cbnr; invs e; reflexivity. } + subst. repeat insum. repeat intimes; tea. (* * left; eexists _, _; intuition auto. *) constructor ; assumption. - apply inversion_Prod in hA as [dom1 [codom1 iA]]; auto. diff --git a/pcuic/theories/PCUICReduction.v b/pcuic/theories/PCUICReduction.v index 7f030bd7a..ab6f0bfe8 100644 --- a/pcuic/theories/PCUICReduction.v +++ b/pcuic/theories/PCUICReduction.v @@ -75,7 +75,7 @@ Instance red_Reflexive Σ Γ : Reflexive (red Σ Γ) Section ReductionCongruence. Context {Σ : global_env}. - Inductive term_context : Set := + Inductive term_context := | tCtxHole : term_context | tCtxEvar : nat -> list_context -> term_context | tCtxProd_l : name -> term_context (* the type *) -> term -> term_context @@ -100,19 +100,19 @@ Section ReductionCongruence. (* | tCtxFix : mfixpoint_context -> nat -> term_context harder because types of fixpoints are necessary *) (* | tCtxCoFix : mfixpoint_context -> nat -> term_context *) - with list_context : Set := + with list_context := | tCtxHead : term_context -> list term -> list_context | tCtxTail : term -> list_context -> list_context - with list_nat_context : Set := + with list_nat_context := | tCtxHead_nat : (nat * term_context) -> list (nat * term) -> list_nat_context | tCtxTail_nat : (nat * term) -> list_nat_context -> list_nat_context. - (* with mfixpoint_context : Set := *) + (* with mfixpoint_context := *) (* | tCtxHead_mfix : def_context -> list (def term) -> mfixpoint_context *) (* | tCtxTail_mfix : def term -> mfixpoint_context -> mfixpoint_context *) - (* with def_context : Set := *) + (* with def_context := *) (* | tCtxType : name -> term_context -> term -> nat -> def_context *) (* | tCtxDef : name -> term -> term_context -> nat -> def_context. *) @@ -1380,4 +1380,4 @@ Proof. rewrite app_context_nil_l in h. econstructor. eapply OnOne2_app. constructor. cbn. intuition auto. -Qed. \ No newline at end of file +Qed. diff --git a/pcuic/theories/PCUICReflect.v b/pcuic/theories/PCUICReflect.v index 3db9fceaf..6da6d5928 100644 --- a/pcuic/theories/PCUICReflect.v +++ b/pcuic/theories/PCUICReflect.v @@ -250,24 +250,34 @@ Next Obligation. cbn. constructor. subst. reflexivity. Defined. -Fixpoint eq_non_empty_list {A : Set} (eqA : A -> A -> bool) (l l' : non_empty_list A) : bool := - match l, l' with - | NEL.sing a, NEL.sing a' => eqA a a' - | NEL.cons a l, NEL.cons a' l' => - eqA a a' && eq_non_empty_list eqA l l' - | _, _ => false - end. +(* TODO: move *) +Lemma eq_universe_iff (u v : Universe.t) : + u = v <-> u = v :> UnivExprSet.t. +Proof. + destruct u, v; cbn; split. now inversion 1. + intros ->. f_equal. apply uip. +Qed. +Lemma eq_universe_iff' (u v : Universe.t) : + u = v <-> UnivExprSet.elements u = UnivExprSet.elements v. +Proof. + etransitivity. apply eq_universe_iff. + destruct u as [[u1 u2] ?], v as [[v1 v2] ?]; cbn; clear; split. + now inversion 1. intros ->. f_equal. apply uip. +Qed. -#[program] Instance reflect_non_empty_list : - forall {A : Set} `{ReflectEq A}, ReflectEq (non_empty_list A) := - { eqb := eq_non_empty_list eqb }. -Next Obligation. - intros A RA. induction x, y; cbn. - destruct (eqb_spec a a0); constructor; congruence. - constructor; congruence. - constructor; congruence. - destruct (eqb_spec a a0), (IHx y); constructor; congruence. -Defined. +(* move in Universes.v ?? *) +Instance eq_dec_UnivExpr : EqDec UnivExpr.t. +Proof. intros e e'. repeat decide equality. Qed. + +Instance eq_dec_univ : EqDec Universe.t. +Proof. + intros u v. + assert (H : {UnivExprSet.elements u = UnivExprSet.elements v} + + {~ UnivExprSet.elements u = UnivExprSet.elements v}). { + repeat decide equality. } + 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. +Qed. Local Ltac finish := let h := fresh "h" in @@ -284,10 +294,10 @@ Local Ltac fcase c := Local Ltac term_dec_tac term_dec := repeat match goal with | t : term, u : term |- _ => fcase (term_dec t u) - | u : universe, u' : universe |- _ => fcase (eq_dec u u') - | x : universe_instance, y : universe_instance |- _ => + | u : Universe.t, u' : Universe.t |- _ => fcase (eq_dec u u') + | x : Instance.t, y : Instance.t |- _ => fcase (eq_dec x y) - | x : list Level.t, y : universe_instance |- _ => + | x : list Level.t, y : Instance.t |- _ => fcase (eq_dec x y) | n : nat, m : nat |- _ => fcase (Nat.eq_dec n m) | i : ident, i' : ident |- _ => fcase (string_dec i i') @@ -302,12 +312,12 @@ Local Ltac term_dec_tac term_dec := Derive NoConfusion NoConfusionHom for term. -Derive EqDec for term. -Next Obligation. - induction x using term_forall_list_ind ; intro t ; +Instance EqDec_term : EqDec term. +Proof. + intro x; induction x using term_forall_list_ind ; intro t ; destruct t ; try (right ; discriminate). all: term_dec_tac term_dec. - - revert l0. induction H ; intro l0. + - revert l0. rename X into H; induction H ; intro l0. + destruct l0. * left. reflexivity. * right. discriminate. diff --git a/pcuic/theories/PCUICRetyping.v b/pcuic/theories/PCUICRetyping.v index b758dccbf..b43210dbc 100644 --- a/pcuic/theories/PCUICRetyping.v +++ b/pcuic/theories/PCUICRetyping.v @@ -20,7 +20,7 @@ Require Import Equations.Prop.DepElim. in particular. *) Axiom reduce_to_sort : - global_env -> context -> term -> typing_result universe. + global_env -> context -> term -> typing_result Universe.t. Axiom reduce_to_prod : global_env -> context -> term -> typing_result (term × term). @@ -106,7 +106,7 @@ Section TypeOf. end end. - Definition sort_of (Γ : context) (t : term) : typing_result universe := + Definition sort_of (Γ : context) (t : term) : typing_result Universe.t := ty <- type_of Γ t;; type_of_as_sort type_of Γ ty. @@ -255,4 +255,4 @@ Section TypeOf. -End TypeOf. \ No newline at end of file +End TypeOf. diff --git a/pcuic/theories/PCUICSN.v b/pcuic/theories/PCUICSN.v index 46b242ab6..7eb694fc2 100644 --- a/pcuic/theories/PCUICSN.v +++ b/pcuic/theories/PCUICSN.v @@ -9,7 +9,7 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICInduction PCUICReflect PCUICLiftSubst PCUICUnivSubst PCUICTyping PCUICPosition PCUICNormal PCUICInversion PCUICCumulativity PCUICSafeLemmata PCUICGeneration PCUICValidity PCUICSR PCUICAlpha PCUICNameless - PCUICEquality PCUICConfluence. + PCUICEquality PCUICConfluence PCUICUnivSubstitution. From Equations Require Import Equations. Require Import Equations.Prop.DepElim. @@ -323,7 +323,7 @@ Section Alpha. - intros ? ? ? []. auto. - intros ? ? ? r. apply Forall2_eq in r. apply map_inj in r. + subst. reflexivity. - + intros ? ? H. inversion H. reflexivity. + + apply Universe.make_inj. Qed. Lemma cored_cored' : diff --git a/pcuic/theories/PCUICSafeLemmata.v b/pcuic/theories/PCUICSafeLemmata.v index f9748a2cb..aa0b1a9c9 100644 --- a/pcuic/theories/PCUICSafeLemmata.v +++ b/pcuic/theories/PCUICSafeLemmata.v @@ -688,7 +688,7 @@ Section Lemmata. left. apply wf_local_app in h2. inversion h2. subst. cbn in *. match goal with - | h : ∃ s : universe, _ |- _ => + | h : ∃ s : Universe.t, _ |- _ => destruct h end. eexists. eassumption. @@ -1297,32 +1297,6 @@ Section Lemmata. inversion e. reflexivity. Qed. - Lemma nleq_term_zipc : - forall u v π, - nleq_term u v -> - nleq_term (zipc u π) (zipc v π). - Proof. - intros u v π h. - eapply ssrbool.introT. - - eapply reflect_nleq_term. - - cbn. rewrite 2!nl_zipc. f_equal. - eapply ssrbool.elimT. - + eapply reflect_nleq_term. - + assumption. - Qed. - - Lemma nleq_term_zipx : - forall Γ u v π, - nleq_term u v -> - nleq_term (zipx Γ u π) (zipx Γ v π). - Proof. - intros Γ u v π h. - unfold zipx. - eapply nleq_term_it_mkLambda_or_LetIn. - eapply nleq_term_zipc. - assumption. - Qed. - Hint Resolve conv_refl conv_alt_red : core. Hint Resolve conv_refl : core. diff --git a/pcuic/theories/PCUICSigmaCalculus.v b/pcuic/theories/PCUICSigmaCalculus.v index 9bfa1e6ae..4ac76891f 100644 --- a/pcuic/theories/PCUICSigmaCalculus.v +++ b/pcuic/theories/PCUICSigmaCalculus.v @@ -67,7 +67,7 @@ Proof. rewrite ?IHt ?IHt1 ?IHt2 ; easy ]. - - simpl. f_equal. induction H. + - simpl. f_equal. induction X. + reflexivity. + simpl. f_equal ; easy. - simpl. rewrite IHt1 IHt2. f_equal. diff --git a/pcuic/theories/PCUICSubstitution.v b/pcuic/theories/PCUICSubstitution.v index cfb7cfa64..3d10cbb4c 100644 --- a/pcuic/theories/PCUICSubstitution.v +++ b/pcuic/theories/PCUICSubstitution.v @@ -35,7 +35,7 @@ Inductive subs {cf:checker_flags} (Σ : global_env_ext) (Γ : context) : list te (** Linking a context (with let-ins), an instance (reversed substitution) for its assumptions and a well-formed substitution for it. *) -Inductive context_subst : context -> list term -> list term -> Set := +Inductive context_subst : context -> list term -> list term -> Type := | context_subst_nil : context_subst [] [] [] | context_subst_ass Γ args s na t a : context_subst Γ args s -> @@ -1098,10 +1098,10 @@ Proof. - intros [= <-]. intros [= ->]. simpl. auto. - intros. destruct decl as [na' [b|] ty]; cbn in *. 2: auto. - specialize (IHuntyped_subslet _ H0 H1). intuition auto. + specialize (IHX _ H H0). intuition auto. - intros [= <-]. intros [= <-]. simpl. split; auto. - - apply IHuntyped_subslet. + - apply IHX. Qed. (* Lemma red1_mkApp Σ Γ M1 N1 M2 : *) diff --git a/pcuic/theories/PCUICTyping.v b/pcuic/theories/PCUICTyping.v index c9345fb9b..2912c3150 100644 --- a/pcuic/theories/PCUICTyping.v +++ b/pcuic/theories/PCUICTyping.v @@ -271,7 +271,7 @@ Lemma red1_ind_all : (forall (Γ : context) (i : nat) (body : term), option_map decl_body (nth_error Γ i) = Some (Some body) -> P Γ (tRel i) ((lift0 (S i)) body)) -> - (forall (Γ : context) (ind : inductive) (pars c : nat) (u : universe_instance) (args : list term) + (forall (Γ : context) (ind : inductive) (pars c : nat) (u : Instance.t) (args : list term) (p : term) (brs : list (nat * term)), P Γ (tCase (ind, pars) p (mkApps (tConstruct ind c u) args) brs) (iota_red pars c args brs)) -> @@ -290,9 +290,9 @@ Lemma red1_ind_all : (forall (Γ : context) (c : ident) (decl : constant_body) (body : term), declared_constant Σ c decl -> - forall u : universe_instance, cst_body decl = Some body -> P Γ (tConst c u) (subst_instance_constr u body)) -> + forall u : Instance.t, cst_body decl = Some body -> P Γ (tConst c u) (subst_instance_constr u body)) -> - (forall (Γ : context) (i : inductive) (pars narg : nat) (args : list term) (k : nat) (u : universe_instance) + (forall (Γ : context) (i : inductive) (pars narg : nat) (args : list term) (k : nat) (u : Instance.t) (arg : term), nth_error args (pars + narg) = Some arg -> P Γ (tProj (i, pars, narg) (mkApps (tConstruct i k u) args)) arg) -> @@ -1007,7 +1007,7 @@ Lemma typing_ind_env `{cf : checker_flags} : LevelSet.In l (global_ext_levels Σ) -> P Σ Γ (tSort (Universe.make l)) (tSort (Universe.super l))) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : name) (t b : term) (s1 s2 : universe), + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : name) (t b : term) (s1 s2 : Universe.t), All_local_env_over typing Pdecl Σ Γ wfΓ -> Σ ;;; Γ |- t : tSort s1 -> P Σ Γ t (tSort s1) -> @@ -1015,14 +1015,14 @@ Lemma typing_ind_env `{cf : checker_flags} : P Σ (Γ,, vass n t) b (tSort s2) -> P Σ Γ (tProd n t b) (tSort (Universe.sort_of_product s1 s2))) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : name) (t b : term) - (s1 : universe) (bty : term), + (s1 : Universe.t) (bty : term), All_local_env_over typing Pdecl Σ Γ wfΓ -> Σ ;;; Γ |- t : tSort s1 -> P Σ Γ t (tSort s1) -> Σ ;;; Γ,, vass n t |- b : bty -> P Σ (Γ,, vass n t) b bty -> P Σ Γ (tLambda n t b) (tProd n t bty)) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : name) (b b_ty b' : term) - (s1 : universe) (b'_ty : term), + (s1 : Universe.t) (b'_ty : term), All_local_env_over typing Pdecl Σ Γ wfΓ -> Σ ;;; Γ |- b_ty : tSort s1 -> P Σ Γ b_ty (tSort s1) -> @@ -1646,4 +1646,4 @@ Section All_local_env. + cbn. eexists. eassumption. Qed. -End All_local_env. \ No newline at end of file +End All_local_env. diff --git a/pcuic/theories/PCUICUnivSubst.v b/pcuic/theories/PCUICUnivSubst.v index 2d554df75..7080c326a 100644 --- a/pcuic/theories/PCUICUnivSubst.v +++ b/pcuic/theories/PCUICUnivSubst.v @@ -1,7 +1,7 @@ (* Distributed under the terms of the MIT license. *) From Coq Require Import Bool String List Program BinPos Compare_dec Arith Lia. -From MetaCoq.Template Require Import utils UnivSubst. +From MetaCoq.Template Require Import utils. From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICInduction PCUICLiftSubst. Require Import String. Local Open Scope string_scope. @@ -10,11 +10,8 @@ Set Asymmetric Patterns. From Equations Require Import Equations. Require Import Equations.Prop.DepElim. - (** * Universe substitution - *WIP* - Substitution of universe levels for universe level variables, used to implement universe polymorphism. *) @@ -103,146 +100,44 @@ Proof. solve_all. now destruct H as [n [-> _]]. Qed. -Section Closedu. - (** Tests that the term is closed over [k] universe variables *) - Context (k : nat). - - Definition closedu_level (l : Level.t) := - match l with - | Level.Var n => n true - end. - - Definition closedu_level_expr (s : Universe.Expr.t) := - closedu_level (fst s). - - Definition closedu_universe (u : universe) := - forallb closedu_level_expr u. - - Definition closedu_instance (u : universe_instance) := - forallb closedu_level u. - - Fixpoint closedu (t : term) : bool := +(** Tests that the term is closed over [k] universe variables *) +Fixpoint closedu (k : nat) (t : term) : bool := match t with - | tSort univ => closedu_universe univ - | tInd _ u => closedu_instance u - | tConstruct _ _ u => closedu_instance u - | tConst _ u => closedu_instance u + | tSort univ => closedu_universe k univ + | tInd _ u => closedu_instance k u + | tConstruct _ _ u => closedu_instance k u + | tConst _ u => closedu_instance k u | tRel i => true - | tEvar ev args => forallb closedu args - | tLambda _ T M | tProd _ T M => closedu T && closedu M - | tApp u v => closedu u && closedu v - (* | tCast c kind t => closedu c && closedu t *) - | tLetIn na b t b' => closedu b && closedu t && closedu b' + | tEvar ev args => forallb (closedu k) args + | tLambda _ T M | tProd _ T M => closedu k T && closedu k M + | tApp u v => closedu k u && closedu k v + | tLetIn na b t b' => closedu k b && closedu k t && closedu k b' | tCase ind p c brs => - let brs' := forallb (test_snd (closedu)) brs in - closedu p && closedu c && brs' - | tProj p c => closedu c + let brs' := forallb (test_snd (closedu k)) brs in + closedu k p && closedu k c && brs' + | tProj p c => closedu k c | tFix mfix idx => - forallb (test_def closedu closedu ) mfix + forallb (test_def (closedu k) (closedu k)) mfix | tCoFix mfix idx => - forallb (test_def closedu closedu) mfix + forallb (test_def (closedu k) (closedu k)) mfix | x => true end. -End Closedu. - -Require Import ssreflect ssrbool. - -(** Universe-closed terms are unaffected by universe substitution. *) - -Section UniverseClosedSubst. - Lemma closedu_subst_instance_level u t : closedu_level 0 t -> subst_instance_level u t = t. - Proof. - destruct t => /=; auto. - move/Nat.ltb_spec0. intro H. inversion H. - Qed. - - Lemma closedu_subst_instance_level_expr u t : closedu_level_expr 0 t -> subst_instance_level_expr u t = t. - Proof. - destruct t as [l n]. - rewrite /closedu_level_expr /subst_instance_level_expr /=. - move/(closedu_subst_instance_level u) => //. congruence. - Qed. - - Lemma closedu_subst_instance_univ u t : closedu_universe 0 t -> subst_instance_univ u t = t. - Proof. - rewrite /closedu_universe /subst_instance_univ => H. - pose proof (proj1 (forallb_forall _ t) H) as HH; clear H. - induction t; cbn; f_equal. - 1-2: now apply closedu_subst_instance_level_expr, HH; cbn. - apply IHt. intros x Hx; apply HH. now right. - Qed. - Hint Resolve closedu_subst_instance_level_expr closedu_subst_instance_level closedu_subst_instance_univ : terms. - - Lemma closedu_subst_instance_instance u t : closedu_instance 0 t -> subst_instance_instance u t = t. - Proof. - rewrite /closedu_instance /subst_instance_instance => H. solve_all. - now apply (closedu_subst_instance_level u). - Qed. - Hint Resolve closedu_subst_instance_instance : terms. - - Lemma closedu_subst_instance_constr u t : closedu 0 t -> subst_instance_constr u t = t. - Proof. - induction t in |- * using term_forall_list_ind; simpl; auto; intros H'; - rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; - try f_equal; eauto with terms; unfold test_def in *; - try solve [f_equal; eauto; repeat (toProp; solve_all)]. - Qed. -End UniverseClosedSubst. - -Section SubstInstanceClosed. - (** Substitution of a universe-closed instance of the right size - produces a universe-closed term. *) - - Context (u : universe_instance) (Hcl : closedu_instance 0 u). - - Lemma subst_instance_level_closedu t : - closedu_level #|u| t -> closedu_level 0 (subst_instance_level u t). - Proof. - destruct t => /=; auto. - move/Nat.ltb_spec0. intro H. - red in Hcl. unfold closedu_instance in Hcl. - eapply forallb_nth in Hcl; eauto. destruct Hcl as [x [Hn Hx]]. now rewrite -> Hn. - Qed. - - Lemma subst_instance_level_expr_closedu t : - closedu_level_expr #|u| t - -> closedu_level_expr 0 (subst_instance_level_expr u t). - Proof. - destruct t as [l n]. - rewrite /closedu_level_expr /subst_instance_level_expr /=. - move/(subst_instance_level_closedu) => //. - Qed. - - Lemma subst_instance_univ_closedu t : - closedu_universe #|u| t -> closedu_universe 0 (subst_instance_univ u t). - Proof. - rewrite /closedu_universe /subst_instance_univ => H. - eapply (forallb_Forall (closedu_level_expr #|u|)) in H; auto. - unfold universe_coercion; rewrite NEL.map_to_list forallb_map. - eapply Forall_forallb; eauto. - now move=> x /(subst_instance_level_expr_closedu). - Qed. - Hint Resolve subst_instance_level_expr_closedu subst_instance_level_closedu subst_instance_univ_closedu : terms. - - Lemma subst_instance_instance_closedu t : - closedu_instance #|u| t -> closedu_instance 0 (subst_instance_instance u t). - Proof. - rewrite /closedu_instance /subst_instance_instance => H. - eapply (forallb_Forall (closedu_level #|u|)) in H; auto. - rewrite forallb_map. eapply Forall_forallb; eauto. - simpl. now move=> x /(subst_instance_level_closedu). - Qed. - Hint Resolve subst_instance_instance_closedu : terms. - - Lemma subst_instance_constr_closedu t : - closedu #|u| t -> closedu 0 (subst_instance_constr u t). - Proof. - induction t in |- * using term_forall_list_ind; simpl; auto; intros H'; +Lemma closedu_subst_instance_constr u t + : closedu 0 t -> subst_instance_constr u t = t. +Proof. + induction t in |- * using term_forall_list_ind; simpl; auto; intros H'; + rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; + try f_equal; eauto with substu; unfold test_def in *; + try solve [f_equal; eauto; repeat (rtoProp; solve_all)]. +Qed. + +Lemma subst_instance_constr_closedu (u : Instance.t) (Hcl : closedu_instance 0 u) t : + closedu #|u| t -> closedu 0 (subst_instance_constr u t). +Proof. + induction t in |- * using term_forall_list_ind; simpl; auto; intros H'; rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length, ?forallb_map; - try f_equal; auto with terms; - unfold test_def, map_def, compose in *; - try solve [f_equal; eauto; repeat (toProp; solve_all); intuition auto with terms]. - Qed. -End SubstInstanceClosed. + try f_equal; auto with substu; + unfold test_def, map_def, compose in *; + try solve [f_equal; eauto; repeat (rtoProp; solve_all); intuition auto with substu]. +Qed. diff --git a/pcuic/theories/PCUICUnivSubstitution.v b/pcuic/theories/PCUICUnivSubstitution.v index dd268c3cd..c8aeb777f 100644 --- a/pcuic/theories/PCUICUnivSubstitution.v +++ b/pcuic/theories/PCUICUnivSubstitution.v @@ -26,14 +26,179 @@ Create HintDb univ_subst. Local Ltac aa := rdest; eauto with univ_subst. -Section CheckerFlags. -Context {cf : checker_flags}. +Lemma subst_instance_level_val u l v v' + (H1 : forall s, valuation_mono v s = valuation_mono v' s) + (H2 : forall n, val v (nth n u Level.lSet) = Z.of_nat (valuation_poly v' n)) + : val v (subst_instance_level u l) = val v' l. +Proof. + destruct l; cbn; try congruence. apply H2. +Qed. + +Lemma eq_val v v' + (H1 : forall s, valuation_mono v s = valuation_mono v' s) + (H2 : forall n, valuation_poly v n = valuation_poly v' n) + : forall u, val v u = val v' u. +Proof. + assert (He : forall e : UnivExpr.t, val v e = val v' e). { + intros [|[[] b]]; cbnr; rewrite ?H1, ?H2; reflexivity. } + intro u. rewrite !val_fold_right. + induction ((List.rev (Universe.exprs u).2)); cbn; congruence. +Qed. + +Lemma is_prop_subst_instance_level u l + (Hu : forallb (negb ∘ Level.is_prop) u) + : Level.is_prop (subst_instance_level u l) = Level.is_prop l. +Proof. + destruct l; cbn; try reflexivity. + destruct (le_lt_dec #|u| n) as [HH|HH]. + + now rewrite nth_overflow. + + eapply (forallb_nth _ _ _ Level.lSet Hu) in HH. + destruct HH as [l [HH1 HH2]]. rewrite HH1. now apply ssrbool.negbTE. +Qed. + +Lemma nth_error_Some' {A} l n : (exists x : A, nth_error l n = Some x) <-> n < length l. +Proof. + revert n. induction l; destruct n; simpl. + - split; [now destruct 1 | inversion 1]. + - split; [now destruct 1 | inversion 1]. + - split; now auto with arith. + - rewrite IHl; split; auto with arith. +Qed. + +Lemma nth_error_forallb {A} P l n : + @forallb A P l -> on_Some_or_None P (nth_error l n). +Proof. + induction l in n |- *. + - intros _. destruct n; constructor. + - intro H; invs H. toProp H1 as [? ?]. destruct n; cbn; auto. +Qed. +Lemma subst_instance_univ_val u l v v' + (Hu : forallb (negb ∘ Level.is_prop) u) + (H1 : forall s, valuation_mono v s = valuation_mono v' s) + (H2 : forall n, val v (nth n u Level.lSet) = Z.of_nat (valuation_poly v' n)) + : val v (subst_instance_univ u l) = val v' l. +Proof. + assert (He: forall e : UnivExpr.t, val v (subst_instance_level_expr u e) = val v' e). { + clear l. intros [|[[] b]]; cbn; rewrite <- ?H1, <- ?H2; try reflexivity. + rewrite nth_nth_error. + destruct (le_lt_dec #|u| n) as [HH|HH]. + + apply nth_error_None in HH; now rewrite HH. + + apply nth_error_Some' in HH. destruct HH as [l HH]; rewrite HH. + destruct l; cbnr. + eapply nth_error_forallb in Hu. rewrite HH in Hu. discriminate. } + symmetry. apply val_caract. split. + - intros e Xe. unfold subst_instance_univ. + apply val_le_caract. eexists; split. + + apply Universe.map_spec. eexists; split; tea. reflexivity. + + now rewrite He. + - destruct ((val_caract (subst_instance_univ u l) v _).p1 eq_refl) + as [_ [e [He1 He2]]]. + apply Universe.map_spec in He1. destruct He1 as [e0 [He0 He1]]; subst. + eexists; split; tea. now rewrite <- He2, He. +Qed. + +Definition subst_instance_valuation (u : Instance.t) (v : valuation) := + {| valuation_mono := valuation_mono v ; + valuation_poly := fun i => Z.to_nat (val v (nth i u Level.lSet)) |}. + + +Lemma subst_instance_univ_val' u l v + (Hu : forallb (negb ∘ Level.is_prop) u) + : val v (subst_instance_univ u l) = val (subst_instance_valuation u v) l. +Proof. + eapply subst_instance_univ_val; auto. + cbn. intro; rewrite Z2Nat.id; auto. + destruct (le_lt_dec #|u| n) as [HH|HH]. + + now rewrite nth_overflow. + + eapply (forallb_nth _ _ _ Level.lSet Hu) in HH. + destruct HH as [?l [HH1 HH2]]. rewrite HH1. + destruct l0; try discriminate; cbn. + apply Zle_0_nat. +Qed. + + +Lemma subst_instance_univ_make l u : + subst_instance_univ u (Universe.make l) + = Universe.make (subst_instance_level u l). +Proof. + destruct l; cbnr. rewrite nth_nth_error. + destruct nth_error; cbnr. +Qed. + + +Class SubstUnivPreserving Re := Build_SubstUnivPreserving : + forall s u1 u2, R_universe_instance Re u1 u2 -> + Re (subst_instance_univ u1 s) (subst_instance_univ u2 s). + +Lemma subst_equal_inst_inst Re : + SubstUnivPreserving Re -> + forall u u1 u2, R_universe_instance Re u1 u2 -> + R_universe_instance Re (subst_instance_instance u1 u) + (subst_instance_instance u2 u). +Proof. + intros hRe u. induction u; cbnr; try now constructor. + intros u1 u2; unfold R_universe_instance; cbn; constructor. + - pose proof (hRe (Universe.make a) u1 u2 H) as HH. + now rewrite !subst_instance_univ_make in HH. + - exact (IHu u1 u2 H). +Qed. + +Lemma eq_term_upto_univ_subst_instance_constr Re : + RelationClasses.Reflexive Re -> + SubstUnivPreserving Re -> + forall t u1 u2, + R_universe_instance Re u1 u2 -> + eq_term_upto_univ Re Re (subst_instance_constr u1 t) + (subst_instance_constr u2 t). +Proof. + intros ref hRe t. + induction t using term_forall_list_ind; intros u1 u2 hu. + all: cbn; try constructor; eauto using subst_equal_inst_inst. + all: eapply All2_map, All_All2; tea; cbn; intros; rdest; eauto. +Qed. + +Instance leq_term_SubstUnivPreserving {cf:checker_flags} φ : + SubstUnivPreserving (eq_universe φ). +Proof. + intros s u1 u2 hu. + unfold eq_universe in *; destruct check_univs; [|trivial]. + intros v Hv; cbn. + assert (He : forall e, val v (subst_instance_level_expr u1 e) + = val v (subst_instance_level_expr u2 e)). { + destruct e as [|[[] b]]; cbnr. + case_eq (nth_error u1 n). + - intros l1 X. eapply Forall2_nth_error_Some_l in hu. + 2: now rewrite nth_error_map, X. + destruct hu as [l2 [H1 H2]]. + rewrite nth_error_map in H1. + destruct (nth_error u2 n) as [l2'|]; [|discriminate]. + apply some_inj in H1; subst. clear u1 u2 X. + specialize (H2 v Hv). + destruct l1, l2'; cbn in *; try lia. + - intros X. eapply Forall2_nth_error_None_l in hu. + 2: now rewrite nth_error_map, X. + rewrite nth_error_map in hu. + destruct (nth_error u2 n); [discriminate|reflexivity]. } + apply val_caract; split. + - intros e Xe. apply Universe.map_spec in Xe as [e' [H1 H2]]; subst. + apply val_le_caract. eexists; split. + + apply Universe.map_spec; eexists; split; tea; reflexivity. + + now rewrite He. + - destruct ((val_caract (subst_instance_univ u2 s) v _).p1 eq_refl) + as [_ [e [He1 He2]]]. rewrite <- He2. + apply Universe.map_spec in He1. destruct He1 as [e0 [He0 He1]]; subst. + eexists; split; [|eapply He]. eapply Universe.map_spec. + now eexists; split; tea. +Qed. + +Section CheckerFlags. + Global Instance subst_instance_list {A} `(UnivSubst A) : UnivSubst (list A) := fun u => map (subst_instance u). -Global Instance subst_instance_def {A : Set} `(UnivSubst A) : UnivSubst (def A) +Global Instance subst_instance_def {A} `(UnivSubst A) : UnivSubst (def A) := fun u => map_def (subst_instance u) (subst_instance u). Global Instance subst_instance_prod {A B} `(UnivSubst A) `(UnivSubst B) @@ -61,14 +226,28 @@ Proof. rewrite <- (map_nth (subst_instance_level u1)); reflexivity. Qed. +Lemma subst_instance_level_expr_two u1 u2 e : + subst_instance_level_expr u1 (subst_instance_level_expr u2 e) + = subst_instance_level_expr (subst_instance_instance u1 u2) e. +Proof. + destruct e as [|[[] b]]; cbnr. + unfold subst_instance_instance. erewrite nth_error_map. + destruct nth_error; cbnr. + destruct t; cbnr; [destruct b; reflexivity|]. + rewrite nth_nth_error. destruct nth_error; cbnr. +Qed. + Lemma subst_instance_univ_two u1 u2 s : subst_instance_univ u1 (subst_instance_univ u2 s) = subst_instance_univ (subst_instance_instance u1 u2) s. Proof. - unfold subst_instance_univ. rewrite NEL.map_map. - apply NEL.map_ext. clear s. - intros [l []]; unfold subst_instance_level_expr; - now rewrite subst_instance_level_two. + unfold subst_instance_univ. apply eq_univ'. + intro l; split; intro Hl; apply Universe.map_spec in Hl as [l' [H1 H2]]; + apply Universe.map_spec; subst. + - apply Universe.map_spec in H1 as [l'' [H1 H2]]; subst. + eexists; split; tea. apply subst_instance_level_expr_two. + - eexists; split. 2: symmetry; eapply subst_instance_level_expr_two. + apply Universe.map_spec. eexists; split; tea; reflexivity. Qed. Lemma subst_instance_instance_two u1 u2 u : @@ -173,8 +352,69 @@ Proof. + now apply In_subst_instance_cstrs'. Qed. +Lemma is_prop_subst_instance_univ u l + (Hu : forallb (negb ∘ Level.is_prop) u) + : Universe.is_prop (subst_instance_univ u l) = Universe.is_prop l. +Proof. + assert (He : forall a, UnivExpr.is_prop (subst_instance_level_expr u a) + = UnivExpr.is_prop a). { + clear l. intros [|[l b]]; cbnr. + destruct l; cbnr. + apply nth_error_forallb with (n0:=n) in Hu. + destruct nth_error; cbnr. + destruct t; cbnr. discriminate. } + apply iff_is_true_eq_bool. + split; intro H; apply UnivExprSet.for_all_spec in H; proper; + apply UnivExprSet.for_all_spec; proper; intros e Xe. + - rewrite <- He. apply H. apply Universe.map_spec. + eexists; split; tea; reflexivity. + - apply Universe.map_spec in Xe as [e' [H1 H2]]; subst. + rewrite He. now apply H. +Qed. + +Lemma is_prop_subst_instance u x0 : + Universe.is_prop x0 -> Universe.is_prop (subst_instance_univ u x0). +Proof. + assert (He : forall a, UnivExpr.is_prop a -> + UnivExpr.is_prop (subst_instance_level_expr u a)). { + intros [|[[][]]]; cbnr; auto. } + intro H; apply UnivExprSet.for_all_spec in H; proper; + apply UnivExprSet.for_all_spec; proper; intros e Xe. + apply Universe.map_spec in Xe as [e' [H1 H2]]; subst. + now apply He, H. +Qed. + +Lemma is_small_subst_instance_univ u l + : Universe.is_small l -> Universe.is_small (subst_instance_univ u l). +Proof. + assert (He : forall a, UnivExpr.is_small a -> + UnivExpr.is_small (subst_instance_level_expr u a)). { + intros [|[[][]]]; cbnr; auto. } + intro H; apply UnivExprSet.for_all_spec in H; proper; + apply UnivExprSet.for_all_spec; proper; intros e Xe. + apply Universe.map_spec in Xe as [e' [H1 H2]]; subst. + now apply He, H. +Qed. + +Lemma sup_subst_instance_univ u s1 s2 : + subst_instance_univ u (Universe.sup s1 s2) + = Universe.sup (subst_instance_univ u s1) (subst_instance_univ u s2). +Proof. + apply eq_univ'. cbn. + intro x; split; intro Hx. + + apply Universe.map_spec in Hx as [y [H H']]; subst. + apply UnivExprSet.union_spec. + apply UnivExprSet.union_spec in H as [H|H]; [left|right]. + all: apply Universe.map_spec; eexists; split; tea; reflexivity. + + apply Universe.map_spec. + apply UnivExprSet.union_spec in Hx as [H|H]; + apply Universe.map_spec in H as [y [H H']]; subst. + all: eexists; split; [eapply UnivExprSet.union_spec|reflexivity]; auto. +Qed. + +Context {cf : checker_flags}. -Lemma consistent_instance_no_prop lvs φ uctx u : +Lemma consistent_instance_no_prop lvs φ uctx u : consistent_instance lvs φ uctx u -> forallb (fun x => negb (Level.is_prop x)) u. Proof. @@ -227,7 +467,7 @@ Hint Resolve subst_instance_cstrs_two Lemma val0_subst_instance_level u l v (Hu : forallb (negb ∘ Level.is_prop) u) - : val0 v (subst_instance_level u l) = val0 (subst_instance_valuation u v) l. + : val v (subst_instance_level u l) = val (subst_instance_valuation u v) l. Proof. destruct l; aa; cbn. rewrite Znat.Z2Nat.id; auto. @@ -326,7 +566,7 @@ Proof. split; apply LS.union_spec; right; apply H. Qed. -Definition is_monomorphic_cstr (c : univ_constraint) +Definition is_monomorphic_cstr (c : UnivConstraint.t) := negb (Level.is_var c.1.1) && negb (Level.is_var c.2). Lemma monomorphic_global_constraint Σ (hΣ : wf Σ) c : @@ -533,9 +773,9 @@ Qed. Hint Resolve consistent_instance_valid_constraints : univ_subst. -Class SubstUnivPreserved {A} `{UnivSubst A} (R : constraints -> crelation A) +Class SubstUnivPreserved {A} `{UnivSubst A} (R : ConstraintSet.t -> crelation A) := Build_SubstUnivPreserved : - forall φ φ' (u : universe_instance), + forall φ φ' (u : Instance.t), forallb (fun x => negb (Level.is_prop x)) u -> valid_constraints φ' (subst_instance_cstrs u φ) -> subrelation (R φ) @@ -587,7 +827,8 @@ Proof. 1: split. 1: apply Forall2_map_inv. 1: apply Forall2_map. - all: rewrite !map_map; apply map_ext; reflexivity. + all: rewrite !map_map; apply map_ext. + all: intro; apply subst_instance_univ_make. Qed. Definition precompose_subst_instance_instance__1 Rle u i i' @@ -598,7 +839,7 @@ Definition precompose_subst_instance_instance__2 Rle u i i' Global Instance eq_term_upto_univ_subst_instance - (Re Rle : constraints -> universe -> universe -> Prop) + (Re Rle : ConstraintSet.t -> Universe.t -> Universe.t -> Prop) {he: SubstUnivPreserved Re} {hle: SubstUnivPreserved Rle} : SubstUnivPreserved (fun φ => eq_term_upto_univ (Re φ) (Rle φ)). Proof. @@ -621,8 +862,6 @@ Proof. exact _. Qed. - - (** Now routine lemmas ... *) Lemma subst_instance_univ_super l u @@ -630,19 +869,10 @@ Lemma subst_instance_univ_super l u : subst_instance_univ u (Universe.super l) = Universe.super (subst_instance_level u l). Proof. - destruct l; cbnr. unfold subst_instance_level_expr; cbn. - destruct (le_lt_dec #|u| n) as [HH|HH]. - + now rewrite nth_overflow. - + eapply (forallb_nth _ _ _ Level.lSet Hu) in HH. - destruct HH as [l [HH1 HH2]]. rewrite HH1. - destruct l; cbn; try reflexivity; discriminate. -Qed. - -Lemma subst_instance_univ_make l u : - subst_instance_univ u (Universe.make l) - = Universe.make (subst_instance_level u l). -Proof. - reflexivity. + destruct l; cbnr. + rewrite nth_nth_error. + destruct nth_error; cbnr. + destruct t; cbnr. Qed. @@ -671,46 +901,6 @@ Proof. Qed. -Lemma is_prop_subst_instance_univ u l - (Hu : forallb (negb ∘ Level.is_prop) u) - : Universe.is_prop (subst_instance_univ u l) = Universe.is_prop l. -Proof. - assert (He : forall a, Universe.Expr.is_prop (subst_instance_level_expr u a) - = Universe.Expr.is_prop a). { - intros [[] b]; cbn; try reflexivity. - destruct (le_lt_dec #|u| n) as [HH|HH]. - + now rewrite nth_overflow. - + eapply (forallb_nth _ _ _ Level.lSet Hu) in HH. - destruct HH as [?l [HH1 HH2]]. rewrite HH1. now apply ssrbool.negbTE. } - induction l. - - apply He. - - cbn. f_equal. - + apply He. - + apply IHl. -Qed. - -Lemma is_small_subst_instance_univ u l - : Universe.is_small l -> Universe.is_small (subst_instance_univ u l). -Proof. - assert (He : forall a, Universe.Expr.is_small a -> - Universe.Expr.is_small (subst_instance_level_expr u a)). { - intros [[] []]; cbn; auto. } - induction l. 1: apply He. - intro HH; cbn in HH; apply andP in HH; destruct HH as [H1 H2]. - cbn. apply andb_and; split. - - now apply He. - - now apply IHl. -Qed. - - -Lemma sup_subst_instance_univ u s1 s2 : - subst_instance_univ u (Universe.sup s1 s2) - = Universe.sup (subst_instance_univ u s1) (subst_instance_univ u s2). -Proof. - unfold subst_instance_univ, Universe.sup. - apply NEL.map_app. -Qed. - Lemma product_subst_instance u s1 s2 (Hu : forallb (negb ∘ Level.is_prop) u) : subst_instance_univ u (Universe.sort_of_product s1 s2) @@ -998,20 +1188,6 @@ Proof. Qed. -(* todo move *) -Lemma option_map_two {A B C} (f : A -> B) (g : B -> C) x - : option_map g (option_map f x) = option_map (g ∘ f) x. -Proof. - destruct x; reflexivity. -Qed. - -Lemma option_map_ext {A B} (f g : A -> B) (H : forall x, f x = g x) - : forall z, option_map f z = option_map g z. -Proof. - intros []; cbn; congruence. -Qed. - - Lemma subst_instance_instantiate_params_subst u0 params pars s ty : option_map (on_pair (map (subst_instance_constr u0)) (subst_instance_constr u0)) (instantiate_params_subst params pars s ty) diff --git a/pcuic/theories/PCUICWcbvEval.v b/pcuic/theories/PCUICWcbvEval.v index f06391da8..c3f02e4cb 100644 --- a/pcuic/theories/PCUICWcbvEval.v +++ b/pcuic/theories/PCUICWcbvEval.v @@ -243,17 +243,17 @@ Section Wcbv. (forall i : nat, option_map decl_body (nth_error Γ i) = Some None -> P (tRel i) (tRel i)) -> (forall (c : ident) (decl : constant_body) (body : term), declared_constant Σ c decl -> - forall (u : universe_instance) (res : term), + forall (u : Instance.t) (res : term), cst_body decl = Some body -> eval (subst_instance_constr u body) res -> P (subst_instance_constr u body) res -> P (tConst c u) res) -> (forall (c : ident) (decl : constant_body), - declared_constant Σ c decl -> forall u : universe_instance, cst_body decl = None -> P (tConst c u) (tConst c u)) -> - (forall (ind : inductive) (pars : nat) (discr : term) (c : nat) (u : universe_instance) + declared_constant Σ c decl -> forall u : Instance.t, cst_body decl = None -> P (tConst c u) (tConst c u)) -> + (forall (ind : inductive) (pars : nat) (discr : term) (c : nat) (u : Instance.t) (args : list term) (p : term) (brs : list (nat × term)) (res : term), eval discr (mkApps (tConstruct ind c u) args) -> P discr (mkApps (tConstruct ind c u) args) -> eval (iota_red pars c args brs) res -> P (iota_red pars c args brs) res -> P (tCase (ind, pars) p discr brs) res) -> - (forall (i : inductive) (pars arg : nat) (discr : term) (args : list term) (k : nat) (u : universe_instance) + (forall (i : inductive) (pars arg : nat) (discr : term) (args : list term) (k : nat) (u : Instance.t) (a res : term), eval discr (mkApps (tConstruct i k u) args) -> P discr (mkApps (tConstruct i k u) args) -> @@ -376,7 +376,7 @@ Section Wcbv. - now eapply atom_mkApps in H. - intros * isapp appeq. move: (value_head_nApp H) => Ht. apply mkApps_eq_inj in appeq; intuition subst; auto. - - intros * isapp appeq. move: (isStuckfix_nApp H1) => Hf. + - intros * isapp appeq. move: (isStuckfix_nApp H) => Hf. apply mkApps_eq_inj in appeq; intuition subst; auto. Qed. @@ -403,7 +403,7 @@ Section Wcbv. destruct decl as [? [b|] ?]; try discriminate. constructor. constructor. - eapply value_stuck_fix => //. - now eapply All2_right in H. + now eapply All2_right in X1. - destruct (mkApps_elim f' [a']). eapply value_mkApps_inv in IHX1 => //. destruct IHX1; intuition subst. @@ -433,8 +433,7 @@ Section Wcbv. Proof. induction 1 using value_values_ind; simpl; auto using value. - now constructor. - - assert (All2 eval l l). - { induction X; constructor; auto. eapply IHX. now depelim H0. } + - apply All_All2_refl in X0 as X1. move/implyP: (value_head_spec t). move/(_ H) => Ht. induction l using rev_ind. simpl. @@ -452,9 +451,9 @@ Section Wcbv. * now eapply eval_atom. * now eapply eval_atom. * rewrite -mkApps_nested. - eapply All_app in H0 as [Hl Hx]. depelim Hx. - eapply All_app in X as [Hl' Hx']. depelim Hx'. - eapply All2_app_inv_r in X0 as [Hl'' [Hx'' [? [? ?]]]]. depelim a0. depelim a0. + eapply All_app in X as [Hl Hx]. depelim Hx. + eapply All_app in X0 as [Hl' Hx']. depelim Hx'. + eapply All2_app_inv_r in X1 as [Hl'' [Hx'' [? [? ?]]]]. depelim a0. depelim a0. eapply eval_app_cong; auto. eapply IHl; auto. now eapply All_All2. @@ -466,8 +465,7 @@ Section Wcbv. rewrite orb_false_r. destruct t; auto. - destruct f; try discriminate. - assert (All2 eval args args). - { clear H0. induction X; constructor; auto. eapply IHX. now depelim H. } + apply All_All2_refl in X0. eapply eval_fix_value => //. constructor; auto. Qed. diff --git a/pcuic/theories/TemplateToPCUICCorrectness.v b/pcuic/theories/TemplateToPCUICCorrectness.v index 3a311fa7e..f3c552bca 100644 --- a/pcuic/theories/TemplateToPCUICCorrectness.v +++ b/pcuic/theories/TemplateToPCUICCorrectness.v @@ -564,7 +564,7 @@ Proof. { inversion wt. assumption. } assert (wargs' : Forall T.wf args'). { inversion wu. assumption. } - induction H in wl, args', wargs', a |- *. + rename X into H; induction H in wl, args', wargs', a |- *. + dependent destruction a. constructor. + dependent destruction a. simpl. constructor. @@ -591,7 +591,7 @@ Proof. * inversion wt. assumption. * inversion wu. assumption. * assumption. - + pose proof (All2_All_mix_left H a) as h. + + pose proof (All2_All_mix_left X a) as h. simpl in h. assert (wl : Forall T.wf l). { inversion wt. assumption. } @@ -1089,7 +1089,7 @@ Proof. - apply Forall_All in H2. clear H H0 H1. revert M1. induction X. simpl. intuition. inv H2. specialize (X H). apply PCUICSubstitution.red1_mkApps_l. apply app_red_r. auto. - inv H2. specialize (IHX H0). + inv H2. specialize (IHX X0). simpl. intros. eapply (IHX (T.tApp M1 [hd])). @@ -1216,7 +1216,7 @@ Definition Tlift_typing (P : Template.Ast.global_env_ext -> Tcontext -> Tterm -> fun Σ Γ t T => match T with | Some T => P Σ Γ t T - | None => { s : universe & P Σ Γ t (T.tSort s) } + | None => { s : Universe.t & P Σ Γ t (T.tSort s) } end. Lemma trans_wf_local: diff --git a/safechecker/_PluginProject.in b/safechecker/_PluginProject.in index 010ef34cc..96600378d 100644 --- a/safechecker/_PluginProject.in +++ b/safechecker/_PluginProject.in @@ -6,15 +6,15 @@ src/classes0.mli src/classes0.ml # From Coq's stdlib -#src/compare_dec.mli -#src/compare_dec.ml -src/mSetList.mli -src/mSetList.ml +src/mSetWeakList.mli +src/mSetWeakList.ml src/eqdepFacts.mli src/eqdepFacts.ml src/monad_utils.ml src/monad_utils.mli +src/utils.ml +src/utils.mli src/uGraph0.ml src/uGraph0.mli src/wGraph.ml diff --git a/safechecker/src/metacoq_safechecker_plugin.mlpack b/safechecker/src/metacoq_safechecker_plugin.mlpack index 3189012f9..eceb3ba03 100644 --- a/safechecker/src/metacoq_safechecker_plugin.mlpack +++ b/safechecker/src/metacoq_safechecker_plugin.mlpack @@ -1,9 +1,9 @@ Monad_utils -MSetList +MSetWeakList EqdepFacts +Utils WGraph -Monad_utils UGraph0 Typing0 diff --git a/safechecker/theories/PCUICSafeChecker.v b/safechecker/theories/PCUICSafeChecker.v index a8969d73e..ea71e64ae 100644 --- a/safechecker/theories/PCUICSafeChecker.v +++ b/safechecker/theories/PCUICSafeChecker.v @@ -129,50 +129,8 @@ Inductive type_error := | Msg (s : string). Local Open Scope string_scope. -Definition string_of_list_aux {A} (f : A -> string) (l : list A) : string := - let fix aux l := - match l return string with - | nil => "" - | cons a nil => f a - | cons a l => f a ++ "," ++ aux l - end - in aux l. - -Definition string_of_list {A} (f : A -> string) (l : list A) : string := - "[" ++ string_of_list_aux f l ++ "]". - -Definition string_of_level (l : Level.t) : string := - match l with - | Level.lProp => "Prop" - | Level.lSet => "Set" - | Level.Level s => s - | Level.Var n => "Var " ++ string_of_nat n - end. - -Definition string_of_level_expr (l : Level.t * bool) : string := - let '(l, b) := l in - string_of_level l ++ (if b then "+1" else ""). - -Definition string_of_sort (u : universe) := - string_of_list string_of_level_expr u. - -Definition string_of_name (na : name) := - match na with - | nAnon => "_" - | nNamed n => n - end. -Definition string_of_universe_instance u := - string_of_list string_of_level u. - -Definition string_of_def {A : Set} (f : A -> string) (def : def A) := - "(" ++ string_of_name (dname def) ++ "," ++ f (dtype def) ++ "," ++ f (dbody def) ++ "," - ++ string_of_nat (rarg def) ++ ")". - -Definition string_of_inductive (i : inductive) := - (inductive_mind i) ++ "," ++ string_of_nat (inductive_ind i). - -Definition print_no_prop_level := string_of_level ∘ level_of_no_prop. +Definition print_no_prop_level := string_of_level ∘ NoPropLevel.to_level. Definition print_edge '(l1, n, l2) := "(" ++ print_no_prop_level l1 ++ ", " ++ string_of_nat n ++ ", " @@ -631,7 +589,7 @@ Section Typecheck. typing_result ({ A : term & ∥ Σ ;;; Γ |- t : A ∥ })). Program Definition infer_type Γ HΓ t - : typing_result ({u : universe & ∥ Σ ;;; Γ |- t : tSort u ∥}) := + : typing_result ({u : Universe.t & ∥ Σ ;;; Γ |- t : tSort u ∥}) := tx <- infer Γ HΓ t ;; u <- reduce_to_sort Γ tx.π1 _ ;; ret (u.π1; _). @@ -684,9 +642,9 @@ Section Typecheck. now apply global_ext_uctx_consistent. Qed. - Lemma is_graph_of_uctx_levels l : + Lemma is_graph_of_uctx_levels (l : NoPropLevel.t) : wGraph.VSet.mem l (uGraph.wGraph.V G) -> - LevelSet.mem l (global_ext_levels Σ). + LevelSet.mem (NoPropLevel.to_level l) (global_ext_levels Σ). Proof. unfold is_graph_of_uctx in HG. case_eq (gc_of_uctx (global_ext_uctx Σ)); [intros [lvs cts] XX|intro XX]; @@ -698,7 +656,7 @@ Section Typecheck. apply LevelSetProp.FM.elements_2. unfold no_prop_levels in H. rewrite LevelSet.fold_spec in H. - cut (SetoidList.InA eq (level_of_no_prop l) (LevelSet.elements lvs) + cut (SetoidList.InA eq (NoPropLevel.to_level l) (LevelSet.elements lvs) \/ wGraph.VSet.In l wGraph.VSet.empty). { intros [|H0]; [trivial|]. now apply wGraph.VSet.empty_spec in H0. } @@ -719,7 +677,7 @@ Section Typecheck. check_eq_nat #|u| 0 (Msg "monomorphic instance should be of length 0") | Polymorphic_ctx (inst, cstrs) => let '(inst, cstrs) := AUContext.repr (inst, cstrs) in - check_eq_true (forallb (fun l => match no_prop_of_level l with + check_eq_true (forallb (fun l => match NoPropLevel.of_level l with | Some l => wGraph.VSet.mem l (uGraph.wGraph.V G) | None => false end) u) @@ -768,7 +726,7 @@ Section Typecheck. : eqb_decl d d' -> eq_decl (global_ext_constraints Σ) d d'. Proof. unfold eqb_decl, eq_decl. - intro H. utils.toProp. apply eqb_opt_term_spec in H. + intro H. rtoProp. apply eqb_opt_term_spec in H. eapply eqb_term_spec in H0; tea. now split. Qed. @@ -801,12 +759,12 @@ Section Typecheck. | tEvar ev args => raise (UnboundEvar ev) | tSort u => - match u with - | NEL.sing (l, false) => + match Universe.get_is_level u with + | Some l => check_eq_true (LevelSet.mem l (global_ext_levels Σ)) (Msg ("undeclared level " ++ string_of_level l));; ret (tSort (Universe.super l); _) - | _ => raise (Msg (string_of_sort u ++ " is not a level")) + | None => raise (Msg (string_of_sort u ++ " is not a level")) end | tProd na A B => @@ -1000,9 +958,15 @@ Section Typecheck. end end. + (* tRel *) Next Obligation. intros; sq; now econstructor. Defined. - Next Obligation. intros; sq; econstructor; tas. - now apply LevelSetFact.mem_2. Defined. + (* tSort *) + Next Obligation. + symmetry in Heq_anonymous. + apply get_is_level_correct in Heq_anonymous. + subst u. sq; econstructor; tas. + now apply LevelSetFact.mem_2. + Defined. (* tProd *) Next Obligation. (* intros Γ HΓ t na A B Heq_t [s ?]; *) @@ -1383,14 +1347,22 @@ Section CheckEnv. easy. Defined. + (* todo move *) + Definition VariableLevel_get_noprop (l : NoPropLevel.t) : option VariableLevel.t + := match l with + | NoPropLevel.lSet => None + | NoPropLevel.Level s => Some (VariableLevel.Level s) + | NoPropLevel.Var n => Some (VariableLevel.Var n) + end. + Definition add_uctx (uctx : wGraph.VSet.t × GoodConstraintSet.t) (G : universes_graph) : universes_graph := let levels := wGraph.VSet.union uctx.1 G.1.1 in let edges := wGraph.VSet.fold (fun l E => - match l with - | lSet => E - | vtn ll => wGraph.EdgeSet.add (edge_of_level ll) E + match VariableLevel_get_noprop l with + | None => E + | Some ll => wGraph.EdgeSet.add (edge_of_level ll) E end) uctx.1 G.1.2 in let edges := GoodConstraintSet.fold diff --git a/safechecker/theories/PCUICSafeConversion.v b/safechecker/theories/PCUICSafeConversion.v index 3b764cdcc..fdca7f7c8 100644 --- a/safechecker/theories/PCUICSafeConversion.v +++ b/safechecker/theories/PCUICSafeConversion.v @@ -36,7 +36,7 @@ Definition wf_global_uctx_invariants {cf:checker_flags} Σ : Proof. intros [HΣ]. split. - cbn. unfold global_levels. - cut (LevelSet.In lSet (LevelSet_pair Level.lSet Level.lProp)). + cut (LevelSet.In Level.lSet (LevelSet_pair Level.lSet Level.lProp)). + generalize (LevelSet_pair Level.lSet Level.lProp). clear HΣ. induction Σ; simpl. 1: easy. intros X H. apply LevelSet.union_spec. now right. @@ -89,7 +89,7 @@ Definition wf_ext_global_uctx_invariants {cf:checker_flags} Σ : Proof. intros [HΣ]. split. - apply LevelSet.union_spec. right. unfold global_levels. - cut (LevelSet.In lSet (LevelSet_pair Level.lSet Level.lProp)). + cut (LevelSet.In Level.lSet (LevelSet_pair Level.lSet Level.lProp)). + generalize (LevelSet_pair Level.lSet Level.lProp). induction Σ.1; simpl. 1: easy. intros X H. apply LevelSet.union_spec. now right. @@ -522,12 +522,12 @@ Section Conversion. pose proof hΣ'. apply eqb_term_upto_univ_impl. - intros u1 u2. - eapply (check_eqb_universe_spec G (global_ext_uctx Σ)). + eapply (check_eqb_universe_spec' G (global_ext_uctx Σ)). + now eapply wf_ext_global_uctx_invariants. + now eapply global_ext_uctx_consistent. + assumption. - intros u1 u2. - eapply (check_leqb_universe_spec G (global_ext_uctx Σ)). + eapply (check_leqb_universe_spec' G (global_ext_uctx Σ)). + now eapply wf_ext_global_uctx_invariants. + now eapply global_ext_uctx_consistent. + assumption. @@ -540,12 +540,12 @@ Section Conversion. pose proof hΣ'. apply eqb_term_upto_univ_impl. - intros u1 u2. - eapply (check_eqb_universe_spec G (global_ext_uctx Σ)). + eapply (check_eqb_universe_spec' G (global_ext_uctx Σ)). + now eapply wf_ext_global_uctx_invariants. + now eapply global_ext_uctx_consistent. + assumption. - intros u1 u2. - eapply (check_eqb_universe_spec G (global_ext_uctx Σ)). + eapply (check_eqb_universe_spec' G (global_ext_uctx Σ)). + now eapply wf_ext_global_uctx_invariants. + now eapply global_ext_uctx_consistent. + assumption. @@ -1353,7 +1353,7 @@ Section Conversion. prog_discr (tCoFix _ _) (tCoFix _ _) := False ; prog_discr _ _ := True. - Inductive prog_view : term -> term -> Set := + Inductive prog_view : term -> term -> Type := | prog_view_App u1 v1 u2 v2 : prog_view (tApp u1 v1) (tApp u2 v2) @@ -1431,9 +1431,9 @@ Section Conversion. Qed. Equations(noeqns) unfold_constants (Γ : context) (leq : conv_pb) - (c : kername) (u : universe_instance) (π1 : stack) + (c : kername) (u : Instance.t) (π1 : stack) (h1 : wtp Γ (tConst c u) π1) - (c' : kername) (u' : universe_instance) (π2 : stack) + (c' : kername) (u' : Instance.t) (π2 : stack) (h2 : wtp Γ (tConst c' u') π2) (hx : conv_stack_ctx Γ π1 π2) (aux : Aux Term Γ (tConst c u) π1 (tConst c' u') π2 h2) @@ -1528,7 +1528,7 @@ Section Conversion. unfold eqb_universe_instance in e. eapply forallb2_Forall2 in e. eapply Forall2_impl. 1: eassumption. - intros. eapply (check_eqb_universe_spec G (global_ext_uctx Σ)). + intros. eapply (check_eqb_universe_spec' G (global_ext_uctx Σ)). all: auto. - eapply wf_ext_global_uctx_invariants. eapply hΣ'. diff --git a/safechecker/theories/PCUICSafeReduce.v b/safechecker/theories/PCUICSafeReduce.v index c68136677..58971f21f 100644 --- a/safechecker/theories/PCUICSafeReduce.v +++ b/safechecker/theories/PCUICSafeReduce.v @@ -274,7 +274,7 @@ Section Reduce. discr_construct (tConstruct ind n ui) := False ; discr_construct _ := True. - Inductive construct_view : term -> Set := + Inductive construct_view : term -> Type := | view_construct : forall ind n ui, construct_view (tConstruct ind n ui) | view_other : forall t, discr_construct t -> construct_view t. @@ -294,7 +294,7 @@ Section Reduce. red_discr (tProj _ _) _ := False ; red_discr _ _ := True. - Inductive red_view : term -> stack -> Set := + Inductive red_view : term -> stack -> Type := | red_view_Rel c π : red_view (tRel c) π | red_view_LetIn A b B c π : red_view (tLetIn A b B c) π | red_view_Const c u π : red_view (tConst c u) π @@ -321,7 +321,7 @@ Section Reduce. discr_construct_cofix (tCoFix mfix idx) := False ; discr_construct_cofix _ := True. - Inductive construct_cofix_view : term -> Set := + Inductive construct_cofix_view : term -> Type := | ccview_construct : forall ind n ui, construct_cofix_view (tConstruct ind n ui) | ccview_cofix : forall mfix idx, construct_cofix_view (tCoFix mfix idx) | ccview_other : forall t, discr_construct_cofix t -> construct_cofix_view t. diff --git a/safechecker/theories/PCUICSafeRetyping.v b/safechecker/theories/PCUICSafeRetyping.v index 684c4f332..0271b65ce 100644 --- a/safechecker/theories/PCUICSafeRetyping.v +++ b/safechecker/theories/PCUICSafeRetyping.v @@ -42,7 +42,7 @@ Section TypeOf. Section SortOf. Context (type_of : forall Γ t, welltyped Σ Γ t -> typing_result (∑ T, ∥ Σ ;;; Γ |- t : T ∥)). - Program Definition type_of_as_sort Γ t (wf : welltyped Σ Γ t) : typing_result universe := + Program Definition type_of_as_sort Γ t (wf : welltyped Σ Γ t) : typing_result Universe.t := tx <- type_of Γ t wf ;; wfs <- @reduce_to_sort cf Σ hΣ Γ (projT1 tx) _ ;; ret (m:=typing_result) (projT1 wfs). diff --git a/template-coq/_PluginProject b/template-coq/_PluginProject index 465e5cb59..d6910f1fd 100644 --- a/template-coq/_PluginProject +++ b/template-coq/_PluginProject @@ -64,24 +64,24 @@ gen-src/mCCompare.ml gen-src/mCCompare.mli gen-src/mCList.ml gen-src/mCList.mli +gen-src/mCOption.ml +gen-src/mCOption.mli gen-src/mCProd.ml gen-src/mCProd.mli gen-src/mCRelations.ml gen-src/mCRelations.mli gen-src/mCString.ml gen-src/mCString.mli -gen-src/mCOption.ml -gen-src/mCOption.mli gen-src/mSetDecide.ml gen-src/mSetDecide.mli gen-src/mSetFacts.ml gen-src/mSetFacts.mli gen-src/mSetInterface.ml gen-src/mSetInterface.mli +gen-src/mSetList.ml +gen-src/mSetList.mli gen-src/mSetProperties.ml gen-src/mSetProperties.mli -gen-src/mSetWeakList.ml -gen-src/mSetWeakList.mli gen-src/nat0.ml gen-src/nat0.mli gen-src/orderedType0.ml @@ -113,8 +113,6 @@ gen-src/universes0.ml gen-src/universes0.mli gen-src/univSubst0.ml gen-src/univSubst0.mli -gen-src/utils.ml -gen-src/utils.mli gen-src/metacoq_template_plugin.mlpack diff --git a/template-coq/gen-src/metacoq_template_plugin.mlpack b/template-coq/gen-src/metacoq_template_plugin.mlpack index 4e8c71253..31e5a795d 100644 --- a/template-coq/gen-src/metacoq_template_plugin.mlpack +++ b/template-coq/gen-src/metacoq_template_plugin.mlpack @@ -21,7 +21,7 @@ OrderedType0 MSetInterface MSetFacts MSetDecide -MSetWeakList +MSetList MSetProperties BinNums Equality0 @@ -34,7 +34,6 @@ MCProd MCString MCCompare All_Forall -Utils Config0 Universes0 BasicAst diff --git a/template-coq/src/ast_denoter.ml b/template-coq/src/ast_denoter.ml index 17d03ed6b..da96ab886 100644 --- a/template-coq/src/ast_denoter.ml +++ b/template-coq/src/ast_denoter.ml @@ -14,7 +14,7 @@ struct type quoted_int = Datatypes.nat type quoted_bool = bool type quoted_name = name - type quoted_sort = Universes0.universe + type quoted_sort = Universes0.Universe.t type quoted_cast_kind = cast_kind type quoted_kernel_name = char list type quoted_inductive = inductive @@ -22,9 +22,9 @@ struct type quoted_global_reference = global_reference type quoted_sort_family = Universes0.sort_family - type quoted_constraint_type = Universes0.constraint_type - type quoted_univ_constraint = Universes0.univ_constraint - type quoted_univ_constraints = Universes0.constraints + type quoted_constraint_type = Universes0.ConstraintType.t + type quoted_univ_constraint = Universes0.UnivConstraint.t + type quoted_univ_constraints = Universes0.ConstraintSet.t type quoted_univ_instance = Universes0.Instance.t type quoted_univ_context = Universes0.UContext.t type quoted_univ_contextset = Universes0.ContextSet.t diff --git a/template-coq/src/ast_quoter.ml b/template-coq/src/ast_quoter.ml index 8e1a603b1..a03c259ad 100644 --- a/template-coq/src/ast_quoter.ml +++ b/template-coq/src/ast_quoter.ml @@ -15,7 +15,7 @@ struct type quoted_int = Datatypes.nat type quoted_bool = bool type quoted_name = name - type quoted_sort = Universes0.universe + type quoted_sort = Universes0.Universe.t type quoted_cast_kind = cast_kind type quoted_kernel_name = char list type quoted_inductive = inductive @@ -23,9 +23,9 @@ struct type quoted_global_reference = global_reference type quoted_sort_family = Universes0.sort_family - type quoted_constraint_type = Universes0.constraint_type - type quoted_univ_constraint = Universes0.univ_constraint - type quoted_univ_constraints = Universes0.constraints + type quoted_constraint_type = Universes0.ConstraintType.t + type quoted_univ_constraint = Universes0.UnivConstraint.t + type quoted_univ_constraints = Universes0.ConstraintSet.t type quoted_univ_instance = Universes0.Instance.t type quoted_univ_context = Universes0.UContext.t type quoted_univ_contextset = Universes0.ContextSet.t @@ -69,15 +69,15 @@ struct let quote_bool x = x let quote_level l = - if Univ.Level.is_prop l then Universes0.Level.prop - else if Univ.Level.is_set l then Universes0.Level.set + if Univ.Level.is_prop l then Universes0.Level.Coq_lProp + else if Univ.Level.is_set l then Universes0.Level.Coq_lSet 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)) - let quote_universe s : Universes0.universe = - let univs = Univ.Universe.map (fun (l,i) -> Universes0.Universe.make' (quote_level l, i > 0)) s in - List.fold_left Universes0.Universe.sup (List.hd univs) (List.tl univs) + let quote_universe s : Universes0.Universe.t = + let univs = Univ.Universe.map (fun (l,i) -> (quote_level l, i > 0)) s in + Universes0.Universe.from_kernel_repr (List.hd univs) (List.tl univs) let quote_sort s = quote_universe (Sorts.univ_of_sort s) @@ -329,18 +329,16 @@ struct Univ.Level.make (Univ.Level.UGlobal.make dp idx) | Universes0.Level.Var n -> Univ.Level.var (unquote_int n) - let unquote_level_expr (trm : Universes0.Level.t) (b : quoted_bool) : Univ.Universe.t = - let l = unquote_level trm in + let unquote_level_expr (trm : Universes0.Level.t * quoted_bool) : Univ.Universe.t = + let l = unquote_level (fst trm) in let u = Univ.Universe.make l in - if b && not (Univ.Level.is_prop l) then Univ.Universe.super u + if snd trm && not (Univ.Level.is_prop l) then Univ.Universe.super u else u let unquote_universe evd (trm : Universes0.Universe.t) = - (* | [] -> Evd.new_univ_variable (Evd.UnivFlexible false) evd *) - let rec aux = function - | Utils.NEL.Coq_sing (l,b) -> unquote_level_expr l b - | Utils.NEL.Coq_cons ((l,b), q) -> Univ.Universe.sup (aux q) (unquote_level_expr l b) - in evd, aux trm + let l = Universes0.Universe.to_kernel_repr trm in + let l = List.map unquote_level_expr l in + evd, List.fold_left Univ.Universe.sup (List.hd l) (List.tl l) let quote_global_reference : Names.GlobRef.t -> quoted_global_reference = function | Names.GlobRef.VarRef _ -> CErrors.user_err (Pp.str "VarRef unsupported") diff --git a/template-coq/src/constr_denoter.ml b/template-coq/src/constr_denoter.ml index 63c35846f..cd72ca815 100644 --- a/template-coq/src/constr_denoter.ml +++ b/template-coq/src/constr_denoter.ml @@ -56,19 +56,6 @@ struct else not_supported_verb trm "unquote_list" - let rec unquote_non_empty_list trm = - let (h,args) = app_full trm [] in - if constr_equall h nel_sing then - match args with - _ :: x :: [] -> [x] - | _ -> bad_term_verb trm "unquote_non_empty_list" - else if constr_equall h nel_cons then - match args with - _ :: x :: xs :: [] -> x :: unquote_non_empty_list xs - | _ -> bad_term_verb trm "unquote_non_empty_list" - else - not_supported_verb trm "unquote_non_empty_list" - (* Unquote Coq nat to OCaml int *) let rec unquote_nat trm = let (h,args) = app_full trm [] in @@ -199,27 +186,68 @@ struct else not_supported_verb trm "unquote_level" - let unquote_level_expr evm trm (* of type level *) b (* of type bool *) : Evd.evar_map * Univ.Universe.t = - let evm, l = unquote_level evm trm in - let u = Univ.Universe.make l in - evm, if unquote_bool b then Univ.Universe.super u else u + let unquote_noproplevel evm trm (* of type noproplevel *) : Evd.evar_map * Univ.Level.t = + let (h,args) = app_full trm [] in + if constr_equall h noprop_tSet then + match args with + | [] -> evm, Univ.Level.set + | _ -> bad_term_verb trm "unquote_noproplevel" + else if constr_equall h noprop_tLevel then + match args with + | s :: [] -> debug (fun () -> str "Unquoting level " ++ Printer.pr_constr_env (Global.env ()) evm trm); + get_level evm (unquote_string s) + | _ -> bad_term_verb trm "unquote_noproplevel" + else if constr_equall h noprop_tLevelVar then + match args with + | l :: [] -> evm, Univ.Level.var (unquote_nat l) + | _ -> bad_term_verb trm "unquote_noproplevel" + else + not_supported_verb trm "unquote_noproplevel" + + let unquote_univ_expr evm trm (* of type UnivExpr.t *) : Evd.evar_map * Univ.Universe.t = + let (h,args) = app_full trm [] in + if constr_equall h univexpr_lProp then + match args with + | [] -> evm, Univ.Universe.type0m + | _ -> bad_term_verb trm "unquote_univ_expr" + else if constr_equall h univexpr_npe then + match args with + | [x] -> + let l, b = unquote_pair x in + let evm, l' = unquote_noproplevel evm l in + let u = Univ.Universe.make l' in + evm, if unquote_bool b then Univ.Universe.super u else u + | _ -> bad_term_verb trm "unquote_univ_expr" + else + not_supported_verb trm "unquote_univ_expr" let unquote_universe evm trm (* of type universe *) = - if constr_equall trm lfresh_universe then + let (h,args) = app_full trm [] in + if constr_equall h lfresh_universe then if !strict_unquote_universe_mode then CErrors.user_err ~hdr:"unquote_universe" (str "It is not possible to unquote a fresh universe in Strict Unquote Universe Mode.") else let evm, u = Evd.new_univ_variable (Evd.UnivFlexible false) evm in Feedback.msg_info (str"Fresh universe " ++ Universe.pr u ++ str" was added to the context."); evm, u + else if constr_equall h tBuild_Universe then + match args with + x :: _ :: [] -> (let (h,args) = app_full x [] in + if constr_equall h tMktUnivExprSet then + match args with + x :: _ :: [] -> (match unquote_list x with + | [] -> assert false + | e::q -> List.fold_left (fun (evm,u) e -> let evm, u' = unquote_univ_expr evm e + in evm, Univ.Universe.sup u u') + (unquote_univ_expr evm e) q) + | _ -> bad_term_verb trm "unquote_universe 0" + else + not_supported_verb trm "unquote_universe 0") + | _ -> bad_term_verb trm "unquote_universe 1" else - let levels = List.map unquote_pair (unquote_non_empty_list trm) in - match levels with - | [] -> assert false - | (l,b)::q -> List.fold_left (fun (evm,u) (l,b) -> let evm, u' = unquote_level_expr evm l b - in evm, Univ.Universe.sup u u') - (unquote_level_expr evm l b) q + not_supported_verb trm "unquote_universe 1" + let unquote_universe_instance evm trm (* of type universe_instance *) = let l = unquote_list trm in diff --git a/template-coq/src/constr_quoted.ml b/template-coq/src/constr_quoted.ml index 79e9e8e04..d625e0d0d 100644 --- a/template-coq/src/constr_quoted.ml +++ b/template-coq/src/constr_quoted.ml @@ -82,9 +82,6 @@ struct let c_nil = resolve "metacoq.list.nil" let c_cons = resolve "metacoq.list.cons" - let nel_sing = resolve "metacoq.nel.sing" - let nel_cons = resolve "metacoq.nel.cons" - let prod_type = resolve "metacoq.prod.type" let c_pair = resolve "metacoq.prod.intro" @@ -135,9 +132,18 @@ struct let tunivLe = ast "constraints.Le" let tunivLt = ast "constraints.Lt" let tunivEq = ast "constraints.Eq" - let tmake_universe = ast "universe.make" + let tMktUnivExprSet = ast "univexprset.mkt" + let tBuild_Universe = ast "universe.build" + let tfrom_kernel_repr = ast "universe.from_kernel_repr" + let tto_kernel_repr = ast "universe.to_kernel_repr" let tLevelSet_of_list = ast "universe.of_list" + let noprop_tSet = ast "noproplevel.lSet" + let noprop_tLevel = ast "noproplevel.Level" + let noprop_tLevelVar = ast "noproplevel.Var" + let univexpr_lProp = ast "univexpr.prop" + let univexpr_npe = ast "univexpr.npe" + (* let tunivcontext = resolve_symbol pkg_univ "universe_context" *) let tVariance = ast "variance.t" let cIrrelevant = ast "variance.Irrelevant" diff --git a/template-coq/src/constr_quoter.ml b/template-coq/src/constr_quoter.ml index 7ba7d4338..f14b7cfc0 100644 --- a/template-coq/src/constr_quoter.ml +++ b/template-coq/src/constr_quoter.ml @@ -107,7 +107,7 @@ struct let levels = Universe.map (fun (l,i) -> pairl tlevel bool_type (quote_level l) (quote_bool (i > 0))) s in let hd = List.hd levels in let tl = to_coq_list (prodl tlevel bool_type) (List.tl levels) in - constr_mkApp (tmake_universe, [| hd ; tl |]) + constr_mkApp (tfrom_kernel_repr, [| hd ; tl |]) let quote_levelset s = let levels = LSet.elements s in diff --git a/template-coq/src/monad_extraction.mlpack b/template-coq/src/monad_extraction.mlpack index c296a9ad3..1279adb78 100644 --- a/template-coq/src/monad_extraction.mlpack +++ b/template-coq/src/monad_extraction.mlpack @@ -19,6 +19,7 @@ Extractable List0 Logic0 MSetWeakList +MSetList MSetProperties MSetFacts Nat0 diff --git a/template-coq/theories/Ast.v b/template-coq/theories/Ast.v index c72633e84..c47cc1fb6 100644 --- a/template-coq/theories/Ast.v +++ b/template-coq/theories/Ast.v @@ -36,19 +36,19 @@ From MetaCoq.Template Require Export Universes. From MetaCoq.Template Require Export BasicAst. -Inductive term : Set := +Inductive term : Type := | tRel (n : nat) | tVar (id : ident) (* For free variables (e.g. in a goal) *) | tEvar (ev : nat) (args : list term) -| tSort (s : universe) +| tSort (s : Universe.t) | tCast (t : term) (kind : cast_kind) (v : term) | tProd (na : name) (ty : term) (body : term) | tLambda (na : name) (ty : term) (body : term) | tLetIn (na : name) (def : term) (def_ty : term) (body : term) | tApp (f : term) (args : list term) -| tConst (c : kername) (u : universe_instance) -| tInd (ind : inductive) (u : universe_instance) -| tConstruct (ind : inductive) (idx : nat) (u : universe_instance) +| tConst (c : kername) (u : Instance.t) +| tInd (ind : inductive) (u : Instance.t) +| tConstruct (ind : inductive) (idx : nat) (u : Instance.t) | tCase (ind_and_nbparams: inductive*nat) (type_info:term) (discr:term) (branches : list (nat * term)) | tProj (proj : projection) (t : term) @@ -164,7 +164,7 @@ Inductive constant_entry := [x1:X1;...;xn:Xn]. *) -Record one_inductive_entry : Set := { +Record one_inductive_entry := { mind_entry_typename : ident; mind_entry_arity : term; mind_entry_template : bool; (* template polymorphism *) diff --git a/template-coq/theories/AstUtils.v b/template-coq/theories/AstUtils.v index 94aff3823..d494e47a8 100644 --- a/template-coq/theories/AstUtils.v +++ b/template-coq/theories/AstUtils.v @@ -29,12 +29,15 @@ Definition string_of_level (l : Level.t) : string := | Level.Var n => "Var" ++ string_of_nat n end. -Definition string_of_level_expr (l : Level.t * bool) : string := - let '(l, b) := l in - string_of_level l ++ (if b then "+1" else ""). +Definition string_of_level_expr (e : UnivExpr.t) : string := + match e with + | UnivExpr.lProp => "Prop" + | UnivExpr.npe (l, b) => string_of_level l ++ (if b then "+1" else "") + end. + +Definition string_of_sort (u : Universe.t) := + string_of_list string_of_level_expr (UnivExprSet.elements u). -Definition string_of_sort (u : universe) := - string_of_list string_of_level_expr u. Definition string_of_name (na : name) := match na with | nAnon => "Anonymous" @@ -43,7 +46,7 @@ Definition string_of_name (na : name) := Definition string_of_universe_instance u := string_of_list string_of_level u. -Definition string_of_def {A : Set} (f : A -> string) (def : def A) := +Definition string_of_def {A} (f : A -> string) (def : def A) := "(" ++ string_of_name (dname def) ++ "," ++ f (dtype def) ++ "," ++ f (dbody def) ++ "," ++ string_of_nat (rarg def) ++ ")". @@ -137,14 +140,14 @@ Proof. destruct decl; reflexivity. Qed. Lemma map_cst_body f decl : option_map f (cst_body decl) = cst_body (map_constant_body f decl). Proof. destruct decl; reflexivity. Qed. -Definition map_def {A B : Set} (tyf bodyf : A -> B) (d : def A) := +Definition map_def {A B} (tyf bodyf : A -> B) (d : def A) := {| dname := d.(dname); dtype := tyf d.(dtype); dbody := bodyf d.(dbody); rarg := d.(rarg) |}. -Lemma map_dtype {A B : Set} (f : A -> B) (g : A -> B) (d : def A) : +Lemma map_dtype {A B} (f : A -> B) (g : A -> B) (d : def A) : f (dtype d) = dtype (map_def f g d). Proof. destruct d; reflexivity. Qed. -Lemma map_dbody {A B : Set} (f : A -> B) (g : A -> B) (d : def A) : +Lemma map_dbody {A B} (f : A -> B) (g : A -> B) (d : def A) : g (dbody d) = dbody (map_def f g d). Proof. destruct d; reflexivity. Qed. @@ -637,13 +640,13 @@ Lemma ind_projs_map f npars_ass arities n oib : map (on_snd (f (S npars_ass))) (ind_projs oib). Proof. destruct oib; simpl. reflexivity. Qed. -Definition test_def {A : Set} (tyf bodyf : A -> bool) (d : def A) := +Definition test_def {A} (tyf bodyf : A -> bool) (d : def A) := tyf d.(dtype) && bodyf d.(dbody). Definition tCaseBrsProp {A} (P : A -> Prop) (l : list (nat * A)) := Forall (fun x => P (snd x)) l. -Definition tFixProp {A : Set} (P P' : A -> Prop) (m : mfixpoint A) := +Definition tFixProp {A} (P P' : A -> Prop) (m : mfixpoint A) := Forall (fun x : def A => P x.(dtype) /\ P' x.(dbody)) m. @@ -652,21 +655,21 @@ Ltac merge_All := repeat toAll. -Lemma map_def_map_def {A B C : Set} (f f' : B -> C) (g g' : A -> B) (d : def A) : +Lemma map_def_map_def {A B C} (f f' : B -> C) (g g' : A -> B) (d : def A) : map_def f f' (map_def g g' d) = map_def (fun x => f (g x)) (fun x => f' (g' x)) d. Proof. destruct d; reflexivity. Qed. -Lemma compose_map_def {A B C : Set} (f f' : B -> C) (g g' : A -> B) : +Lemma compose_map_def {A B C} (f f' : B -> C) (g g' : A -> B) : compose (A:=def A) (map_def f f') (map_def g g') = map_def (compose f g) (compose f' g'). Proof. reflexivity. Qed. -Lemma map_def_id {t : Set} x : map_def (@id t) (@id t) x = id x. +Lemma map_def_id {t} x : map_def (@id t) (@id t) x = id x. Proof. now destruct x. Qed. Hint Rewrite @map_def_id @map_id : map. -Lemma map_def_spec {A B : Set} (P P' : A -> Prop) (f f' g g' : A -> B) (x : def A) : +Lemma map_def_spec {A B} (P P' : A -> Prop) (f f' g g' : A -> B) (x : def A) : P' x.(dbody) -> P x.(dtype) -> (forall x, P x -> f x = g x) -> (forall x, P' x -> f' x = g' x) -> map_def f f' x = map_def g g' x. @@ -675,7 +678,7 @@ Proof. rewrite !H1 // !H2 //. Qed. -Lemma case_brs_map_spec {A B : Set} {P : A -> Prop} {l} {f g : A -> B} : +Lemma case_brs_map_spec {A B} {P : A -> Prop} {l} {f g : A -> B} : tCaseBrsProp P l -> (forall x, P x -> f x = g x) -> map (on_snd f) l = map (on_snd g) l. Proof. @@ -684,7 +687,7 @@ Proof. apply on_snd_eq_spec; eauto. Qed. -Lemma tfix_map_spec {A B : Set} {P P' : A -> Prop} {l} {f f' g g' : A -> B} : +Lemma tfix_map_spec {A B} {P P' : A -> Prop} {l} {f f' g g' : A -> B} : tFixProp P P' l -> (forall x, P x -> f x = g x) -> (forall x, P' x -> f' x = g' x) -> map (map_def f f') l = map (map_def g g') l. @@ -696,7 +699,7 @@ Proof. Qed. -Lemma map_def_test_spec {A B : Set} +Lemma map_def_test_spec {A B} (P P' : A -> Prop) (p p' : A -> bool) (f f' g g' : A -> B) (x : def A) : P x.(dtype) -> P' x.(dbody) -> (forall x, P x -> p x -> f x = g x) -> (forall x, P' x -> p' x -> f' x = g' x) -> @@ -708,7 +711,7 @@ Proof. rewrite !H1 // !H2 //; intuition auto. Qed. -Lemma case_brs_forallb_map_spec {A B : Set} {P : A -> Prop} {p : A -> bool} +Lemma case_brs_forallb_map_spec {A B} {P : A -> Prop} {p : A -> bool} {l} {f g : A -> B} : tCaseBrsProp P l -> forallb (test_snd p) l -> @@ -721,7 +724,7 @@ Proof. eapply on_snd_test_spec; eauto. Qed. -Lemma tfix_forallb_map_spec {A B : Set} {P P' : A -> Prop} {p p'} {l} {f f' g g' : A -> B} : +Lemma tfix_forallb_map_spec {A B} {P P' : A -> Prop} {p p'} {l} {f f' g g' : A -> B} : tFixProp P P' l -> forallb (test_def p p') l -> (forall x, P x -> p x -> f x = g x) -> diff --git a/template-coq/theories/BasicAst.v b/template-coq/theories/BasicAst.v index 7d7c17aa4..352f5c356 100644 --- a/template-coq/theories/BasicAst.v +++ b/template-coq/theories/BasicAst.v @@ -17,7 +17,7 @@ Definition projection : Set := inductive * nat (* params *) * nat (* argument *) (** Parametrized by term because term is not yet defined *) -Record def (term : Set) : Set := mkdef { +Record def term := mkdef { dname : name; (* the name **) dtype : term; dbody : term; (* the body (a lambda term). Note, this may mention other (mutually-defined) names **) @@ -28,7 +28,7 @@ Arguments dtype {term} _. Arguments dbody {term} _. Arguments rarg {term} _. -Definition mfixpoint (term : Set) : Set := +Definition mfixpoint term := list (def term). (** The kind of a cast *) diff --git a/template-coq/theories/Constants.v b/template-coq/theories/Constants.v index c6b51ca85..8cb416027 100644 --- a/template-coq/theories/Constants.v +++ b/template-coq/theories/Constants.v @@ -42,9 +42,6 @@ Register MetaCoq.Template.TemplateMonad.Common.existT_typed_term as metacoq.sigm (* Ast *) -Register MetaCoq.Template.utils.NEL.sing as metacoq.nel.sing. -Register MetaCoq.Template.utils.NEL.cons as metacoq.nel.cons. - 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. @@ -78,7 +75,8 @@ Register MetaCoq.Template.Universes.InType as metacoq.ast.InType. Register MetaCoq.Template.Universes.ConstraintType.Le as metacoq.ast.constraints.Le. Register MetaCoq.Template.Universes.ConstraintType.Lt as metacoq.ast.constraints.Lt. Register MetaCoq.Template.Universes.ConstraintType.Eq as metacoq.ast.constraints.Eq. -Register MetaCoq.Template.Universes.Universe.make'' as metacoq.ast.universe.make. +Register MetaCoq.Template.Universes.Universe.from_kernel_repr as metacoq.ast.universe.from_kernel_repr. +Register MetaCoq.Template.Universes.Universe.to_kernel_repr as metacoq.ast.universe.to_kernel_repr. Register MetaCoq.Template.Universes.LevelSetProp.of_list as metacoq.ast.universe.of_list. Register MetaCoq.Template.Universes.Level.t as metacoq.ast.level.t. Register MetaCoq.Template.Universes.Level.Level as metacoq.ast.level.Level. @@ -86,6 +84,15 @@ Register MetaCoq.Template.Universes.Level.lProp as metacoq.ast.level.lProp. Register MetaCoq.Template.Universes.Level.lSet as metacoq.ast.level.lSet. Register MetaCoq.Template.Universes.Level.Var as metacoq.ast.level.Var. +Register MetaCoq.Template.Universes.NoPropLevel.Level as metacoq.ast.noproplevel.Level. +Register MetaCoq.Template.Universes.NoPropLevel.lSet as metacoq.ast.noproplevel.lSet. +Register MetaCoq.Template.Universes.NoPropLevel.Var as metacoq.ast.noproplevel.Var. +Register MetaCoq.Template.Universes.UnivExpr.lProp as metacoq.ast.univexpr.prop. +Register MetaCoq.Template.Universes.UnivExpr.npe as metacoq.ast.univexpr.npe. + +Register MetaCoq.Template.Universes.UnivExprSet.Mkt as metacoq.ast.univexprset.mkt. +Register MetaCoq.Template.Universes.Universe.Build_t as metacoq.ast.universe.build. + Register MetaCoq.Template.Universes.Variance.t as metacoq.ast.variance.t. Register MetaCoq.Template.Universes.Variance.Irrelevant as metacoq.ast.variance.Irrelevant. Register MetaCoq.Template.Universes.Variance.Covariant as metacoq.ast.variance.Covariant. @@ -105,7 +112,7 @@ Register MetaCoq.Template.Universes.AUContext.t as metacoq.ast.AUContext.t. Register MetaCoq.Template.Universes.AUContext.make as metacoq.ast.AUContext.make. Register MetaCoq.Template.Universes.LevelSet.t_ as metacoq.ast.LevelSet.t. -Register MetaCoq.Template.Universes.make_univ_constraint as metacoq.ast.make_univ_constraint. +Register MetaCoq.Template.Universes.UnivConstraint.make as metacoq.ast.make_univ_constraint. Register MetaCoq.Template.common.uGraph.init_graph as metacoq.ast.graph.init. (* FIXME wrong! *) diff --git a/template-coq/theories/Environment.v b/template-coq/theories/Environment.v index ac36dfbff..f007f5932 100644 --- a/template-coq/theories/Environment.v +++ b/template-coq/theories/Environment.v @@ -8,13 +8,13 @@ Set Asymmetric Patterns. Module Type Term. - Parameter (term : Set). + Parameter (term : Type). Parameter (tRel : nat -> term). - Parameter (tSort : universe -> term). + Parameter (tSort : Universe.t -> term). Parameter (tProd : name -> term -> term -> term). Parameter (tLetIn : name -> term -> term -> term -> term). - Parameter (tInd : inductive -> universe_instance -> term). + Parameter (tInd : inductive -> Instance.t -> term). Parameter (mkApps : term -> list term -> term). @@ -145,4 +145,4 @@ End Environment. Module Type EnvironmentSig (T : Term). Include Environment T. -End EnvironmentSig. \ No newline at end of file +End EnvironmentSig. diff --git a/template-coq/theories/EnvironmentTyping.v b/template-coq/theories/EnvironmentTyping.v index 8d73f13af..e635e1cf5 100644 --- a/template-coq/theories/EnvironmentTyping.v +++ b/template-coq/theories/EnvironmentTyping.v @@ -86,7 +86,7 @@ Module Lookup (T : Term) (E : EnvironmentSig T). be ensured if we added [global_constraints] as well as a coercion, as it would forget the extension's constraints. *) - Definition global_constraints (Σ : global_env) : constraints := + Definition global_constraints (Σ : global_env) : ConstraintSet.t := fold_right (fun decl ctrs => ConstraintSet.union (monomorphic_constraints_decl decl.2) ctrs ) ConstraintSet.empty Σ. @@ -97,18 +97,18 @@ Module Lookup (T : Term) (E : EnvironmentSig T). Definition global_ext_levels (Σ : global_env_ext) : LevelSet.t := LevelSet.union (levels_of_udecl (snd Σ)) (global_levels Σ.1). - Definition global_ext_constraints (Σ : global_env_ext) : constraints := + Definition global_ext_constraints (Σ : global_env_ext) : ConstraintSet.t := ConstraintSet.union (constraints_of_udecl (snd Σ)) (global_constraints Σ.1). - Coercion global_ext_constraints : global_env_ext >-> constraints. + Coercion global_ext_constraints : global_env_ext >-> ConstraintSet.t. Definition global_ext_uctx (Σ : global_env_ext) : ContextSet.t := (global_ext_levels Σ, global_ext_constraints Σ). - Lemma prop_global_ext_levels Σ : LevelSet.In Level.prop (global_ext_levels Σ). + Lemma prop_global_ext_levels Σ : LevelSet.In Level.lProp (global_ext_levels Σ). Proof. destruct Σ as [Σ φ]; cbn. apply LevelSetFact.union_3. cbn -[global_levels]; clear φ. @@ -120,7 +120,7 @@ Module Lookup (T : Term) (E : EnvironmentSig T). (** Check that [uctx] instantiated at [u] is consistent with the current universe graph. *) - Definition consistent_instance `{checker_flags} (lvs : LevelSet.t) (φ : constraints) uctx (u : universe_instance) := + Definition consistent_instance `{checker_flags} (lvs : LevelSet.t) (φ : ConstraintSet.t) uctx (u : Instance.t) := match uctx with | Monomorphic_ctx c => List.length u = 0 | Polymorphic_ctx c => @@ -177,7 +177,7 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T). fun Σ Γ t T => match T with | Some T => P Σ Γ t T - | None => { s : universe & P Σ Γ t (tSort s) } + | None => { s : Universe.t & P Σ Γ t (tSort s) } end. Definition on_local_decl (P : context -> term -> option term -> Type) Γ d := @@ -255,7 +255,7 @@ Module DeclarationTyping (T : Term) (E : EnvironmentSig T) (** For well-formedness of inductive declarations we need a way to check that a assumptions of a given context is typable in a sort [u]. *) - Fixpoint type_local_ctx Σ (Γ Δ : context) (u : universe) : Type := + Fixpoint type_local_ctx Σ (Γ Δ : context) (u : Universe.t) : Type := match Δ with | [] => True | {| decl_body := None ; decl_type := t |} :: Δ => @@ -373,7 +373,7 @@ Module DeclarationTyping (T : Term) (E : EnvironmentSig T) { (** The type of the inductive must be an arity, sharing the same params as the rest of the block, and maybe having a context of indices. *) ind_indices : context; - ind_sort : universe; + ind_sort : Universe.t; ind_arity_eq : idecl.(ind_type) = it_mkProd_or_LetIn mdecl.(ind_params) (it_mkProd_or_LetIn ind_indices (tSort ind_sort)); @@ -382,7 +382,7 @@ Module DeclarationTyping (T : Term) (E : EnvironmentSig T) onArity : on_type Σ [] idecl.(ind_type); (** There is a sort bounding the indices and arguments for each constructor *) - ind_ctors_sort : list universe; + ind_ctors_sort : list Universe.t; (** Constructors are well-typed *) onConstructors : @@ -594,4 +594,4 @@ Module DeclarationTyping (T : Term) (E : EnvironmentSig T) end. End wf_local_size. -End DeclarationTyping. \ No newline at end of file +End DeclarationTyping. diff --git a/template-coq/theories/Induction.v b/template-coq/theories/Induction.v index 9f8335c27..29eb2c017 100644 --- a/template-coq/theories/Induction.v +++ b/template-coq/theories/Induction.v @@ -139,7 +139,7 @@ Qed. Definition tCaseBrsType {A} (P : A -> Type) (l : list (nat * A)) := All (fun x => P (snd x)) l. -Definition tFixType {A : Set} (P P' : A -> Type) (m : mfixpoint A) := +Definition tFixType {A} (P P' : A -> Type) (m : mfixpoint A) := All (fun x : def A => P x.(dtype) * P' x.(dbody))%type m. Lemma term_forall_list_rect : @@ -178,4 +178,4 @@ Proof. destruct arg; constructor; [|apply aux_arg]; try split; apply auxt end. -Defined. \ No newline at end of file +Defined. diff --git a/template-coq/theories/LiftSubst.v b/template-coq/theories/LiftSubst.v index 1e960f513..3c467b6df 100644 --- a/template-coq/theories/LiftSubst.v +++ b/template-coq/theories/LiftSubst.v @@ -196,7 +196,7 @@ Qed. Hint Resolve -> on_snd_eq_id_spec : all. Hint Resolve -> on_snd_eq_spec : all. -Lemma map_def_eq_spec {A B : Set} (f f' g g' : A -> B) (x : def A) : +Lemma map_def_eq_spec {A B} (f f' g g' : A -> B) (x : def A) : f (dtype x) = g (dtype x) -> f' (dbody x) = g' (dbody x) -> map_def f f' x = map_def g g' x. @@ -205,7 +205,7 @@ Proof. Qed. Hint Resolve map_def_eq_spec : all. -Lemma map_def_id_spec {A : Set} (f f' : A -> A) (x : def A) : +Lemma map_def_id_spec {A} (f f' : A -> A) (x : def A) : f (dtype x) = (dtype x) -> f' (dbody x) = (dbody x) -> map_def f f' x = x. @@ -512,7 +512,7 @@ Proof. elim t using term_forall_list_ind; intros; try easy; rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; unfold test_def in *; - simpl closed in *; try solve [simpl lift; simpl closed; f_equal; auto; toProp; solve_all]; try easy. + simpl closed in *; try solve [simpl lift; simpl closed; f_equal; auto; rtoProp; solve_all]; try easy. - rewrite lift_rel_lt; auto. revert H. elim (Nat.ltb_spec n0 k); intros; try easy. - simpl lift. f_equal. solve_all. unfold test_def in b. toProp. solve_all. @@ -525,7 +525,7 @@ Proof. elim t using term_forall_list_ind; intros; try lia; rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; simpl closed in *; unfold test_snd, test_def in *; - try solve [(try f_equal; simpl; repeat (toProp; solve_all); eauto)]. + try solve [(try f_equal; simpl; repeat (rtoProp; solve_all); eauto)]. - elim (ltb_spec n k'); auto. intros. apply ltb_lt in H. lia. diff --git a/template-coq/theories/Pretty.v b/template-coq/theories/Pretty.v index 9d464efa8..424c361a0 100644 --- a/template-coq/theories/Pretty.v +++ b/template-coq/theories/Pretty.v @@ -42,7 +42,7 @@ Section print_term. string_of_level l1 ++ " " ++ print_constraint_type d ++ " " ++ string_of_level l2) " /\ " (ConstraintSet.elements t). - Definition print_def {A : Set} (f : A -> string) (g : A -> string) (def : def A) := + Definition print_def {A} (f : A -> string) (g : A -> string) (def : def A) := string_of_name (dname def) ++ " { struct " ++ string_of_nat (rarg def) ++ " }" ++ " : " ++ f (dtype def) ++ " := " ++ nl ++ g (dbody def). diff --git a/template-coq/theories/TemplateMonad/Common.v b/template-coq/theories/TemplateMonad/Common.v index 775b0366d..df8f3037f 100644 --- a/template-coq/theories/TemplateMonad/Common.v +++ b/template-coq/theories/TemplateMonad/Common.v @@ -34,7 +34,7 @@ Record TMInstance@{t u r} := (* Quote the body of a definition or inductive. Its name need not be fully quaified *) ; tmQuoteInductive : kername -> TemplateMonad mutual_inductive_body -; tmQuoteUniverses : TemplateMonad constraints +; tmQuoteUniverses : TemplateMonad ConstraintSet.t ; tmQuoteConstant : kername -> bool (* bypass opacity? *) -> TemplateMonad constant_body (* unquote before making the definition *) (* FIXME take an optional universe context as well *) diff --git a/template-coq/theories/TemplateMonad/Core.v b/template-coq/theories/TemplateMonad/Core.v index bb72a0744..3bd27802e 100644 --- a/template-coq/theories/TemplateMonad/Core.v +++ b/template-coq/theories/TemplateMonad/Core.v @@ -45,7 +45,7 @@ Cumulative Inductive TemplateMonad@{t u} : Type@{t} -> Prop := | tmQuoteRec : forall {A:Type@{t}}, A -> TemplateMonad program (* Quote the body of a definition or inductive. Its name need not be fully qualified *) | tmQuoteInductive : qualid -> TemplateMonad mutual_inductive_body -| tmQuoteUniverses : TemplateMonad constraints +| tmQuoteUniverses : TemplateMonad ConstraintSet.t | tmQuoteConstant : qualid -> bool (* bypass opacity? *) -> TemplateMonad constant_body (* unquote before making the definition *) (* FIXME take an optional universe context as well *) diff --git a/template-coq/theories/TemplateMonad/Extractable.v b/template-coq/theories/TemplateMonad/Extractable.v index 944bc3242..b63672517 100644 --- a/template-coq/theories/TemplateMonad/Extractable.v +++ b/template-coq/theories/TemplateMonad/Extractable.v @@ -52,7 +52,7 @@ Cumulative Inductive TM@{t} : Type@{t} -> Type := | tmQuoteConstant (nm : kername) (bypass_opacity : bool) : TM constant_body -| tmQuoteUniverses : TM constraints +| tmQuoteUniverses : TM ConstraintSet.t (* unquote before making the definition *) (* FIXME take an optional universe context as well *) diff --git a/template-coq/theories/Typing.v b/template-coq/theories/Typing.v index b975fa593..5d0f59b5d 100644 --- a/template-coq/theories/Typing.v +++ b/template-coq/theories/Typing.v @@ -272,7 +272,7 @@ Lemma red1_ind_all : (forall (Γ : context) (i : nat) (body : term), option_map decl_body (nth_error Γ i) = Some (Some body) -> P Γ (tRel i) ((lift0 (S i)) body)) -> - (forall (Γ : context) (ind : inductive) (pars c : nat) (u : universe_instance) (args : list term) + (forall (Γ : context) (ind : inductive) (pars c : nat) (u : Instance.t) (args : list term) (p : term) (brs : list (nat * term)), P Γ (tCase (ind, pars) p (mkApps (tConstruct ind c u) args) brs) (iota_red pars c args brs)) -> @@ -290,9 +290,9 @@ Lemma red1_ind_all : (forall (Γ : context) (c : ident) (decl : constant_body) (body : term), declared_constant Σ c decl -> - forall u : universe_instance, cst_body decl = Some body -> P Γ (tConst c u) (subst_instance_constr u body)) -> + forall u : Instance.t, cst_body decl = Some body -> P Γ (tConst c u) (subst_instance_constr u body)) -> - (forall (Γ : context) (i : inductive) (pars narg : nat) (args : list term) (k : nat) (u : universe_instance) + (forall (Γ : context) (i : inductive) (pars narg : nat) (args : list term) (k : nat) (u : Instance.t) (arg : term), nth_error args (pars + narg) = Some arg -> P Γ (tProj (i, pars, narg) (mkApps (tConstruct i k u) args)) arg) -> @@ -441,7 +441,7 @@ Inductive red Σ Γ M : term -> Type := Definition R_universe_instance R := fun u u' => Forall2 R (List.map Universe.make u) (List.map Universe.make u'). -Inductive eq_term_upto_univ (Re Rle : universe -> universe -> Prop) : term -> term -> Type := +Inductive eq_term_upto_univ (Re Rle : Universe.t -> Universe.t -> Prop) : term -> term -> Type := | eq_Rel n : eq_term_upto_univ Re Rle (tRel n) (tRel n) @@ -554,10 +554,10 @@ Fixpoint strip_casts t := | tRel _ | tVar _ | tSort _ | tConst _ _ | tInd _ _ | tConstruct _ _ _ => t end. -Definition eq_term_nocast `{checker_flags} (φ : constraints) (t u : term) := +Definition eq_term_nocast `{checker_flags} (φ : ConstraintSet.t) (t u : term) := eq_term φ (strip_casts t) (strip_casts u). -Definition leq_term_nocast `{checker_flags} (φ : constraints) (t u : term) := +Definition leq_term_nocast `{checker_flags} (φ : ConstraintSet.t) (t u : term) := leq_term φ (strip_casts t) (strip_casts u). (** ** Utilities for typing *) @@ -1122,10 +1122,10 @@ Lemma typing_ind_env `{cf : checker_flags} : P Σ Γ (tSort (Universe.make l)) (tSort (Universe.super l))) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (c : term) (k : cast_kind) - (t : term) (s : universe), + (t : term) (s : Universe.t), Σ ;;; Γ |- t : tSort s -> P Σ Γ t (tSort s) -> Σ ;;; Γ |- c : t -> P Σ Γ c t -> P Σ Γ (tCast c k t) t) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : name) (t b : term) (s1 s2 : universe), + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : name) (t b : term) (s1 s2 : Universe.t), All_local_env_over typing Pdecl Σ Γ wfΓ -> Σ ;;; Γ |- t : tSort s1 -> P Σ Γ t (tSort s1) -> @@ -1133,14 +1133,14 @@ Lemma typing_ind_env `{cf : checker_flags} : P Σ (Γ,, vass n t) b (tSort s2) -> P Σ Γ (tProd n t b) (tSort (Universe.sort_of_product s1 s2))) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : name) (t b : term) - (s1 : universe) (bty : term), + (s1 : Universe.t) (bty : term), All_local_env_over typing Pdecl Σ Γ wfΓ -> Σ ;;; Γ |- t : tSort s1 -> P Σ Γ t (tSort s1) -> Σ ;;; Γ,, vass n t |- b : bty -> P Σ (Γ,, vass n t) b bty -> P Σ Γ (tLambda n t b) (tProd n t bty)) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : name) (b b_ty b' : term) - (s1 : universe) (b'_ty : term), + (s1 : Universe.t) (b'_ty : term), All_local_env_over typing Pdecl Σ Γ wfΓ -> Σ ;;; Γ |- b_ty : tSort s1 -> P Σ Γ b_ty (tSort s1) -> diff --git a/template-coq/theories/TypingWf.v b/template-coq/theories/TypingWf.v index 72ea97680..b6082eb73 100644 --- a/template-coq/theories/TypingWf.v +++ b/template-coq/theories/TypingWf.v @@ -300,7 +300,7 @@ Proof. Qed. Lemma declared_projection_wf {cf:checker_flags}: - forall (Σ : global_env) (p : projection) (u : universe_instance) + forall (Σ : global_env) (p : projection) (u : Instance.t) (mdecl : mutual_inductive_body) (idecl : one_inductive_body) (pdecl : ident * term), declared_projection Σ mdecl idecl p pdecl -> Forall_decls_typing (fun (_ : global_env_ext) (_ : context) (t T : term) => Ast.wf t /\ Ast.wf T) Σ -> diff --git a/template-coq/theories/UnivSubst.v b/template-coq/theories/UnivSubst.v index 298095f59..28e40b0b2 100644 --- a/template-coq/theories/UnivSubst.v +++ b/template-coq/theories/UnivSubst.v @@ -5,52 +5,22 @@ From MetaCoq Require Import utils Ast AstUtils Induction LiftSubst. Require Import String Lia. Local Open Scope string_scope. Set Asymmetric Patterns. -Require Import ssreflect ssrbool. (** * Universe substitution - *WIP* - Substitution of universe levels for universe level variables, used to implement universe polymorphism. *) -(** Substitutable type *) - -Class UnivSubst A := subst_instance : universe_instance -> A -> A. - - -Instance subst_instance_level : UnivSubst Level.t := - fun u l => match l with - | Level.lProp | Level.lSet | Level.Level _ => l - | Level.Var n => List.nth n u Level.lSet - end. - -Instance subst_instance_cstr : UnivSubst univ_constraint := - fun u c => (subst_instance_level u c.1.1, c.1.2, subst_instance_level u c.2). - -Instance subst_instance_cstrs : UnivSubst constraints := - fun u ctrs => ConstraintSet.fold (fun c => ConstraintSet.add (subst_instance_cstr u c)) - ctrs ConstraintSet.empty. - -Instance subst_instance_level_expr : UnivSubst Universe.Expr.t := - fun u e => (subst_instance_level u e.1, e.2). - -Instance subst_instance_univ : UnivSubst universe := - fun u s => NEL.map (subst_instance_level_expr u) s. - -Instance subst_instance_instance : UnivSubst universe_instance := - fun u u' => List.map (subst_instance_level u) u'. - Instance subst_instance_constr : UnivSubst term := fix subst_instance_constr u c {struct c} : term := match c with - | tRel _ | tVar _ => c + | tRel _ | tVar _ => c + | tEvar ev args => tEvar ev (List.map (subst_instance_constr u) args) | tSort s => tSort (subst_instance_univ u s) | tConst c u' => tConst c (subst_instance_instance u u') | tInd i u' => tInd i (subst_instance_instance u u') | tConstruct ind k u' => tConstruct ind k (subst_instance_instance u u') - | tEvar ev args => tEvar ev (List.map (subst_instance_constr u) args) | tLambda na T M => tLambda na (subst_instance_constr u T) (subst_instance_constr u M) | tApp f v => tApp (subst_instance_constr u f) (List.map (subst_instance_constr u) v) | tProd na A B => tProd na (subst_instance_constr u A) (subst_instance_constr u B) @@ -69,14 +39,17 @@ Instance subst_instance_constr : UnivSubst term := tCoFix mfix' idx end. -Instance subst_instance_context : UnivSubst context := - fun u c => AstUtils.map_context (subst_instance_constr u) c. +Instance subst_instance_decl : UnivSubst context_decl + := map_decl ∘ subst_instance_constr. + +Instance subst_instance_context : UnivSubst context + := map_context ∘ subst_instance_constr. Lemma lift_subst_instance_constr u c n k : lift n k (subst_instance_constr u c) = subst_instance_constr u (lift n k c). Proof. induction c in k |- * using term_forall_list_ind; simpl; auto; - rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; + rewrite ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; try solve [f_equal; eauto; solve_all; eauto]. elim (Nat.leb k n0); reflexivity. @@ -96,7 +69,8 @@ Lemma subst_instance_constr_it_mkProd_or_LetIn u ctx t : it_mkProd_or_LetIn (subst_instance_context u ctx) (subst_instance_constr u t). Proof. induction ctx in u, t |- *; simpl; try congruence. - rewrite IHctx /= /mkProd_or_LetIn /=. f_equal. destruct (decl_body a); eauto. + rewrite IHctx. unfold mkProd_or_LetIn; cbn. f_equal. + destruct (decl_body a); eauto. Qed. Lemma subst_instance_context_length u ctx @@ -111,7 +85,7 @@ Lemma subst_subst_instance_constr u c N k : subst_instance_constr u (subst N k c). Proof. induction c in k |- * using term_forall_list_ind; simpl; auto; - rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; + rewrite ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; try solve [f_equal; eauto; solve_all; eauto]. elim (Nat.leb k n). rewrite nth_error_map. @@ -130,148 +104,45 @@ Proof. solve_all. now destruct H as [n [-> _]]. Qed. -Section Closedu. - (** Tests that the term is closed over [k] universe variables *) - Context (k : nat). - - Definition closedu_level (l : Level.t) := - match l with - | Level.Var n => n true - end. - - Definition closedu_level_expr (s : Universe.Expr.t) := - closedu_level (fst s). - - Definition closedu_universe (u : universe) := - forallb closedu_level_expr u. - - Definition closedu_instance (u : universe_instance) := - forallb closedu_level u. - - Fixpoint closedu (t : term) : bool := +(** Tests that the term is closed over [k] universe variables *) +Fixpoint closedu (k : nat) (t : term) : bool := match t with - | tSort univ => closedu_universe univ - | tInd _ u => closedu_instance u - | tConstruct _ _ u => closedu_instance u - | tConst _ u => closedu_instance u + | tSort univ => closedu_universe k univ + | tInd _ u => closedu_instance k u + | tConstruct _ _ u => closedu_instance k u + | tConst _ u => closedu_instance k u | tRel i => true - | tEvar ev args => forallb closedu args - | tLambda _ T M | tProd _ T M => closedu T && closedu M - | tApp u v => closedu u && forallb (closedu) v - | tCast c kind t => closedu c && closedu t - | tLetIn na b t b' => closedu b && closedu t && closedu b' + | tEvar ev args => forallb (closedu k) args + | tLambda _ T M | tProd _ T M => closedu k T && closedu k M + | tApp u v => closedu k u && forallb (closedu k) v + | tCast c kind t => closedu k c && closedu k t + | tLetIn na b t b' => closedu k b && closedu k t && closedu k b' | tCase ind p c brs => - let brs' := forallb (test_snd (closedu)) brs in - closedu p && closedu c && brs' - | tProj p c => closedu c + let brs' := forallb (test_snd (closedu k)) brs in + closedu k p && closedu k c && brs' + | tProj p c => closedu k c | tFix mfix idx => - forallb (test_def closedu closedu ) mfix + forallb (test_def (closedu k) (closedu k)) mfix | tCoFix mfix idx => - forallb (test_def closedu closedu) mfix + forallb (test_def (closedu k) (closedu k)) mfix | x => true end. -End Closedu. - -(** Universe-closed terms are unaffected by universe substitution. *) - -Section UniverseClosedSubst. - Lemma closedu_subst_instance_level u t - : closedu_level 0 t -> subst_instance_level u t = t. - Proof. - destruct t => /=; auto. - move/Nat.ltb_spec0. intro H. inversion H. - Qed. - - Lemma closedu_subst_instance_level_expr u t - : closedu_level_expr 0 t -> subst_instance_level_expr u t = t. - Proof. - destruct t as [l n]. - rewrite /closedu_level_expr /subst_instance_level_expr /=. - move/(closedu_subst_instance_level u) => //. congruence. - Qed. - - Lemma closedu_subst_instance_univ u t - : closedu_universe 0 t -> subst_instance_univ u t = t. - Proof. - rewrite /closedu_universe /subst_instance_univ => H. - pose proof (proj1 (forallb_forall _ t) H) as HH; clear H. - induction t; cbn; f_equal. - 1-2: now apply closedu_subst_instance_level_expr, HH; cbn. - apply IHt. intros x Hx; apply HH. now right. - Qed. - Hint Resolve closedu_subst_instance_level_expr closedu_subst_instance_level closedu_subst_instance_univ : terms. - - Lemma closedu_subst_instance_instance u t - : closedu_instance 0 t -> subst_instance_instance u t = t. - Proof. - rewrite /closedu_instance /subst_instance_instance => H. solve_all. - now apply (closedu_subst_instance_level u). - Qed. - Hint Resolve closedu_subst_instance_instance : terms. - - Lemma closedu_subst_instance_constr u t - : closedu 0 t -> subst_instance_constr u t = t. - Proof. - induction t in |- * using term_forall_list_ind; simpl; auto; intros H'; - rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; - try f_equal; eauto with terms; unfold test_def in *; - try solve [f_equal; eauto; repeat (toProp; solve_all)]. - Qed. -End UniverseClosedSubst. - -Section SubstInstanceClosed. - (** Substitution of a universe-closed instance of the right size - produces a universe-closed term. *) - - Context (u : universe_instance) (Hcl : closedu_instance 0 u). - - Lemma subst_instance_level_closedu t - : closedu_level #|u| t -> closedu_level 0 (subst_instance_level u t). - Proof. - destruct t => /=; auto. - move/Nat.ltb_spec0. intro H. - red in Hcl. unfold closedu_instance in Hcl. - eapply forallb_nth in Hcl; eauto. destruct Hcl as [x [Hn Hx]]. now rewrite -> Hn. - Qed. - - Lemma subst_instance_level_expr_closedu t : - closedu_level_expr #|u| t -> closedu_level_expr 0 (subst_instance_level_expr u t). - Proof. - destruct t as [l n]. - rewrite /closedu_level_expr /subst_instance_level_expr /=. - move/(subst_instance_level_closedu) => //. - Qed. - - Lemma subst_instance_univ_closedu t - : closedu_universe #|u| t -> closedu_universe 0 (subst_instance_univ u t). - Proof. - rewrite /closedu_universe /subst_instance_univ => H. - eapply (forallb_Forall (closedu_level_expr #|u|)) in H; auto. - unfold universe_coercion; rewrite NEL.map_to_list forallb_map. - eapply Forall_forallb; eauto. - now move=> x /(subst_instance_level_expr_closedu). - Qed. - Hint Resolve subst_instance_level_expr_closedu subst_instance_level_closedu subst_instance_univ_closedu : terms. - - Lemma subst_instance_instance_closedu t : - closedu_instance #|u| t -> closedu_instance 0 (subst_instance_instance u t). - Proof. - rewrite /closedu_instance /subst_instance_instance => H. - eapply (forallb_Forall (closedu_level #|u|)) in H; auto. - rewrite forallb_map. eapply Forall_forallb; eauto. - simpl. now move=> x /(subst_instance_level_closedu). - Qed. - Hint Resolve subst_instance_instance_closedu : terms. +Lemma closedu_subst_instance_constr u t + : closedu 0 t -> subst_instance_constr u t = t. +Proof. + induction t in |- * using term_forall_list_ind; simpl; auto; intros H'; + rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; + try f_equal; eauto with substu; unfold test_def in *; + try solve [f_equal; eauto; repeat (rtoProp; solve_all)]. +Qed. - Lemma subst_instance_constr_closedu t : - closedu #|u| t -> closedu 0 (subst_instance_constr u t). - Proof. - induction t in |- * using term_forall_list_ind; simpl; auto; intros H'; +Lemma subst_instance_constr_closedu (u : Instance.t) (Hcl : closedu_instance 0 u) t : + closedu #|u| t -> closedu 0 (subst_instance_constr u t). +Proof. + induction t in |- * using term_forall_list_ind; simpl; auto; intros H'; rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length, ?forallb_map; - try f_equal; auto with terms; - unfold test_def, map_def, compose in *; - try solve [f_equal; eauto; repeat (toProp; solve_all); intuition auto with terms]. - Qed. -End SubstInstanceClosed. + try f_equal; auto with substu; + unfold test_def, map_def, compose in *; + try solve [f_equal; eauto; repeat (rtoProp; solve_all); intuition auto with substu]. +Qed. diff --git a/template-coq/theories/Universes.v b/template-coq/theories/Universes.v index 8aaedd70c..dcd00a38f 100644 --- a/template-coq/theories/Universes.v +++ b/template-coq/theories/Universes.v @@ -1,20 +1,29 @@ -From Coq Require Import Ascii String ZArith List Bool. -From Coq Require Import MSetWeakList MSetFacts MSetProperties RelationClasses. +From Coq Require Import Ascii String ZArith List Bool Lia. +From Coq Require Import MSetList MSetFacts MSetProperties RelationClasses. From MetaCoq.Template Require Import utils BasicAst config. Import ListNotations. -Declare Scope univ_scope. -Delimit Scope univ_scope with u. +Local Open Scope Z_scope. + +(** * Valuations *) + +(** A valuation is a universe level (nat) given for each + universe variable (Level.t). + It is >= for polymorphic universes and > 0 for monomorphic universes. *) +Record valuation := + { valuation_mono : string -> positive ; + valuation_poly : nat -> nat }. + +Class Evaluable (A : Type) := val : valuation -> A -> Z. Module Level. - Inductive t : Set := + Inductive t_ : Set := | lProp | lSet | Level (_ : string) | Var (_ : nat) (* these are debruijn indices *). - Definition set : t := lSet. - Definition prop : t := lProp. + Definition t := t_. Definition is_small (x : t) := match x with @@ -40,6 +49,15 @@ Module Level. | _ => false end. + Global Instance Evaluable : Evaluable t + := fun v l => match l with + | lProp => -1 + | lSet => 0 + | Level s => Z.pos (v.(valuation_mono) s) + | Var x => Z.of_nat (v.(valuation_poly) x) + end. + + Definition compare (l1 l2 : t) : comparison := match l1, l2 with | lProp, lProp => Eq @@ -53,41 +71,88 @@ Module Level. | _, Level _ => Gt | Var n, Var m => Nat.compare n m end. - (** Comparison function *) + + 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. - Definition equal (l1 l2 : Level.t) : bool + Inductive lt_ : t -> t -> Prop := + | ltPropSet : lt_ lProp lSet + | ltPropLevel s : lt_ lProp (Level s) + | ltPropVar n : lt_ lProp (Var n) + | ltSetLevel s : lt_ lSet (Level s) + | ltSetVar n : lt_ lSet (Var n) + | ltLevelLevel s s' : string_lt s s' -> lt_ (Level s) (Level s') + | ltLevelVar s n : lt_ (Level s) (Var n) + | ltVarVar n n' : Nat.lt n n' -> lt_ (Var n) (Var n'). + + Definition lt := lt_. + + Definition lt_strorder : StrictOrder lt. + Proof. + constructor. + - intros [| | |] X; inversion X. + eapply string_lt_irreflexive; tea. + eapply Nat.lt_irrefl; tea. + - intros [| | |] [| | |] [| | |] X1 X2; + inversion X1; inversion X2; constructor. + eapply transitive_string_lt; tea. + etransitivity; tea. + Qed. + + Definition lt_compat : Proper (eq ==> eq ==> iff) lt. + Proof. + intros x y e z t e'. unfold eq in *; subst. reflexivity. + Qed. + + Definition compare_spec : + forall x y : t, CompareSpec (x = y) (lt x y) (lt y x) (compare x y). + Proof. + intros [] []; repeat constructor. + - eapply CompareSpec_Proper. + 5: eapply CompareSpec_string. 4: reflexivity. + all: split; [now inversion 1|]. now intros ->. + all: intro; now constructor. + - eapply CompareSpec_Proper. + 5: eapply Nat.compare_spec. 4: reflexivity. + all: split; [now inversion 1|]. now intros ->. + all: intro; now constructor. + Qed. + + (* Bonus *) + Definition eqb (l1 l2 : Level.t) : bool := match compare l1 l2 with Eq => true | _ => false end. - Lemma equal_refl : - forall l, equal l l. + Global Instance eqb_refl : Reflexive eqb. Proof. - intros []. - all: unfold equal. - all: simpl. - 1-2: reflexivity. + intros []; unfold eqb; cbnr. - rewrite (ssreflect.iffRL (string_compare_eq s s)). all: auto. - rewrite Nat.compare_refl. reflexivity. Qed. + Lemma eqb_spec l l' : reflect (eq l l') (eqb l l'). + 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; + try reflexivity; try discriminate; try congruence. + all: inversion HH; subst; now apply string_lt_irreflexive in H. + - apply iff_reflect. unfold eqb; cbn. + destruct (Nat.compare_spec n n0); split; intro HH; + try reflexivity; try discriminate; try congruence. + all: inversion HH; subst; now apply Nat.lt_irrefl in H. + Qed. + + Definition eq_leibniz (x y : t) : eq x y -> x = y := id. End Level. -Module LevelDecidableType. - Definition t : Type := Level.t. - Definition eq : t -> t -> Prop := eq. - Definition eq_refl : forall x : t, eq x x := @eq_refl _. - Definition eq_sym : forall x y : t, eq x y -> eq y x := @eq_sym _. - Definition eq_trans : forall x y z : t, eq x y -> eq y z -> eq x z := @eq_trans _. - Definition eq_equiv : RelationClasses.Equivalence eq := _. - Definition eq_dec : forall x y : t, {eq x y} + {~ eq x y} := Level.eq_dec. -End LevelDecidableType. -Module LevelSet <: (MSetInterface.WSetsOn LevelDecidableType) := MSets.MSetWeakList.Make LevelDecidableType. -Module LevelSetFact := WFactsOn LevelDecidableType LevelSet. -Module LevelSetProp := WPropertiesOn LevelDecidableType LevelSet. +Module LevelSet := MSetList.MakeWithLeibniz Level. +Module LevelSetFact := WFactsOn Level LevelSet. +Module LevelSetProp := WPropertiesOn Level LevelSet. Definition LevelSet_pair x y := LevelSet.add y (LevelSet.singleton x). @@ -107,143 +172,841 @@ Proof. apply orb_true_iff. Qed. -Definition universe_level := Level.t. +(** no prop levels are levels which are Set or Level or Var *) +Module NoPropLevel. + Inductive t_ := lSet | Level (_ : string) | Var (_ : nat). -Module Universe. - Module Expr. - (* level+1 if true. Except if Prop -> boolean ignored *) - Definition t : Set := Level.t * bool. - - (* Warning: (lProp, true) represents also Prop *) - Definition is_small (e : t) : bool := - match e with - | (Level.lProp, _) - | (Level.lSet, false) => true - | _ => false - end. - - Definition is_level (e : t) : bool := - match e with - | (Level.lProp, _) - | (_, false) => true - | _ => false - end. - - Definition is_prop (e : t) : bool := - match e with - | (Level.lProp, _) => true - | _ => false - end. - - Definition equal (e1 e2 : t) : bool - := match e1, e2 with - | (Level.lProp, _), (Level.lProp, _) => true - | (l1, b1), (l2, b2) => Level.equal l1 l2 && Bool.eqb b1 b2 - end. - - Definition get_level (e : t) : Level.t := fst e. - - Definition prop : t := (Level.prop, false). - Definition set : t := (Level.set, false). - Definition type1 : t := (Level.set, true). - - Lemma equal_refl : - forall e, equal e e. - Proof. - intro e. destruct e as [[] b]. all: simpl. - - reflexivity. - - apply eqb_reflx. - - rewrite eqb_reflx. rewrite Level.equal_refl. reflexivity. - - rewrite Level.equal_refl, eqb_reflx. reflexivity. - Qed. - - End Expr. - - Definition t : Set := non_empty_list Expr.t. - - Definition equal (u1 u2 : t) : bool := - NEL.forallb2 Expr.equal u1 u2. - - Definition make (l : Level.t) : t - := NEL.sing (l, false). - (** Create a universe representing the given level. *) + Definition t : Set := t_. - Definition make' (e : Expr.t) : t - := NEL.sing e. + Global Instance Level_Evaluable : Evaluable t + := fun v l => match l with + | lSet => 0 + | Level s => Z.pos (v.(valuation_mono) s) + | Var x => Z.of_nat (v.(valuation_poly) x) + end. - Definition make'' (e : Expr.t) (u : list Expr.t) : t - := NEL.cons' e u. - (* FIXME: take duplicates in account *) - Definition is_level (u : t) : bool := - match u with - | NEL.sing e => Expr.is_level e + Definition compare (l1 l2 : t) : comparison := + match l1, l2 with + | lSet, lSet => Eq + | lSet, _ => Lt + | _, lSet => Gt + | Level s1, Level s2 => string_compare s1 s2 + | Level _, _ => Lt + | _, Level _ => Gt + | Var n, Var m => Nat.compare n m + end. + + Definition eq : t -> t -> Prop := Logic.eq. + + Definition eq_equiv : RelationClasses.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_ lSet (Level s) + | ltSetVar n : lt_ lSet (Var n) + | ltLevelLevel s s' : string_lt s s' -> lt_ (Level s) (Level s') + | ltLevelVar s n : lt_ (Level s) (Var n) + | ltVarVar n n' : Nat.lt n n' -> lt_ (Var n) (Var n'). + + Definition lt := lt_. + + Global Instance lt_strorder : StrictOrder lt. + split. + - intros [| |] X; inversion X. + eapply string_lt_irreflexive; tea. + eapply Nat.lt_irrefl; tea. + - intros [| |] [| |] [| |] X1 X2; + inversion X1; inversion X2; constructor. + eapply transitive_string_lt; tea. + etransitivity; tea. + Qed. + + Definition lt_compat : Proper (eq ==> eq ==> iff) lt. + Proof. + intros x y e z t e'. unfold eq in *; subst. reflexivity. + Qed. + + Definition compare_spec : + forall x y : t, CompareSpec (x = y) (lt x y) (lt y x) (compare x y). + Proof. + intros [] []; repeat constructor. + - eapply CompareSpec_Proper. + 5: eapply CompareSpec_string. 4: reflexivity. + all: split; [now inversion 1|]. now intros ->. + all: intro; now constructor. + - eapply CompareSpec_Proper. + 5: eapply Nat.compare_spec. 4: reflexivity. + all: split; [now inversion 1|]. now intros ->. + all: intro; now constructor. + Qed. + + Definition eq_leibniz (x y : t) : eq x y -> x = y := id. + + Definition of_level (l : Level.t) : option t := + match l with + | Level.lProp => None + | Level.lSet => Some lSet + | Level.Level s => Some (Level s) + | Level.Var n => Some (Var n) + end. + + Definition to_level (l : t) : Level.t := + match l with + | lSet => Level.lSet + | Level s => Level.Level s + | Var n => Level.Var n + end. + + Lemma of_to_level l + : of_level (to_level l) = Some l. + Proof. + destruct l; reflexivity. + Qed. + + Definition is_small (e : t) : bool := + match e with + | lSet => true | _ => false end. - (** Test if the universe is a level or an algebraic universe. *) - Definition is_levels (u : t) : bool := - NEL.forallb Expr.is_level u. - (** Test if the universe is a lub of levels or contains +n's. *) + Lemma val_to_level v l : + val v (to_level l) = val v l. + Proof. + destruct l; cbnr. + Qed. - Definition level (u : t) : option Level.t := - match u with - | NEL.sing (Level.lProp, _) => Some Level.lProp - | NEL.sing (l, false) => Some l - | _ => None + Lemma val_zero l v : 0 <= val v l. + Proof. + destruct l; cbn; lia. + Qed. +End NoPropLevel. + +Coercion NoPropLevel.to_level : NoPropLevel.t >-> Level.t. + + +Module UnivExpr. + (* npe = no prop expression, +1 if true *) + Inductive t_ := lProp | npe (e : NoPropLevel.t * bool). + + Definition t := t_. + + Global Instance Evaluable : Evaluable t + := fun v l => match l with + | lProp => -1 + | npe (l, b) => (if b then 1 else 0) + val v l + end. + + + Definition is_small (e : t) : bool := + match e with + | lProp + | npe (NoPropLevel.lSet, false) => true + | _ => false end. - (** Try to get a level out of a universe, returns [None] if it - is an algebraic universe. *) - (* Definition levels (u : t) : list Level.t := *) - (* LevelSet.elements (NEL.fold_left (fun s '(l, _) => LevelSet.add l s) *) - (* u LevelSet.empty). *) - (* (** Get the levels inside the universe, forgetting about increments *) *) + Definition is_level (e : t) : bool := + match e with + | lProp + | npe (_, false) => true + | _ => false + end. - Definition is_small (u : t) : bool := - NEL.forallb Expr.is_small u. + Definition is_prop (e : t) : bool := + match e with + | lProp => true + | _ => false + end. - Definition is_prop (u : t) : bool := - NEL.forallb Expr.is_prop u. + Definition get_level (e : t) : Level.t := + match e with + | lProp => Level.lProp + | npe (l, _) => l + end. - Definition type0m : t := make Level.prop. - Definition type0 : t := make Level.set. - Definition type1 :t := make' Expr.type1. + Definition get_noprop (e : UnivExpr.t) := + match e with + | lProp => None + | npe (l, _) => Some l + end. + + Definition make (l : Level.t) : t := + match NoPropLevel.of_level l with + | None => lProp + | Some l => npe (l, false) + end. + + Definition prop : t := lProp. + Definition set : t := npe (NoPropLevel.lSet, false). + Definition type1 : t := npe (NoPropLevel.lSet, true). + + (* Used for (un)quoting. *) + Definition from_kernel_repr (e : Level.t * bool) : t + := match NoPropLevel.of_level e.1 with + | Some l => npe (l, e.2) + | None => lProp + end. + Definition to_kernel_repr (e : t) : Level.t * bool + := match e with + | lProp => (Level.lProp, false) + | npe (l, b) => (NoPropLevel.to_level l, b) + end. + + Definition eq : t -> t -> Prop := eq. + + Definition eq_equiv : Equivalence eq := _. + + Inductive lt_ : t -> t -> Prop := + | ltUnivExpr0 e : lt_ lProp (npe e) + | ltUnivExpr1 l : lt_ (npe (l, false)) (npe (l, true)) + | ltUnivExpr2 l l' b b' : NoPropLevel.lt l l' -> lt_ (npe (l, b)) (npe (l', b')). + + Definition lt := lt_. + + Global Instance lt_strorder : StrictOrder lt. + Proof. + constructor. + - intros x X; inversion X. + eapply NoPropLevel.lt_strorder; eassumption. + - intros x y z X1 X2; invs X1; invs X2; constructor; tea. + etransitivity; tea. + Qed. + + Definition lt_compat : Proper (Logic.eq ==> Logic.eq ==> iff) lt. + intros x x' H1 y y' H2; now rewrite H1, H2. + Qed. + + Definition compare (x y : t) : comparison := + match x, y with + | lProp, lProp => Eq + | lProp, _ => Lt + | _, lProp => Gt + | npe (l1, b1), npe (l2, b2) => + match NoPropLevel.compare l1 l2 with + | Eq => bool_compare b1 b2 + | x => x + end + end. + + Definition compare_spec : + forall x y : t, CompareSpec (x = y) (lt x y) (lt y x) (compare x y). + Proof. + intros [|[]] [|[]]; cbn; repeat constructor. + destruct (NoPropLevel.compare_spec t0 t1); repeat constructor; tas. + subst. destruct b, b0; repeat constructor. + Qed. + + Definition eq_dec (l1 l2 : t) : {l1 = l2} + {l1 <> l2}. + Proof. + repeat decide equality. + Defined. + + Definition eq_leibniz (x y : t) : eq x y -> x = y := id. + + + Lemma val_make v l + : val v (UnivExpr.make l) = val v l. + Proof. + destruct l; cbnr. + Qed. + + Lemma val_make_npl v (l : NoPropLevel.t) + : val v (UnivExpr.make l) = val v l. + Proof. + destruct l; cbnr. + Qed. + + Lemma val_minus_one e v : -1 <= val v e. + Proof. + destruct e as [|[[] b]]; cbn; try destruct b; lia. + Qed. + + Lemma is_prop_val e : is_prop e -> forall v, val v e = -1. + Proof. + destruct e as [|[[] []]]; cbn; try reflexivity; try discriminate. + Qed. + + Lemma val_is_prop e v : val v e <= -1 -> is_prop e. + Proof. + destruct e as [|[[] b]]; cbnr; try destruct b; lia. + Qed. + + Lemma is_prop_val_false e : + is_prop e = false -> forall v, 0 <= val v e. + Proof. + intros He v. + pose proof (val_is_prop e v). + pose proof (val_minus_one e v). + destruct (Z_le_gt_dec (val v e) (-1)); [|lia]. + forward H; tas. rewrite He in H; discriminate. + Qed. + + Lemma val_is_prop_false e v : + 0 <= val v e -> is_prop e = false. + Proof. + pose proof (is_prop_val e) as H. destruct (is_prop e); cbnr. + rewrite (H eq_refl v). lia. + Qed. +End UnivExpr. + +Module UnivExprSet := MSetList.MakeWithLeibniz UnivExpr. +Module UnivExprSetFact := WFactsOn UnivExpr UnivExprSet. +Module UnivExprSetProp := WPropertiesOn UnivExpr UnivExprSet. + + +Module Universe. + (** A universe is a list of universe expressions which is: + - sorted + - without duplicate + - non empty *) + Record t := { t_set : UnivExprSet.t ; + t_ne : UnivExprSet.is_empty t_set = false }. + + Coercion t_set : t >-> UnivExprSet.t. + + Definition eqb (u1 u2 : t) : bool := + UnivExprSet.equal u1.(t_set) u2.(t_set). + + Definition make' (e : UnivExpr.t) : t + := {| t_set := UnivExprSet.singleton e; + t_ne := eq_refl |}. + + (** Create a universe representing the given level. *) + Definition make (l : Level.t) : t := make' (UnivExpr.make l). + + Lemma not_Empty_is_empty s : + ~ UnivExprSet.Empty s -> UnivExprSet.is_empty s = false. + Proof. + intro H. apply not_true_is_false. intro H'. + apply H. now apply UnivExprSetFact.is_empty_2 in H'. + Qed. + + Program Definition add (e : UnivExpr.t) (u : t) : t + := {| t_set := UnivExprSet.add e u |}. + Next Obligation. + apply not_Empty_is_empty; intro H. + eapply H. eapply UnivExprSet.add_spec. + left; reflexivity. + Qed. + + Lemma add_spec e u e' : + UnivExprSet.In e' (add e u) <-> e' = e \/ UnivExprSet.In e' u. + Proof. + apply UnivExprSet.add_spec. + Qed. + + Definition add_list : list UnivExpr.t -> t -> t + := List.fold_left (fun u e => add e u). + + Lemma add_list_spec l u e : + UnivExprSet.In e (add_list l u) <-> In e l \/ UnivExprSet.In e u. + Proof. + unfold add_list. rewrite <- fold_left_rev_right. + etransitivity. 2:{ eapply or_iff_compat_r. etransitivity. + 2: apply @InA_In_eq with (A:=UnivExpr.t). + eapply InA_rev. } + induction (List.rev l); cbn. + - split. intuition. intros [H|H]; tas. invs H. + - split. + + intro H. apply add_spec in H. destruct H as [H|H]. + * left. now constructor. + * apply IHl0 in H. destruct H as [H|H]; [left|now right]. + now constructor 2. + + intros [H|H]. inv H. + * apply add_spec; now left. + * apply add_spec; right. apply IHl0. now left. + * apply add_spec; right. apply IHl0. now right. + Qed. + + (** Test if the universe is a lub of levels or contains +n's. *) + Definition is_levels : t -> bool + := UnivExprSet.for_all UnivExpr.is_level. + + (** Test if the universe is a level or an algebraic universe. *) + Definition is_level (u : t) : bool := + (UnivExprSet.cardinal u =? 1)%nat && is_levels u. + + Definition is_small : t -> bool := + UnivExprSet.for_all UnivExpr.is_small. + + Definition is_prop : t -> bool := + UnivExprSet.for_all UnivExpr.is_prop. + + Definition type0m : t := make Level.lProp. + Definition type0 : t := make Level.lSet. + Definition type1 : t := make' UnivExpr.type1. + + (* Used for (un)quoting. *) + Definition from_kernel_repr (e : Level.t * bool) (es : list (Level.t * bool)) : t + := add_list (map UnivExpr.from_kernel_repr es) + (make' (UnivExpr.from_kernel_repr e)). + Definition to_kernel_repr (u : t) : list (Level.t * bool) + := map (UnivExpr.to_kernel_repr) (UnivExprSet.elements u). - Definition super (l : Level.t) : t := - if Level.is_small l then type1 - else make' (l, true). (** The universe strictly above FOR TYPING (not cumulativity) *) + Definition super (l : Level.t) : t := + match NoPropLevel.of_level l with + | None => type1 + | Some NoPropLevel.lSet => type1 + | Some l => make' (UnivExpr.npe (l, true)) + end. + + (** The non empty / sorted / without dup list of univ expressions, the + components of the pair are the head and the tail of the (non empty) list *) + Program Definition exprs (u : t) : UnivExpr.t * list UnivExpr.t + := match UnivExprSet.elements u with + | [] => False_rect _ _ + | e :: l => (e, l) + end. + Next Obligation. + destruct u as [u1 u2]; cbn in *. revert u2. + apply eq_true_false_abs. + unfold UnivExprSet.is_empty, UnivExprSet.Raw.is_empty, + UnivExprSet.elements, UnivExprSet.Raw.elements in *. + rewrite <- Heq_anonymous; reflexivity. + Qed. + + Lemma exprs_make' e : Universe.exprs (Universe.make' e) = (e, []). + Proof. reflexivity. Defined. + + Lemma exprs_make l : Universe.exprs (Universe.make l) = (UnivExpr.make l, []). + Proof. reflexivity. Defined. + + + Lemma exprs_spec u : + let '(e, u') := Universe.exprs u in + e :: u' = UnivExprSet.elements u. + Proof. + destruct u as [u1 u2]. + pose (l := UnivExprSet.elements u1). + change (let '(e, u') := + match l return l = UnivExprSet.elements u1 -> _ with + | [] => fun Heq_anonymous : [] = _ u1 => + False_rect (UnivExpr.t × list UnivExpr.t) + (exprs_obligation_1 + {| Universe.t_set := u1; Universe.t_ne := u2 |} + Heq_anonymous) + | e :: l0 => fun _ : e :: l0 = _ u1 => (e, l0) + end eq_refl in e :: u' = UnivExprSet.elements u1). + set (e := eq_refl). clearbody e. change (l = UnivExprSet.elements u1) in e. + destruct l. + - exfalso. revert u2. apply eq_true_false_abs. + unfold UnivExprSet.is_empty, UnivExprSet.Raw.is_empty, + UnivExprSet.elements, UnivExprSet.Raw.elements in *. + rewrite <- e; reflexivity. + - assumption. + Qed. + + Lemma exprs_spec' u : + (Universe.exprs u).1 :: (Universe.exprs u).2 = UnivExprSet.elements u. + Proof. + pose proof (exprs_spec u). + now destruct (Universe.exprs u). + Qed. + + Lemma In_exprs (u : t) (e : UnivExpr.t) : + UnivExprSet.In e u + <-> e = (Universe.exprs u).1 \/ In e (Universe.exprs u).2. + Proof. + etransitivity. symmetry. apply UnivExprSet.elements_spec1. + pose proof (Universe.exprs_spec' u) as H. + destruct (Universe.exprs u) as [e' l]; cbn in *. + rewrite <- H; clear. etransitivity. apply InA_cons. + eapply or_iff_compat_l. apply InA_In_eq. + Qed. + + Lemma In_exprs_rev (u : t) (e : UnivExpr.t) : + UnivExprSet.In e u + <-> e = (Universe.exprs u).1 \/ In e (List.rev (Universe.exprs u).2). + Proof. + etransitivity. eapply In_exprs. + apply or_iff_compat_l. apply in_rev. + Qed. - Definition try_suc (u : t) : t := - NEL.map (fun '(l, b) => (l, true)) u. + Definition map (f : UnivExpr.t -> UnivExpr.t) (u : t) : t := + let '(e, l) := exprs u in + add_list (List.map f l) (make' (f e)). - (* FIXME: don't duplicate *) - Definition sup (u1 u2 : t) : t := NEL.app u1 u2. - (** The l.u.b. of 2 universes (naive because of duplicates) *) + Lemma map_spec f u e : + UnivExprSet.In e (map f u) <-> exists e0, UnivExprSet.In e0 u /\ e = (f e0). + Proof. + unfold map. symmetry. etransitivity. + { eapply iff_ex; intro. eapply and_iff_compat_r. eapply In_exprs. } + destruct (exprs u) as [e' l]; cbn in *. + symmetry. etransitivity. eapply add_list_spec. + etransitivity. eapply or_iff_compat_l. apply UnivExprSet.singleton_spec. + etransitivity. eapply or_iff_compat_r. + apply in_map_iff. clear u. split. + - intros [[e0 []]|H]. + + exists e0. split. right; tas. congruence. + + exists e'. split; tas. left; reflexivity. + - intros [xx [[H|H] ?]]. + + right. congruence. + + left. exists xx. split; tas; congruence. + Qed. - (* Definition existsb (P : Expr.t -> bool) (u : t) : bool := NEL.existsb P u. *) - Definition for_all (P : Expr.t -> bool) (u : t) : bool := NEL.forallb P u. + Definition try_suc + := map (fun l => match l with + | UnivExpr.lProp => UnivExpr.type1 + | UnivExpr.npe (l, _) => UnivExpr.npe (l, true) + end). + + (** The l.u.b. of 2 universes *) + Program Definition sup (u v : t) : t := + {| t_set := UnivExprSet.union u v |}. + Next Obligation. + apply not_Empty_is_empty; intro H. + assert (HH: UnivExprSet.Empty u). { + intros x Hx. apply (H x). + eapply UnivExprSet.union_spec. now left. } + apply UnivExprSetFact.is_empty_1 in HH. + rewrite t_ne in HH; discriminate. + Qed. - (** Type of product *) + (** Type of a product *) Definition sort_of_product (domsort rangsort : t) := (* Prop impredicativity *) if Universe.is_prop rangsort then rangsort else Universe.sup domsort rangsort. - Lemma equal_refl : - forall u, equal u u. + Lemma eqb_refl u : eqb u u. Proof. - intro u. unfold equal. eapply NEL.forallb2_refl. - apply Expr.equal_refl. + apply UnivExprSet.equal_spec. reflexivity. Qed. + Lemma elements_not_empty (u : Universe.t) : UnivExprSet.elements u <> []. + Proof. + destruct u as [u1 u2]; cbn; intro e. + unfold UnivExprSet.is_empty, UnivExprSet.elements, + UnivExprSet.Raw.elements in *. + rewrite e in u2; discriminate. + Qed. + + Definition get_is_level (u : t) : option Level.t := + match UnivExprSet.elements u with + | [UnivExpr.lProp] => Some Level.lProp + | [UnivExpr.npe (l, false)] => Some (NoPropLevel.to_level l) + | _ => None + end. + + Global Instance Evaluable : Evaluable t + := fun v u => let '(e, u) := Universe.exprs u in + List.fold_left (fun n e => Z.max (val v e) n) u (val v e). + + Lemma val_make' v e + : val v (make' e) = val v e. + Proof. + reflexivity. + Qed. + + Lemma val_make v l + : val v (make l) = val v l. + Proof. + destruct l; cbnr. + Qed. + + Lemma val_make_npl v (l : NoPropLevel.t) + : val v (make l) = val v l. + Proof. + destruct l; cbnr. + Qed. + + Lemma make_inj x y : + make x = make y -> x = y. + Proof. + destruct x, y; now inversion 1. + Qed. End Universe. -Definition universe := Universe.t. -Definition universe_coercion : universe -> list Universe.Expr.t := NEL.to_list. -Coercion universe_coercion : universe >-> list. +(** This coercion allows to see the universes as a [UnivExprSet.t] *) +Coercion Universe.t_set : Universe.t >-> UnivExprSet.t. + + +Ltac u := + change LevelSet.elt with Level.t in *; + change UnivExprSet.elt with UnivExpr.t in *. + (* change ConstraintSet.elt with UnivConstraint.t in *. *) + + +Lemma val_fold_right (u : Universe.t) v : + val v u = fold_right (fun e x => Z.max (val v e) x) (val v (Universe.exprs u).1) + (List.rev (Universe.exprs u).2). +Proof. + unfold val at 1, Universe.Evaluable. + destruct (Universe.exprs u). + now rewrite fold_left_rev_right. +Qed. + +Lemma val_In_le (u : Universe.t) v e : + UnivExprSet.In e u -> val v e <= val v u. +Proof. + intro H. rewrite val_fold_right. + apply Universe.In_exprs_rev in H. destruct (Universe.exprs u); cbn in *. + clear -H. destruct H as [H|H]. + - subst. induction (List.rev l); cbnr; lia. + - induction (List.rev l); cbn; invs H. + + u; lia. + + specialize (IHl0 H0). lia. +Qed. + +Lemma val_In_max (u : Universe.t) v : + exists e, UnivExprSet.In e u /\ val v e = val v u. +Proof. + eapply iff_ex. { + intro. eapply and_iff_compat_r. apply Universe.In_exprs_rev. } + rewrite val_fold_right. destruct (Universe.exprs u) as [e l]; cbn in *. + clear. induction (List.rev l); cbn. + - exists e. split; cbnr. left; reflexivity. + - destruct IHl0 as [e' [H1 H2]]. + destruct (Z.max_dec (val v a) (fold_right (fun e0 x0 => Z.max (val v e0) x0) + (val v e) l0)) as [H|H]; + rewrite H; clear H. + + exists a. split; cbnr. right. now constructor. + + rewrite <- H2. exists e'. split; cbnr. + destruct H1 as [H1|H1]; [now left|right]. now constructor 2. +Qed. + +Lemma val_ge_caract (u : Universe.t) v k : + (forall e, UnivExprSet.In e u -> val v e <= k) <-> val v u <= k. +Proof. + split. + - eapply imp_iff_compat_r. { + eapply iff_forall; intro. eapply imp_iff_compat_r. + apply Universe.In_exprs_rev. } + rewrite val_fold_right. + destruct (Universe.exprs u) as [e l]; cbn; clear. + induction (List.rev l); cbn. + + intros H. apply H. left; reflexivity. + + intros H. + destruct (Z.max_dec (val v a) (fold_right (fun e0 x => Z.max (val v e0) x) + (val v e) l0)) as [H'|H']; + rewrite H'; clear H'. + * apply H. right. now constructor. + * apply IHl0. intros x [H0|H0]; apply H. now left. + right; now constructor 2. + - intros H e He. eapply val_In_le in He. etransitivity; eassumption. +Qed. + +Lemma val_le_caract (u : Universe.t) v k : + (exists e, UnivExprSet.In e u /\ k <= val v e) <-> k <= val v u. +Proof. + split. + - eapply imp_iff_compat_r. { + eapply iff_ex; intro. eapply and_iff_compat_r. apply Universe.In_exprs_rev. } + rewrite val_fold_right. + destruct (Universe.exprs u) as [e l]; cbn; clear. + induction (List.rev l); cbn. + + intros H. destruct H as [e' [[H1|H1] H2]]. + * now subst. + * invs H1. + + intros [e' [[H1|H1] H2]]. + * forward IHl0; [|lia]. exists e'. split; tas. left; assumption. + * invs H1. + -- u; lia. + -- forward IHl0; [|lia]. exists e'. split; tas. right; assumption. + - intros H. destruct (val_In_max u v) as [e [H1 H2]]. + exists e. rewrite H2; split; assumption. +Qed. + + +Lemma val_caract (u : Universe.t) v k : + val v u = k + <-> (forall e, UnivExprSet.In e u -> val v e <= k) + /\ exists e, UnivExprSet.In e u /\ val v e = k. +Proof. + split. + - intros <-; split. + + apply val_In_le. + + apply val_In_max. + - intros [H1 H2]. + apply val_ge_caract in H1. + assert (k <= val v u); [clear H1|lia]. + apply val_le_caract. destruct H2 as [e [H2 H2']]. + exists e. split; tas. lia. +Qed. + +Lemma val_add v e s + : val v (Universe.add e s) = Z.max (val v e) (val v s). +Proof. + apply val_caract. split. + - intros e' H. apply UnivExprSet.add_spec in H. destruct H as [H|H]. + + subst. u; lia. + + eapply val_In_le with (v:=v) in H. lia. + - destruct (Z.max_dec (val v e) (val v s)) as [H|H]; rewrite H; clear H. + + exists e. split; cbnr. apply UnivExprSetFact.add_1. reflexivity. + + destruct (val_In_max s v) as [e' [H1 H2]]. + exists e'. split; tas. now apply UnivExprSetFact.add_2. +Qed. + +Lemma val_sup v s1 s2 : + val v (Universe.sup s1 s2) = Z.max (val v s1) (val v s2). +Proof. + eapply val_caract. cbn. split. + - intros e' H. eapply UnivExprSet.union_spec in H. destruct H as [H|H]. + + eapply val_In_le with (v:=v) in H. rewrite H; lia. + + eapply val_In_le with (v:=v) in H. rewrite H; lia. + - destruct (Z.max_dec (val v s1) (val v s2)) as [H|H]; rewrite H; clear H. + + destruct (val_In_max s1 v) as [e' [H1 H2]]. + exists e'. split; tas. apply UnivExprSet.union_spec. now left. + + destruct (val_In_max s2 v) as [e' [H1 H2]]. + exists e'. split; tas. apply UnivExprSet.union_spec. now right. +Qed. + +Ltac proper := let H := fresh in try (intros ? ? H; destruct H; reflexivity). + +Lemma for_all_elements (P : UnivExpr.t -> bool) u : + UnivExprSet.for_all P u = forallb P (UnivExprSet.elements u). +Proof. + apply UnivExprSetFact.for_all_b; proper. +Qed. + +Lemma val_minus_one u v : + (-1 <= val v u)%Z. +Proof. + rewrite val_fold_right. + destruct (Universe.exprs u) as [e u']; clear u; cbn. + induction (List.rev u'); simpl. + - apply UnivExpr.val_minus_one. + - pose proof (UnivExpr.val_minus_one a v). lia. +Qed. + +Lemma is_prop_val u : + Universe.is_prop u -> forall v, val v u = -1. +Proof. + unfold Universe.is_prop. rewrite for_all_elements. + intros H v. rewrite val_fold_right. + rewrite <- (Universe.exprs_spec' u) in H. + destruct (Universe.exprs u) as [e l]; simpl in *. + apply andb_true_iff in H; destruct H as [He H]; revert H. + rewrite <- forallb_rev. induction (List.rev l); simpl. + - intros _. apply UnivExpr.is_prop_val; assumption. + - intro H; apply andb_true_iff in H; destruct H as [H1 H2]. + apply UnivExpr.is_prop_val with (v:=v) in H1. now rewrite IHl0, H1. +Qed. + +Lemma val_is_prop u v : + val v u <= -1 -> Universe.is_prop u. +Proof. + unfold Universe.is_prop. rewrite for_all_elements. + rewrite val_fold_right. + rewrite <- (Universe.exprs_spec' u). + destruct (Universe.exprs u) as [e l]; simpl. + rewrite <- forallb_rev. + induction (List.rev l); cbn. + - rewrite andb_true_r. apply UnivExpr.val_is_prop. + - intro H. forward IHl0; [lia|]. + specialize (UnivExpr.val_is_prop a v) as H'. forward H'; [lia|]. + apply andb_true_iff. apply andb_true_iff in IHl0. + intuition. +Qed. + +Lemma is_prop_val_false u : + Universe.is_prop u = false -> forall v, 0 <= val v u. +Proof. + intros Hu v. + pose proof (val_is_prop u v). + pose proof (val_minus_one u v). + destruct (Z_le_gt_dec (val v u) (-1)); [|lia]. + forward H; tas. rewrite Hu in H; discriminate. +Qed. + +Lemma val_is_prop_false u v : + 0 <= val v u -> Universe.is_prop u = false. +Proof. + pose proof (is_prop_val u) as H. destruct (Universe.is_prop u); cbnr. + rewrite (H eq_refl v). lia. +Qed. + + +Lemma eq_univ (u v : Universe.t) : + u = v :> UnivExprSet.t -> u = v. +Proof. + destruct u as [u1 u2], v as [v1 v2]; cbn. intros X; destruct X. + now rewrite (uip_bool _ _ u2 v2). +Qed. + +Lemma eq_univ' (u v : Universe.t) : + UnivExprSet.Equal u v -> u = v. +Proof. + intro H. now apply eq_univ, UnivExprSet.eq_leibniz. +Qed. + +Lemma eq_univ'' (u v : Universe.t) : + UnivExprSet.elements u = UnivExprSet.elements v -> u = v. +Proof. + intro H. apply eq_univ. + destruct u as [u1 u2], v as [v1 v2]; cbn in *; clear u2 v2. + destruct u1 as [u1 u2], v1 as [v1 v2]; cbn in *. + destruct H. now rewrite (uip_bool _ _ u2 v2). +Qed. + +Lemma get_is_level_correct u l : + Universe.get_is_level u = Some l -> u = Universe.make l. +Proof. + intro H. apply eq_univ''; cbn. unfold Universe.get_is_level in *. + destruct (UnivExprSet.elements u) as [|l0 L]; [discriminate|]. + destruct l0, L; try discriminate. + apply some_inj in H; now subst. + destruct e as [? []]; try discriminate. + apply some_inj in H; subst. now destruct t. + destruct e as [? []]; try discriminate. +Qed. + +Lemma eqb_true_iff u v : + Universe.eqb u v <-> u = v. +Proof. + split. 2: intros ->; apply Universe.eqb_refl. + intro H. apply eq_univ'. now apply UnivExprSet.equal_spec. +Qed. + +Lemma UnivExprSet_for_all_false f u : + UnivExprSet.for_all f u = false -> UnivExprSet.exists_ (negb ∘ f) u. +Proof. + intro H. rewrite UnivExprSetFact.exists_b. + rewrite UnivExprSetFact.for_all_b in H. + all: try now intros x y []. + induction (UnivExprSet.elements u); cbn in *; [discriminate|]. + apply andb_false_iff in H; apply orb_true_iff; destruct H as [H|H]. + left; now rewrite H. + right; now rewrite IHl. +Qed. + +Lemma UnivExprSet_For_all_exprs (P : UnivExpr.t -> Prop) (u : Universe.t) + : UnivExprSet.For_all P u + <-> P (Universe.exprs u).1 /\ Forall P (Universe.exprs u).2. +Proof. + etransitivity. + - eapply iff_forall; intro e. eapply imp_iff_compat_r. + apply Universe.In_exprs. + - cbn; split. + + intro H. split. apply H. now left. + apply Forall_forall. intros x H0. apply H; now right. + + intros [H1 H2] e [He|He]. subst e; tas. + eapply Forall_forall in H2; tea. +Qed. + +Lemma is_prop_sort_sup x1 x2 : + Universe.is_prop (Universe.sup x1 x2) -> Universe.is_prop x2. +Proof. + intro H. apply UnivExprSet.for_all_spec; proper. + apply UnivExprSet.for_all_spec in H; proper. + intros x HX; apply H. + apply UnivExprSet.union_spec; now right. +Qed. + +Lemma is_prop_sort_prod x2 x3 : + Universe.is_prop (Universe.sort_of_product x2 x3) -> Universe.is_prop x3. +Proof. + unfold Universe.sort_of_product. + case_eq (Universe.is_prop x3); try reflexivity. + intros <-. apply is_prop_sort_sup. +Qed. + (** Sort families *) @@ -259,58 +1022,126 @@ Definition leb_sort_family x y := end. (** Family of a universe [u]. *) - -Definition universe_family (u : universe) := +Definition universe_family (u : Universe.t) := if Universe.is_prop u then InProp else if Universe.is_small u then InSet else InType. Module ConstraintType. - Inductive t : Set := Lt | Le | Eq. -End ConstraintType. + Inductive t_ : Set := Lt | Le | Eq. + Definition t := t_. + Definition eq : t -> t -> Prop := eq. + Definition eq_equiv : RelationClasses.Equivalence eq := _. -Definition constraint_type := ConstraintType.t. -Definition univ_constraint : Set := universe_level * ConstraintType.t * universe_level. + Inductive lt_ : t -> t -> Prop := + | LtLe : lt_ Lt Le + | LtEq : lt_ Lt Eq + | LeEq : lt_ Le Eq. + Definition lt := lt_. + + Lemma lt_strorder : StrictOrder lt. + Proof. + constructor. + - intros []; intro X; inversion X. + - intros ? ? ? X Y; invs X; invs Y; constructor. + Qed. + + Lemma lt_compat : Proper (eq ==> eq ==> iff) lt. + Proof. + intros ? ? X ? ? Y; invs X; invs Y. reflexivity. + Qed. + + Definition compare (x y : t) : comparison := + match x, y with + | Lt, Lt => Datatypes.Eq + | Lt, _ => Datatypes.Lt + | Le, Lt => Datatypes.Gt + | Le, Le => Datatypes.Eq + | Le, Eq => Datatypes.Lt + | Eq, Eq => Datatypes.Eq + | Eq, _ => Datatypes.Gt + end. + + Lemma compare_spec x y : CompareSpec (eq x y) (lt x y) (lt y x) (compare x y). + Proof. + destruct x, y; repeat constructor. + Qed. + + Lemma eq_dec x y : {eq x y} + {~ eq x y}. + Proof. + unfold eq. decide equality. + Qed. +End ConstraintType. -Module UnivConstraintDec. - Definition t : Set := univ_constraint. +Module UnivConstraint. + Definition t : Set := Level.t * ConstraintType.t * Level.t. Definition eq : t -> t -> Prop := eq. Definition eq_equiv : RelationClasses.Equivalence eq := _. - Definition eq_dec : forall x y : t, {eq x y} + {~ eq x y}. - unfold eq. - decide equality. decide equality. apply string_dec. apply Peano_dec.eq_nat_dec. - decide equality. decide equality. apply LevelDecidableType.eq_dec. - Defined. -End UnivConstraintDec. -Module ConstraintSet <: MSetInterface.WSetsOn UnivConstraintDec := MSets.MSetWeakList.Make UnivConstraintDec. -Module ConstraintSetFact := MSetFacts.WFactsOn UnivConstraintDec ConstraintSet. -Module ConstraintSetProp := MSetProperties.WPropertiesOn UnivConstraintDec - ConstraintSet. -Definition make_univ_constraint : universe_level -> constraint_type -> universe_level -> univ_constraint - := fun x y z => (x, y, z). + Definition make l1 ct l2 : t := (l1, ct, l2). + + Inductive lt_ : t -> t -> Prop := + | lt_Level2 l1 t l2 l2' : Level.lt l2 l2' -> lt_ (l1, t, l2) (l1, t, l2') + | lt_Cstr l1 t t' l2 l2' : ConstraintType.lt t t' -> lt_ (l1, t, l2) (l1, t', l2') + | lt_Level1 l1 l1' t t' l2 l2' : Level.lt l1 l1' -> lt_ (l1, t, l2) (l1', t', l2'). + Definition lt := lt_. + + Lemma lt_strorder : StrictOrder lt. + Proof. + constructor. + - intros []; intro X; inversion X; subst; + try (eapply Level.lt_strorder; eassumption). + eapply ConstraintType.lt_strorder; eassumption. + - intros ? ? ? X Y; invs X; invs Y; constructor; tea. + etransitivity; eassumption. + 2: etransitivity; eassumption. + eapply ConstraintType.lt_strorder; eassumption. + Qed. + + Lemma lt_compat : Proper (eq ==> eq ==> iff) lt. + Proof. + intros ? ? X ? ? Y; invs X; invs Y. reflexivity. + Qed. + + Definition compare : t -> t -> comparison := + fun '(l1, t, l2) '(l1', t', l2') => + Pos.switch_Eq (Pos.switch_Eq (Level.compare l2 l2') + (ConstraintType.compare t t')) + (Level.compare l1 l1'). + + Lemma compare_spec x y + : CompareSpec (eq x y) (lt x y) (lt y x) (compare x y). + Proof. + destruct x as [[l1 t] l2], y as [[l1' t'] l2']; cbn. + destruct (Level.compare_spec l1 l1'); cbn; repeat constructor; tas. + invs H. + destruct (ConstraintType.compare_spec t t'); cbn; repeat constructor; tas. + invs H. + destruct (Level.compare_spec l2 l2'); cbn; repeat constructor; tas. + invs H. reflexivity. + Qed. -(** Needs to be in Type because template polymorphism of MSets does not allow setting - the lowest universe *) -Definition constraints : Type := ConstraintSet.t. (* list univ_constraint *) + Lemma eq_dec x y : {eq x y} + {~ eq x y}. + Proof. + unfold eq. repeat decide equality. + Defined. -Definition empty_constraint : constraints := ConstraintSet.empty. + Definition eq_leibniz (x y : t) : eq x y -> x = y := id. +End UnivConstraint. -(* Definition constraint_function A : Type := A -> A -> constraints -> constraints. *) -(* val enforce_eq : universe constraint_function *) -(* val enforce_leq : universe constraint_function *) -(* val enforce_eq_level : universe_level constraint_function *) -(* val enforce_leq_level : universe_level constraint_function *) +Module ConstraintSet := MSetList.MakeWithLeibniz UnivConstraint. +Module ConstraintSetFact := WFactsOn UnivConstraint ConstraintSet. +Module ConstraintSetProp := WPropertiesOn UnivConstraint ConstraintSet. (** {6 Universe instances} *) Module Instance. - Definition t : Set := list Level.t. (** A universe instance represents a vector of argument universes to a polymorphic definition (constant, inductive or constructor). *) + Definition t : Set := list Level.t. Definition empty : t := []. Definition is_empty (i : t) : bool := @@ -319,36 +1150,32 @@ Module Instance. | _ => false end. - Definition equal (i j : t) := - forallb2 Level.equal i j. + Definition eqb (i j : t) := + forallb2 Level.eqb i j. Definition equal_upto (f : Level.t -> Level.t -> bool) (i j : t) := forallb2 f i j. End Instance. -Definition universe_instance := Instance.t. - -(* val enforce_eq_instances : universe_instance constraint_function *) Module UContext. - Definition t := universe_instance × constraints. + Definition t := Instance.t × ConstraintSet.t. Definition make : Instance.t -> ConstraintSet.t -> t := pair. Definition empty : t := (Instance.empty, ConstraintSet.empty). - (* val is_empty : t -> bool *) Definition instance : t -> Instance.t := fst. - Definition constraints : t -> constraints := snd. + Definition constraints : t -> ConstraintSet.t := snd. Definition dest : t -> Instance.t * ConstraintSet.t := fun x => x. End UContext. Module AUContext. - Definition t := list name × constraints. + Definition t := list name × ConstraintSet.t. - Definition make (ids : list name) (ctrs : constraints) : t := (ids, ctrs). + Definition make (ids : list name) (ctrs : ConstraintSet.t) : t := (ids, ctrs). Definition repr '((u, cst) : t) : UContext.t := (mapi (fun i _ => Level.Var i) u, cst). @@ -357,7 +1184,7 @@ Module AUContext. End AUContext. Module ContextSet. - Definition t := LevelSet.t × constraints. + Definition t := LevelSet.t × ConstraintSet.t. Definition empty : t := (LevelSet.empty, ConstraintSet.empty). @@ -407,277 +1234,460 @@ Definition constraints_of_udecl u := end. -(** * Valuations *) -Import Level ConstraintType. +Definition llt {cf:checker_flags} (x y : Z) : Prop := + if prop_sub_type then x < y else 0 <= x /\ x < y. -(** A valuation is a universe level (nat) given for each - universe variable (Level.t). - It is >= for polymorphic universes and > 0 for monomorphic universes. *) +Declare Scope univ_scope. +Delimit Scope univ_scope with u. +Notation "x < y" := (llt x y) : univ_scope. -Record valuation := - { valuation_mono : string -> positive ; - valuation_poly : nat -> nat }. +Definition lle {cf:checker_flags} (x y : Z) : Prop := + (x < y)%u \/ x = y. -Definition val0 (v : valuation) (l : Level.t) : Z := - match l with - | lProp => -1 - | lSet => 0 - | Level s => Z.pos (v.(valuation_mono) s) - | Var x => Z.of_nat (v.(valuation_poly) x) - end. +Notation "x <= y" := (lle x y) : univ_scope. -Definition val1 v (e : Universe.Expr.t) : Z := - let n := val0 v (fst e) in - if snd e && negb (is_prop (fst e)) then n + 1 else n. +Ltac lle := unfold lle, llt in *. +Ltac lled := lle; match goal with + | H : prop_sub_type = true |- _ => rewrite H in * + | H : prop_sub_type = false |- _ => rewrite H in * + | H : is_true prop_sub_type |- _ => rewrite H in * + | _ => destruct prop_sub_type eqn:?Hb + end. -Definition val (v : valuation) (u : universe) : Z := - match u with - | NEL.sing e => val1 v e - | NEL.cons e u => NEL.fold_left (fun n e => Z.max (val1 v e) n) u (val1 v e) - end. -Definition llt `{checker_flags} (x y : Z) : Prop := - if prop_sub_type - then (x < y)%Z - else (0 <= x /\ x < y)%Z. +Section Univ. - Notation "x < y" := (llt x y) : univ_scope. +Context {cf:checker_flags}. + Global Instance lle_refl : Reflexive lle. + Proof. + intro x. now right. + Qed. -Definition lle `{checker_flags} (x y : Z) : Prop := - (x < y)%u \/ x = y. + Global Instance llt_trans : Transitive llt. + Proof. + intros x y z; unfold llt; destruct prop_sub_type; lia. + Qed. -Notation "x <= y" := (lle x y) : univ_scope. + Global Instance lle_trans : Transitive lle. + Proof. + intros x y z [] []; subst; try (right; reflexivity); left; tas. + etransitivity; eassumption. + Qed. -Section Univ. + Lemma llt_lt n m : (n < m)%u -> (n < m)%Z. + Proof. lled; lia. Qed. -Context `{cf : checker_flags}. + Lemma lle_le n m : (n <= m)%u -> (n <= m)%Z. + Proof. lled; lia. Qed. -Inductive satisfies0 (v : valuation) : univ_constraint -> Prop := -| satisfies0_Lt l l' : (val0 v l < val0 v l')%u -> satisfies0 v (l, Lt, l') -| satisfies0_Le l l' : (val0 v l <= val0 v l')%u -> satisfies0 v (l, Le, l') -| satisfies0_Eq l l' : val0 v l = val0 v l' -> satisfies0 v (l, Eq, l'). + Lemma lt_llt n m : prop_sub_type -> (n < m)%Z -> (n < m)%u. + Proof. unfold llt. now intros ->. Qed. -Definition satisfies v : constraints -> Prop := - ConstraintSet.For_all (satisfies0 v). + Lemma le_lle n m : prop_sub_type -> (n <= m)%Z -> (n <= m)%u. + Proof. lled; [lia|discriminate]. Qed. -Definition consistent ctrs := exists v, satisfies v ctrs. + Lemma lt_llt' n m : (0 <= n)%Z -> (n < m)%Z -> (n < m)%u. + Proof. lled; lia. Qed. -Definition eq_universe0 (φ : constraints) u u' := - forall v, satisfies v φ -> val v u = val v u'. + Lemma le_lle' n m : (0 <= n)%Z -> (n <= m)%Z -> (n <= m)%u. + Proof. lled; lia. Qed. -Definition leq_universe_n n (φ : constraints) u u' := - forall v, satisfies v φ -> (Z.of_nat n + val v u <= val v u')%Z. -Definition leq_universe0 := leq_universe_n 0. -Definition lt_universe := leq_universe_n 1. + Inductive satisfies0 (v : valuation) : UnivConstraint.t -> Prop := + | satisfies0_Lt (l l' : Level.t) : (val v l < val v l')%u + -> satisfies0 v (l, ConstraintType.Lt, l') + | satisfies0_Le (l l' : Level.t) : (val v l <= val v l')%u + -> satisfies0 v (l, ConstraintType.Le, l') + | satisfies0_Eq (l l' : Level.t) : val v l = val v l' + -> satisfies0 v (l, ConstraintType.Eq, l'). -Definition eq_universe φ u u' - := if check_univs then eq_universe0 φ u u' else True. + Definition satisfies v : ConstraintSet.t -> Prop := + ConstraintSet.For_all (satisfies0 v). -Definition leq_universe φ u u' - := if check_univs then leq_universe0 φ u u' else True. + Definition consistent ctrs := exists v, satisfies v ctrs. -(* ctrs are "enforced" by φ *) + Definition eq_universe0 (φ : ConstraintSet.t) u u' := + forall v, satisfies v φ -> val v u = val v u'. -Definition valid_constraints0 φ ctrs - := forall v, satisfies v φ -> satisfies v ctrs. + Definition leq_universe_n n (φ : ConstraintSet.t) u u' := + forall v, satisfies v φ -> (Z.of_nat n + val v u <= val v u')%u. -Definition valid_constraints φ ctrs - := if check_univs then valid_constraints0 φ ctrs else True. + Definition leq_universe0 := leq_universe_n 0. + Definition lt_universe := leq_universe_n 1. -Lemma valid_subset φ φ' ctrs - : ConstraintSet.Subset φ φ' -> valid_constraints φ ctrs - -> valid_constraints φ' ctrs. -Proof. - unfold valid_constraints. - destruct check_univs; [|trivial]. - intros Hφ H v Hv. apply H. - intros ctr Hc. apply Hv. now apply Hφ. -Qed. + Definition eq_universe φ u u' + := if check_univs then eq_universe0 φ u u' else True. + Definition leq_universe φ u u' + := if check_univs then leq_universe0 φ u u' else True. -(** **** Lemmas about eq and leq **** *) + (* ctrs are "enforced" by φ *) -(** We show that equality and inequality of universes form an equivalence and - a partial order (one w.r.t. the other). - We use classes from [CRelationClasses] for consistency with the rest of the - development which uses relations in [Type] rather than [Prop]. - These definitions hence use [Prop <= Type]. *) + Definition valid_constraints0 φ ctrs + := forall v, satisfies v φ -> satisfies v ctrs. -Global Instance eq_universe0_refl φ : Reflexive (eq_universe0 φ). -Proof. - intros vH s; reflexivity. -Qed. + Definition valid_constraints φ ctrs + := if check_univs then valid_constraints0 φ ctrs else True. -Global Instance eq_universe_refl φ : Reflexive (eq_universe φ). -Proof. - intro s. - unfold eq_universe; destruct check_univs; - [apply eq_universe0_refl|constructor]. -Qed. + Lemma valid_subset φ φ' ctrs + : ConstraintSet.Subset φ φ' -> valid_constraints φ ctrs + -> valid_constraints φ' ctrs. + Proof. + unfold valid_constraints. + destruct check_univs; [|trivial]. + intros Hφ H v Hv. apply H. + intros ctr Hc. apply Hv. now apply Hφ. + Qed. -Global Instance leq_universe0_refl φ : Reflexive (leq_universe0 φ). -Proof. - intros s vH; reflexivity. -Qed. -Global Instance leq_universe_refl φ : Reflexive (leq_universe φ). -Proof. - intro s. - unfold leq_universe; destruct check_univs; - [apply leq_universe0_refl|constructor]. -Qed. + (** **** Lemmas about eq and leq **** *) -Global Instance eq_universe0_sym φ : Symmetric (eq_universe0 φ). -Proof. - intros s s' e vH. symmetry ; eauto. -Qed. + (** We show that equality and inequality of universes form an equivalence and + a partial order (one w.r.t. the other). + We use classes from [CRelationClasses] for consistency with the rest of the + development which uses relations in [Type] rather than [Prop]. + These definitions hence use [Prop <= Type]. *) -Global Instance eq_universe_sym φ : Symmetric (eq_universe φ). -Proof. - unfold eq_universe. destruct check_univs ; eauto. - eapply eq_universe0_sym. -Qed. + Global Instance eq_universe0_refl φ : Reflexive (eq_universe0 φ). + Proof. + intros vH s; reflexivity. + Qed. -Global Instance eq_universe0_trans φ : Transitive (eq_universe0 φ). -Proof. - intros s1 s2 s3 h1 h2 v h. - etransitivity ; try eapply h1 ; eauto. -Qed. + Global Instance eq_universe_refl φ : Reflexive (eq_universe φ). + Proof. + intro s. + unfold eq_universe; destruct check_univs; + [apply eq_universe0_refl|constructor]. + Qed. -Global Instance eq_universe_trans φ : Transitive (eq_universe φ). -Proof. - intros s1 s2 s3. - unfold eq_universe. destruct check_univs ; auto. - intros h1 h2. - eapply eq_universe0_trans ; eauto. -Qed. + Global Instance leq_universe0_refl φ : Reflexive (leq_universe0 φ). + Proof. + intros s vH; reflexivity. + Qed. -Global Instance leq_universe0_trans φ : Transitive (leq_universe0 φ). -Proof. - intros s1 s2 s3 h1 h2 v h. etransitivity. - - eapply h1. assumption. - - eapply h2. assumption. -Qed. + Global Instance leq_universe_refl φ : Reflexive (leq_universe φ). + Proof. + intro s. + unfold leq_universe; destruct check_univs; + [apply leq_universe0_refl|constructor]. + Qed. -Global Instance leq_universe_trans φ : Transitive (leq_universe φ). -Proof. - intros s1 s2 s3. - unfold leq_universe. destruct check_univs ; auto. - intros h1 h2. - eapply leq_universe0_trans ; eauto. -Qed. + Global Instance eq_universe0_sym φ : Symmetric (eq_universe0 φ). + Proof. + intros s s' e vH. symmetry ; eauto. + Qed. + Global Instance eq_universe_sym φ : Symmetric (eq_universe φ). + Proof. + unfold eq_universe. destruct check_univs ; eauto. + eapply eq_universe0_sym. + Qed. -Lemma val_cons v e s - : val v (NEL.cons e s) = Z.max (val1 v e) (val v s). -Proof. - cbn. generalize (val1 v e); clear e. - induction s. - intro; cbn. apply Z.max_comm. - intro; cbn in *. rewrite !IHs. Lia.lia. -Defined. + Global Instance eq_universe0_trans φ : Transitive (eq_universe0 φ). + Proof. + intros s1 s2 s3 h1 h2 v h. + etransitivity ; try eapply h1 ; eauto. + Qed. -Lemma val_sup v s1 s2 : - val v (Universe.sup s1 s2) = Z.max (val v s1) (val v s2). -Proof. - induction s1. - unfold Universe.sup, NEL.app. now rewrite val_cons. - change (Universe.sup (NEL.cons a s1) s2) with (NEL.cons a (Universe.sup s1 s2)). - rewrite !val_cons. rewrite IHs1. Lia.lia. -Qed. + Global Instance eq_universe_trans φ : Transitive (eq_universe φ). + Proof. + intros s1 s2 s3. + unfold eq_universe. destruct check_univs ; auto. + intros h1 h2. + eapply eq_universe0_trans ; eauto. + Qed. -Lemma leq_universe0_sup_l φ s1 s2 : leq_universe0 φ s1 (Universe.sup s1 s2). -Proof. - intros v H. rewrite val_sup. Lia.lia. -Qed. + Global Instance leq_universe0_trans φ : Transitive (leq_universe0 φ). + Proof. + intros s1 s2 s3 h1 h2 v h. etransitivity. + - eapply h1. assumption. + - eapply h2. assumption. + Qed. -Lemma leq_universe0_sup_r φ s1 s2 : leq_universe0 φ s2 (Universe.sup s1 s2). -Proof. - intros v H. rewrite val_sup. Lia.lia. -Qed. + Global Instance leq_universe_trans φ : Transitive (leq_universe φ). + Proof. + intros s1 s2 s3. + unfold leq_universe. destruct check_univs ; auto. + intros h1 h2. + eapply leq_universe0_trans ; eauto. + Qed. -Lemma leq_universe_product φ s1 s2 - : leq_universe φ s2 (Universe.sort_of_product s1 s2). -Proof. - unfold leq_universe; destruct check_univs; [cbn|constructor]. - unfold Universe.sort_of_product; case_eq (Universe.is_prop s2); intro eq. - apply leq_universe0_refl. - apply leq_universe0_sup_r. -Qed. + Lemma eq_leq_universe φ u u' : + eq_universe0 φ u u' <-> leq_universe0 φ u u' /\ leq_universe0 φ u' u. + Proof. + split. + intro H; split; intros v Hv; specialize (H v Hv); lled; lia. + intros [H1 H2] v Hv; specialize (H1 v Hv); specialize (H2 v Hv); lled; lia. + Qed. -(* Rk: [leq_universe φ s1 (sort_of_product s1 s2)] does not hold due - impredicativity. *) -Global Instance eq_universe_leq_universe φ : subrelation (eq_universe φ) (leq_universe φ). -Proof. - unfold eq_universe, leq_universe; destruct check_univs; [|intuition]. - intros u u' HH v Hv. rewrite (HH v Hv). apply BinInt.Z.le_refl. -Qed. + Lemma leq_universe0_sup_l φ s1 s2 : + Universe.is_prop s1 = false -> leq_universe0 φ s1 (Universe.sup s1 s2). + Proof. + intros H v Hv. apply is_prop_val_false with (v:=v) in H. + rewrite val_sup. unfold lle, llt. destruct prop_sub_type; lia. + Qed. + + Lemma leq_universe0_sup_r φ s1 s2 : + Universe.is_prop s2 = false -> leq_universe0 φ s2 (Universe.sup s1 s2). + Proof. + intros H v Hv. apply is_prop_val_false with (v:=v) in H. + rewrite val_sup. unfold lle, llt. destruct prop_sub_type; lia. + Qed. + + Lemma leq_universe0_sup_l' φ s1 s2 : + prop_sub_type -> leq_universe0 φ s1 (Universe.sup s1 s2). + Proof. + intros H v Hv. rewrite val_sup. unfold lle, llt; rewrite H. lia. + Qed. -Global Instance eq_universe0_equivalence φ : Equivalence (eq_universe0 φ) := - {| Equivalence_Reflexive := _ ; - Equivalence_Symmetric := _; - Equivalence_Transitive := _ |}. + Lemma leq_universe0_sup_r' φ s1 s2 : + prop_sub_type -> leq_universe0 φ s2 (Universe.sup s1 s2). + Proof. + intros H v Hv. rewrite val_sup. unfold lle, llt; rewrite H. lia. + Qed. -Global Instance eq_universe_equivalence φ : Equivalence (eq_universe φ) := - {| Equivalence_Reflexive := eq_universe_refl _ ; - Equivalence_Symmetric := eq_universe_sym _; - Equivalence_Transitive := eq_universe_trans _ |}. + Lemma leq_universe_product φ s1 s2 + : leq_universe φ s2 (Universe.sort_of_product s1 s2). + Proof. + unfold leq_universe; destruct check_univs; [cbn|constructor]. + unfold Universe.sort_of_product; case_eq (Universe.is_prop s2); intro eq. + apply leq_universe0_refl. + now apply leq_universe0_sup_r. + Qed. -Global Instance leq_universe_preorder φ : PreOrder (leq_universe φ) := - {| PreOrder_Reflexive := leq_universe_refl _ ; - PreOrder_Transitive := leq_universe_trans _ |}. + (* Rk: [leq_universe φ s1 (sort_of_product s1 s2)] does not hold due to + impredicativity. *) + Global Instance eq_universe_leq_universe φ : subrelation (eq_universe φ) (leq_universe φ). + Proof. + unfold eq_universe, leq_universe; destruct check_univs; [|intuition]. + intros u u' HH v Hv. rewrite (HH v Hv). reflexivity. + Qed. -Global Instance leq_universe0_antisym φ - : Antisymmetric _ (eq_universe0 φ) (leq_universe0 φ). -Proof. - intros t u tu ut. unfold leq_universe0, eq_universe0 in *. - red in tu, ut. - intros v sat. - specialize (tu _ sat). - specialize (ut _ sat). - simpl in tu, ut. Lia.lia. -Qed. + Global Instance eq_universe0_equivalence φ : Equivalence (eq_universe0 φ) := + {| Equivalence_Reflexive := _ ; + Equivalence_Symmetric := _; + Equivalence_Transitive := _ |}. -Global Instance leq_universe_antisym φ - : Antisymmetric _ (eq_universe φ) (leq_universe φ). -Proof. - intros t u tu ut. unfold leq_universe, eq_universe in *. - destruct check_univs; [|trivial]. eapply leq_universe0_antisym; auto. -Qed. + Global Instance eq_universe_equivalence φ : Equivalence (eq_universe φ) := + {| Equivalence_Reflexive := eq_universe_refl _ ; + Equivalence_Symmetric := eq_universe_sym _; + Equivalence_Transitive := eq_universe_trans _ |}. -Global Instance leq_universe_partial_order φ - : PartialOrder (eq_universe φ) (leq_universe φ). -Proof. - intros x y; split. intros eqxy; split. now eapply eq_universe_leq_universe. red. - now eapply eq_universe_leq_universe, symmetry. - intros [l r]. now eapply leq_universe_antisym. -Defined. + Global Instance leq_universe_preorder φ : PreOrder (leq_universe φ) := + {| PreOrder_Reflexive := leq_universe_refl _ ; + PreOrder_Transitive := leq_universe_trans _ |}. + + Global Instance llt_str_order : StrictOrder llt. + Proof. + split. + - intros x H. unfold llt in *; destruct prop_sub_type; lia. + - exact _. + Qed. + Global Instance lle_antisym : Antisymmetric _ eq lle. + Proof. + intros t u [] []; subst; try reflexivity. + exfalso; eapply llt_str_order; etransitivity; eassumption. + Qed. + + Global Instance leq_universe0_antisym φ + : Antisymmetric _ (eq_universe0 φ) (leq_universe0 φ). + Proof. + intros t u tu ut. unfold leq_universe0, eq_universe0 in *. + red in tu, ut. + intros v sat. + specialize (tu _ sat). + specialize (ut _ sat). + simpl in tu, ut. + apply lle_antisym; tea. + Qed. + + Global Instance leq_universe_antisym φ + : Antisymmetric _ (eq_universe φ) (leq_universe φ). + Proof. + intros t u tu ut. unfold leq_universe, eq_universe in *. + destruct check_univs; [|trivial]. eapply leq_universe0_antisym; auto. + Qed. -Definition eq_universe_leq_universe' {cf} φ u u' - := @eq_universe_leq_universe cf φ u u'. -Definition leq_universe_refl' φ u - := @leq_universe_refl φ u. + Global Instance leq_universe_partial_order φ + : PartialOrder (eq_universe φ) (leq_universe φ). + Proof. + intros x y; split. intros eqxy; split. now eapply eq_universe_leq_universe. red. + now eapply eq_universe_leq_universe, symmetry. + intros [l r]. now eapply leq_universe_antisym. + Defined. + + + Definition eq_universe_leq_universe' {cf} φ u u' + := @eq_universe_leq_universe cf φ u u'. + Definition leq_universe_refl' φ u + := @leq_universe_refl φ u. + + Hint Resolve eq_universe_leq_universe' leq_universe_refl' : core. + +End Univ. -Hint Resolve eq_universe_leq_universe' leq_universe_refl' : core. (* This universe is a hack used in plugings to generate fresh universes *) -Definition fresh_universe : universe. exact Universe.type0. Qed. +Definition fresh_universe : Universe.t. exact Universe.type0. Qed. (* This level is a hack used in plugings to generate fresh levels *) -Definition fresh_level : Level.t. exact Level.set. Qed. +Definition fresh_level : Level.t. exact Level.lSet. Qed. -End Univ. -Definition is_prop_sort s := - match Universe.level s with - | Some l => Level.is_prop l - | None => false - end. +(** * Universe substitution + + Substitution of universe levels for universe level variables, used to + implement universe polymorphism. *) + + +(** Substitutable type *) + +Class UnivSubst A := subst_instance : Instance.t -> A -> A. + + +Instance subst_instance_level : UnivSubst Level.t := + fun u l => match l with + | Level.lProp | Level.lSet | Level.Level _ => l + | Level.Var n => List.nth n u Level.lSet + end. + +Instance subst_instance_cstr : UnivSubst UnivConstraint.t := + fun u c => (subst_instance_level u c.1.1, c.1.2, subst_instance_level u c.2). + +Instance subst_instance_cstrs : UnivSubst ConstraintSet.t := + fun u ctrs => ConstraintSet.fold (fun c => ConstraintSet.add (subst_instance_cstr u c)) + ctrs ConstraintSet.empty. + +Instance subst_instance_level_expr : UnivSubst UnivExpr.t := + fun u e => match e with + | UnivExpr.lProp => UnivExpr.lProp + | UnivExpr.npe (NoPropLevel.lSet, _) + | UnivExpr.npe (NoPropLevel.Level _, _) => e + | UnivExpr.npe (NoPropLevel.Var n, b) => + match nth_error u n with + | Some l => match NoPropLevel.of_level l with + | Some l => UnivExpr.npe (l, b) + | None => if b then UnivExpr.type1 else UnivExpr.lProp + end + | None => UnivExpr.npe (NoPropLevel.lSet, b) + end + end. + +Instance subst_instance_univ : UnivSubst Universe.t := + fun u s => Universe.map (subst_instance_level_expr u) s. + +Instance subst_instance_instance : UnivSubst Instance.t := + fun u u' => List.map (subst_instance_level u) u'. + +(** Tests that the term is closed over [k] universe variables *) +Section Closedu. + Context (k : nat). + + Definition closedu_level (l : Level.t) := + match l with + | Level.Var n => (n true + end. + + Definition closedu_level_expr (s : UnivExpr.t) := + closedu_level (UnivExpr.get_level s). + + Definition closedu_universe (u : Universe.t) := + UnivExprSet.for_all closedu_level_expr u. + + Definition closedu_instance (u : Instance.t) := + forallb closedu_level u. +End Closedu. + +(** Universe-closed terms are unaffected by universe substitution. *) +Section UniverseClosedSubst. + Lemma closedu_subst_instance_level u l + : closedu_level 0 l -> subst_instance_level u l = l. + Proof. + destruct l; cbnr. discriminate. + Qed. + + Lemma closedu_subst_instance_level_expr u e + : closedu_level_expr 0 e -> subst_instance_level_expr u e = e. + Proof. + destruct e as [|[[] b]]; cbnr. discriminate. + Qed. + + Lemma closedu_subst_instance_univ u s + : closedu_universe 0 s -> subst_instance_univ u s = s. + Proof. + intro H. apply eq_univ'. + intro e; split; intro He. + - apply Universe.map_spec in He. destruct He as [e' [He' X]]. + rewrite closedu_subst_instance_level_expr in X. + now subst. apply UnivExprSet.for_all_spec in H; proper. + exact (H _ He'). + - apply Universe.map_spec. exists e; split; tas. + symmetry; apply closedu_subst_instance_level_expr. + apply UnivExprSet.for_all_spec in H; proper. now apply H. + Qed. + + Lemma closedu_subst_instance_instance u t + : closedu_instance 0 t -> subst_instance_instance u t = t. + Proof. + intro H. apply forall_map_id_spec. + apply Forall_forall; intros l Hl. + apply closedu_subst_instance_level. + eapply forallb_forall in H; eassumption. + Qed. +End UniverseClosedSubst. + +Hint Resolve closedu_subst_instance_level closedu_subst_instance_level_expr + closedu_subst_instance_univ closedu_subst_instance_instance : substu. + +(** Substitution of a universe-closed instance of the right size + produces a universe-closed term. *) +Section SubstInstanceClosed. + Context (u : Instance.t) (Hcl : closedu_instance 0 u). + + Lemma subst_instance_level_closedu l + : closedu_level #|u| l -> closedu_level 0 (subst_instance_level u l). + Proof. + destruct l; cbnr. + unfold closedu_instance in Hcl. + destruct (nth_in_or_default n u Level.lSet). + - intros _. eapply forallb_forall in Hcl; tea. + - rewrite e; reflexivity. + Qed. + + Lemma subst_instance_level_expr_closedu e : + closedu_level_expr #|u| e -> closedu_level_expr 0 (subst_instance_level_expr u e). + Proof. + destruct e as [|[[] b]]; cbnr. + case_eq (nth_error u n); cbnr. intros [] Hl X; cbnr. + now destruct b. apply nth_error_In in Hl. + eapply forallb_forall in Hcl; tea. discriminate. + Qed. + + Lemma subst_instance_univ_closedu s + : closedu_universe #|u| s -> closedu_universe 0 (subst_instance_univ u s). + Proof. + intro H. apply UnivExprSet.for_all_spec; proper. + intros e He. eapply Universe.map_spec in He. + destruct He as [e' [He' X]]; subst. + apply subst_instance_level_expr_closedu. + apply UnivExprSet.for_all_spec in H; proper. + now apply H. + Qed. + + Lemma subst_instance_instance_closedu t : + closedu_instance #|u| t -> closedu_instance 0 (subst_instance_instance u t). + Proof. + intro H. etransitivity. eapply forallb_map. + eapply forallb_impl; tea. + intros l Hl; cbn. apply subst_instance_level_closedu. + Qed. +End SubstInstanceClosed. -(* Check (_ : forall (cf : checker_flags) Σ, Reflexive (eq_universe Σ)). *) +Hint Resolve subst_instance_level_closedu subst_instance_level_expr_closedu + subst_instance_univ_closedu subst_instance_instance_closedu : substu. diff --git a/template-coq/theories/common/uGraph.v b/template-coq/theories/common/uGraph.v index 23ad1fa52..de266b266 100644 --- a/template-coq/theories/common/uGraph.v +++ b/template-coq/theories/common/uGraph.v @@ -1,74 +1,102 @@ Require Import Nat Bool String BinInt List Relations Lia ssrbool. -Require Import MSetFacts MSetProperties. +Require Import MSetWeakList MSetFacts MSetProperties. From MetaCoq.Template Require Import utils config Universes monad_utils wGraph. Import ListNotations. Import ConstraintType MonadNotation. Local Open Scope nat_scope. -Definition on_Some {A} (P : A -> Prop) : option A -> Prop := - fun x => match x with - | Some x => P x - | None => False - end. - -Definition on_Some_or_None {A} (P : A -> Prop) : option A -> Prop := - fun x => match x with - | Some x => P x - | None => True - end. - -Ltac toProp := - match goal with - | |- is_true (_ apply PeanoNat.Nat.ltb_lt - | |- is_true (_ <=? _) => apply PeanoNat.Nat.leb_le - | |- is_true (_ =? _) => apply PeanoNat.Nat.eqb_eq - | H : is_true (_ apply PeanoNat.Nat.ltb_lt in H - | H : is_true (_ <=? _) |- _ => apply PeanoNat.Nat.leb_le in H - | H : is_true (_ =? _) |- _ => apply PeanoNat.Nat.eqb_eq in H - end. - -Lemma InA_In_eq {A} x (l : list A) : InA eq x l <-> In x l. -Proof. - etransitivity. eapply InA_alt. - firstorder. now subst. -Qed. - +Arguments Z.add : simpl nomatch. +Arguments Nat.leb : simpl nomatch. +Arguments Nat.eqb : simpl nomatch. -Module ConstraintSetFact := WFactsOn UnivConstraintDec ConstraintSet. -Module ConstraintSetProp := WPropertiesOn UnivConstraintDec ConstraintSet. +(** variable levels are levels which are Level or Var *) +Module VariableLevel. + Inductive t := Level (_ : string) | Var (_ : nat). + Definition lt : t -> t -> Prop := + fun x y => match x, y with + | Level _, Var _ => True + | Level s, Level s' => string_lt s s' + | Var n, Var n' => n < n' + | Var _, Level _ => False + end. + Global Instance lt_strorder : StrictOrder lt. + split. + - intros [s|n] H; cbn in H. + unfold string_lt in H. + pose proof (string_compare_eq s s). intuition. + rewrite H in *. discriminate. intuition. + - intros [s1|n1] [s2|n2] [s3|n3]; cbn; intuition. + eapply transitive_string_lt; eassumption. + Qed. + Definition lt_compat : Proper (Logic.eq ==> Logic.eq ==> iff) lt. + intros x y [] z t []; reflexivity. + Qed. + Definition compare : t -> t -> comparison := + fun x y => match x, y with + | Level _, Var _ => Datatypes.Lt + | Level s, Level s' => string_compare s s' + | Var n, Var n' => Nat.compare n n' + | Var _, Level _ => Datatypes.Gt + end. + Definition compare_spec : + forall x y : t, CompareSpec (x = y) (lt x y) (lt y x) (compare x y). + intros [s|n] [s'|n']; cbn; try now constructor. + - eapply CompareSpec_Proper. 2-4: reflexivity. + 2: apply CompareSpec_string. + split; congruence. + - eapply CompareSpec_Proper. 2-4: reflexivity. + 2: apply PeanoNat.Nat.compare_spec. + split; congruence. + Qed. + 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 (PeanoNat.Nat.eq_dec n n'); [left|right]; congruence. + Defined. -Inductive variable_level := mLevel (_ : string) | mVar (_ : nat). + Definition to_noprop (l : t) : NoPropLevel.t := + match l with + | Level s => NoPropLevel.Level s + | Var n => NoPropLevel.Var n + end. -Inductive good_constraint := -(* l <= l' *) -| gc_le : variable_level -> variable_level -> good_constraint -(* l < l' *) -| gc_lt : variable_level -> variable_level -> good_constraint -(* Set < Var n *) -| gc_lt_set : nat -> good_constraint -(* Set = Var n *) -| gc_eq_set : nat -> good_constraint. + Definition to_level (l : t) : Level.t := to_noprop l. + Global Instance Evaluable : Evaluable t + := fun v l => match l with + | Level s => Z.pos (v.(valuation_mono) s) + | Var x => Z.of_nat (v.(valuation_poly) x) + end. +End VariableLevel. -Module GoodConstraintDec. - Definition t : Set := good_constraint. - Definition eq : t -> t -> Prop := eq. +Coercion VariableLevel.to_noprop : VariableLevel.t >-> NoPropLevel.t. + + +Module GoodConstraint. + Inductive t_ := + (* l <= l' *) + | gc_le : VariableLevel.t -> VariableLevel.t -> t_ + (* l < l' *) + | gc_lt : VariableLevel.t -> VariableLevel.t -> t_ + (* Set < Var n *) + | gc_lt_set : nat -> t_ + (* Set = Var n *) + | gc_eq_set : nat -> t_. + Definition t : Set := t_. + Definition eq : t -> t -> Prop := Logic.eq. Definition eq_equiv : RelationClasses.Equivalence eq := _. - Definition variable_level_dec : forall x y : variable_level, {x = y} + {x <> y}. - decide equality. apply string_dec. apply Peano_dec.eq_nat_dec. - Defined. Definition eq_dec : forall x y : t, {eq x y} + {~ eq x y}. unfold eq. - decide equality. all: try apply variable_level_dec. + decide equality. all: try apply VariableLevel.eq_dec. all: apply Peano_dec.eq_nat_dec. Defined. -End GoodConstraintDec. -Module GoodConstraintSet := MSets.MSetWeakList.Make GoodConstraintDec. -Module GoodConstraintSetFact := WFactsOn GoodConstraintDec GoodConstraintSet. -Module GoodConstraintSetProp := WPropertiesOn GoodConstraintDec GoodConstraintSet. +End GoodConstraint. +Module GoodConstraintSet := Make GoodConstraint. +Module GoodConstraintSetFact := WFactsOn GoodConstraint GoodConstraintSet. +Module GoodConstraintSetProp := WPropertiesOn GoodConstraint GoodConstraintSet. Definition GoodConstraintSet_pair x y := GoodConstraintSet.add y (GoodConstraintSet.singleton x). @@ -82,16 +110,13 @@ Proof. apply GoodConstraintSetFact.singleton_1 in H; intuition. Qed. -Definition gc_val0 (v : valuation) (l : variable_level) : nat := - match l with - | mLevel s => Pos.to_nat (v.(valuation_mono) s) - | mVar x => v.(valuation_poly) x - end. +Import VariableLevel GoodConstraint. -Definition gc_satisfies0 v (gc : good_constraint) : bool := + +Definition gc_satisfies0 v (gc : GoodConstraint.t) : bool := match gc with - | gc_le l l' => gc_val0 v l <=? gc_val0 v l' - | gc_lt l l' => gc_val0 v l (val v l <=? val v l')%Z + | gc_lt l l' => (val v l 0 0 =? v.(valuation_poly) l end. @@ -105,7 +130,7 @@ Lemma gc_satisfies_pair v gc1 gc2 : (gc_satisfies0 v gc1 /\ gc_satisfies0 v gc2) <-> gc_satisfies v (GoodConstraintSet_pair gc1 gc2). Proof. - cbn; destruct (GoodConstraintDec.eq_dec gc2 gc1); cbn; + cbn; destruct (GoodConstraint.eq_dec gc2 gc1); cbn; rewrite if_true_false. now destruct e. symmetry. apply andb_and. Defined. @@ -113,7 +138,8 @@ Defined. (* None -> not satisfiable *) (* Some empty -> useless *) (* else: singleton or two elements set (l = l' -> {l<=l', l'<=l}) *) -Definition gc_of_constraint `{checker_flags} (uc : univ_constraint) : option GoodConstraintSet.t +Definition gc_of_constraint `{checker_flags} (uc : UnivConstraint.t) + : option GoodConstraintSet.t := let empty := Some GoodConstraintSet.empty in let singleton := fun x => Some (GoodConstraintSet.singleton x) in let pair := fun x y => Some (GoodConstraintSet_pair x y) in @@ -142,57 +168,38 @@ Definition gc_of_constraint `{checker_flags} (uc : univ_constraint) : option Goo | (Level.Level _, Le, Level.lProp) => None | (Level.Level _, Le, Level.lSet) => None | (Level.Level l, Le, Level.Level l') - => singleton (gc_le (mLevel l) (mLevel l')) - | (Level.Level l, Le, Level.Var n) => singleton (gc_le (mLevel l) (mVar n)) + => singleton (gc_le (Level l) (Level l')) + | (Level.Level l, Le, Level.Var n) => singleton (gc_le (Level l) (Var n)) | (Level.Level _, Eq, Level.lProp) => None | (Level.Level _, Eq, Level.lSet) => None | (Level.Level l, Eq, Level.Level l') - => pair (gc_le (mLevel l) (mLevel l')) (gc_le (mLevel l') (mLevel l)) + => pair (gc_le (Level l) (Level l')) (gc_le (Level l') (Level l)) | (Level.Level l, Eq, Level.Var n) - => pair (gc_le (mLevel l) (mVar n)) (gc_le (mVar n) (mLevel l)) + => pair (gc_le (Level l) (Var n)) (gc_le (Var n) (Level l)) | (Level.Level _, Lt, Level.lProp) => None | (Level.Level _, Lt, Level.lSet) => None | (Level.Level l, Lt, Level.Level l') - => singleton (gc_lt (mLevel l) (mLevel l')) - | (Level.Level l, Lt, Level.Var n) => singleton (gc_lt (mLevel l) (mVar n)) + => singleton (gc_lt (Level l) (Level l')) + | (Level.Level l, Lt, Level.Var n) => singleton (gc_lt (Level l) (Var n)) (* Var _ _ *) | (Level.Var _, Le, Level.lProp) => None | (Level.Var n, Le, Level.lSet) => singleton (gc_eq_set n) - | (Level.Var n, Le, Level.Level l) => singleton (gc_le (mVar n) (mLevel l)) - | (Level.Var n, Le, Level.Var n') => singleton (gc_le (mVar n) (mVar n')) + | (Level.Var n, Le, Level.Level l) => singleton (gc_le (Var n) (Level l)) + | (Level.Var n, Le, Level.Var n') => singleton (gc_le (Var n) (Var n')) | (Level.Var _, Eq, Level.lProp) => None | (Level.Var n, Eq, Level.lSet) => singleton (gc_eq_set n) | (Level.Var n, Eq, Level.Level l) - => pair (gc_le (mVar n) (mLevel l)) (gc_le (mLevel l) (mVar n)) + => pair (gc_le (Var n) (Level l)) (gc_le (Level l) (Var n)) | (Level.Var n, Eq, Level.Var n') - => pair (gc_le (mVar n) (mVar n')) (gc_le (mVar n') (mVar n)) + => pair (gc_le (Var n) (Var n')) (gc_le (Var n') (Var n)) | (Level.Var _, Lt, Level.lProp) => None | (Level.Var _, Lt, Level.lSet) => None - | (Level.Var n, Lt, Level.Level l) => singleton (gc_lt (mVar n) (mLevel l)) - | (Level.Var n, Lt, Level.Var n') => singleton (gc_lt (mVar n) (mVar n')) + | (Level.Var n, Lt, Level.Level l) => singleton (gc_lt (Var n) (Level l)) + | (Level.Var n, Lt, Level.Var n') => singleton (gc_lt (Var n) (Var n')) end. - -Ltac gc_of_constraint_tac := - match goal with - | |- is_true (if _ then true else false) => rewrite if_true_false - | |- is_true (_ apply PeanoNat.Nat.ltb_lt - | |- is_true (_ <=? _) => apply PeanoNat.Nat.leb_le - | |- is_true (_ =? _) => apply PeanoNat.Nat.eqb_eq - | |- is_true (gc_satisfies _ (GoodConstraintSet_pair _ _)) - => apply gc_satisfies_pair; cbn -[Nat.leb Nat.eqb]; split - | H : is_true false |- _ => discriminate H - | H : is_true (if _ then true else false) |- _ => rewrite if_true_false in H - | H : is_true (_ apply PeanoNat.Nat.ltb_lt in H - | H : is_true (_ <=? _) |- _ => apply PeanoNat.Nat.leb_le in H - | H : is_true (_ =? _) |- _ => apply PeanoNat.Nat.eqb_eq in H - | H : is_true (gc_satisfies _ (GoodConstraintSet_pair _ _)) |- _ - => apply gc_satisfies_pair in H; cbn -[Nat.leb Nat.eqb] in H; - destruct H - end. - Section GC. Context `{cf : checker_flags}. @@ -201,28 +208,17 @@ Lemma gc_of_constraint_spec v uc : satisfies0 v uc <-> on_Some (gc_satisfies v) (gc_of_constraint uc). Proof. split. - - destruct 1 ; destruct l, l' ; try constructor ; try reflexivity. - all: cbn -[Nat.leb Nat.eqb GoodConstraintSet_pair] in *. - all: unfold lle, llt in *. - (* all: unfold gc_of_constraint in *. *) - all: destruct prop_sub_type. - all: try solve [ repeat gc_of_constraint_tac ; lia ]. - all: try solve [ cbn ; easy ]. - - destruct uc as [[[] []] []]. - all: cbn - [Nat.leb Nat.eqb GoodConstraintSet_pair]. - all: try (now inversion 1). - all: constructor. - all: unfold lle, llt in *. - all: destruct prop_sub_type. - all: try solve [ - cbn - [Nat.leb Nat.eqb GoodConstraintSet_pair] in * ; lia - ]. - all: try solve [ - cbn - [Nat.leb Nat.eqb GoodConstraintSet_pair] in * ; - try lia ; - repeat gc_of_constraint_tac ; - lia - ]. + - destruct 1; destruct l, l'; try constructor. + all: cbn -[GoodConstraintSet_pair] in *. + all: lled; cbn -[GoodConstraintSet_pair]; try reflexivity. + all: rewrite ?if_true_false; repeat toProp ; try lia. + all: apply gc_satisfies_pair; split; cbn; toProp; try lia. + - destruct uc as [[[] []] []]; intro H; constructor. + all: cbn -[GoodConstraintSet_pair] in *; try contradiction. + all: rewrite ?if_true_false in *; lled; cbn -[GoodConstraintSet_pair] in *; + try contradiction; repeat toProp; try lia. + all: apply gc_satisfies_pair in H; destruct H as [H1 H2]; cbn in *; + repeat toProp; try lia. Qed. @@ -230,15 +226,8 @@ Definition add_gc_of_constraint uc (S : option GoodConstraintSet.t) := S1 <- S ;; S2 <- gc_of_constraint uc ;; ret (GoodConstraintSet.union S1 S2). - (* := match S with *) - (* | None => None *) - (* | Some S1 => match gc_of_constraint uc with *) - (* | None => None *) - (* | Some S2 => Some (GoodConstraintSet.union S1 S2) *) - (* end *) - (* end. *) - -Definition gc_of_constraints (ctrs : constraints) : option GoodConstraintSet.t + +Definition gc_of_constraints (ctrs : ConstraintSet.t) : option GoodConstraintSet.t := ConstraintSet.fold add_gc_of_constraint ctrs (Some GoodConstraintSet.empty). @@ -252,7 +241,7 @@ Proof. etransitivity. eapply iff_forall. intro; eapply imp_iff_compat_r. eapply ConstraintSetFact.elements_iff. set (l := ConstraintSet.elements ctrs). simpl. - transitivity ((forall uc, InA eq uc l -> satisfies0 v uc) /\ + transitivity ((forall uc, InA Logic.eq uc l -> satisfies0 v uc) /\ (forall gc, GoodConstraintSet.In gc S -> gc_satisfies0 v gc)). { intuition. inversion H0. } clearbody S; revert S; induction l; intro S; cbn. @@ -310,16 +299,16 @@ Qed. -Definition gc_leq_universe_n n ctrs u u' - := forall v, gc_satisfies v ctrs -> (Z.of_nat n + val v u <= val v u')%Z. +Definition gc_leq_universe_n n ctrs (u u' : Universe.t) + := forall v, gc_satisfies v ctrs -> (Z.of_nat n + val v u <= val v u')%u. -Definition gc_eq_universe0 φ u u' := +Definition gc_eq_universe0 φ (u u' : Universe.t) := forall v, gc_satisfies v φ -> val v u = val v u'. -Definition gc_leq_universe φ u u' +Definition gc_leq_universe φ (u u' : Universe.t) := if check_univs then gc_leq_universe_n 0 φ u u' else True. -Definition gc_eq_universe φ u u' +Definition gc_eq_universe φ (u u' : Universe.t) := if check_univs then gc_eq_universe0 φ u u' else True. @@ -388,99 +377,14 @@ Qed. End GC. -(* no_prop_levels of the graph are levels which are not Prop *) -(* vtn : "variable to no_prop" *) -Inductive no_prop_level := lSet | vtn (l : variable_level). - -Coercion vtn : variable_level >-> no_prop_level. - -Module VariableLevel. - Definition t := variable_level. - Definition lt : t -> t -> Prop := - fun x y => match x, y with - | mLevel _, mVar _ => True - | mLevel s, mLevel s' => string_lt s s' - | mVar n, mVar n' => n < n' - | mVar _, mLevel _ => False - end. - Definition lt_strorder : StrictOrder lt. - split. - - intros [s|n] H; cbn in H. - unfold string_lt in H. - pose proof (string_compare_eq s s). intuition. - rewrite H in *. discriminate. intuition. - - intros [s1|n1] [s2|n2] [s3|n3]; cbn; intuition. - eapply transitive_string_lt; eassumption. - Qed. - Definition lt_compat : Proper (Logic.eq ==> Logic.eq ==> iff) lt. - intros x y [] z t []; reflexivity. - Qed. - Definition compare : t -> t -> comparison := - fun x y => match x, y with - | mLevel _, mVar _ => Datatypes.Lt - | mLevel s, mLevel s' => string_compare s s' - | mVar n, mVar n' => Nat.compare n n' - | mVar _, mLevel _ => Datatypes.Gt - end. - Definition compare_spec : - forall x y : t, CompareSpec (x = y) (lt x y) (lt y x) (compare x y). - intros [s|n] [s'|n']; cbn; try now constructor. - - eapply CompareSpec_Proper. 2-4: reflexivity. - 2: apply CompareSpec_string. - split; congruence. - - eapply CompareSpec_Proper. 2-4: reflexivity. - 2: apply PeanoNat.Nat.compare_spec. - split; congruence. - Qed. - 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 (PeanoNat.Nat.eq_dec n n'); [left|right]; congruence. - Defined. -End VariableLevel. - -Module NoPropLevel. - Definition t : Set := no_prop_level. - Definition eq : t -> t -> Prop := eq. - Definition eq_equiv : RelationClasses.Equivalence eq := _. - Definition lt : t -> t -> Prop := - fun x y => match x, y with - | lSet, lSet => False - | lSet, vtn _ => True - | vtn v, vtn v' => VariableLevel.lt v v' - | vtn _, lSet => False - end. - Definition lt_strorder : StrictOrder lt. - split. - - intros [|v] H; cbn in H; intuition. - now apply VariableLevel.lt_strorder in H. - - intros [|v1] [|v2] [|v3]; cbn; intuition. - eapply VariableLevel.lt_strorder; eassumption. - Qed. - Definition lt_compat : Proper (Logic.eq ==> Logic.eq ==> iff) lt. - intros x x' H1 y y' H2; now rewrite H1, H2. - Qed. - Definition compare : t -> t -> comparison := - fun x y => match x, y with - | lSet, lSet => Datatypes.Eq - | lSet, vtn _ => Datatypes.Lt - | vtn v, vtn v' => VariableLevel.compare v v' - | vtn _, lSet => Datatypes.Gt - end. - Definition compare_spec : - forall x y : t, CompareSpec (x = y) (lt x y) (lt y x) (compare x y). - intros [|v] [|v']; cbn; try now constructor. - destruct (VariableLevel.compare_spec v v'); constructor; congruence. - Qed. - Definition eq_dec : forall x y : t, {x = y} + {x <> y}. - decide equality. apply VariableLevel.eq_dec. - Defined. -End NoPropLevel. - Module Import wGraph := wGraph.WeightedGraph NoPropLevel. +Local Notation lSet := NoPropLevel.lSet. +(* vtn = variable to noprop *) +Local Notation vtn := VariableLevel.to_noprop. + Definition universes_graph := wGraph.t. Definition init_graph : universes_graph := (VSet.singleton lSet, EdgeSet.empty, lSet). @@ -494,23 +398,15 @@ Proof. rewrite H; constructor. Defined. - -Definition no_prop_of_level l := - match l with - | Level.lProp => None - | Level.lSet => Some lSet - | Level.Level s => Some (vtn (mLevel s)) - | Level.Var n => Some (vtn (mVar n)) - end. - Definition no_prop_levels (X : LevelSet.t) : VSet.t - := LevelSet.fold (fun l X => match no_prop_of_level l with + := LevelSet.fold (fun l X => match NoPropLevel.of_level l with | Some l => VSet.add l X | None => X end) X VSet.empty. -Definition declared := LevelSet.In. + +Definition declared : Level.t -> LevelSet.t -> Prop := LevelSet.In. Definition uctx_invariants (uctx : ContextSet.t) := ConstraintSet.For_all (fun '(l, _, l') => declared l uctx.1 /\ declared l' uctx.1) @@ -519,12 +415,14 @@ Definition uctx_invariants (uctx : ContextSet.t) Definition global_uctx_invariants (uctx : ContextSet.t) := LevelSet.In Level.lSet uctx.1 /\ uctx_invariants uctx. + Definition global_gc_uctx_invariants (uctx : VSet.t * GoodConstraintSet.t) := VSet.In lSet uctx.1 /\ GoodConstraintSet.For_all (fun gc => match gc with | gc_le l l' - | gc_lt l l' => VSet.In l uctx.1 /\ VSet.In l' uctx.1 + | gc_lt l l' => VSet.In (vtn l) uctx.1 + /\ VSet.In (vtn l') uctx.1 | gc_lt_set n - | gc_eq_set n => VSet.In (mVar n) uctx.1 + | gc_eq_set n => VSet.In (NoPropLevel.Var n) uctx.1 end) uctx.2. Definition gc_of_uctx `{checker_flags} (uctx : ContextSet.t) @@ -533,32 +431,11 @@ Definition gc_of_uctx `{checker_flags} (uctx : ContextSet.t) ret (no_prop_levels uctx.1, ctrs). -Definition level_of_variable l := - match l with - | mLevel s => Level.Level s - | mVar n => Level.Var n - end. - -Definition level_of_no_prop l := - match l with - | lSet => Level.lSet - | vtn l => level_of_variable l - end. - -Coercion level_of_no_prop : no_prop_level >-> Level.t. - -Lemma no_prop_of_level_no_prop l - : no_prop_of_level (level_of_no_prop l) = Some l. -Proof. - destruct l; try reflexivity. - destruct l; try reflexivity. -Qed. - -Lemma no_prop_levels_no_prop_level (l : no_prop_level) levels +Lemma no_prop_levels_no_prop_level (l : NoPropLevel.t) levels : declared l levels -> VSet.In l (no_prop_levels levels). Proof. unfold no_prop_levels, declared. rewrite LevelSet.fold_spec. - intro H. apply LevelSetFact.elements_1, InAeq_In in H. + intro H. apply LevelSetFact.elements_1, InA_In_eq in H. cut (In (l : Level.t) (LevelSet.elements levels) \/ VSet.In l VSet.empty); [|now left]. clear H; generalize VSet.empty. @@ -566,11 +443,11 @@ Proof. - intuition. - intros X [[HH|HH]|HH]. + subst a; cbn. apply IHl0. - right. rewrite no_prop_of_level_no_prop. + right. rewrite NoPropLevel.of_to_level. apply VSet.add_spec; intuition. + apply IHl0. now left. + apply IHl0. right. - destruct (no_prop_of_level a); tas. + destruct (NoPropLevel.of_level a); tas. apply VSet.add_spec; intuition. Qed. @@ -649,10 +526,10 @@ Proof. Qed. -Definition edge_of_level (l : variable_level) : EdgeSet.elt := +Definition edge_of_level (l : VariableLevel.t) : EdgeSet.elt := match l with - | mLevel l => (lSet, 1, vtn (mLevel l)) - | mVar n => (lSet, 0, vtn (mVar n)) + | Level l => (lSet, 1, NoPropLevel.Level l) + | Var n => (lSet, 0, NoPropLevel.Var n) end. Definition EdgeSet_pair x y @@ -660,31 +537,21 @@ Definition EdgeSet_pair x y Definition EdgeSet_triple x y z := EdgeSet.add z (EdgeSet_pair x y). -Definition edge_of_constraint (gc : good_constraint) : EdgeSet.elt := +Definition edge_of_constraint (gc : GoodConstraint.t) : EdgeSet.elt := match gc with | gc_le l l' => (vtn l, 0, vtn l') | gc_lt l l' => (vtn l, 1, vtn l') - | gc_lt_set n => (lSet, 1, vtn (mVar n)) - | gc_eq_set n => (vtn (mVar n), 0, lSet) + | gc_lt_set n => (lSet, 1, vtn (Var n)) + | gc_eq_set n => (vtn (Var n), 0, lSet) end. -Definition vtn_inj x y : vtn x = vtn y -> x = y. -Proof. - now inversion 1. -Defined. - -Definition vtn_lSet x : vtn x <> lSet. -Proof. - now inversion 1. -Defined. - Lemma source_edge_of_level g : (edge_of_level g)..s = lSet. Proof. destruct g; reflexivity. Qed. -Lemma target_edge_of_level g : (edge_of_level g)..t = g. +Lemma target_edge_of_level g : (edge_of_level g)..t = vtn g. Proof. destruct g; reflexivity. Qed. @@ -692,7 +559,10 @@ Qed. Definition make_graph (uctx : VSet.t * GoodConstraintSet.t) : wGraph.t := let init_edges := VSet.fold (fun l E => match l with - | vtn l => EdgeSet.add (edge_of_level l) E + | NoPropLevel.Level s => + EdgeSet.add (edge_of_level (Level s)) E + | NoPropLevel.Var n => + EdgeSet.add (edge_of_level (Var n)) E | lSet => E end) uctx.1 EdgeSet.empty in let edges := GoodConstraintSet.fold (fun ctr => EdgeSet.add (edge_of_constraint ctr)) uctx.2 init_edges in @@ -740,13 +610,14 @@ Proof. + intro E. etransitivity. apply IHl. split. * intro HH. destruct HH as [[l' Hl]|HH]. left. exists l'. intuition. - destruct a as [|l']. right; tas. - apply EdgeSet.add_spec in HH. destruct HH. - left. exists l'. intuition. right; tas. + destruct a as [|l'|l']. right; tas. + all: apply EdgeSet.add_spec in HH; destruct HH; + [left|right; tas]. + exists (Level l'); intuition. exists (Var l'); intuition. * intros [[l' [[H1|H1] H2]]|H]. - right. subst a. apply EdgeSet.add_spec. left; tas. - left. exists l'; intuition. - right. destruct a; tas. apply EdgeSet.add_spec. right; tas. + right. subst a. destruct l'; apply EdgeSet.add_spec; left; tas. + destruct l'; left; [exists (Level s)|exists (Var n)]; intuition. + right. destruct a; tas; apply EdgeSet.add_spec; right; tas. Qed. @@ -758,14 +629,18 @@ Proof. destruct He as [[l [Hl He]]|[gc [Hgc He]]]. + subst e. split. rewrite source_edge_of_level. apply Hi. rewrite target_edge_of_level; tas. - + subst e. split. destruct gc; try apply (Hi..2 _ Hgc). apply Hi. - destruct gc; try apply (Hi..2 _ Hgc). apply Hi. + + subst e. split. destruct gc; try apply (Hi.p2 _ Hgc). apply Hi. + destruct gc; try apply (Hi.p2 _ Hgc). apply Hi. - apply Hi. - cbn. intros l Hl. sq. destruct l. constructor. econstructor. 2: constructor. - assert (He: EdgeSet.In (edge_of_level l) (wGraph.E (make_graph uctx))). { - apply make_graph_E. left. exists l. intuition. } - destruct l; eexists; exact He. + assert (He: EdgeSet.In (edge_of_level (Level s)) (wGraph.E (make_graph uctx))). { + apply make_graph_E. left. exists (Level s). intuition. } + eexists; exact He. + econstructor. 2: constructor. + assert (He: EdgeSet.In (edge_of_level (Var n)) (wGraph.E (make_graph uctx))). { + apply make_graph_E. left. exists (Var n). intuition. } + eexists; exact He. Qed. @@ -823,13 +698,13 @@ Ltac simplify_sets := Definition labelling_of_valuation (v : valuation) : labelling := fun x => match x with | lSet => 0 - | vtn (mLevel l) => Pos.to_nat (v.(valuation_mono) l) - | vtn (mVar n) => v.(valuation_poly) n + | NoPropLevel.Level l => Pos.to_nat (v.(valuation_mono) l) + | NoPropLevel.Var n => v.(valuation_poly) n end. Definition valuation_of_labelling (l : labelling) : valuation - := {| valuation_mono := fun s => Pos.of_nat (l (vtn (mLevel s))); - valuation_poly := fun n => l (vtn (mVar n)) |}. + := {| valuation_mono := fun s => Pos.of_nat (l (vtn (Level s))); + valuation_poly := fun n => l (vtn (Var n)) |}. Section MakeGraph. @@ -841,18 +716,17 @@ Section MakeGraph. : forall x, VSet.In x uctx.1 -> labelling_of_valuation (valuation_of_labelling l) x = l x. Proof. - destruct x; cbn. + destruct x; cbnr. - intros _. now apply proj1 in Hl; cbn in Hl. - - destruct l0; cbn. 2: reflexivity. - intro Hs. apply Nat2Pos.id. - assert (HH: EdgeSet.In (lSet, 1, vtn (mLevel s)) (wGraph.E G)). { + - intro Hs. apply Nat2Pos.id. + assert (HH: EdgeSet.In (lSet, 1, vtn (Level s)) (wGraph.E G)). { subst G. apply make_graph_E. left. - exists (mLevel s). intuition. } + exists (Level s). intuition. } apply (proj2 Hl) in HH; cbn in HH. lia. Qed. Lemma make_graph_spec v : - gc_satisfies v ctrs <-> correct_labelling G (labelling_of_valuation v). + gc_satisfies v uctx.2 <-> correct_labelling G (labelling_of_valuation v). Proof. unfold gc_satisfies, correct_labelling. split; intro H. - split. reflexivity. @@ -864,25 +738,19 @@ Section MakeGraph. apply GoodConstraintSet.for_all_spec in H. 2: intros x y []; reflexivity. specialize (H _ Hc). cbn in *. - destruct ctr as [g g0|g g0|n|n]; try apply PeanoNat.Nat.leb_le in H; - unfold gc_val0 in *; cbn in *; try lia. - destruct (valuation_poly v n); [reflexivity|discriminate]. + destruct ctr as [[] []|[] []|n|n]; cbn in *; toProp H; try lia. - apply GoodConstraintSet.for_all_spec. intros x y []; reflexivity. intros gc Hgc. - pose proof (proj2 (make_graph_E uctx (edge_of_constraint gc)) (or_intror (ex_intro _ gc (conj Hgc eq_refl)))) as XX. - specialize (H..2 _ XX). - destruct gc; intro HH; try apply PeanoNat.Nat.leb_le. - 4: apply PeanoNat.Nat.eqb_eq. - cbn in *; unfold gc_val0; lia. - cbn in *; unfold gc_val0; lia. - assumption. - cbn in HH. lia. + pose proof (proj2 (make_graph_E uctx (edge_of_constraint gc)) + (or_intror (ex_intro _ gc (conj Hgc eq_refl)))) as XX. + specialize (H.p2 _ XX). + destruct gc as [[] []|[] []|n|n]; intro HH; cbn in *; toProp; lia. Qed. Corollary make_graph_spec' l : - (* gc_satisfies (valuation_of_labelling l) ctrs <-> correct_labelling G l. *) - correct_labelling G l -> gc_satisfies (valuation_of_labelling l) ctrs. + (* gc_satisfies (valuation_of_labelling l) uctx.2 <-> correct_labelling G l. *) + correct_labelling G l -> gc_satisfies (valuation_of_labelling l) uctx.2. Proof. intro H. apply (make_graph_spec (valuation_of_labelling l)). unfold correct_labelling. intuition. @@ -891,7 +759,7 @@ Section MakeGraph. Qed. Corollary make_graph_spec2 : - gc_consistent ctrs <-> exists l, correct_labelling G l. + gc_consistent uctx.2 <-> exists l, correct_labelling G l. Proof. split. - intros [v H]. exists (labelling_of_valuation v). @@ -945,32 +813,82 @@ Section CheckLeq. uctx (Huctx: global_gc_uctx_invariants uctx) (HC : gc_consistent uctx.2) (HG : G = make_graph uctx). - (** ** Check of leq ** *) + Definition gc_level_declared l + := on_Some_or_None (fun l => VSet.In l uctx.1) (NoPropLevel.of_level l). + + Lemma gc_level_declared_make_graph (l : NoPropLevel.t) : + gc_level_declared l -> VSet.In l (wGraph.V G). + Proof. + intros Hl. red in Hl. rewrite NoPropLevel.of_to_level in Hl; cbn in Hl. + subst G; assumption. + Qed. + + Definition gc_expr_declared e + := on_Some_or_None (fun l => VSet.In l uctx.1) (UnivExpr.get_noprop e). + + Definition gc_levels_declared (u : Universe.t) + := UnivExprSet.For_all gc_expr_declared u. - Lemma val_level_of_variable_level v (l : variable_level) - : val0 v l = Z.of_nat (gc_val0 v l). + + Lemma val_level_of_variable_level v (l : VariableLevel.t) + : val v (l : Level.t) = val v l. Proof. destruct l; cbn; lia. Qed. - Lemma labelling_of_valuation_val0 v (l : no_prop_level) - : Z.of_nat (labelling_of_valuation v l) = val0 v l. + Lemma val_labelling_of_valuation v (l : NoPropLevel.t) + : val v (Universe.make l) = Z.of_nat (labelling_of_valuation v l). + Proof. + destruct l; cbnr. lia. + Qed. + + Lemma val_labelling_of_valuation' v (l : NoPropLevel.t) b : + val v (Universe.make' (UnivExpr.npe (l, b))) + = ((if b then 1 else 0) + Z.of_nat (labelling_of_valuation v l))%Z. Proof. - destruct l; cbn. reflexivity. - destruct l; cbn; rewrite ?Z_of_pos_alt; reflexivity. + destruct l; cbnr. lia. Qed. - Lemma leq_universe_vertices0 n (l l' : no_prop_level) + Lemma val_valuation_of_labelling' L (l : NoPropLevel.t) b + (e := UnivExpr.npe (l, b)) : + gc_level_declared l -> + correct_labelling G L -> + val (valuation_of_labelling L) e = ((if b then 1 else 0) + Z.of_nat (L l))%Z. + Proof. + intros Hl [HG1 HG2]. subst G. simpl in HG1. + destruct l as [|l|l]; rewrite ?HG1; cbnr. + pose proof (make_graph_E uctx (edge_of_level (VariableLevel.Level l))).p2 as H. + forward H. { + left. eexists; split; try reflexivity; tas. } + specialize (HG2 _ H); cbn in HG2. rewrite HG1 in HG2; cbn in HG2. + f_equal. clear -HG2. set (L (NoPropLevel.Level l)) in *; clearbody n. + destruct n; try lia. + rewrite <- Pos.of_nat_succ. lia. + Qed. + + Lemma val_valuation_of_labelling L (l : NoPropLevel.t) : + gc_level_declared l -> + correct_labelling G L -> + val (valuation_of_labelling L) l = Z.of_nat (L l)%Z. + Proof. + intros Hl HL. + exact (val_valuation_of_labelling' L l false Hl HL). + Qed. + + + (** ** Check of leq ** *) + + Lemma leq_universe_vertices0 n (l l' : NoPropLevel.t) : leq_vertices G n l l' -> gc_leq_universe_n n uctx.2 (Universe.make l) (Universe.make l'). Proof. intros H v Hv. subst G. apply make_graph_spec in Hv; tas. - specialize (H _ Hv). cbn. - rewrite <- !labelling_of_valuation_val0. lia. + specialize (H _ Hv). + rewrite !val_labelling_of_valuation. lled; lia. Qed. - Lemma leq_universe_vertices1 n l l' + Lemma leq_universe_vertices1 n (l l' : NoPropLevel.t) (Hl : VSet.In l (wGraph.V G)) (Hl' : VSet.In l' (wGraph.V G)) : gc_leq_universe_n n uctx.2 (Universe.make l) (Universe.make l') -> leq_vertices G n l l'. @@ -979,27 +897,12 @@ Section CheckLeq. pose proof (H _ (make_graph_spec' _ Huctx _ Hv)) as HH. rewrite <- (valuation_labelling_eq _ _ Hv l Hl). rewrite <- (valuation_labelling_eq _ _ Hv l' Hl'). - pose proof (labelling_of_valuation_val0 (valuation_of_labelling v) l). - pose proof (labelling_of_valuation_val0 (valuation_of_labelling v) l'). - cbn in *; lia. + pose proof (val_labelling_of_valuation (valuation_of_labelling v) l). + pose proof (val_labelling_of_valuation (valuation_of_labelling v) l'). + cbn in *; lled; lia. Qed. - (* Definition update_valuation v l k := *) - (* match l with *) - (* | mLevel s => {| valuation_mono := fun s' => if string_dec s s' then k else *) - (* v.(valuation_mono) s'; *) - (* valuation_poly := v.(valuation_poly)|} *) - (* | mVar n => {| valuation_mono := v.(valuation_mono); *) - (* valuation_poly := fun n' => if n =? n' then Pos.to_nat k else *) - (* v.(valuation_poly) n' |} *) - (* end. *) - - (* Lemma Zpos_to_pos z : (z <= Z.pos (Z.to_pos z))%Z. *) - (* Proof. *) - (* destruct z; cbn; lia. *) - (* Qed. *) - - Lemma leq_universe_vertices n (l l' : no_prop_level) + Lemma leq_universe_vertices n (l l' : NoPropLevel.t) (Hl : VSet.In l (wGraph.V G)) (Hl' : VSet.In l' (wGraph.V G)) : gc_leq_universe_n n uctx.2 (Universe.make l) (Universe.make l') <-> leq_vertices G n l l'. @@ -1010,10 +913,9 @@ Section CheckLeq. Qed. - Definition leqb_no_prop_n n (l l' : no_prop_level) + Definition leqb_no_prop_n n (l l' : NoPropLevel.t) := leqb_vertices G n l l'. - Lemma leqb_no_prop_n_spec0 n l l' : leqb_no_prop_n n l l' -> gc_leq_universe_n n uctx.2 (Universe.make l) (Universe.make l'). @@ -1022,7 +924,7 @@ Section CheckLeq. apply leqb_vertices_correct; tas; subst G; exact _. Qed. - Lemma leqb_no_prop_n_spec n l l' + Lemma leqb_no_prop_n_spec n (l l' : NoPropLevel.t) (Hl : VSet.In l uctx.1) (Hl' : VSet.In l' uctx.1) : leqb_no_prop_n n l l' <-> gc_leq_universe_n n uctx.2 (Universe.make l) (Universe.make l'). @@ -1033,295 +935,617 @@ Section CheckLeq. Qed. - Definition leqb_level_n n l l' := - match no_prop_of_level l, no_prop_of_level l' with - | None, None => n =? 0 - | None, Some l' => match n with - | O => true - | S n => leqb_no_prop_n n lSet l' - end - | Some l, None => false - | Some l, Some l' => leqb_no_prop_n n l l' + (* this is function [check_smaller_expr] of kernel/uGraph.ml *) + Definition leqb_expr_n n (e1 e2 : UnivExpr.t) := + match e1, e2 with + | UnivExpr.lProp, UnivExpr.lProp => n =? 0 + | UnivExpr.npe _, UnivExpr.lProp => false + | UnivExpr.lProp, UnivExpr.npe (l, false) => match n with + | O => prop_sub_type + | S n => leqb_no_prop_n n lSet l + end + | UnivExpr.lProp, UnivExpr.npe (l, true) => match n with + | O => prop_sub_type + | S O => true + | S (S n) => leqb_no_prop_n n lSet l + end + | UnivExpr.npe (l1, false), UnivExpr.npe (l2, false) + | UnivExpr.npe (l1, true), UnivExpr.npe (l2, true) => leqb_no_prop_n n l1 l2 + | UnivExpr.npe (l1, true), UnivExpr.npe (l2, false) => leqb_no_prop_n (S n) l1 l2 + | UnivExpr.npe (l1, false), UnivExpr.npe (l2, true) + => leqb_no_prop_n (Peano.pred n) l1 l2 (* surprising! *) end. - Lemma no_prop_of_level_inv {l l0} - : no_prop_of_level l = Some l0 -> level_of_no_prop l0 = l. + (* Non trivial lemma *) + Lemma constraint_strengthening (l1 l2 : NoPropLevel.t) : + gc_level_declared l1 -> gc_level_declared l2 -> + (forall v, gc_satisfies v uctx.2 -> (val v l1 <= 1 + val v l2)%u) -> + forall v, gc_satisfies v uctx.2 -> (val v l1 <= val v l2)%u. Proof. - destruct l; inversion 1; reflexivity. + intros Hl1 Hl2 Hl12. + assert (gc_leq_universe_n 0 uctx.2 (Universe.make l1) (Universe.make l2)). + 2:{ intros v Hv; specialize (H v Hv). now rewrite !Universe.val_make_npl in H. } + apply gc_level_declared_make_graph in Hl1 as Hl1'. + apply gc_level_declared_make_graph in Hl2 as Hl2'. + assert (HG1 : invariants G) by (subst; exact _). + assert (HG2 : acyclic_no_loop G) by (subst; exact _). + apply leq_universe_vertices0. + apply leq_vertices_caract0; tas. + case_eq (lsp G l1 l2); [cbn; intros; lia|]. + pose (K := 12). intro HH. exfalso. + apply correct_labelling_lsp_G' with (K:=K) in HH as HH1; tas. + set (G' := wGraph.G' G l2 l1 K) in HH1; cbn in HH1. + assert (Hs : wGraph.s G = lSet) by now subst G. + rewrite Hs in HH1. + pose proof (make_graph_spec' _ Huctx + (fun x => option_get 0 (lsp G' lSet x))) as HH2. + forward HH2; [now subst G|]. + specialize (Hl12 _ HH2); cbn in Hl12; clear HH2. + assert (HG'1 : invariants G'). { + subst G'; apply HI'; subst G; auto. } + assert (HG'2 : acyclic_no_loop G'). { + subst G'; apply HG'; subst G; auto. } + apply lsp_s in Hl1' as Hl1G; tas. + destruct Hl1G as [n1 Hn1]; rewrite Hs in Hn1. + apply lsp_s in Hl2' as Hl2G; tas. + destruct Hl2G as [n2 Hn2]; rewrite Hs in Hn2. + pose proof (lsp_G'_incr G l2 l1 K lSet l1) as Xl1. + rewrite Hn1 in Xl1. + replace (wGraph.G' G l2 l1 K) with G' in Xl1; [|now subst G]. + case_eq (lsp G' lSet l1); + [|intro XX; now rewrite XX in Xl1]. + intros n1' Hn1'; rewrite Hn1' in *; cbn in *. + pose proof (lsp_G'_incr G l2 l1 K lSet l2) as Xl2. + rewrite Hn2 in Xl2. + replace (wGraph.G' G l2 l1 K) with G' in Xl2; [|now subst G]. + case_eq (lsp G' lSet l2); + [|intro XX; now rewrite XX in Xl2]. + intros n2' Hn2'; rewrite Hn2' in *; cbn in *. + pose proof (lsp_G'_sx G l2 l1 Hl1' Hl2' HH K) as XX. + replace (lsp (wGraph.G' G l2 l1 K) (wGraph.s G) l2) with (Some n2') in XX; + [|subst G G'; congruence]. + replace (lsp (wGraph.G' G l2 l1 K) (wGraph.s G) l1) with (Some n1') in XX; + [|subst G G'; congruence]; cbn in XX. + apply lle_le in Hl12. + rewrite !val_valuation_of_labelling in Hl12; tas. + rewrite Hn1', Hn2' in Hl12; cbn in Hl12; lia. Qed. - Definition gc_level_declared l - := on_Some_or_None (fun l => VSet.In l uctx.1) (no_prop_of_level l). - - Lemma leqb_level_n_spec0 n l l' - : leqb_level_n n l l' - -> gc_leq_universe_n n uctx.2 (Universe.make l) (Universe.make l'). + Lemma leqb_expr_n_spec0 n e e' + : leqb_expr_n n e e' + -> gc_leq_universe_n n uctx.2 (Universe.make' e) (Universe.make' e'). Proof. - unfold leqb_level_n. - case_eq (no_prop_of_level l). - - intros l0 Hl. case_eq (no_prop_of_level l'). - + intros l0' Hl'. intro HH. apply leqb_no_prop_n_spec0 in HH. - now rewrite !no_prop_of_level_inv in HH by eassumption. - + destruct l'; inversion 1. discriminate. - - destruct l; inversion 1. case_eq (no_prop_of_level l'). - + intros l0' Hl'. destruct n. - * intros _ v Hv; cbn. - destruct l'; inversion Hl'; cbn; lia. - * intro HH; apply leqb_no_prop_n_spec0 in HH. - intros v Hv; specialize (HH v Hv); cbn - [Z.of_nat] in *. - erewrite <- (no_prop_of_level_inv Hl'); lia. - + destruct l'; inversion 1. intro HH; toProp. - subst. intros v Hv; lia. + unfold leqb_expr_n. + destruct e as [|[l []]], e' as [|[l' []]]; cbn in *; try discriminate; + intros H v Hv; cbn; + try (apply leqb_no_prop_n_spec0 in H; specialize (H v Hv); cbn in H; + rewrite ?UnivExpr.val_make_npl in H; try (lled; lia)). + - toProp; subst; reflexivity. + - destruct n as [|[|n]]. + * pose proof (NoPropLevel.val_zero l' v); lled; lia. + * pose proof (NoPropLevel.val_zero l' v); lled; lia. + * apply leqb_no_prop_n_spec0 in H. specialize (H v Hv). + cbn in H. rewrite ?UnivExpr.val_make_npl in H. lled; lia. + - destruct n. + * pose proof (NoPropLevel.val_zero l' v); lled; lia. + * apply leqb_no_prop_n_spec0 in H. specialize (H v Hv). + cbn in H. rewrite ?UnivExpr.val_make_npl in H. lled; lia. + - pose proof (NoPropLevel.val_zero l v). + destruct n; lled; lia. Qed. - Lemma leqb_level_n_spec n l l' - (HHl : gc_level_declared l) - (HHl' : gc_level_declared l') - : leqb_level_n n l l' - <-> gc_leq_universe_n n uctx.2 (Universe.make l) (Universe.make l'). + Local Ltac tac0 v := + repeat match goal with + | l : NoPropLevel.t |- _ + => pose proof (NoPropLevel.val_zero l v); + change (id NoPropLevel.t) in l + end; + unfold id in *. + Local Ltac tac1 + := match goal with + | H : gc_leq_universe_n _ _ _ _ |- is_true (leqb_no_prop_n _ _ _) + => let v := fresh "v" in + apply leqb_no_prop_n_spec; tas; try apply Huctx; + [..|intros v Hv; specialize (H v Hv)]; tac0 v; + cbn in *; rewrite ?UnivExpr.val_make_npl in * + | v0 : valuation |- _ => tac0 v0 + end. + + Lemma leqb_expr_n_spec n e e' + (HHl : gc_expr_declared e) + (HHl' : gc_expr_declared e') + : leqb_expr_n n e e' + <-> gc_leq_universe_n n uctx.2 (Universe.make' e) (Universe.make' e'). Proof. - unfold leqb_level_n. unfold gc_level_declared in *. - case_eq (no_prop_of_level l). - - intros l0 Hl. rewrite Hl in HHl; cbn in HHl. - case_eq (no_prop_of_level l'). - + intros l0' Hl'. rewrite Hl' in HHl'; cbn in HHl'. - etransitivity. eapply leqb_no_prop_n_spec; tea. - now erewrite !no_prop_of_level_inv by eassumption. - + destruct l'; inversion 1. - split; [discriminate|]. - intros HH. destruct HC as [v Hv]. - specialize (HH v Hv); cbn in HH. - destruct l; try (now inversion Hl); cbn in HH; lia. - - destruct l; inversion 1. clear HHl. case_eq (no_prop_of_level l'). - + intros l0' Hl'. rewrite Hl' in HHl'; cbn in HHl'. destruct n. - * split; [intros _|reflexivity]. - intros v Hv; cbn. - destruct l'; inversion Hl'; cbn; lia. - * etransitivity. eapply leqb_no_prop_n_spec; tas. - apply Huctx. - split; intros HH v Hv; specialize (HH v Hv); cbn - [Z.of_nat] in *. - -- erewrite <- (no_prop_of_level_inv Hl'); lia. - -- erewrite (no_prop_of_level_inv Hl'); lia. - + destruct l'; inversion 1. split; intro HH; toProp. - subst. intros v Hv; lia. - destruct HC as [v Hv]. specialize (HH v Hv); cbn in HH; lia. + split; [apply leqb_expr_n_spec0|]. + unfold leqb_expr_n. + destruct e as [|[l []]], e' as [|[l' []]]; cbn; intro H; + destruct HC as [v0 Hv0]; pose proof (H v0 Hv0) as H0; cbn in H0. + - toProp. lled; lia. + - destruct n as [|[|n]]; tac1; lled; try reflexivity; try lia. + - destruct n as [|n]; tac1; lled; try reflexivity; try lia. + - tac1; lled; try reflexivity; try lia. + - tac1; lled; try reflexivity; try lia. + - tac1; lled; try reflexivity; try lia. + - tac1; lled; try reflexivity; try lia. + - destruct n as [|n]. + + apply leqb_no_prop_n_spec; tas. + intros v Hv. + simple refine (let HH := constraint_strengthening l l' _ _ H v Hv in _). + * red. rewrite NoPropLevel.of_to_level; tas. + * red. rewrite NoPropLevel.of_to_level; tas. + * clearbody HH. rewrite !Universe.val_make_npl; assumption. + + tac1; lled; try reflexivity; try lia. + - tac1; lled; try reflexivity; try lia. Qed. - Definition leqb_level l l' := negb check_univs || leqb_level_n 0 l l'. - - Definition eqb_level l l' := leqb_level l l' && leqb_level l' l. - Definition eqb_univ_instance (u1 u2 : universe_instance) - := forallb2 eqb_level u1 u2. + (* this is function [exists_bigger] of kernel/uGraph.ml *) + Definition leqb_expr_univ_n n (e1 : UnivExpr.t ) (u : Universe.t) := + if UnivExpr.is_prop e1 && (n =? 0) then + prop_sub_type || Universe.is_prop u + else + let '(e2, u) := Universe.exprs u in + List.fold_left (fun b e2 => leqb_expr_n n e1 e2 || b) + u (leqb_expr_n n e1 e2). - Definition leqb_expr_n n (e1 e2 : Universe.Expr.t) := - match e1.2 && negb (Level.is_prop e1.1), e2.2 && negb (Level.is_prop e2.1) with - | false, false - | true, true => leqb_level_n n e1.1 e2.1 - | true, false => leqb_level_n (S n) e1.1 e2.1 - | false, true => leqb_level_n (Peano.pred n) e1.1 e2.1 (* surprising! *) - end. - - Lemma no_prop_of_level_not_is_prop {l l0} - : no_prop_of_level l = Some l0 -> Level.is_prop l = false. + Lemma no_prop_not_zero_le_zero e n : + UnivExpr.is_prop e && (n =? 0) = false + -> forall v, (0 <= Z.of_nat n + val v e)%Z. Proof. - destruct l; inversion 1; reflexivity. + intros Hp v. apply andb_false_iff in Hp; destruct Hp as [H|H]. + apply (UnivExpr.is_prop_val_false e) with (v:=v) in H. lia. + pose proof (UnivExpr.val_minus_one e v). + destruct n; [discriminate|lia]. Qed. - Lemma no_prop_of_level_is_prop {l} - : no_prop_of_level l = None -> Level.is_prop l = true. + + Lemma leqb_expr_univ_n_spec0 n e1 u + : leqb_expr_univ_n n e1 u + -> gc_leq_universe_n n uctx.2 (Universe.make' e1) u. Proof. - destruct l; inversion 1; reflexivity. + unfold leqb_expr_univ_n, gc_leq_universe_n; cbn. + case_eq (UnivExpr.is_prop e1 && (n =? 0)); intro Hp. { + apply andb_true_iff in Hp; destruct Hp as [He1 Hn]. + apply PeanoNat.Nat.eqb_eq in Hn; subst n; cbn. + intros H v Hv; apply UnivExpr.is_prop_val with (v:=v) in He1; rewrite He1. + apply orb_true_iff in H; destruct H as [H|H]. + pose proof (val_minus_one u v); lled; lia. + apply is_prop_val with (v:=v) in H; lled; lia. } + intros HH v Hv. rewrite val_fold_right. + apply no_prop_not_zero_le_zero with (v:=v) in Hp. + destruct (Universe.exprs u) as [e u']. + rewrite <- !fold_left_rev_right in HH; cbn in *. + induction (List.rev u'); cbn in *. + - apply leqb_expr_n_spec0; tas. + - apply orb_true_iff in HH. destruct HH as [HH|HH]. + + apply leqb_expr_n_spec0 in HH. specialize (HH v Hv); cbn in HH. + lled; lia. + + apply IHl in HH; clear IHl. lled; lia. Qed. - Ltac rew_no_prop := + Import Nbar Datatypes. + + (* Non trivial lemma *) + Lemma gc_leq_universe_n_sup n (l : NoPropLevel.t) b (u : Universe.t) + (e := UnivExpr.npe (l, b)) : + gc_level_declared l -> + gc_levels_declared u -> + gc_leq_universe_n n uctx.2 (Universe.make' e) u -> + exists (e' : UnivExpr.t), UnivExprSet.In e' u + /\ gc_leq_universe_n n uctx.2 (Universe.make' e) (Universe.make' e'). + Proof. + intros Hl Hu H. + assert (HG1 : invariants G) by (subst; exact _). + assert (HG2 : acyclic_no_loop G) by (subst; exact _). + assert (Hs : wGraph.s G = lSet) by now subst G. + + case_eq (lsp G l lSet). + (* case where there is a path from l to Set *) + - intros m Hm. red in H. + assert (Hl' : forall v, gc_satisfies v uctx.2 -> val v l = 0%Z). { + intros v Hv. apply make_graph_spec in Hv. + enough (val v (Universe.make l) <= 0)%Z as HH. { + rewrite Universe.val_make_npl in HH. + pose proof (NoPropLevel.val_zero l v); lia. } + rewrite <- HG in Hv. + eapply correct_labelling_lsp in Hm; tea. + cbn in Hm. rewrite val_labelling_of_valuation; lia. } + assert (UnivExprSet.for_all + (fun ei => match ei with + | UnivExpr.lProp => true + | UnivExpr.npe (li, bi) => + le_lt_dec (Some (if bi then 1 else 0) + + Some (if b then 0 else 1) + + lsp G (wGraph.s G) li) + (Some n) + end)%nbar + u = false) as HH. { + apply not_true_iff_false; intro HH. + pose proof (lsp_correctness G) as Hlab. + set (lab := fun x => option_get 0 (lsp G (wGraph.s G) x)) in *. + pose proof (make_graph_spec' _ Huctx lab) as Hv. + forward Hv; [now rewrite <- HG|]. + specialize (H _ Hv); cbn in H. rewrite Hl' in H; tas. + subst e; clear Hl' l Hl m Hm Hv. + apply lle_le in H. + apply UnivExprSet.for_all_spec in HH; proper. + apply val_le_caract in H. + destruct H as [ei [Hei H]]. specialize (HH ei Hei); cbn in HH. + specialize (Hu ei Hei). + destruct ei as [|[li bi]]; cbn in *; [destruct b; lia|]. + destruct (lsp_s G li) as [ni Hni]. + { now rewrite HG. } + rewrite Hni in HH; cbn in HH. + match goal with + | H : ssrbool.is_left ?X = true |- _ => + destruct X as [HH'|?]; [|discriminate]; clear H + end. + assert (val (valuation_of_labelling lab) li = Z.of_nat ni) as XX. { + rewrite val_valuation_of_labelling; tas. + subst lab; cbn; now rewrite Hni. + { red. now rewrite NoPropLevel.of_to_level. } } + rewrite XX in H; clear Hni XX. + cbn in H; destruct b, bi; lia. } + apply UnivExprSet_for_all_false in HH. + apply UnivExprSet.exists_spec in HH; proper. + destruct HH as [[|[li bi]] [He' HH]]; [discriminate|]. + eexists; split; tea. + match goal with + | H : ~~ ssrbool.is_left ?X = true |- _ => + destruct X as [HH'|HH']; try discriminate; clear H + end. + cbn in HH'. case_eq (lsp G (wGraph.s G) li). + 2: intros X; rewrite X in HH'; destruct bi, b; contradiction. + intros nl Hnl v Hv; rewrite Hnl in HH'. + rewrite (val_labelling_of_valuation' _ li); cbn. + rewrite (Hl' _ Hv); clear Hl'. + apply make_graph_spec in Hv; tas. rewrite <- HG in Hv. + apply (correct_labelling_lsp _ Hnl) in Hv. + rewrite Hs in Hv; cbn in *. + destruct bi, b; lled; lia. + + (* case where there is no path from l to Set *) + - intros HlSet. subst e. + assert (UnivExprSet.for_all + (fun ei => match ei with + | UnivExpr.lProp => true + | UnivExpr.npe (li, bi) => + le_lt_dec (Some (if bi then 1 else 0) + + Some (if b then 0 else 1) + + lsp G l li) + (Some n) + end)%nbar + u = false) as HH. { + apply not_true_iff_false; intro HH. + assert (Hl' : VSet.In l (wGraph.V G)). { + red in Hl; rewrite NoPropLevel.of_to_level in Hl; cbn in Hl; + now subst G. } + destruct (lsp_s G _ Hl') as [nl Hnl]; cbn in Hnl. + + assert (exists K, nl <= K /\ + UnivExprSet.For_all + (fun ei => match ei with + | UnivExpr.lProp => True + | UnivExpr.npe (li, bi) => + match lsp G (wGraph.s G) li with + | None => True + | Some ni => (if bi then 1 else 0) + ni < K + end + end) u) as XX. { + exists (UnivExprSet.fold + (fun ei K => match ei with + | UnivExpr.lProp => K + | UnivExpr.npe (li, bi) => + match lsp G (wGraph.s G) li with + | None => K + | Some ni => Nat.max K (S + ((if bi then 1 else 0) + ni)) + end + end) u nl). + clear -Hu HG HG1. split. + - rewrite UnivExprSet.fold_spec. rewrite <- fold_left_rev_right. + induction (List.rev (UnivExprSet.elements u)). reflexivity. + cbn. destruct a as [|[li bi]]; tas. + destruct (lsp G (wGraph.s G) li); tas; lia. + - intros [|[li bi]] Hei; trivial. + specialize (Hu _ Hei); cbn in Hu. + destruct (lsp_s G li) as [ni' Hni']. + { now subst G. } + rewrite Hni'. + rewrite UnivExprSet.fold_spec. rewrite <- fold_left_rev_right. + apply UnivExprSetFact.elements_1, InA_In_eq, in_rev in Hei. + change (In (UnivExpr.npe (li, bi)) + (@List.rev UnivExprSet.elt (UnivExprSet.elements u))) in Hei. + induction (List.rev (UnivExprSet.elements u)); inv Hei. + + subst a; cbn. rewrite Hni'. lia. + + specialize (IHl H). cbn. destruct a as [|[li' bi']]; [lia|]. + destruct (lsp G (wGraph.s G) li'); lia. } + destruct XX as [K [HK1 HK2]]. + assert (Hs' : VSet.In lSet (wGraph.V G)). { + rewrite <- Hs; apply HG1. } + set (G' := wGraph.G' G lSet l K) in *. + assert (HG'1 : invariants G'). { + subst G'; apply HI'; tas. } + assert (HG'2 : acyclic_no_loop G'). { + subst G'; apply HG'; tas. } + apply correct_labelling_lsp_G' with (K:=K) in HlSet as Hlab; tas. + fold G' in Hlab; cbn in Hlab. + set (lab := fun x => option_get 0 (lsp G' (wGraph.s G) x)) in *. + pose proof (make_graph_spec' _ Huctx lab) as Hv. + forward Hv; [now rewrite <- HG|]. + specialize (H _ Hv); clear Hv. + rewrite Universe.val_make' in H. + rewrite val_valuation_of_labelling' in H; tas. + + apply lle_le in H. apply val_le_caract in H. + destruct H as [ei [Hei H]]. + apply UnivExprSet.for_all_spec in HH; proper. + specialize (HH _ Hei); cbn in HH. + specialize (Hu _ Hei). + destruct ei as [|[li bi]]; cbn in H; [destruct b; lia|]. + rewrite val_valuation_of_labelling in H; tas. + 2:{ red. now rewrite NoPropLevel.of_to_level. } + match goal with + | H : ssrbool.is_left ?X = true |- _ => + destruct X as [HH'|HH']; try discriminate; clear H + end. + assert (lab l = K) as XX. { + subst lab; cbn. subst G'. rewrite Hs in *. + rewrite lsp_G'_spec_left; tas. rewrite Hnl. + unfold lsp. rewrite acyclic_lsp0_xx; tas. cbn; lia. } + rewrite XX in H; clear XX Hnl Hl Hl'. + destruct (lsp_s G li) as [ni Hni]. + { now subst G. } + specialize (HK2 _ Hei); cbn in HK2. rewrite Hni in HK2. + + case_eq (lsp G l li). + - intros ki Hki. rewrite Hki in HH'; cbn in HH'. + assert (lab li = K + ki) as XX. { + subst lab; cbn. subst G'. rewrite Hs in *. + rewrite lsp_G'_spec_left; tas. rewrite Hki. + rewrite Hni; cbn; lia. } + rewrite XX in H. destruct b, bi; cbn in *; lia. + + - intro Hki. + assert (lab li = ni) as XX. { + subst lab; cbn. subst G'. rewrite Hs in *. + rewrite lsp_G'_spec_left; tas. now rewrite Hki, Hni. } + rewrite XX in H. destruct b, bi; lia. } + + apply UnivExprSet_for_all_false in HH. + apply UnivExprSet.exists_spec in HH; proper. + destruct HH as [[|[li bi]] [He' HH]]; [discriminate|]. + eexists; split; tea. match goal with - | H : no_prop_of_level _ = Some _ |- _ - => rewrite !(no_prop_of_level_not_is_prop H) in * - | H : no_prop_of_level _ = None |- _ - => rewrite !(no_prop_of_level_is_prop H) in * + | H : ~~ ssrbool.is_left ?X = true |- _ => + destruct X as [HH'|HH']; try discriminate; clear H end. + cbn in HH'. case_eq (lsp G l li). + 2: intros X; rewrite X in HH'; destruct bi, b; contradiction. + intros nl Hnl v Hv; rewrite Hnl in HH'. + apply make_graph_spec in Hv; tas. rewrite <- HG in Hv. + apply (correct_labelling_lsp _ Hnl) in Hv. + rewrite !val_labelling_of_valuation'. + destruct bi, b; cbn in *; lled; lia. + Qed. - Opaque Z.of_nat. - - Conjecture constraint_strengthening : forall l l', - (forall v, gc_satisfies v uctx.2 -> (val0 v l <= val0 v l' + 1)%Z) -> - forall v, gc_satisfies v uctx.2 -> (val0 v l <= val0 v l')%Z. - - Lemma leqb_expr_n_spec0 n e e' - : leqb_expr_n n e e' - -> gc_leq_universe_n n uctx.2 (Universe.make' e) (Universe.make' e'). + Lemma leqb_expr_S_n_Prop_Set n e + (He : gc_expr_declared e) + : leqb_expr_n (S n) UnivExpr.prop e = leqb_expr_n n UnivExpr.set e. Proof. - unfold leqb_expr_n. - destruct e as [l []], e' as [l' []]; cbn. - - case_eq (no_prop_of_level l); [intros l0 Hl|intros Hl]; rew_no_prop; - (case_eq (no_prop_of_level l'); [intros l0' Hl'|intros Hl']); rew_no_prop; cbn; - intro HH; apply leqb_level_n_spec0 in HH; tas. - + intros v Hv; specialize (HH v Hv); cbn in *; - unfold val1 in *; cbn in *; repeat rew_no_prop; cbn in *; try lia. - + intros v Hv; specialize (HH v Hv); cbn in *; - unfold val1 in *; cbn in *; repeat rew_no_prop; cbn in *; try lia. - + destruct n. - * unfold gc_leq_universe_n. cbn. - unfold val1; cbn. repeat rew_no_prop; cbn. - intros v Hv; specialize (HH v Hv); cbn in *; lia. - * intros v Hv; specialize (HH v Hv); cbn in *; - unfold val1 in *; cbn in *; repeat rew_no_prop; cbn in *; try lia. - + intros v Hv; specialize (HH v Hv); cbn in *; - unfold val1 in *; cbn in *; repeat rew_no_prop; cbn in *; try lia. - - case_eq (no_prop_of_level l); [intros l0 Hl|intros Hl]; rew_no_prop; - cbn; (intro HH; eapply leqb_level_n_spec0 in HH; tas). - + intros v Hv; specialize (HH v Hv); cbn in *; - unfold val1 in *; cbn in *; repeat rew_no_prop; cbn in *; try lia. - + intros v Hv; specialize (HH v Hv); cbn in *; - unfold val1 in *; cbn in *; repeat rew_no_prop; cbn in *; try lia. - - case_eq (no_prop_of_level l'); [intros l0 Hl|intros Hl]; rew_no_prop; - cbn; (intro HH; eapply leqb_level_n_spec0 in HH; tas). - + destruct n. - * unfold gc_leq_universe_n. cbn. - unfold val1; cbn. repeat rew_no_prop; cbn. - intros v Hv; specialize (HH v Hv); cbn in *; lia. - * intros v Hv; specialize (HH v Hv); cbn in *; - unfold val1 in *; cbn in *; repeat rew_no_prop; cbn in *; try lia. - + intros v Hv; specialize (HH v Hv); cbn in *; - unfold val1 in *; cbn in *; repeat rew_no_prop; cbn in *; try lia. - - intro HH; eapply leqb_level_n_spec0 in HH; tas. + destruct e as [|[l []]]; cbnr. + destruct n; cbnr. + symmetry. apply leqb_no_prop_n_spec; tas. apply Huctx. + intros v Hv; cbn. + destruct l; cbn; lled; lia. Qed. - Lemma leqb_expr_n_spec n e e' - (HHl : gc_level_declared e.1) - (HHl' : gc_level_declared e'.1) - : leqb_expr_n n e e' - <-> gc_leq_universe_n n uctx.2 (Universe.make' e) (Universe.make' e'). + Lemma leqb_expr_univ_n_spec n e1 u + (He1 : gc_expr_declared e1) + (Hu : gc_levels_declared u) + : leqb_expr_univ_n n e1 u + <-> gc_leq_universe_n n uctx.2 (Universe.make' e1) u. Proof. - unfold leqb_expr_n. - destruct e as [l []], e' as [l' []]; cbn. - - case_eq (no_prop_of_level l); [intros l0 Hl|intros Hl]; rew_no_prop; - (case_eq (no_prop_of_level l'); [intros l0' Hl'|intros Hl']); rew_no_prop; - cbn; (etransitivity; [eapply leqb_level_n_spec; tas|]). - + split; intros HH v Hv; specialize (HH v Hv); cbn in *; - unfold val1 in *; cbn in *; repeat rew_no_prop; cbn in *; try lia. - + split; intros HH v Hv; specialize (HH v Hv); cbn in *; - unfold val1 in *; cbn in *; repeat rew_no_prop; cbn in *; try lia. - + destruct n. - * unfold gc_leq_universe_n. cbn. - unfold val1; cbn. repeat rew_no_prop; cbn. split. - -- intros HH v Hv; specialize (HH v Hv); lia. - -- apply constraint_strengthening. - * split; intros HH v Hv; specialize (HH v Hv); cbn in *; - unfold val1 in *; cbn in *; repeat rew_no_prop; cbn in *; try lia. - + split; intros HH v Hv; specialize (HH v Hv); cbn in *; - unfold val1 in *; cbn in *; repeat rew_no_prop; cbn in *; try lia. - - case_eq (no_prop_of_level l); [intros l0 Hl|intros Hl]; rew_no_prop; - cbn; (etransitivity; [eapply leqb_level_n_spec; tas|]). - + split; intros HH v Hv; specialize (HH v Hv); cbn in *; - unfold val1 in *; cbn in *; repeat rew_no_prop; cbn in *; try lia. - + split; intros HH v Hv; specialize (HH v Hv); cbn in *; - unfold val1 in *; cbn in *; repeat rew_no_prop; cbn in *; try lia. - - case_eq (no_prop_of_level l'); [intros l0 Hl|intros Hl]; rew_no_prop; - cbn; (etransitivity; [eapply leqb_level_n_spec; tas|]). - + destruct n. - * unfold gc_leq_universe_n. cbn. - unfold val1; cbn. repeat rew_no_prop; cbn. split. - -- intros HH v Hv; specialize (HH v Hv); lia. - -- apply constraint_strengthening. - * split; intros HH v Hv; specialize (HH v Hv); cbn in *; - unfold val1 in *; cbn in *; repeat rew_no_prop; cbn in *; try lia. - + split; intros HH v Hv; specialize (HH v Hv); cbn in *; - unfold val1 in *; cbn in *; repeat rew_no_prop; cbn in *; try lia. - - etransitivity; [eapply leqb_level_n_spec; tas|]. - split; intros HH v Hv; specialize (HH v Hv); cbn in *; lia. + split; [apply leqb_expr_univ_n_spec0|]. + unfold leqb_expr_univ_n; intro HH. + case_eq (UnivExpr.is_prop e1 && (n =? 0)); intro Hp. { + apply andb_true_iff in Hp; destruct Hp as [Hp1 Hn]. + apply PeanoNat.Nat.eqb_eq in Hn; subst n. + destruct HC as [v0 Hv0]. specialize (HH v0 Hv0); cbn in HH. + apply UnivExpr.is_prop_val with (v:=v0) in Hp1; rewrite Hp1 in HH. + lled; cbnr. apply val_is_prop with (v:=v0). lia. } + case_eq (Universe.exprs u). intros e u' ee. + assert (Hu': gc_expr_declared e /\ Forall gc_expr_declared u'). { + split. apply Hu. apply Universe.In_exprs. left. now rewrite ee. + apply Forall_forall. intros e' He'. apply Hu. + apply Universe.In_exprs. right. now rewrite ee. } + (* We replace the Prop+1 case by Set *) + assert (exists l1 b1 n', let e1' := UnivExpr.npe (l1, b1) in + gc_expr_declared e1' /\ + gc_leq_universe_n n' uctx.2 (Universe.make' e1') u /\ + fold_left (fun b e2 => leqb_expr_n n e1 e2 || b) u' + (leqb_expr_n n e1 e) = + fold_left (fun b e2 => leqb_expr_n n' e1' e2 || b) u' + (leqb_expr_n n' e1' e)) as XX. { + destruct e1 as [|[l1 b1]]; [|exists l1, b1, n; repeat split; tas]. + destruct n as [|n]; [discriminate|]; clear Hp. + exists NoPropLevel.lSet, false, n. repeat split. + - apply Huctx. + - intros v Hv; specialize (HH v Hv); cbn in *. lled; lia. + - clear ee u Hu HH. rewrite <- !fold_left_rev_right. + destruct Hu' as [He Hu']. apply rev_Forall in Hu'. + induction (List.rev u'); cbn -[leqb_expr_n]. + + apply leqb_expr_S_n_Prop_Set; tas. + + inv Hu'. f_equal. apply leqb_expr_S_n_Prop_Set; tas. + apply IHl; tas. } + clear He1 HH Hp. + destruct XX as [l1 [b1 [n' [He1 [HH XX]]]]]; rewrite XX; clear XX e1. + apply gc_leq_universe_n_sup in HH; tas. + 2:{ red. now rewrite NoPropLevel.of_to_level. } + destruct HH as [e' [He' HH]]. apply leqb_expr_n_spec in HH; tas. + 2:{ now apply Hu. } + apply Universe.In_exprs in He'. rewrite ee in He'; cbn in He'. + rewrite <- !fold_left_rev_right. + clear -He' HH. destruct He' as [H|H]; [subst|]. + - induction (List.rev u'); tas. cbn -[leqb_expr_n]. now rewrite IHl, orb_true_r. + - apply In_rev in H. change (@In UnivExpr.t e' (List.rev u')) in H. + induction (List.rev u'); tas; cbn -[leqb_expr_n]; invs H. + now rewrite HH. now rewrite IHl, orb_true_r. Qed. - Fixpoint leqb_univ_expr_n n (u : universe) (e2 : Universe.Expr.t) := - match u with - | NEL.sing e1 => leqb_expr_n n e1 e2 - | NEL.cons e1 u => leqb_expr_n n e1 e2 && leqb_univ_expr_n n u e2 - end. - - Definition gc_levels_declared (u : universe) - := List.Forall (fun e => gc_level_declared e.1) (NEL.to_list u). + (* this is function [real_check_leq] of kernel/uGraph.ml *) + Definition leqb_universe_n n (u1 u2 : Universe.t) := + if Universe.is_prop u1 && (n =? 0) then + prop_sub_type || Universe.is_prop u2 + else + let '(e1, u1) := Universe.exprs u1 in + List.fold_left (fun b e1 => ((UnivExpr.is_prop e1 && (n =? 0)) + || leqb_expr_univ_n n e1 u2) && b) + u1 ((UnivExpr.is_prop e1 && (n =? 0)) + || leqb_expr_univ_n n e1 u2). + + Lemma no_prop_not_zero_le_zero' u n : + Universe.is_prop u && (n =? 0) = false + -> forall v, (0 <= Z.of_nat n + val v u)%Z. + Proof. + intros Hp v. apply andb_false_iff in Hp; destruct Hp as [H|H]. + apply (is_prop_val_false u) with (v:=v) in H. lia. + pose proof (val_minus_one u v). + destruct n; [discriminate|lia]. + Qed. - Lemma leqb_univ_expr_n_spec0 n u e2 - : leqb_univ_expr_n n u e2 - -> gc_leq_universe_n n uctx.2 u (Universe.make' e2). + Lemma leqb_universe_n_spec0 n u1 u2 + : leqb_universe_n n u1 u2 -> gc_leq_universe_n n uctx.2 u1 u2. Proof. - induction u; cbn. - - apply leqb_expr_n_spec0; tas. - - intro HH. apply andb_true_iff in HH. - destruct HH as [H1 H2]. apply IHu in H2; tas. - eapply leqb_expr_n_spec0 in H1; tas. - intros v Hv; specialize (H1 v Hv); specialize (H2 v Hv). - rewrite val_cons; cbn in *; lia. + unfold leqb_universe_n, gc_leq_universe_n. + case_eq (Universe.is_prop u1 && (n =? 0)); intro Hp. { + intros H v Hv. + apply andb_true_iff in Hp; destruct Hp as [Hu1 Hn]. + apply PeanoNat.Nat.eqb_eq in Hn; subst n. + apply is_prop_val with (v:=v) in Hu1; rewrite Hu1; cbn. + apply orb_true_iff in H; destruct H as [H|H]. + pose proof (val_minus_one u2 v); lled; lia. + apply is_prop_val with (v:=v) in H; now rewrite H. } + intros HH v Hv. + apply no_prop_not_zero_le_zero' with (v:=v) in Hp. + rewrite (val_fold_right u1) in *. + destruct (Universe.exprs u1) as [e1 u1']; clear u1. + rewrite <- !fold_left_rev_right in *; cbn in *. + induction (List.rev u1'); cbn in *. + - apply orb_true_iff in HH; destruct HH as [H|H]. + + exfalso. apply andb_true_iff in H; destruct H as [He1 Hn]. + apply UnivExpr.is_prop_val with (v:=v) in He1. + destruct n; [|discriminate]. lia. + + apply leqb_expr_univ_n_spec0 in H. specialize (H v Hv); cbn in H. + lled; lia. + - set (z := (fold_right (fun e x => Z.max (val v e) x) (val v e1) l)) in *. + apply andb_true_iff in HH; destruct HH as [H HH]. + apply orb_true_iff in H; destruct H as [Han|H]. + + assert (ee : Z.max (val v a) z = z). { + apply andb_true_iff in Han; destruct Han as [Ha Hn]. + apply UnivExpr.is_prop_val with (v:=v) in Ha. + destruct n; [|discriminate]. lia. } + rewrite ee in *; clear ee. apply IHl; tas. + + apply leqb_expr_univ_n_spec0 in H. specialize (H v Hv). cbn in H. + destruct (Z.max_dec (val v a) z) as [ee|ee]; rewrite ee in *. + * lled; lia. + * apply IHl; tas. Qed. - Lemma leqb_univ_expr_n_spec n u e2 - (Hu : gc_levels_declared u) - (He2 : gc_level_declared e2.1) - : leqb_univ_expr_n n u e2 - <-> gc_leq_universe_n n uctx.2 u (Universe.make' e2). + Lemma leqb_universe_n_spec n u1 u2 + (Hu1 : gc_levels_declared u1) + (Hu2 : gc_levels_declared u2) + : leqb_universe_n n u1 u2 + <-> gc_leq_universe_n n uctx.2 u1 u2. Proof. - induction u; cbn. - - apply leqb_expr_n_spec; tas. now inversion Hu. - - etransitivity. apply andb_true_iff. - inversion_clear Hu. - etransitivity. eapply and_iff_compat_l. apply IHu; tas. - etransitivity. eapply and_iff_compat_r. eapply leqb_expr_n_spec; tas. - split. - + intros [H1 H2] v Hv; specialize (H1 v Hv); specialize (H2 v Hv). - rewrite val_cons; cbn in *; lia. - + intro HH; split; intros v Hv; specialize (HH v Hv); - rewrite val_cons in HH; cbn in *; lia. + split; [apply leqb_universe_n_spec0|]. + unfold leqb_universe_n; intro HH. + case_eq (Universe.is_prop u1 && (n =? 0)); intro Hp. { + toProp Hp as [Hp1 Hp2]. toProp Hp2; subst n. + destruct HC as [v0 Hv0]. specialize (HH v0 Hv0); cbn in HH. + apply is_prop_val with (v:=v0) in Hp1; rewrite Hp1 in HH. + lled; cbnr. apply val_is_prop with (v:=v0). lia. } + + case_eq (Universe.exprs u1); intros e1 uu1 Huu1. + rewrite (fold_left_andb_forallb (fun e => _)). + pose proof (Universe.exprs_spec' u1) as X; rewrite Huu1 in X; cbn in X. + rewrite X. apply forallb_Forall. apply Forall_forall. + intros ei Hei. + case_eq (UnivExpr.is_prop ei && (n =? 0)); cbnr; intro Hpi. + apply InA_In_eq, UnivExprSetFact.elements_2 in Hei. + specialize (Hu1 _ Hei). + eapply leqb_expr_univ_n_spec; tas. + intros v Hv. specialize (HH v Hv). rewrite Universe.val_make'. + apply no_prop_not_zero_le_zero with (v:=v) in Hpi. + apply lle_le in HH. apply le_lle'; tas. + apply Z.le_add_le_sub_l. apply Z.le_add_le_sub_l in HH. + eapply (val_ge_caract u1 v (val v u2 - Z.of_nat n)%Z).p2 in HH; tea. Qed. - Definition leqb_univ_expr u e2 := - negb check_univs || leqb_univ_expr_n 0 u e2. - (* This function is correct but not complete! *) - Fixpoint try_leqb_universe_n n (u1 u2 : universe) := - match u2 with - | NEL.sing e => leqb_univ_expr_n n u1 e - | NEL.cons e u2 => leqb_univ_expr_n n u1 e || try_leqb_universe_n n u1 u2 - end. + Definition check_leqb_universe (u1 u2 : Universe.t) := + ~~ check_univs + || (prop_sub_type && Universe.is_prop u1) + || Universe.eqb u1 u2 + || leqb_universe_n 0 u1 u2. - Lemma try_leqb_universe_n_spec n u1 u2 - : try_leqb_universe_n n u1 u2 -> gc_leq_universe_n n uctx.2 u1 u2. + Lemma check_leqb_universe_spec u1 u2 + (Hu1 : gc_levels_declared u1) + (Hu2 : gc_levels_declared u2) + : check_leqb_universe u1 u2 <-> gc_leq_universe uctx.2 u1 u2. Proof. - induction u2; cbn. - - apply leqb_univ_expr_n_spec0; tas. - - intro HH; apply orb_true_iff in HH; destruct HH as [HH|HH]; - [apply leqb_univ_expr_n_spec0 in HH; tas|apply IHu2 in HH; tas]; - intros v Hv; specialize (HH v Hv); rewrite val_cons; cbn in *; lia. + unfold check_leqb_universe, gc_leq_universe. + destruct check_univs; split; cbn; trivial; intro H. + - toProp H as [H|H]; [|now apply leqb_universe_n_spec0]. + toProp H as [H|H]. + + toProp H as [H1 H2]. + intros v Hv. rewrite is_prop_val; tas. + apply le_lle; tas. apply val_minus_one. + + apply eqb_true_iff in H; subst. intros v Hv; reflexivity. + - apply leqb_universe_n_spec in H; tas. rewrite H. + now rewrite orb_true_r. Qed. - Definition check_leqb_universe (u1 u2 : universe) := - negb check_univs - || Universe.is_prop u1 - || Universe.equal u1 u2 - || try_leqb_universe_n 0 u1 u2. + Definition check_eqb_universe (u1 u2 : Universe.t) := + ~~ check_univs + || Universe.eqb u1 u2 + || (leqb_universe_n 0 u1 u2 && leqb_universe_n 0 u2 u1). - Definition check_eqb_universe (u1 u2 : universe) := - negb check_univs - || Universe.equal u1 u2 - || (try_leqb_universe_n 0 u1 u2 && try_leqb_universe_n 0 u2 u1). + Lemma check_eqb_universe_spec u1 u2 + (Hu1 : gc_levels_declared u1) + (Hu2 : gc_levels_declared u2) + : check_eqb_universe u1 u2 <-> gc_eq_universe uctx.2 u1 u2. + Proof. + unfold check_eqb_universe, gc_eq_universe. + destruct check_univs; split; cbn; trivial; intro H. + - toProp H as [H|H]. + + apply eqb_true_iff in H; subst. intros v Hv; reflexivity. + + toProp H as [H1 H2]. + apply leqb_universe_n_spec0 in H1. apply leqb_universe_n_spec0 in H2. + intros v Hv. specialize (H1 v Hv); specialize (H2 v Hv). + lled; lia. + - toProp; right. + toProp; apply leqb_universe_n_spec; tas; intros v Hv; specialize (H v Hv); lled; lia. + Qed. Lemma check_eqb_universe_refl : forall u, check_eqb_universe u u. Proof. intro u. unfold check_eqb_universe. - rewrite Universe.equal_refl. + rewrite Universe.eqb_refl. rewrite ssrbool.orbT. reflexivity. Qed. - Definition check_gc_constraint (gc : good_constraint) := + Definition check_gc_constraint (gc : GoodConstraint.t) := negb check_univs || match gc with | gc_le l l' => leqb_no_prop_n 0 l l' | gc_lt l l' => leqb_no_prop_n 1 l l' - | gc_lt_set n => leqb_no_prop_n 1 lSet (mVar n) - | gc_eq_set n => leqb_no_prop_n 0 (mVar n) lSet + | gc_lt_set n => leqb_no_prop_n 1 lSet (Var n) + | gc_eq_set n => leqb_no_prop_n 0 (Var n) lSet end. Definition check_gc_constraints @@ -1333,12 +1557,6 @@ Section CheckLeq. | None => false end. - Lemma toto l : level_of_no_prop (vtn l) = level_of_variable l. - destruct l; reflexivity. - Qed. - - Transparent Z.of_nat. - Lemma check_gc_constraint_spec gc : check_gc_constraint gc -> if check_univs then forall v, gc_satisfies v uctx.2 -> gc_satisfies0 v gc else True. @@ -1349,19 +1567,20 @@ Section CheckLeq. specialize (HH v Hv). cbn in *. toProp. pose proof (val_level_of_variable_level v l). pose proof (val_level_of_variable_level v l'). - destruct l, l'; cbn in *; lia. + destruct l, l'; cbn in *; lled; lia. - intros HH v Hv; apply leqb_no_prop_n_spec0 in HH. specialize (HH v Hv). cbn -[Z.of_nat] in HH. unfold gc_satisfies0. toProp. pose proof (val_level_of_variable_level v l) as H1. pose proof (val_level_of_variable_level v l') as H2. - rewrite !toto in *. - rewrite H1, H2 in HH. clear -HH. lia. + cbn in *. rewrite !UnivExpr.val_make_npl in *. + rewrite NoPropLevel.val_to_level in *. + rewrite H1, H2 in HH. clear -HH. lled; lia. - intros HH v Hv; apply leqb_no_prop_n_spec0 in HH. specialize (HH v Hv). cbn in HH. unfold gc_satisfies0. toProp. - lia. + lled; lia. - intros HH v Hv; apply leqb_no_prop_n_spec0 in HH. specialize (HH v Hv). cbn in HH. unfold gc_satisfies0. toProp. - lia. + lled; lia. Qed. Lemma check_gc_constraints_spec ctrs @@ -1377,8 +1596,16 @@ Section CheckLeq. apply XX; assumption. Qed. + Definition eqb_univ_instance (u1 u2 : Instance.t) : bool + := forallb2 (fun l1 l2 => check_eqb_universe + (Universe.make l1) (Universe.make l2)) u1 u2. + End CheckLeq. + + + + (* This section: specif in term of raw uctx *) Section CheckLeq2. Context {cf:checker_flags}. @@ -1420,11 +1647,14 @@ Section CheckLeq2. symmetry; exact HG. contradiction HG. Qed. + Let level_declared (l : Level.t) := LevelSet.In l uctx.1. - Definition level_declared l := LevelSet.In l uctx.1. + Let expr_declared (e : UnivExpr.t) + := on_Some_or_None (fun l : NoPropLevel.t => level_declared l) + (UnivExpr.get_noprop e). - Definition levels_declared (u : universe) - := List.Forall (fun e => level_declared e.1) (NEL.to_list u). + Let levels_declared (u : Universe.t) + := UnivExprSet.For_all expr_declared u. Lemma level_gc_declared_declared l : level_declared l -> gc_level_declared uctx' l. @@ -1433,30 +1663,36 @@ Section CheckLeq2. unfold is_graph_of_uctx, gc_of_uctx in HG. destruct (gc_of_constraints uctx.2); [|contradiction HG]. cbn; clear HG. unfold level_declared, gc_level_declared; cbn. - case_eq (no_prop_of_level l); cbn; [|intuition]. - intros l' HH Hl. apply no_prop_levels_no_prop_level. - now rewrite (no_prop_of_level_inv HH). + destruct l; cbn; trivial; intro; + now apply no_prop_levels_no_prop_level. + Qed. + + Lemma expr_gc_declared_declared e + : expr_declared e -> gc_expr_declared uctx' e. + Proof. + destruct e as [|[l b]]; cbn; trivial. + intro; apply (level_gc_declared_declared l) in H. + red in H. now rewrite NoPropLevel.of_to_level in H. Qed. Lemma levels_gc_declared_declared u : levels_declared u -> gc_levels_declared uctx' u. Proof. unfold levels_declared, gc_levels_declared. - intro HH. eapply Forall_impl. eassumption. - intro. apply level_gc_declared_declared. + intros HH e He; specialize (HH e He). + now apply expr_gc_declared_declared. Qed. - - Lemma leqb_univ_expr_n_spec' n u e2 + Lemma leqb_univ_expr_n_spec' n e1 u + (He1 : expr_declared e1) (Hu : levels_declared u) - (He2 : level_declared e2.1) - : leqb_univ_expr_n G n u e2 - <-> leq_universe_n n uctx.2 u (Universe.make' e2). + : leqb_expr_univ_n G n e1 u + <-> leq_universe_n n uctx.2 (Universe.make' e1) u. Proof. etransitivity. - apply (leqb_univ_expr_n_spec G uctx' Huctx' HC' HG'); tas. + apply (leqb_expr_univ_n_spec G uctx' Huctx' HC' HG'); tas. + - apply expr_gc_declared_declared; tas. - apply levels_gc_declared_declared; tas. - - apply level_gc_declared_declared; tas. - symmetry. etransitivity. apply gc_leq_universe_n_iff. subst uctx'; cbn; clear -HG. unfold is_graph_of_uctx, gc_of_uctx in *. @@ -1464,87 +1700,22 @@ Section CheckLeq2. reflexivity. contradiction HG. Qed. - Lemma val_is_prop u v : - Universe.is_prop u -> (val v u = -1)%Z. - Proof. - clear. - induction u. - - destruct a as [[] []]; inversion 1; reflexivity. - - intro H. cbn in H. apply andb_true_iff in H as [H1 H2]. - rewrite val_cons. rewrite IHu; tas. - destruct a as [[] []]; inversion H1; reflexivity. - Qed. - - Lemma val_minus_one u v : - (-1 <= val v u)%Z. - Proof. - clear. - induction u. cbn. - - destruct a as [[] []]; cbn; lia. - - rewrite val_cons. - destruct a as [[] []]; cbn; lia. - Qed. - - Lemma val_is_prop' u v : - (val v u = -1)%Z -> Universe.is_prop u. - Proof. - clear. - induction u. - - destruct a as [[] []]; cbnr; lia. - - intro H. cbn. rewrite val_cons in H. - apply andb_true_iff. split. - destruct a as [[] []]; cbn in *; try reflexivity; try lia. - apply IHu. - pose proof (val_minus_one u v). - lia. - Qed. - - Lemma val0_equal l1 l2 v : - Level.equal l1 l2 -> val0 v l1 = val0 v l2. - Proof. - clear. - destruct l1, l2; cbnr; try now inversion 1. - all: unfold Level.equal; cbn. - destruct (CompareSpec_string s s0); try now inversion 1. now subst. - case_eq (n ?= n0); intro H; try now inversion 1. - apply PeanoNat.Nat.compare_eq in H. now subst. - Qed. - - Lemma val1_equal e1 e2 v : - Universe.Expr.equal e1 e2 -> val1 v e1 = val1 v e2. - Proof. - destruct e1 as [[] []], e2 as [[] []]; cbn; try now inversion 1. - all: try (rewrite andb_false_r; now inversion 1). - all: rewrite andb_true_r. - all: intro H; apply val0_equal with (v:=v) in H; cbn in H; lia. - Qed. - - Lemma val_equal u1 u2 v : - Universe.equal u1 u2 -> val v u1 = val v u2. - Proof. - clear. - induction u1 in u2 |- *. - - destruct u2 as [a2|]; [|inversion 1]. cbn. apply val1_equal. - - destruct u2 as [|a2 u2]; [inversion 1|]. intro H. cbn in H. - apply andb_true_iff in H as [H1 H2]. - rewrite !val_cons. erewrite IHu1; tea. - erewrite val1_equal; tea. reflexivity. - Qed. - - Lemma check_leqb_universe_spec u1 u2 + (* todo complete *) + Lemma check_leqb_universe_spec' u1 u2 : check_leqb_universe G u1 u2 -> leq_universe uctx.2 u1 u2. Proof. unfold check_leqb_universe, leq_universe. destruct check_univs; cbn; [|trivial]. - case_eq (Universe.is_prop u1). { - intros e _ v Hv. rewrite val_is_prop; tas. + case_eq (prop_sub_type && Universe.is_prop u1). { + intros e _ v Hv. toProp e as [Hp Hu1]. + rewrite (is_prop_val _ Hu1 v). apply le_lle; tas. apply val_minus_one. } intros _. - case_eq (Universe.equal u1 u2). { - intros e _ v Hv. now erewrite val_equal; tea. } + case_eq (Universe.eqb u1 u2). { + intros e _ v Hv. apply eqb_true_iff in e. now destruct e. } intros _; cbn. - intro H; unshelve eapply (try_leqb_universe_n_spec G uctx' Huctx' HC' HG' - _ _ _) in H. + intro H; unshelve eapply (leqb_universe_n_spec0 G uctx' Huctx' HC' HG' + _ _ _) in H. eapply gc_leq_universe_n_iff. unfold uctx' in H. unfold is_graph_of_uctx, gc_of_uctx in HG. @@ -1552,28 +1723,19 @@ Section CheckLeq2. exact I. Qed. - Lemma eq_leq_universe φ u u' : - eq_universe0 φ u u' <-> leq_universe0 φ u u' /\ leq_universe0 φ u' u. - Proof. - split. - intro H; split; intros v Hv; specialize (H v Hv); lia. - intros [H1 H2] v Hv; specialize (H1 v Hv); specialize (H2 v Hv); lia. - Qed. - - - Lemma check_eqb_universe_spec u1 u2 + (* todo complete *) + Lemma check_eqb_universe_spec' u1 u2 : check_eqb_universe G u1 u2 -> eq_universe uctx.2 u1 u2. Proof. unfold check_eqb_universe, eq_universe. destruct check_univs; cbn; [|trivial]. - case_eq (Universe.equal u1 u2). { - intros e _ v Hv. now erewrite val_equal; tea. } + case_eq (Universe.eqb u1 u2). { + intros e _ v Hv. apply eqb_true_iff in e. now destruct e. } intros _; cbn. intro H. apply andb_prop in H. destruct H as [H1 H2]. - unshelve eapply (try_leqb_universe_n_spec G uctx' Huctx' HC' HG' - _ _ _) in H1. - unshelve eapply (try_leqb_universe_n_spec G uctx' Huctx' HC' HG' - _ _ _) in H2. + unshelve eapply (leqb_universe_n_spec0 G uctx' Huctx' HC' HG' _ _ _) in H1. + unshelve eapply (leqb_universe_n_spec0 G uctx' Huctx' HC' HG' + _ _ _) in H2. unfold uctx' in H1, H2. unfold is_graph_of_uctx, gc_of_uctx in HG. apply eq_leq_universe. @@ -1581,20 +1743,7 @@ Section CheckLeq2. (destruct (gc_of_constraints uctx.2); [cbn in *|contradiction HG]); tas. Qed. - - (* Lemma check_gc_constraints_spec ctrs *) - (* : check_gc_constraints ctrs *) - (* -> if check_univs then forall v, gc_satisfies v uctx.2 -> gc_satisfies v ctrs else True. *) - (* Proof. *) - (* pose proof check_gc_constraint_spec as XX. *) - (* unfold check_gc_constraint. destruct check_univs; [cbn|trivial]. *) - (* intros HH v Hv. *) - (* apply GoodConstraintSet.for_all_spec. now intros x y []. *) - (* apply GoodConstraintSet.for_all_spec in HH. 2: now intros x y []. *) - (* intros gc Hgc. specialize (HH gc Hgc). *) - (* apply XX; assumption. *) - (* Qed. *) - + (* todo complete *) Lemma check_constraints_spec ctrs : check_constraints G ctrs -> valid_constraints uctx.2 ctrs. Proof. @@ -1611,69 +1760,12 @@ Section CheckLeq2. unfold is_graph_of_uctx, gc_of_uctx in HG. now destruct (gc_of_constraints uctx.2). Qed. - - End CheckLeq2. -(* Extract Constant constrained_dec => "(fun g ctrs l -> assert false)". *) - - -(* Definition G := make_graph *) -(* (GoodConstraintSet.add (gc_lt (mVar 0) (mVar 1)) *) -(* (GoodConstraintSet_pair (gc_lt_set 0) (gc_eq_set 1))). *) - -(* Compute (lsp G lSet (mVar 0)). *) -(* Compute (lsp G lSet (mVar 1)). *) -(* Compute (lsp G lSet lSet). *) - -(* Section Test. *) - -(* Definition get_graph G0 := match G0 with *) -(* | Some φ => φ *) -(* | None => init_graph *) -(* end. *) - -(* Definition G0 := make_graph ConstraintSet.empty. *) -(* Check (eq_refl : G0 = Some _). *) -(* Definition G := get_graph G0. *) - -(* Local Existing Instance default_checker_flags. *) - -(* Check (eq_refl : check_leq G Universe.type0m Universe.type0 = true). *) -(* Check (eq_refl : check_lt G Universe.type0m Universe.type0 = true). *) -(* Check (eq_refl : check_lt G Universe.type0m Universe.type0m = false). *) -(* Check (eq_refl : check_lt G Universe.type0 Universe.type0m = false). *) -(* Check (eq_refl : check_lt G Universe.type0 Universe.type0 = false). *) -(* Check (eq_refl : no_universe_inconsistency G = true). *) - -(* Definition ctrs0 := ConstraintSet.add (Level.Level "Top.0", Lt, Level.Level "Top.1") *) -(* (ConstraintSet.singleton (Level.Var 0, Lt, Level.Var 1)). *) -(* Definition G'0 := make_graph ctrs0. *) -(* Check (eq_refl : G'0 = Some _). *) -(* Definition G' := get_graph G'0. *) - -(* Check (eq_refl : check_lt G' (Universe.make (Level.Level "Top.0")) (Universe.make (Level.Var 0)) = false). *) -(* Check (eq_refl : check_leq G' (Universe.make (Level.Level "Top.1")) (Universe.make (Level.lProp)) = false). *) -(* Check (eq_refl : check_leq G' (Universe.super (Level.Level "Top.1")) (Universe.make (Level.Level "Top.1")) = false). *) -(* Check (eq_refl : check_lt G' (Universe.make (Level.Level "Top.1")) (Universe.super (Level.Level "Top.1")) = true). *) -(* Check (eq_refl : check_lt G' (Universe.make (Level.Level "Top.1")) (Universe.make (Level.Level "Top.1")) = false). *) -(* Check (eq_refl : check_eq G' (Universe.make (Level.Level "Top.1")) (Universe.make (Level.Level "Top.1")) = true). *) -(* Check (eq_refl : no_universe_inconsistency G = true). *) - -(* Check (eq_refl : check_lt G' (Universe.make (Level.Var 1)) (Universe.make (Level.Var 0)) = false). *) -(* Check (eq_refl : check_leq G' (Universe.make (Level.Var 0)) (Universe.make (Level.Var 1)) = true). *) -(* Check (eq_refl : check_lt G' (Universe.make (Level.Var 0)) (Universe.make (Level.Var 1)) = true). *) -(* Check (eq_refl : check_leq G' Universe.type1 Universe.type0 = false). *) -(* Check (eq_refl : check_lt G' Universe.type1 Universe.type1 = false). *) - - -(* Definition ctrs1 := ConstraintSet.add (Level.Var 1, Eq, Level.Var 2) *) -(* (ConstraintSet.add (Level.Var 2, Le, Level.lSet) ctrs0). *) -(* Definition G''0 := make_graph ctrs1. *) -(* Check (eq_refl : G''0 = Some _). *) -(* Definition G'' := get_graph G''0. *) - -(* Check (eq_refl : no_universe_inconsistency G'' = false). *) -(* End Test. *) +(* If I would restart: *) +(* - maybe use nat everywhere instead of Z and Pos *) +(* - maybe not put Prop in UnivExprs but in Universes only *) +(* - more generally keep the graph closer to the spec (valuation = labelling ?) + but the the quoting is more complicated ... *) diff --git a/template-coq/theories/utils.v b/template-coq/theories/utils.v index 04039a8b5..4f7214099 100644 --- a/template-coq/theories/utils.v +++ b/template-coq/theories/utils.v @@ -1,4 +1,4 @@ -From Coq Require Import Bool List String Program. +From Coq Require Import Nat ZArith Bool List String Program. Global Set Asymmetric Patterns. @@ -63,12 +63,79 @@ Ltac rdest := (** We cannot use ssrbool as it breaks extraction. *) Coercion is_true : bool >-> Sortclass. -Ltac toProp := +Tactic Notation "toProp" ident(H) := + match type of H with + | is_true (_ apply PeanoNat.Nat.ltb_lt in H + | is_true (_ <=? _)%nat => apply PeanoNat.Nat.leb_le in H + | is_true (_ =? _)%nat => apply PeanoNat.Nat.eqb_eq in H + | (_ apply PeanoNat.Nat.ltb_lt in H + | (_ <=? _)%nat = true => apply PeanoNat.Nat.leb_le in H + | (_ =? _)%nat = true => apply PeanoNat.Nat.eqb_eq in H + + | is_true (_ apply Z.ltb_lt in H + | is_true (_ <=? _)%Z => apply Z.leb_le in H + | is_true (_ =? _)%Z => apply Z.eqb_eq in H + | (_ apply Z.ltb_lt in H + | (_ <=? _)%Z = true => apply Z.leb_le in H + | (_ =? _)%Z = true => apply Z.eqb_eq in H + + | is_true (_ && _) => apply andb_true_iff in H + | (_ && _) = true => apply andb_true_iff in H + | (_ && _) = false => apply andb_false_iff in H + + | is_true (_ || _) => apply orb_true_iff in H + | (_ || _) = true => apply orb_true_iff in H + | (_ || _) = false => apply orb_false_iff in H + end. + +Tactic Notation "toProp" := + match goal with + | |- is_true (_ apply PeanoNat.Nat.ltb_lt + | |- is_true (_ <=? _)%nat => apply PeanoNat.Nat.leb_le + | |- is_true (_ =? _)%nat => apply PeanoNat.Nat.eqb_eq + | |- (_ apply PeanoNat.Nat.ltb_lt + | |- (_ <=? _)%nat = true => apply PeanoNat.Nat.leb_le + | |- (_ =? _)%nat = true => apply PeanoNat.Nat.eqb_eq + + | |- is_true (_ apply Z.ltb_lt + | |- is_true (_ <=? _)%Z => apply Z.leb_le + | |- is_true (_ =? _)%Z => apply Z.eqb_eq + | |- (_ apply Z.ltb_lt + | |- (_ <=? _)%Z = true => apply Z.leb_le + | |- (_ =? _)%Z = true => apply Z.eqb_eq + + | |- is_true (_ && _) => apply andb_true_iff; split + | |- (_ && _) = true => apply andb_true_iff; split + + | |- is_true (_ || _) => apply orb_true_iff + | |- (_ || _) = true => apply orb_true_iff + + | H : _ |- _ => toProp H + end. + +Tactic Notation "toProp" ident(H) "as" simple_intropattern(X) := + match type of H with + | is_true (_ && _) => apply andb_true_iff in H; destruct H as X + | (_ && _) = true => apply andb_true_iff in H; destruct H as X + | (_ && _) = false => apply andb_false_iff in H; destruct H as X + + | is_true (_ || _) => apply orb_true_iff in H; destruct H as X + | (_ || _) = true => apply orb_true_iff in H; destruct H as X + | (_ || _) = false => apply orb_false_iff in H; destruct H as X + end. + +Tactic Notation "toProp" "as" simple_intropattern(X) := + match goal with + | H : _ |- _ => toProp H as X + end. + +Ltac rtoProp := repeat match goal with | H : is_true (_ && _) |- _ => apply andb_and in H; destruct H | |- context [is_true (_ && _)] => rewrite andb_and end. + (** "Incoherent" notion of equivalence, that we only apply to hProps actually. *) Record isEquiv (A B : Type) := { equiv : A -> B; @@ -94,153 +161,6 @@ Hint Resolve Peano_dec.eq_nat_dec : eq_dec. Ltac invs H := inversion H; subst; clear H. -(* Sorted lists without duplicates *) -Class ComparableType A := { compare : A -> A -> comparison }. -Arguments compare {A} {_} _ _. - -Fixpoint insert {A} `{ComparableType A} (x : A) (l : list A) := - match l with - | [] => [x] - | y :: l' => match compare x y with - | Eq => l - | Lt => x :: l - | Gt => y :: (insert x l') - end - end. - -Definition list_union {A} `{ComparableType A} (l l' : list A) : list A - := fold_left (fun l' x => insert x l') l l'. - - - - -(** * Non Empty List : will go away *) -Module NEL. - - Inductive t (A : Set) := - | sing : A -> t A - | cons : A -> t A -> t A. - - Arguments sing {A} _. - Arguments cons {A} _ _. - - Declare Scope nel_scope. - Infix "::" := cons : nel_scope. - Notation "[ x ]" := (sing x) : nel_scope. - Delimit Scope nel_scope with nel. - Bind Scope nel_scope with t. - - Local Open Scope nel_scope. - - Fixpoint length {A} (l : t A) : nat := - match l with - | [ _ ] => 1 - | _ :: l' => S (length l') - end. - - Notation "[ x ; y ; .. ; z ; e ]" := - (cons x (cons y .. (cons z (sing e)) ..)) : nel_scope. - - Fixpoint to_list {A} (l : t A) : list A := - match l with - | [x] => [x] - | x :: l => x :: (to_list l) - end. - - Fixpoint map {A B : Set} (f : A -> B) (l : t A) : t B := - match l with - | [x] => [f x] - | x :: l => f x :: map f l - end. - - Fixpoint app {A} (l l' : t A) : t A := - match l with - | [x] => x :: l' - | x :: l => x :: (app l l') - end. - - Infix "++" := app : nel_scope. - - Fixpoint app_r {A : Set} (l : list A) (l' : t A) : t A := - match l with - | [] => l' - | (x :: l)%list => x :: app_r l l' - end. - - Fixpoint cons' {A : Set} (x : A) (l : list A) : t A := - match l with - | [] => [x] - | (y :: l)%list => x :: cons' y l - end. - - Lemma cons'_spec {A : Set} (x : A) l - : to_list (cons' x l) = (x :: l)%list. - Proof. - revert x; induction l; simpl. - reflexivity. - intro x; now rewrite IHl. - Qed. - - Fixpoint fold_left {A} {B : Set} (f : A -> B -> A) (l : t B) (a0 : A) : A := - match l with - | [b] => f a0 b - | b :: l => fold_left f l (f a0 b) - end. - - Fixpoint forallb {A : Set} (f : A -> bool) (l : t A) := - match l with - | [x] => f x - | hd :: tl => f hd && forallb f tl - end. - - Fixpoint forallb2 {A : Set} (f : A -> A -> bool) (l l' : t A) := - match l, l' with - | [x], [x'] => f x x' - | hd :: tl, hd' :: tl' => f hd hd' && forallb2 f tl tl' - | _, _ => false - end. - - Lemma map_to_list {A B : Set} (f : A -> B) (l : t A) - : to_list (map f l) = List.map f (to_list l). - Proof. - induction l; cbn; congruence. - Qed. - - Lemma map_app {A B : Set} (f : A -> B) l l' : - map f (l ++ l') = map f l ++ map f l'. - Proof. - induction l; cbn; congruence. - Qed. - - Lemma map_map {A B C : Set} (f : A -> B) (g : B -> C) l : - map g (map f l) = map (fun x => g (f x)) l. - Proof. - induction l; simpl; auto. - rewrite IHl; auto. - Qed. - - Lemma map_ext {A B : Set} (f g : A -> B) l : - (forall x, f x = g x) -> map f l = map g l. - Proof. - intros. - induction l; cbn; rewrite ?H; congruence. - Defined. - - Lemma forallb2_refl : - forall (A : Set) (R : A -> A -> bool), - (forall x, is_true (R x x)) -> - forall l, is_true (forallb2 R l l). - Proof. - intros A R R_refl l. - induction l. - - simpl. apply R_refl. - - simpl. rewrite R_refl. assumption. - Qed. - -End NEL. - -Definition non_empty_list := NEL.t. - Lemma iff_forall {A} B C (H : forall x : A, B x <-> C x) : (forall x, B x) <-> (forall x, C x). @@ -256,9 +176,25 @@ Lemma if_true_false (b : bool) : (if b then true else false) = b. destruct b; reflexivity. Qed. +Lemma iff_is_true_eq_bool (b b' : bool) : + (b <-> b') -> b = b'. +Proof. + destruct b, b'; cbnr; intros [H1 H2]; + try specialize (H1 eq_refl); try specialize (H2 eq_refl); + discriminate. +Qed. + +Lemma uip_bool (b1 b2 : bool) (p q : b1 = b2) : p = q. +Proof. + destruct q. apply Eqdep_dec.UIP_refl_bool. +Qed. + Ltac tas := try assumption. Ltac tea := try eassumption. Axiom todo : String.string -> forall {A}, A. Ltac todo s := exact (todo s). Extract Constant todo => "fun s -> failwith (String.concat """" (List.map (String.make 1) s))". + + +Coercion is_left A B (u : {A} + {B}) := match u with left _ => true | _ => false end. diff --git a/template-coq/theories/utils/All_Forall.v b/template-coq/theories/utils/All_Forall.v index c86aa2e70..14ee3e76b 100644 --- a/template-coq/theories/utils/All_Forall.v +++ b/template-coq/theories/utils/All_Forall.v @@ -76,7 +76,7 @@ Proof. induction l => /= // [=] Ha Hl; constructor; auto. Qed. -Lemma forall_map_id_spec {A} {P : A -> Prop} {l} {f : A -> A} : +Lemma forall_map_id_spec {A} {l} {f : A -> A} : Forall (fun x => f x = x) l <-> map f l = l. Proof. rewrite -{3}(map_id l). apply forall_map_spec. @@ -810,14 +810,6 @@ Proof. intros. apply IHl. rewrite -> andb_and in H; intuition. Qed. -Lemma forallb_rev {A} (p : A -> bool) l : - forallb p (List.rev l) = forallb p l. -Proof. - induction l using rev_ind; simpl; try congruence. - rewrite rev_unit forallb_app. simpl. rewrite <- IHl. - now rewrite andb_comm andb_true_r. -Qed. - Lemma Forall_forallb_eq_forallb {A} (P : A -> Prop) (p q : A -> bool) l : Forall P l -> (forall x, P x -> p x = q x) -> @@ -1638,3 +1630,55 @@ Proof. depelim l'; depelim X. destruct l; depelim X. specialize (IHl _ _ X). intuition auto. Qed. + +Lemma Forall2_eq {A} l l' + : @Forall2 A A eq l l' -> l = l'. +Proof. + induction 1; congruence. +Qed. + +Lemma Forall2_map' {A B A' B'} (R : A' -> B' -> Prop) (f : A -> A') (g : B -> B') l l' + : Forall2 R (map f l) (map g l') -> Forall2 (fun x y => R (f x) (g y)) l l'. +Proof. + induction l in l' |- *. + - destruct l'; inversion 1. constructor. + - destruct l'; inversion 1. constructor; auto. +Qed. + +Lemma Forall2_same : + forall A (P : A -> A -> Prop) l, + (forall x, P x x) -> + Forall2 P l l. +Proof. + intros A P l h. + induction l. + - constructor. + - constructor. + + eapply h. + + assumption. +Qed. + +Lemma Forall2_sym : + forall A (P : A -> A -> Prop) l l', + Forall2 P l l' -> + Forall2 (fun x y => P y x) l' l. +Proof. + intros A P l l' h. + induction h. + - constructor. + - constructor. all: auto. +Qed. + +Lemma forallb2_Forall2 : + forall A (p : A -> A -> bool) l l', + forallb2 p l l' -> + Forall2 (fun x y => p x y) l l'. +Proof. + intros A p l l' h. + induction l in l', h |- *. + - destruct l'. 2: discriminate. + constructor. + - destruct l'. 1: discriminate. + simpl in h. apply andP in h as [? ?]. + constructor. all: auto. +Qed. diff --git a/template-coq/theories/utils/MCList.v b/template-coq/theories/utils/MCList.v index a0a33cd39..f9f2c9d8f 100644 --- a/template-coq/theories/utils/MCList.v +++ b/template-coq/theories/utils/MCList.v @@ -1,4 +1,4 @@ -From Coq Require Import List Program Arith Lia. +From Coq Require Import Bool List Program Arith Lia SetoidList. Import ListNotations. @@ -756,3 +756,36 @@ Proof. intros. induction l; simpl; auto. Defined. + +Lemma InA_In_eq {A} x (l : list A) : InA Logic.eq x l <-> In x l. +Proof. + etransitivity. eapply InA_alt. + firstorder. now subst. +Qed. + +Lemma forallb_rev {A} (p : A -> bool) l : + forallb p (List.rev l) = forallb p l. +Proof. + induction l using rev_ind; simpl; try congruence. + rewrite rev_unit forallb_app. simpl. rewrite <- IHl. + now rewrite andb_comm andb_true_r. +Qed. + +Lemma fold_left_andb_forallb {A} P l x : + fold_left (fun b x => P x && b) l (P x) = @forallb A P (x :: l). +Proof. + cbn. rewrite <- fold_left_rev_right. rewrite <- forallb_rev. + induction (List.rev l); cbn. + - now rewrite andb_true_r. + - rewrite IHl0. rewrite !andb_assoc. + f_equal. now rewrite andb_comm. +Qed. + +Lemma nth_nth_error {A} n l (d : A) : + nth n l d = match nth_error l n with + | Some x => x + | None => d + end. +Proof. + symmetry. apply nth_default_eq. +Qed. diff --git a/template-coq/theories/utils/MCOption.v b/template-coq/theories/utils/MCOption.v index 07c828d21..8d632f070 100644 --- a/template-coq/theories/utils/MCOption.v +++ b/template-coq/theories/utils/MCOption.v @@ -12,6 +12,18 @@ Definition on_some {A} (P : A -> Type) (o : option A) := | None => False end. +Definition on_Some {A} (P : A -> Prop) : option A -> Prop := + fun x => match x with + | Some x => P x + | None => False + end. + +Definition on_Some_or_None {A} (P : A -> Prop) : option A -> Prop := + fun x => match x with + | Some x => P x + | None => True + end. + Definition option_default {A B} (f : A -> B) (o : option A) (b : B) := match o with Some x => f x | None => b end. @@ -41,3 +53,15 @@ Proof. simpl. destruct a; reflexivity. destruct a; simpl in *; congruence. Qed. + +Lemma option_map_two {A B C} (f : A -> B) (g : B -> C) x + : option_map g (option_map f x) = option_map (fun x => g (f x)) x. +Proof. + destruct x; reflexivity. +Qed. + +Lemma option_map_ext {A B} (f g : A -> B) (H : forall x, f x = g x) + : forall z, option_map f z = option_map g z. +Proof. + intros []; cbn; congruence. +Qed. diff --git a/template-coq/theories/utils/MCProd.v b/template-coq/theories/utils/MCProd.v index 4f4108d55..04ece32dd 100644 --- a/template-coq/theories/utils/MCProd.v +++ b/template-coq/theories/utils/MCProd.v @@ -25,6 +25,10 @@ Open Scope pair_scope. Notation "x × y" := (prod x y )(at level 80, right associativity). +Notation "p .p1" := (proj1 p) (at level 2, left associativity, format "p .p1"). +Notation "p .p2" := (proj2 p) (at level 2, left associativity, format "p .p2"). + + Definition on_snd {A B C} (f : B -> C) (p : A * B) := (fst p, f (snd p)). @@ -74,3 +78,15 @@ Proof. apply andb_true_iff. Qed. Lemma andP {b b'} : is_true (b && b') -> is_true b /\ is_true b'. Proof. apply andb_and. Qed. + +Definition fst_eq {A B} {x x' : A} {y y' : B} + : (x, y) = (x', y') -> x = x'. +Proof. + inversion 1; reflexivity. +Qed. + +Definition snd_eq {A B} {x x' : A} {y y' : B} + : (x, y) = (x', y') -> y = y'. +Proof. + inversion 1; reflexivity. +Qed. diff --git a/template-coq/theories/utils/wGraph.v b/template-coq/theories/utils/wGraph.v index 28b49eb16..a02a81722 100644 --- a/template-coq/theories/utils/wGraph.v +++ b/template-coq/theories/utils/wGraph.v @@ -2,38 +2,8 @@ Require Import Peano_dec Nat Bool List Structures.Equalities Lia MSets.MSetList MSetFacts MSetProperties. From MetaCoq.Template Require Import utils monad_utils. -Notation "p .1" := (fst p) - (at level 2, left associativity, format "p .1") : pair_scope. -Notation "p .2" := (snd p) - (at level 2, left associativity, format "p .2") : pair_scope. -Notation "p ..1" := (proj1 p) - (at level 2, left associativity, format "p ..1") : pair_scope. -Notation "p ..2" := (proj2 p) - (at level 2, left associativity, format "p ..2") : pair_scope. -Open Scope pair_scope. -(* Coercion pair_of_and P Q (PandQ : P /\ Q) := (proj1 PandQ, proj2 PandQ). *) -(* Coercion isSome T (u : option T) := if u is Some _ then true else false. *) -(* Coercion is_inl A B (u : A + B) := if u is inl _ then true else false. *) -Coercion is_left A B (u : {A} + {B}) := match u with left _ => true | _ => false end. -(* Coercion is_inleft A B (u : A + {B}) := if u is inleft _ then true else false. *) - -Definition fst_eq {A B} {x x' : A} {y y' : B} - : (x, y) = (x', y') -> x = x'. -Proof. - inversion 1; reflexivity. -Qed. - -Definition snd_eq {A B} {x x' : A} {y y' : B} - : (x, y) = (x', y') -> y = y'. -Proof. - inversion 1; reflexivity. -Qed. - -Lemma InAeq_In {A} (l : list A) x : - InA eq x l <-> In x l. -Proof. - etransitivity. eapply InA_alt. firstorder. now subst. -Defined. +From Equations Require Import Equations. +Require Import Equations.Prop.DepElim. Lemma fold_max_In n m l (H : fold_left max l n = m) : n = m \/ In m l. @@ -62,8 +32,6 @@ Proof. eexists. split. eassumption. reflexivity. Qed. -(* Definition is_Some {A} (x : option A) := exists a, x = Some a. *) - Definition eq_max n m k : max n m = k -> n = k \/ m = k. intro; lia. Qed. @@ -83,6 +51,9 @@ Module Nbar. | Some n, Some m => Some (n + m) | _, _ => None end. + + Definition S : t -> t := option_map S. + Definition le (n m : t) : Prop := match n, m with | Some n, Some m => n <= m @@ -90,13 +61,22 @@ Module Nbar. | None, _ => True end. + Definition lt (n m : t) : Prop := + match n, m with + | Some n, Some m => n < m + | None, Some _ => True + | _, None => False + end. + Arguments max _ _ : simpl nomatch. Arguments add _ _ : simpl nomatch. Arguments le _ _ : simpl nomatch. + Arguments lt _ _ : simpl nomatch. Declare Scope nbar_scope. Infix "+" := add : nbar_scope. Infix "<=" := le : nbar_scope. + Infix "<" := lt : nbar_scope. Delimit Scope nbar_scope with nbar. Bind Scope nbar_scope with t. @@ -245,11 +225,24 @@ Module Nbar. all: intuition. Defined. + Lemma le_lt_dec n m : ({n <= m} + {m < n})%nbar. + Proof. + destruct n as [n|], m as [m|]; cbn. + apply Compare_dec.le_lt_dec. + all: (right; constructor) || (left; constructor). + Defined. + Lemma le_plus_r n m : m <= Some n + m. Proof. destruct m; cbn; lia. Qed. + Lemma le_antisymm {n m} : n <= m -> m <= n -> n = m. + Proof. + destruct n, m; cbn; try easy. + intros. f_equal. now apply PeanoNat.Nat.le_antisymm. + Qed. + End Nbar. Import Nbar. @@ -372,7 +365,7 @@ Module WeightedGraph (V : UsualOrderedType). Fixpoint length {x y} (p : Paths x y) := match p with | paths_refl x => 0 - | paths_step x y z e p => S (length p) + | paths_step x y z e p => Datatypes.S (length p) end. (* Global Instance Paths_refl : CRelationClasses.Reflexive Paths := paths_refl. *) @@ -561,14 +554,14 @@ Module WeightedGraph (V : UsualOrderedType). repeat split. 2: intros [H0|H0]. - intro H0. apply VSet.remove_spec in H0. destruct H0 as [H0 H1]. - pose proof ((H..1 y0)..1 H0) as H2. + pose proof ((H.p1 y0).p1 H0) as H2. destruct H2; [now left|right]. apply VSetFact.remove_2; intuition. - subst. apply VSet.remove_spec. split; [|assumption]. - apply H..1. left; reflexivity. + apply H.p1. left; reflexivity. - apply VSet.remove_spec in H0; destruct H0 as [H0 H1]. apply VSet.remove_spec; split; [|assumption]. - apply H..1. right; assumption. + apply H.p1. right; assumption. - intro H0. apply VSet.remove_spec in H0; destruct H0 as [H0 _]. apply H; assumption. Qed. @@ -767,7 +760,7 @@ Module WeightedGraph (V : UsualOrderedType). (DisjointAdd_remove1 (VSetFact.mem_2 XX))) (simplify_aux2 XX Hs)) | false => fun XX => (simplify q (add_end p e - (DisjointAdd_add2 ((VSetFact.not_mem_iff _ _)..2 XX))) + (DisjointAdd_add2 ((VSetFact.not_mem_iff _ _).p2 XX))) (simplify_aux3 Hs)) end eq_refl end. @@ -780,7 +773,7 @@ Module WeightedGraph (V : UsualOrderedType). (* | paths_step y y' _ e q => *) (* fun p => match VSet.mem y s with *) (* | true => let '(p1, p2) := split p in *) - (* if 0 @simplify _ _ _ q (@add_end _ _ _ p _ e _) *) (* end *) @@ -793,7 +786,7 @@ Module WeightedGraph (V : UsualOrderedType). (* Defined. *) (* Lemma weight_simplify (HG : acyclic_no_loop) {s x y} q (p : SimplePaths s x y) *) - (* : 0 < weight q \/ 0 < sweight p -> 0 < sweight (simplify q p)..2..2. *) + (* : 0 < weight q \/ 0 < sweight p -> 0 < sweight (simplify q p).p2.p2. *) (* Proof. *) (* revert s p; induction q. *) (* - cbn. intuition. *) @@ -801,7 +794,7 @@ Module WeightedGraph (V : UsualOrderedType). (* set (F := proj2 (VSetFact.not_mem_iff s x)); clearbody F. *) (* destruct (VSet.mem x s). *) (* + case_eq (split p); intros p1 p2 Hp. *) - (* case_eq (0 base - | S fuel => + | Datatypes.S fuel => match VSet.mem x s with | true => let ds := List.map @@ -947,14 +940,14 @@ Module WeightedGraph (V : UsualOrderedType). assert (XX: VSet.Equal (VSet.remove x s') s0). { clear -d. intro a; split; intro Ha. - * apply VSet.remove_spec in Ha. pose proof (d..1 a). + * apply VSet.remove_spec in Ha. pose proof (d.p1 a). intuition. * apply VSet.remove_spec. split. apply d. right; assumption. intro H. apply proj2 in d. apply d. subst; assumption. } rewrite (lsp0_VSet_Equal XX); reflexivity. + apply filter_In. split. - apply InAeq_In, EdgeSet.elements_spec1. exact e.π2. + apply InA_In_eq, EdgeSet.elements_spec1. exact e.π2. cbn. destruct (V.eq_dec x x); [reflexivity|contradiction]. Qed. @@ -991,7 +984,7 @@ Module WeightedGraph (V : UsualOrderedType). apply VSetProp.Add_remove. apply VSet.mem_spec; assumption. -- eexists. - apply (EdgeSet.elements_spec1 _ _)..1, InAeq_In; eassumption. + apply (EdgeSet.elements_spec1 _ _).p1, InA_In_eq; eassumption. -- cbn. now apply some_inj in H1. * subst. clear -Hx. apply VSet.mem_spec in Hx. apply VSetProp.remove_cardinal_1 in Hx. lia. @@ -1008,6 +1001,14 @@ Module WeightedGraph (V : UsualOrderedType). specialize (Hl (x, e.π1, y) e.π2). cbn in *; lia. Qed. + Lemma correct_labelling_lsp {x y n} (e : lsp x y = Some n) : + forall l, correct_labelling l -> l x + n <= l y. + Proof. + eapply lsp0_spec_eq in e as [p Hp]. + intros l Hl; eapply correct_labelling_Paths with (p:= to_paths p) in Hl. + now rewrite <- sweight_weight, Hp in Hl. + Qed. + Lemma acyclic_labelling l : correct_labelling l -> acyclic_no_loop. Proof. intros Hl x p. @@ -1287,24 +1288,27 @@ Module WeightedGraph (V : UsualOrderedType). (* Qed. *) End graph. + Definition spaths_one {G s x y k} (Hx : VSet.In x s) + (Hk : EdgeSet.In (x, k, y) (E G)) + : SimplePaths G s x y. + Proof. + econstructor. 3: constructor. now apply DisjointAdd_remove1. + exists k. assumption. + Defined. + + + Section graph2. Context (G : t) {HI : invariants G} {HG : acyclic_no_loop G}. Section subgraph. - Context (n : nat) {x_0 y_0 : V.t} (Vx : VSet.In x_0 (V G)) - (Vy : VSet.In y_0 (V G)) (kx ky : nat) - (Hkx : lsp G (s G) x_0 = Some kx) - (Hky : lsp G (s G) y_0 = Some ky) - (Hle : leq_vertices G n x_0 y_0) - (p_0 : SimplePaths G (V G) (s G) y_0) - (Hp_0 : lsp G (s G) y_0 = Some (sweight G p_0)) - (Hxs : lsp G x_0 (s G) = None) - (Hx_0_0 : VSet.mem x_0 (snodes G p_0) = false). + Context (y_0 x_0 : V.t) (Vx : VSet.In x_0 (V G)) + (Vy : VSet.In y_0 (V G)) + (Hxs : lsp G x_0 y_0 = None) + (K : nat). - Definition K := Peano.max kx (S ky). - - Let G' : t - := (V G, EdgeSet.add (s G, K, x_0) (E G), s G). + Local Definition G' : t + := (V G, EdgeSet.add (y_0, K, x_0) (E G), s G). Definition to_G' {u v} (q : Paths G u v) : Paths G' u v. Proof. @@ -1314,23 +1318,23 @@ Module WeightedGraph (V : UsualOrderedType). exists e.π1. cbn. apply EdgeSet.add_spec; right. exact e.π2. Defined. - Local Instance HI' : invariants G'. + Global Instance HI' : invariants G'. Proof. split. - cbn. intros e He. apply EdgeSet.add_spec in He; destruct He as [He|He]. - subst; cbn. split. apply HI. assumption. + subst; cbn. split; assumption. now apply HI. - split. apply HI. - intros z Hz. pose proof (HI..2..2 z Hz). + intros z Hz. pose proof (HI.p2.p2 z Hz). sq; now apply to_G'. Qed. Definition from_G' {S u v} (q : SimplePaths G' S u v) - : SimplePaths G S u v + (SimplePaths G S u (s G) * SimplePaths G S x_0 v). + : SimplePaths G S u v + (SimplePaths G S u y_0 * SimplePaths G S x_0 v). Proof. clear -q. induction q. - left; constructor. - - induction (Edge.eq_dec (s G, K, x_0) (x, e.π1, y)) as [XX|XX]. + - induction (Edge.eq_dec (y_0, K, x_0) (x, e.π1, y)) as [XX|XX]. + right. split. * rewrite (fst_eq (fst_eq XX)); constructor. * eapply SimplePaths_sub. @@ -1362,23 +1366,23 @@ Module WeightedGraph (V : UsualOrderedType). induction q. - reflexivity. - simpl. - destruct (Edge.eq_dec (s G, K, x_0) (x, e.π1, y)) as [XX|XX]; simpl. + destruct (Edge.eq_dec (y_0, K, x_0) (x, e.π1, y)) as [XX|XX]; simpl. + destruct (fst_eq (fst_eq XX)). simpl. inversion XX. rewrite weight_SimplePaths_sub. destruct (from_G' q) as [q'|[q1 q2]]; simpl. * destruct (snd_eq XX); cbn. destruct e as [e He]; cbn in *; lia. - * inversion XX; subst; clear XX. + * subst y. simple refine (let XX := lsp0_spec_le G (simplify2' G (to_paths G q1)) in _). - unfold lsp in *; rewrite Hxs in XX; inversion XX. + unfold lsp in *. rewrite Hxs in XX; inversion XX. + destruct (from_G' q) as [q'|[q1 q2]]; simpl. * lia. * rewrite weight_SimplePaths_sub; lia. Qed. - Local Instance HG' : acyclic_no_loop G'. + Global Instance HG' : acyclic_no_loop G'. Proof. apply acyclic_caract2. exact _. intros x Hx. pose proof (lsp0_spec_le G' (spaths_refl G' (V G') x)) as H; cbn in H. @@ -1395,42 +1399,100 @@ Module WeightedGraph (V : UsualOrderedType). Qed. - Lemma lsp_xy_le_n : (Some n <= lsp G x_0 y_0)%nbar. + Lemma lsp_G'_yx : (Some K <= lsp G' y_0 x_0)%nbar. Proof. - assert (XX : correct_labelling G (option_get 0 ∘ lsp G' (s G'))). { - pose proof (lsp_correctness G') as XX. split. - exact XX..1. - intros e He; apply XX; cbn. - apply EdgeSet.add_spec; now right. } - specialize (Hle _ XX); clear XX; cbn in Hle. - assert (XX: (Some K <= lsp G' (s G) x_0)%nbar). { - etransitivity. shelve. unshelve eapply lsp0_spec_le. - unshelve econstructor. 5: constructor. exact (VSet.remove (s G) (V G')). - apply DisjointAdd_remove1. apply HI. - exists K. apply EdgeSet.add_spec; now left. - Unshelve. simpl. lia. } - destruct (lsp_s G' _ Vx) as [kx' Hkx']. - cbn in Hkx'; rewrite Hkx' in *; cbn in *. - destruct (lsp_s G' _ Vy) as [ky' Hky']. - cbn in Hky'; rewrite Hky' in *; cbn in *. - destruct (lsp0_spec_eq _ _ Hky') as [py Hpy]. - pose proof (from_G'_weight py) as H. - destruct (from_G' py) as [py'|[py1 py2]]. - - pose proof (lsp0_spec_le _ py') as HH. - cbn in HH; unfold lsp in Hky; rewrite Hky in HH; cbn in *. - unfold K in *; lia. - - pose proof (lsp0_spec_le _ py2) as HH; cbn in HH. - case_eq (lsp G x_0 y_0); - [|intro ZZ; unfold lsp in ZZ; rewrite ZZ in HH; inversion HH]. - intros kxy Hkxy. unfold lsp in Hkxy; rewrite Hkxy in HH; cbn in *. - assert (YY : sweight py1 = 0). { - rewrite sweight_weight. - exact (HG (s G) (to_paths G py1)). } - lia. + unshelve refine (transport (fun K => (Some K <= _)%nbar) _ (lsp0_spec_le _ _)). + - eapply spaths_one; tas. + apply EdgeSet.add_spec. now left. + - cbn. lia. + Qed. + + Lemma lsp_G'_sx : (lsp G' (s G) y_0 + Some K <= lsp G' (s G) x_0)%nbar. + Proof. + etransitivity. 2: eapply lsp_codistance; try exact _. + apply plus_le_compat_l. apply lsp_G'_yx. + Qed. + + Lemma correct_labelling_lsp_G' + : correct_labelling G (option_get 0 ∘ lsp G' (s G')). + Proof. + pose proof (lsp_correctness G') as XX. + split. exact XX.p1. + intros e He; apply XX; cbn. + apply EdgeSet.add_spec; now right. + Qed. + + Definition sto_G' {S u v} (p : SimplePaths G S u v) + : SimplePaths G' S u v. + Proof. + clear -p. induction p. + - constructor. + - econstructor; tea. exists e.π1. + apply EdgeSet.add_spec. right. exact e.π2. Defined. + + Lemma sto_G'_weight {S u v} (p : SimplePaths G S u v) + : sweight (sto_G' p) = sweight p. + Proof. + clear. + induction p. reflexivity. + simpl. now rewrite IHp. + Qed. + + Lemma lsp_G'_incr x y : (lsp G x y <= lsp G' x y)%nbar. + Proof. + case_eq (lsp G x y); cbn; [|trivial]. + intros kxy Hkxy. apply lsp0_spec_eq in Hkxy. + destruct Hkxy as [pxy Hkxy]. + etransitivity. 2: eapply (lsp0_spec_le _ (sto_G' pxy)). + rewrite sto_G'_weight. now rewrite Hkxy. + Qed. + + Lemma lsp_G'_spec_left x : + lsp G' y_0 x = max (lsp G y_0 x) (Some K + lsp G x_0 x)%nbar. + Proof. + apply Nbar.le_antisymm. + - case_eq (lsp G' y_0 x); [|cbn; trivial]. + intros k Hk. + apply lsp0_spec_eq in Hk; destruct Hk as [p' Hk]. + subst k. + rewrite (from_G'_weight p'). + destruct (from_G' p') as [p|[p1 p2]]. + + apply max_le'. left. apply lsp0_spec_le. + + apply max_le'. right. + apply (plus_le_compat (Some (sweight p1 + K)) _ (Some (sweight p2))). + * cbn. now rewrite sweight_weight, HG. + * apply lsp0_spec_le. + - apply max_lub. apply lsp_G'_incr. + case_eq (lsp G x_0 x); cbn; [intros k Hk|trivial]. + apply lsp0_spec_eq in Hk. destruct Hk as [p Hk]. + unshelve refine (transport (fun K => (Some K <= _)%nbar) + _ (lsp0_spec_le _ _)). + eapply SimplePaths_sub. shelve. + pose (p' := sto_G' p). + unshelve econstructor. + + exact (snodes G' p'). + + shelve. + + apply DisjointAdd_add2. + intro HH. eapply left, split in HH. + destruct HH as [p1 _]. + apply from_G' in p1. destruct p1 as [p1|[p1 p2]]. + all: epose proof (lsp0_spec_le _ (SimplePaths_sub _ _ p1)) as HH; + unfold lsp in Hxs; now erewrite Hxs in HH. + + exists K. apply EdgeSet.add_spec. now left. + + exact (reduce _ p'). + + rewrite weight_SimplePaths_sub. cbn. + now rewrite weight_reduce, sto_G'_weight, Hk. + Unshelve. + 2-3: now apply VSetProp.subset_remove_3. + apply VSetProp.subset_add_3. assumption. + apply snodes_Subset. + Qed. + End subgraph. + Lemma leq_vertices_caract0 {n x y} (Vy : VSet.In y (V G)) : leq_vertices G n x y <-> (Some n <= lsp G x y)%nbar. Proof. @@ -1438,7 +1500,7 @@ Module WeightedGraph (V : UsualOrderedType). - assert (Vx : VSet.In x (V G)). { case_eq (VSet.mem x (V G)); intro Vx; [now apply VSet.mem_spec in Vx|]. apply False_rect. apply VSetFact.not_mem_iff in Vx. - pose (K := S (option_get 0 (lsp G (s G) y))). + pose (K := Datatypes.S (option_get 0 (lsp G (s G) y))). pose (l := fun z => if V.eq_dec z x then K else option_get 0 (lsp G (s G) z)). unshelve refine (let XX := Hle l _ in _); subst l K. @@ -1456,30 +1518,35 @@ Module WeightedGraph (V : UsualOrderedType). destruct (V.eq_dec x x) as [Hx|Hx]; [|now apply Hx]. destruct (V.eq_dec y x) as [Hy|Hy]. now subst. lia. } - destruct (lsp_s G y Vy) as [k Hk]. + destruct (lsp_s G y Vy) as [ky Hky]. case_eq (lsp G x (s G)). - * intros m Hm. etransitivity. + + intros m Hm. etransitivity. 2: exact (lsp_codistance G x (s G) y). rewrite Hm. etransitivity. 2: eapply le_plus_r. specialize (Hle _ (lsp_correctness G)); cbn in Hle. - rewrite Hk in *. cbn in *; lia. - * intros Hxs. destruct (lsp0_spec_eq G _ Hk) as [p Hp]; subst. - case_eq (VSet.mem x (snodes G p)); intro Hx. - -- apply VSet.mem_spec in Hx. - pose proof (weight_split G p (Hu:=left Hx)). - set (pp := split G p x (left Hx)) in *; - destruct pp as [p1 p2]; cbn in *. - specialize (Hle _ (lsp_correctness G)); cbn in Hle. - rewrite Hk in Hle; cbn in Hle. etransitivity. - 2: eapply (lsp0_spec_le G p2). cbn. - destruct (lsp_s G x Vx) as [m Hm]. - rewrite Hm in Hle. - simple refine (let XX := lsp0_spec_le G (SimplePaths_sub G _ p1) in _). - exact (V G). apply VSetProp.subset_remove_3; reflexivity. - rewrite weight_SimplePaths_sub in XX. - unfold lsp in Hm; rewrite Hm in XX; cbn in XX. - cbn in Hle; lia. - -- eapply lsp_xy_le_n; try eassumption. + rewrite Hky in *. cbn in *; lia. + + intros Hxs. + destruct (lsp_s G _ Vx) as [kx Hkx]. + pose (K := Peano.max kx (Datatypes.S ky)). + unshelve epose proof (correct_labelling_lsp_G' _ _ _ _ Hxs K) as XX; + tas; try apply HI. + set (G' := G' (s G) x K) in XX. + assert (HI' : invariants G'). { + eapply HI'; tas. apply HI. } + specialize (Hle _ XX); clear XX; cbn in Hle. + assert (XX: (Some K <= lsp G' (s G) x)%nbar). { + apply lsp_G'_yx. apply HI. } + destruct (lsp_s G' _ Vx) as [kx' Hkx']. + cbn in Hkx'; rewrite Hkx' in *; cbn in *. + destruct (lsp_s G' _ Vy) as [ky' Hky']. + cbn in Hky'; rewrite Hky' in *; cbn in *. + assert (Hle' : K + n <= ky') by lia; clear Hle. + enough ((Some ky' <= Some K + lsp G x y)%nbar) as HH. { + revert HH. destruct (lsp G x y); cbn; lia. } + unshelve erewrite (lsp_G'_spec_left (s G) x _ Hxs K) in Hky'. apply HI. + apply eq_max in Hky'. destruct Hky' as [H|H]. + -- rewrite Hky in H. apply some_inj in H; subst; lia. + -- rewrite <- H. reflexivity. - case_eq (lsp G x y). * intros m Hm l Hl. rewrite Hm in Hle. apply (lsp0_spec_eq G) in Hm. destruct Hm as [p Hp]. @@ -1488,6 +1555,7 @@ Module WeightedGraph (V : UsualOrderedType). * intro X; rewrite X in Hle; inversion Hle. Defined. + Lemma leq_vertices_caract {n x y} : leq_vertices G n x y <-> (if VSet.mem y (V G) then Some n <= lsp G x y else n = 0 /\ (x = y \/ Some 0 <= lsp G x (s G)))%nbar. @@ -1518,13 +1586,17 @@ Module WeightedGraph (V : UsualOrderedType). intros Hxs; apply False_rect. case_eq (VSet.mem x (V G)); intro Vx. * apply VSet.mem_spec in Vx. - pose (G' := (V G, EdgeSet.add (s G, K 0 0, x) (E G), s G)). - pose proof (HI' Vx 0 0 : invariants G') as HI'. - pose proof (HG' Vx 0 0 Hxs : acyclic_no_loop G') as HG'. + pose (K := 1). + unshelve epose proof (correct_labelling_lsp_G' _ _ _ _ Hxs K) as X; + tas; try apply HI. + set (G' := G' (s G) x K) in X. + assert (HI' : invariants G'). { + eapply HI'; tas. apply HI. } + assert (HG' : acyclic_no_loop G'). { + eapply HG'; tas. apply HI. } pose (l := fun z => if V.eq_dec z y then 0 else option_get 0 (lsp G' (s G) z)). assert (XX : correct_labelling G l). { - pose proof (lsp_correctness G') as XX. subst l; split. -- destruct (V.eq_dec (s G) y). apply False_rect. apply Vy. subst; apply HI. @@ -1534,8 +1606,7 @@ Module WeightedGraph (V : UsualOrderedType). apply False_rect, Vy; subst; now apply HI. destruct (V.eq_dec e..t y). apply False_rect, Vy; subst; now apply HI. - cbn in XX. apply XX. - apply EdgeSet.add_spec. now right. } + now apply X. } specialize (Hle _ XX); subst l; cbn in *. destruct (V.eq_dec y y) as [_|?]; [|contradiction]. destruct (V.eq_dec x y) as [Hy|Hy]. contradiction. diff --git a/test-suite/erasure_test.v b/test-suite/erasure_test.v index 271027dd4..b74f2ac27 100644 --- a/test-suite/erasure_test.v +++ b/test-suite/erasure_test.v @@ -17,7 +17,7 @@ Environment is well-formed and Construct(Coq.Init.Datatypes.bool,0,0,[]) erases Construct(Coq.Init.Datatypes.bool,0,0) *) -MetaCoq Erase (exist _ 0 (eq_refl 0) : {x : nat | x = 0}). +MetaCoq Erase (exist _ 0 (eq_refl) : {x : nat | x = 0}). (* (* *) (* Environment is well-formed and exist nat (fun x : nat => eq nat x O) O (eq_refl nat O):sig nat (fun x : nat => eq nat x O) erases to: *) (* (fun f => f) (exist ∎ ∎ O ∎) *) diff --git a/test-suite/univ.v b/test-suite/univ.v index 0bc0642ec..4e38f8eb7 100644 --- a/test-suite/univ.v +++ b/test-suite/univ.v @@ -24,7 +24,7 @@ Monomorphic Universe i j. Set Strict Unquote Universe Mode. Test Quote (Type@{j} -> Type@{i}). -Make Definition T'' := (tSort (Universe.make' (Level.Level "j", false))). +Make Definition T'' := (tSort (Universe.make (Level.Level "j"))). Unset Strict Unquote Universe Mode. @@ -170,7 +170,7 @@ Print qtest. Make Definition bla := qtest. Unset Strict Unquote Universe Mode. -Make Definition bla' := (tLambda (nNamed "T") (tSort (Universe.make'' (Level.Level "Top.46", false) [])) (tLambda (nNamed "T2") (tSort (Universe.make'' (Level.Level "Top.477", false) [])) (tProd nAnon (tRel 1) (tRel 1)))). +Make Definition bla' := (tLambda (nNamed "T") (tSort (Universe.make (Level.Level "Top.46"))) (tLambda (nNamed "T2") (tSort (Universe.make (Level.Level "Top.477"))) (tProd nAnon (tRel 1) (tRel 1)))). Set Printing Universes. Print bla. @@ -224,11 +224,10 @@ Definition f' := (forall (A:Type@{i1}) (B: Type@{j1}), A -> B -> A). Quote Recursively Definition ff := f'. Require Import MetaCoq.Checker.All. +Compute (infer' (empty_ext (fst ff)) [] (snd ff)). Check (eq_refl : infer' (empty_ext (fst ff)) [] (snd ff) = - Checked (tSort - (NEL.cons (Level.Level _, true) - (NEL.sing (Level.Level _, true))))). + Checked (tSort (Universe.from_kernel_repr (Level.Level _, true) [(Level.Level _, true)]))). Open Scope string_scope. -Check (eq_refl : infer [] init_graph [] ((tProd (nNamed "A") (tSort (Universe.make' (Level.Level "Toto.85", false))) (tProd (nNamed "B") (tSort (Universe.make' (Level.Level "Toto.86", false))) (tProd nAnon (tRel 1) (tProd nAnon (tRel 1) (tRel 3)))))) = Checked (tSort _)). +Check (eq_refl : infer [] init_graph [] ((tProd (nNamed "A") (tSort (Universe.make (Level.Level _))) (tProd (nNamed "B") (tSort (Universe.make (Level.Level _))) (tProd nAnon (tRel 1) (tProd nAnon (tRel 1) (tRel 3)))))) = Checked (tSort _)). -Make Definition t4 := (tSort (Universe.make (fresh_level))). \ No newline at end of file +(* Make Definition t4 := (tSort (Universe.make (fresh_level))). *)