diff --git a/common/theories/BasicAst.v b/common/theories/BasicAst.v index 6bb51924a..f364e0d79 100644 --- a/common/theories/BasicAst.v +++ b/common/theories/BasicAst.v @@ -213,20 +213,17 @@ Proof. eapply map_def_spec; eauto. Qed. -Variant typ_or_sort_ {term} := Typ (T : term) | Sort. -Arguments typ_or_sort_ : clear implicits. - -Definition typ_or_sort_map {T T'} (f: T -> T') t := - match t with - | Typ T => Typ (f T) - | Sort => Sort - end. - -Definition typ_or_sort_default {T A} (f: T -> A) t d := - match t with - | Typ T => f T - | Sort => d - end. +Record judgment_ {universe Term} := Judge { + j_term : option Term; + j_typ : Term; + j_univ : option universe; + (* j_rel : option relevance; *) +}. +Arguments judgment_ : clear implicits. +Arguments Judge {universe Term} _ _ _. + +Definition judgment_map {univ T A} (f: T -> A) (j : judgment_ univ T) := + Judge (option_map f (j_term j)) (f (j_typ j)) (j_univ j) (* (j_rel j) *). Section Contexts. Context {term : Type}. @@ -242,6 +239,18 @@ End Contexts. Arguments context_decl : clear implicits. +Notation Typ typ := (Judge None typ None). +Notation TermTyp tm ty := (Judge (Some tm) ty None). +Notation TermoptTyp tm typ := (Judge tm typ None). +Notation TypUniv ty u := (Judge None ty (Some u)). +Notation TermTypUniv tm ty u := (Judge (Some tm) ty (Some u)). + +Notation j_vass na ty := (Typ ty (* na.(binder_relevance) *)). +Notation j_vass_s na ty s := (TypUniv ty s (* na.(binder_relevance) *)). +Notation j_vdef na b ty := (TermTyp b ty (* na.(binder_relevance) *)). +Notation j_decl d := (TermoptTyp (decl_body d) (decl_type d) (* (decl_name d).(binder_relevance) *)). +Notation j_decl_s d s := (Judge (decl_body d) (decl_type d) s (* (decl_name d).(binder_relevance) *)). + Definition map_decl {term term'} (f : term -> term') (d : context_decl term) : context_decl term' := {| decl_name := d.(decl_name); decl_body := option_map f d.(decl_body); @@ -308,8 +317,46 @@ Definition snoc {A} (Γ : list A) (d : A) := d :: Γ. Notation " Γ ,, d " := (snoc Γ d) (at level 20, d at next level). +Definition app_context {A} (Γ Γ': list A) := Γ' ++ Γ. + +Notation "Γ ,,, Γ'" := (app_context Γ Γ') (at level 25, Γ' at next level, left associativity). + +Lemma app_context_nil_l {T} Γ : [] ,,, Γ = Γ :> list T. +Proof. + unfold app_context. rewrite app_nil_r. reflexivity. +Qed. + +Lemma app_context_assoc {T} Γ Γ' Γ'' : Γ ,,, (Γ' ,,, Γ'') = Γ ,,, Γ' ,,, Γ'' :> list T. +Proof. unfold app_context; now rewrite app_assoc. Qed. + +Lemma app_context_cons {T} Γ Γ' A : Γ ,,, (Γ' ,, A) = (Γ ,,, Γ') ,, A :> list T. +Proof. exact (app_context_assoc _ _ [A]). Qed. + +Lemma app_context_push {T} Γ Δ Δ' d : (Γ ,,, Δ ,,, Δ') ,, d = (Γ ,,, Δ ,,, (Δ' ,, d)) :> list T. +Proof using Type. + reflexivity. +Qed. + +Lemma snoc_app_context {T Γ Δ d} : (Γ ,,, (d :: Δ)) = (Γ ,,, Δ) ,,, [d] :> list T. +Proof using Type. + reflexivity. +Qed. + +Lemma app_context_length {T} (Γ Γ' : list T) : #|Γ ,,, Γ'| = #|Γ'| + #|Γ|. +Proof. unfold app_context. now rewrite app_length. Qed. +#[global] Hint Rewrite @app_context_length : len. + +Lemma nth_error_app_context_ge {T} v Γ Γ' : + #|Γ'| <= v -> nth_error (Γ ,,, Γ') v = nth_error Γ (v - #|Γ'|) :> option T. +Proof. apply nth_error_app_ge. Qed. + +Lemma nth_error_app_context_lt {T} v Γ Γ' : + v < #|Γ'| -> nth_error (Γ ,,, Γ') v = nth_error Γ' v :> option T. +Proof. apply nth_error_app_lt. Qed. + + Definition ondecl {A} (P : A -> Type) (d : context_decl A) := - P d.(decl_type) × option_default P d.(decl_body) unit. + option_default P d.(decl_body) unit × P d.(decl_type). Notation onctx P := (All (ondecl P)). diff --git a/common/theories/Environment.v b/common/theories/Environment.v index debe1d4aa..bd47913f8 100644 --- a/common/theories/Environment.v +++ b/common/theories/Environment.v @@ -9,7 +9,7 @@ Module Type Term. Parameter Inline term : Type. Parameter Inline tRel : nat -> term. - Parameter Inline tSort : Universe.t -> term. + Parameter Inline tSort : Sort.t -> term. Parameter Inline tProd : aname -> term -> term -> term. Parameter Inline tLambda : aname -> term -> term -> term. Parameter Inline tLetIn : aname -> term -> term -> term -> term. @@ -129,7 +129,7 @@ Module Environment (T : Term). Import T. #[global] Existing Instance subst_instance_constr. - Definition typ_or_sort := typ_or_sort_ term. + Definition judgment := judgment_ Sort.t term. (** ** Declarations *) Notation context_decl := (context_decl term). @@ -344,7 +344,7 @@ Module Environment (T : Term). Record one_inductive_body := { ind_name : ident; ind_indices : context; (* Indices of the inductive types, under params *) - ind_sort : Universe.t; (* Sort of the inductive. *) + ind_sort : Sort.t; (* Sort of the inductive. *) ind_type : term; (* Closed arity = forall mind_params, ind_indices, tSort ind_sort *) ind_kelim : allowed_eliminations; (* Allowed eliminations *) ind_ctors : list constructor_body; @@ -856,10 +856,10 @@ Module Environment (T : Term). Definition primitive_invariants (p : prim_tag) (cdecl : constant_body) := match p with | primInt | primFloat => - [/\ cdecl.(cst_type) = tSort Universe.type0, cdecl.(cst_body) = None & + [/\ cdecl.(cst_type) = tSort Sort.type0, cdecl.(cst_body) = None & cdecl.(cst_universes) = Monomorphic_ctx] | primArray => - let s := Universe.make (Level.lvar 0) in + let s := sType (Universe.make' (Level.lvar 0)) in [/\ cdecl.(cst_type) = tImpl (tSort s) (tSort s), cdecl.(cst_body) = None & cdecl.(cst_universes) = Polymorphic_ctx array_uctx] end. @@ -882,12 +882,6 @@ Module Environment (T : Term). Definition program : Type := global_env * term. - (* TODO MOVE AstUtils factorisation *) - - Definition app_context (Γ Γ' : context) : context := Γ' ++ Γ. - Notation "Γ ,,, Γ'" := - (app_context Γ Γ') (at level 25, Γ' at next level, left associativity). - (** Make a lambda/let-in string of abstractions from a context [Γ], ending with term [t]. *) Definition mkLambda_or_LetIn d t := @@ -1008,30 +1002,6 @@ Module Environment (T : Term). Proof. unfold arities_context. now rewrite rev_map_length. Qed. #[global] Hint Rewrite arities_context_length : len. - Lemma app_context_nil_l Γ : [] ,,, Γ = Γ. - Proof. - unfold app_context. rewrite app_nil_r. reflexivity. - Qed. - - Lemma app_context_assoc Γ Γ' Γ'' : Γ ,,, (Γ' ,,, Γ'') = Γ ,,, Γ' ,,, Γ''. - Proof. unfold app_context; now rewrite app_assoc. Qed. - - Lemma app_context_cons Γ Γ' A : Γ ,,, (Γ' ,, A) = (Γ ,,, Γ') ,, A. - Proof. exact (app_context_assoc _ _ [A]). Qed. - - Lemma app_context_length Γ Γ' : #|Γ ,,, Γ'| = #|Γ'| + #|Γ|. - Proof. unfold app_context. now rewrite app_length. Qed. - #[global] Hint Rewrite app_context_length : len. - - Lemma nth_error_app_context_ge v Γ Γ' : - #|Γ'| <= v -> nth_error (Γ ,,, Γ') v = nth_error Γ (v - #|Γ'|). - Proof. apply nth_error_app_ge. Qed. - - Lemma nth_error_app_context_lt v Γ Γ' : - v < #|Γ'| -> nth_error (Γ ,,, Γ') v = nth_error Γ' v. - Proof. apply nth_error_app_lt. Qed. - - Definition map_mutual_inductive_body f m := match m with | Build_mutual_inductive_body finite ind_npars ind_pars ind_bodies ind_universes ind_variance => @@ -1269,7 +1239,7 @@ End EnvironmentDecideReflectInstances. Module Type TermUtils (T: Term) (E: EnvironmentSig T). Import T E. - Parameter Inline destArity : context -> term -> option (context × Universe.t). + Parameter Inline destArity : context -> term -> option (context × Sort.t). Parameter Inline inds : kername -> Instance.t -> list one_inductive_body -> list term. End TermUtils. diff --git a/common/theories/EnvironmentTyping.v b/common/theories/EnvironmentTyping.v index 7701aa5f4..8e7f8757b 100644 --- a/common/theories/EnvironmentTyping.v +++ b/common/theories/EnvironmentTyping.v @@ -280,16 +280,16 @@ Module Lookup (T : Term) (E : EnvironmentSig T). now rewrite H; cbn; autorewrite with len. Qed. - Definition wf_universe Σ s : Prop := - Universe.on_sort - (fun u => forall l, LevelExprSet.In l u -> LevelSet.In (LevelExpr.get_level l) (global_ext_levels Σ)) - True s. + Definition wf_universe Σ (u : Universe.t) : Prop := + forall l, LevelExprSet.In l u -> LevelSet.In (LevelExpr.get_level l) (global_ext_levels Σ). - Definition wf_universe_dec Σ s : {@wf_universe Σ s} + {~@wf_universe Σ s}. + Definition wf_sort Σ (s : sort) : Prop := + Sort.on_sort (wf_universe Σ) True s. + + Definition wf_universe_dec Σ u : {wf_universe Σ u} + {~wf_universe Σ u}. Proof. - destruct s; try (left; exact I). - cbv [wf_universe Universe.on_sort LevelExprSet.In LevelExprSet.this t_set]. - destruct t as [[t _] _]. + cbv [wf_universe LevelExprSet.In LevelExprSet.this t_set]. + destruct u as [[t _] _]. induction t as [|t ts [IHt|IHt]]; [ left | | right ]. { inversion 1. } { destruct (LevelSetProp.In_dec (LevelExpr.get_level t) (global_ext_levels Σ)) as [H|H]; [ left | right ]. @@ -298,6 +298,12 @@ Module Lookup (T : Term) (E : EnvironmentSig T). { intro H; apply IHt; intros; apply H; now constructor. } Defined. + Definition wf_sort_dec Σ s : {@wf_sort Σ s} + {~@wf_sort Σ s}. + Proof. + destruct s; try (left; exact I). + apply wf_universe_dec. + Defined. + Lemma declared_ind_declared_constructors `{cf : checker_flags} {Σ ind mib oib} : declared_inductive Σ ind mib oib -> Alli (fun i => declared_constructor Σ (ind, i) mib oib) 0 (ind_ctors oib). @@ -316,8 +322,286 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). Import T E TU. + (** Well-formedness of local environments embeds a sorting for each variable *) + + Notation on_local_decl P Γ d := + (P Γ (j_decl d)). + + Definition on_def_type (P : context -> judgment -> Type) Γ d := + P Γ (Typ d.(dtype)). + + Definition on_def_body (P : context -> judgment -> Type) types Γ d := + P (Γ ,,, types) (TermTyp d.(dbody) (lift0 #|types| d.(dtype))). + + (* Various kinds of lifts *) + + Definition lift_wf_term wf_term (j : judgment) := option_default wf_term (j_term j) (unit : Type) × wf_term (j_typ j). + Notation lift_wf_term1 wf_term := (fun (Γ : context) => lift_wf_term (wf_term Γ)). + + Definition lift_wfu_term wf_term wf_univ (j : judgment) := option_default wf_term (j_term j) unit × wf_term (j_typ j) × option_default wf_univ (j_univ j) unit. + + Definition lift_wfb_term wfb_term (j : judgment) := option_default wfb_term (j_term j) true && wfb_term (j_typ j). + Notation lift_wfb_term1 wfb_term := (fun (Γ : context) => lift_wfb_term (wfb_term Γ)). + + Definition lift_wfbu_term wfb_term wfb_univ (j : judgment) := option_default wfb_term (j_term j) true && wfb_term (j_typ j) && option_default wfb_univ (j_univ j) true. + + Definition lift_sorting checking sorting : judgment -> Type := + fun j => option_default (fun tm => checking tm (j_typ j)) (j_term j) (unit : Type) × + ∑ s, sorting (j_typ j) s × option_default (fun u => (u = s : Type)) (j_univ j) unit. + + Notation lift_sorting1 checking sorting := (fun Γ => lift_sorting (checking Γ) (sorting Γ)). + Notation lift_sorting2 checking sorting := (fun Σ Γ => lift_sorting (checking Σ Γ) (sorting Σ Γ)). + + Notation typing_sort typing := (fun T s => typing T (tSort s)). + Notation typing_sort1 typing := (fun Γ T s => typing Γ T (tSort s)). + Notation typing_sort2 typing := (fun Σ Γ T s => typing Σ Γ T (tSort s)). + + Definition lift_typing0 typing := lift_sorting typing (typing_sort typing). + Notation lift_typing1 typing := (fun Γ => lift_typing0 (typing Γ)). + Notation lift_typing typing := (fun Σ Γ => lift_typing0 (typing Σ Γ)). + + Notation Prop_local_conj P Q := (fun Γ t T => P Γ t T × Q Γ t T). + Notation Prop_conj P Q := (fun Σ Γ t T => P Σ Γ t T × Q Σ Γ t T). + + Definition lift_typing_conj (P Q : context -> _) := lift_typing1 (Prop_local_conj P Q). + + Lemma lift_wf_term_it_impl {P Q} {tm tm' : option term} {t t' : term} {u u'} : + forall tu: lift_wf_term P (Judge tm t u), + match tm', tm with None, _ => unit | Some tm', Some tm => P tm -> Q tm' | _, _ => False end -> + (P t -> Q t') -> + lift_wf_term Q (Judge tm' t' u'). + Proof. + intros (Htm & Hs) HPQc HPQs. + split; auto. + destruct tm, tm' => //=. now apply HPQc. + Qed. + + Lemma lift_wf_term_f_impl P Q tm t u u' : + forall f, + lift_wf_term P (Judge tm t u) -> + (forall t, P t -> Q (f t)) -> + lift_wf_term Q (Judge (option_map f tm) (f t) u'). + Proof. + unfold lift_wf_term; cbn. + intros f (Hb & Ht) HPQ. + split; auto. + destruct tm; cbn in *; auto. + Defined. + + Lemma lift_wf_term_impl P Q j : + lift_wf_term P j -> + (forall t, P t -> Q t) -> + lift_wf_term Q j. + Proof. + unfold lift_wf_term; cbn. + intros (Hb & Ht) HPQ. + split; auto. + destruct j_term; cbn in *; auto. + Defined. + + Lemma lift_wfbu_term_f_impl (P Q : term -> bool) tm t u : + forall f fu, + lift_wfbu_term P (P ∘ tSort) (Judge tm t u) -> + (forall u, f (tSort u) = tSort (fu u)) -> + (forall t, P t -> Q (f t)) -> + lift_wfbu_term Q (Q ∘ tSort) (Judge (option_map f tm) (f t) (option_map fu u)). + Proof. + unfold lift_wfbu_term; cbn. + intros. rtoProp. + repeat split; auto. + 1: destruct tm; cbn in *; auto. + destruct u; rewrite //= -H0 //. auto. + Defined. + + Lemma unlift_TermTyp {Pc Ps tm ty u} : + lift_sorting Pc Ps (Judge (Some tm) ty u) -> + Pc tm ty. + Proof. + apply fst. + Defined. + + Definition unlift_TypUniv {Pc Ps tm ty u} : + lift_sorting Pc Ps (Judge tm ty (Some u)) -> + Ps ty u + := fun H => eq_rect_r _ H.2.π2.1 H.2.π2.2. + + Definition lift_sorting_extract {c s tm ty u} (w : lift_sorting c s (Judge tm ty u)) : + lift_sorting c s (Judge tm ty (Some w.2.π1)) := + (w.1, existT _ w.2.π1 (w.2.π2.1, eq_refl)). + + Lemma lift_sorting_forget_univ {Pc Ps tm ty u} : + lift_sorting Pc Ps (Judge tm ty u) -> + lift_sorting Pc Ps (TermoptTyp tm ty). + Proof. + intros (? & ? & ? & ?). + repeat (eexists; tea). + Qed. + + Lemma lift_sorting_forget_body {Pc Ps tm ty u} : + lift_sorting Pc Ps (Judge tm ty u) -> + lift_sorting Pc Ps (Judge None ty u). + Proof. + intros (? & ? & ? & ?). + repeat (eexists; tea). + Qed. + + Lemma lift_sorting_ex_it_impl_gen {Pc Qc Ps Qs} {tm tm' : option term} {t t' : term} : + forall tu: lift_sorting Pc Ps (TermoptTyp tm t), + let s := tu.2.π1 in + match tm', tm with None, _ => unit | Some tm', Some tm => Pc tm t -> Qc tm' t' | _, _ => False end -> + (Ps t s -> ∑ s', Qs t' s') -> + lift_sorting Qc Qs (TermoptTyp tm' t'). + Proof. + intros (? & ? & Hs & e) s HPQc HPQs. + split. + - destruct tm, tm' => //=. now apply HPQc. + - specialize (HPQs Hs) as (s' & H). + eexists. split; eassumption. + Qed. + + Lemma lift_sorting_it_impl_gen {Pc Qc Ps Qs} {tm tm' : option term} {t t' : term} {u} : + forall tu: lift_sorting Pc Ps (Judge tm t u), + let s := tu.2.π1 in + match tm', tm with None, _ => unit | Some tm', Some tm => Pc tm t -> Qc tm' t' | _, _ => False end -> + (Ps t s -> Qs t' s) -> + lift_sorting Qc Qs (Judge tm' t' u). + Proof. + intros (? & ? & Hs & e) s HPQc HPQs. + split. + - destruct tm, tm' => //=. now apply HPQc. + - eexists. split; [now apply HPQs|]. + destruct u => //. + Qed. + + Lemma lift_sorting_fu_it_impl {Pc Qc Ps Qs} {tm : option term} {t : term} {u} : + forall f fu, forall tu: lift_sorting Pc Ps (Judge tm t u), + let s := tu.2.π1 in + option_default (fun tm => Pc tm t -> Qc (f tm) (f t)) tm unit -> + (Ps t s -> Qs (f t) (fu s)) -> + lift_sorting Qc Qs (Judge (option_map f tm) (f t) (option_map fu u)). + Proof. + intros ?? (? & ? & Hs & e) s HPQc HPQs. + split. + - destruct tm => //=. now apply HPQc. + - eexists. split; [now apply HPQs|]. + destruct u => //. + cbn. f_equal => //. + Qed. + + Lemma lift_sorting_f_it_impl {Pc Qc Ps Qs} {j : judgment} : + forall f, forall tu: lift_sorting Pc Ps j, + let s := tu.2.π1 in + option_default (fun tm => Pc tm (j_typ j) -> Qc (f tm) (f (j_typ j))) (j_term j) unit -> + (Ps (j_typ j) s -> Qs (f (j_typ j)) s) -> + lift_sorting Qc Qs (judgment_map f j). + Proof. + intro f. + replace (judgment_map f j) with (Judge (option_map f (j_term j)) (f (j_typ j)) (option_map id (j_univ j))). + 2: unfold judgment_map; destruct j_univ => //. + apply lift_sorting_fu_it_impl with (fu := id). + Qed. + + Lemma lift_sorting_it_impl {Pc Qc Ps Qs} {j} : + forall tu: lift_sorting Pc Ps j, + let s := tu.2.π1 in + option_default (fun tm => Pc tm (j_typ j) -> Qc tm (j_typ j)) (j_term j) unit -> + (Ps (j_typ j) s -> Qs (j_typ j) s) -> + lift_sorting Qc Qs j. + Proof. + relativize (lift_sorting Qc Qs j). + 1: apply lift_sorting_f_it_impl with (f := id). + destruct j, j_term => //. + Qed. + + Lemma lift_sorting_fu_impl {Pc Qc Ps Qs tm t u} : + forall f fu, + lift_sorting Pc Ps (Judge tm t u) -> + (forall t T, Pc t T -> Qc (f t) (f T)) -> + (forall t u, Ps t u -> Qs (f t) (fu u)) -> + lift_sorting Qc Qs (Judge (option_map f tm) (f t) (option_map fu u)). + Proof. + intros ?? tu ??. + apply lift_sorting_fu_it_impl with (tu := tu); auto. + destruct tm => //=. auto. + Qed. + + Lemma lift_typing_fu_impl {P Q tm t u} : + forall f fu, + lift_typing0 P (Judge tm t u) -> + (forall t T, P t T -> Q (f t) (f T)) -> + (forall u, f (tSort u) = tSort (fu u)) -> + lift_typing0 Q (Judge (option_map f tm) (f t) (option_map fu u)). + Proof. + intros ?? HT HPQ Hf. + apply lift_sorting_fu_impl with (1 := HT); tas. + intros; rewrite -Hf; now apply HPQ. + Qed. + + Lemma lift_sorting_f_impl {Pc Qc Ps Qs j} : + forall f, + lift_sorting Pc Ps j -> + (forall t T, Pc t T -> Qc (f t) (f T)) -> + (forall t u, Ps t u -> Qs (f t) u) -> + lift_sorting Qc Qs (judgment_map f j). + Proof. + intros ? tu ??. + apply lift_sorting_f_it_impl with (tu := tu); auto. + destruct j_term => //=. auto. + Qed. + + Lemma lift_typing_f_impl {P Q j} : + forall f, + lift_typing0 P j -> + (forall t T, P t T -> Q (f t) (f T)) -> + (forall u, f (tSort u) = tSort u) -> + lift_typing0 Q (judgment_map f j). + Proof. + intros ? HT HPQ Hf. + apply lift_sorting_f_impl with (1 := HT); tas. + intros; rewrite -Hf; now apply HPQ. + Qed. + + Lemma lift_typing_map {P} f j : + lift_typing0 (fun t T => P (f t) (f T)) j -> + (forall u, f (tSort u) = tSort u) -> + lift_typing0 P (judgment_map f j). + Proof. + intros HT Hf. + apply lift_typing_f_impl with (1 := HT) => //. + Qed. + + Lemma lift_typing_mapu {P} f fu {tm ty u} : + lift_typing0 (fun t T => P (f t) (f T)) (Judge tm ty u) -> + (forall u, f (tSort u) = tSort (fu u)) -> + lift_typing0 P (Judge (option_map f tm) (f ty) (option_map fu u)). + Proof. + intros HT. + eapply lift_typing_fu_impl with (1 := HT) => //. + Qed. + + Lemma lift_sorting_impl {Pc Qc Ps Qs j} : + lift_sorting Pc Ps j -> + (forall t T, Pc t T -> Qc t T) -> + (forall t u, Ps t u -> Qs t u) -> + lift_sorting Qc Qs j. + Proof. + intros tu ??. + apply lift_sorting_it_impl with (tu := tu); auto. + destruct j_term => //=. auto. + Qed. + + Lemma lift_typing_impl {P Q j} : + lift_typing0 P j -> + (forall t T, P t T -> Q t T) -> + lift_typing0 Q j. + Proof. + intros HT HPQ. + apply lift_sorting_impl with (1 := HT); tas. + intros; now apply HPQ. + Qed. + Section TypeLocal. - Context (typing : forall (Γ : context), term -> typ_or_sort -> Type). + Context (typing : forall (Γ : context), judgment -> Type). Inductive All_local_env : context -> Type := | localenv_nil : @@ -325,14 +609,14 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). | localenv_cons_abs Γ na t : All_local_env Γ -> - typing Γ t Sort -> + typing Γ (j_vass na t) -> All_local_env (Γ ,, vass na t) | localenv_cons_def Γ na b t : All_local_env Γ -> - typing Γ t Sort -> - typing Γ b (Typ t) -> + typing Γ (j_vdef na b t) -> All_local_env (Γ ,, vdef na b t). + Derive Signature NoConfusion for All_local_env. End TypeLocal. @@ -340,8 +624,38 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). Arguments localenv_cons_def {_ _ _ _ _} _ _. Arguments localenv_cons_abs {_ _ _ _} _ _. + Definition localenv_cons {typing Γ na bo t} : + All_local_env typing Γ -> typing Γ (TermoptTyp bo t) -> All_local_env typing (Γ ,, mkdecl na bo t) + := match bo with None => localenv_cons_abs | Some b => localenv_cons_def end. + + Definition All_local_env_snoc {P Γ decl} : All_local_env P Γ -> on_local_decl P Γ decl -> All_local_env P (Γ ,, decl) := + match decl with mkdecl na bo t => localenv_cons end. + + Lemma All_local_env_tip {typing Γ decl} : All_local_env typing (Γ ,, decl) -> All_local_env typing Γ × on_local_decl typing Γ decl. + Proof. + intros wfΓ; depelim wfΓ. + all: split; assumption. + Defined. + + Lemma All_local_env_ind1 typing P : + P [] -> + (forall Γ decl, P Γ -> on_local_decl typing Γ decl -> P (Γ ,, decl)) -> + forall Γ, All_local_env typing Γ -> P Γ. + Proof. + induction Γ => //. + move/All_local_env_tip => [] ??. + now apply X0. + Defined. + + Lemma All_local_env_map (P : context -> judgment -> Type) f Γ : + All_local_env (fun Γ j => P (map (map_decl f) Γ) (judgment_map f j)) Γ -> + All_local_env P (map (map_decl f) Γ). + Proof using Type. + induction 1; econstructor; eauto. + Qed. + Lemma All_local_env_fold P f Γ : - All_local_env (fun Γ t T => P (fold_context_k f Γ) (f #|Γ| t) (typ_or_sort_map (f #|Γ|) T)) Γ <~> + All_local_env (fun Γ j => P (fold_context_k f Γ) (judgment_map (f #|Γ|) j)) Γ <~> All_local_env P (fold_context_k f Γ). Proof. split. @@ -351,48 +665,101 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). * destruct a as [na [b|] ty]; depelim H; specialize (IHΓ H); constructor; simpl; auto. Qed. - Lemma All_local_env_impl (P Q : context -> term -> typ_or_sort -> Type) l : + Lemma All_local_env_impl_gen (P Q : context -> judgment -> Type) l : All_local_env P l -> - (forall Γ t T, P Γ t T -> Q Γ t T) -> + (forall Γ decl, P Γ (j_decl decl) -> Q Γ (j_decl decl)) -> All_local_env Q l. Proof. - induction 1; intros; simpl; econstructor; eauto. + intros H X. + induction H using All_local_env_ind1. 1: constructor. + apply All_local_env_snoc; auto. Qed. - Lemma All_local_env_impl_ind {P Q : context -> term -> typ_or_sort -> Type} {l} : + Lemma All_local_env_impl (P Q : context -> judgment -> Type) l : All_local_env P l -> - (forall Γ t T, All_local_env Q Γ -> P Γ t T -> Q Γ t T) -> + (forall Γ j, P Γ j -> Q Γ j) -> All_local_env Q l. Proof. induction 1; intros; simpl; econstructor; eauto. Qed. - Lemma All_local_env_skipn P Γ : All_local_env P Γ -> forall n, All_local_env P (skipn n Γ). + Lemma All_local_env_impl_ind {P Q : context -> judgment -> Type} {l} : + All_local_env P l -> + (forall Γ j, All_local_env Q Γ -> P Γ j -> Q Γ j) -> + All_local_env Q l. Proof. - induction 1; simpl; intros; destruct n; simpl; try econstructor; eauto. + induction 1; intros; simpl; econstructor; eauto. Qed. + + Lemma All_local_env_skipn {P Γ} n : All_local_env P Γ -> All_local_env P (skipn n Γ). + Proof. + intros hΓ. + induction n in Γ, hΓ |- * => //. + destruct Γ; cbn; eauto. + apply All_local_env_tip in hΓ as []. + eauto. + Defined. #[global] Hint Resolve All_local_env_skipn : wf. + Lemma All_local_env_app_skipn {P Γ Δ} n : All_local_env P (Γ ,,, Δ) -> + All_local_env P (Γ ,,, skipn n Δ). + Proof. + intros hΓ. + induction n in Δ, hΓ |- * => //. + destruct Δ; cbn; eauto. + apply All_local_env_tip in hΓ as []. + eauto. + Qed. + + Lemma All_local_env_nth_error {P Γ n decl} : All_local_env P Γ -> nth_error Γ n = Some decl -> on_local_decl P (skipn (S n) Γ) decl. + Proof. + induction Γ in n |- *; destruct n => //= /All_local_env_tip [] wfΓ ondecl Hnth //=. + - injection Hnth as [= ->]. assumption. + - now eapply IHΓ. + Defined. + + Lemma All_local_env_cst {P Γ} : All_local_env (fun _ => P) Γ <~> All (fun d => P (j_decl d)) Γ. + Proof. + split. + - induction 1 using All_local_env_ind1; constructor => //. + - induction 1. 1: constructor. + apply All_local_env_snoc => //. + Defined. + Section All_local_env_rel. Definition All_local_rel P Γ Γ' - := (All_local_env (fun Γ0 t T => P (Γ ,,, Γ0) t T) Γ'). + := (All_local_env (fun Δ j => P (Γ ,,, Δ) j) Γ'). Definition All_local_rel_nil {P Γ} : All_local_rel P Γ [] := localenv_nil. + Definition All_local_rel_snoc {P Γ Γ' decl} : + All_local_rel P Γ Γ' -> on_local_decl P (Γ ,,, Γ') decl -> + All_local_rel P Γ (Γ' ,, decl) + := All_local_env_snoc. + Definition All_local_rel_abs {P Γ Γ' A na} : - All_local_rel P Γ Γ' -> P (Γ ,,, Γ') A Sort + All_local_rel P Γ Γ' -> P (Γ ,,, Γ') (j_vass na A) -> All_local_rel P Γ (Γ',, vass na A) - := localenv_cons_abs. + := localenv_cons. Definition All_local_rel_def {P Γ Γ' t A na} : All_local_rel P Γ Γ' -> - P (Γ ,,, Γ') A Sort -> - P (Γ ,,, Γ') t (Typ A) -> + P (Γ ,,, Γ') (j_vdef na t A) -> All_local_rel P Γ (Γ',, vdef na t A) - := localenv_cons_def. + := localenv_cons. + + Definition All_local_rel_tip {P Γ Γ' decl} : + All_local_rel P Γ (Γ' ,, decl) -> All_local_rel P Γ Γ' × on_local_decl P (Γ ,,, Γ') decl + := All_local_env_tip. + + Definition All_local_rel_ind1 typing Γ P : + P [] -> + (forall Δ decl, P Δ -> on_local_decl typing (Γ ,,, Δ) decl -> P (Δ ,, decl)) -> + forall Δ, All_local_rel typing Γ Δ -> P Δ + := All_local_env_ind1 _ P. Lemma All_local_rel_local : forall P Γ, @@ -401,17 +768,15 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). Proof. intros P Γ h. eapply All_local_env_impl. - exact h. - - intros Δ t [] h'. - all: cbn. - + rewrite app_context_nil_l. assumption. - + rewrite app_context_nil_l. assumption. + - intros. + rewrite app_context_nil_l. assumption. Defined. Lemma All_local_local_rel P Γ : All_local_rel P [] Γ -> All_local_env P Γ. Proof. intro X. eapply All_local_env_impl. exact X. - intros Γ0 t [] XX; cbn in XX; rewrite app_context_nil_l in XX; assumption. + intros Γ0 j XX; cbn in XX; rewrite app_context_nil_l in XX; assumption. Defined. Lemma All_local_app_rel {P Γ Γ'} : @@ -422,168 +787,172 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). split. 1: exact hΓ. constructor. - - inversion 1 ; subst. - all: edestruct IHΓ' ; auto. - all: split ; auto. - all: constructor ; auto. + - move => /= /All_local_env_tip [] hΓ ona. + edestruct IHΓ' ; auto. + split ; auto. + apply All_local_rel_snoc; auto. Defined. - Lemma All_local_rel_app {P Γ Γ'} : - All_local_env P Γ -> All_local_rel P Γ Γ' -> All_local_env P (Γ ,,, Γ'). + Definition All_local_env_app_inv {P Γ Γ'} := @All_local_app_rel P Γ Γ'. + + Lemma All_local_rel_app_inv {P Γ Γ' Γ''} : + All_local_rel P Γ (Γ' ,,, Γ'') -> All_local_rel P Γ Γ' × All_local_rel P (Γ ,,, Γ') Γ''. Proof. - intros ? hΓ'. - induction hΓ'. - - assumption. - - cbn. - constructor ; auto. - - cbn. - constructor ; auto. + intro H. + eapply All_local_env_app_inv in H as [H H']. + split; tas. + apply All_local_env_impl with (1 := H'). + intros; now rewrite -app_context_assoc. Defined. - End All_local_env_rel. - - (** Well-formedness of local environments embeds a sorting for each variable *) - - Definition on_local_decl (P : context -> term -> typ_or_sort -> Type) Γ d := - match d.(decl_body) with - | Some b => P Γ b (Typ d.(decl_type)) - | None => P Γ d.(decl_type) Sort - end. - - Definition on_def_type (P : context -> term -> typ_or_sort -> Type) Γ d := - P Γ d.(dtype) Sort. - - Definition on_def_body (P : context -> term -> typ_or_sort -> Type) types Γ d := - P (Γ ,,, types) d.(dbody) (Typ (lift0 #|types| d.(dtype))). - - Definition lift_judgment - (check : global_env_ext -> context -> term -> term -> Type) - (infer_sort : global_env_ext -> context -> term -> Type) : - (global_env_ext -> context -> term -> typ_or_sort -> Type) := - fun Σ Γ t T => - match T with - | Typ T => check Σ Γ t T - | Sort => infer_sort Σ Γ t - end. - - Lemma lift_judgment_impl {P Ps Q Qs Σ Σ' Γ Γ' t t' T} : - lift_judgment P Ps Σ Γ t T -> - (forall T, P Σ Γ t T -> Q Σ' Γ' t' T) -> - (Ps Σ Γ t -> Qs Σ' Γ' t') -> - lift_judgment Q Qs Σ' Γ' t' T. - Proof. - intros HT HPQ HPsQs. - destruct T; simpl. - * apply HPQ, HT. - * apply HPsQs, HT. - Qed. - - (* Common uses *) - - Definition lift_wf_term wf_term := (lift_judgment (fun Σ Γ t T => wf_term Σ Γ t × wf_term Σ Γ T) wf_term). - - Definition infer_sort (sorting : global_env_ext -> context -> term -> Universe.t -> Type) := (fun Σ Γ T => { s : Universe.t & sorting Σ Γ T s }). - Notation typing_sort typing := (fun Σ Γ T s => typing Σ Γ T (tSort s)). - - Definition lift_typing typing := lift_judgment typing (infer_sort (typing_sort typing)). - Definition lift_sorting checking sorting := lift_judgment checking (infer_sort sorting). - - Notation Prop_conj P Q := (fun Σ Γ t T => P Σ Γ t T × Q Σ Γ t T). + Lemma All_local_env_app {P Γ Γ'} : + All_local_env P Γ -> All_local_rel P Γ Γ' -> All_local_env P (Γ ,,, Γ'). + Proof. + induction 2 using All_local_rel_ind1 => //=. + apply All_local_env_snoc; tas. + Defined. - Definition lift_typing2 P Q := lift_typing (Prop_conj P Q). + Lemma All_local_rel_app {P Γ Γ' Γ''} : + All_local_rel P Γ Γ' -> All_local_rel P (Γ ,,, Γ') Γ'' -> All_local_rel P Γ (Γ' ,,, Γ''). + Proof. + induction 2 using All_local_rel_ind1 => //=. + apply All_local_rel_snoc; tas. now rewrite app_context_assoc. + Defined. - Lemma infer_sort_impl {P Q} {Σ Σ' : global_env_ext} {Γ Γ' : context} {t t' : term} : - forall f, forall tu: infer_sort P Σ Γ t, - let s := tu.π1 in - (P Σ Γ t s -> Q Σ' Γ' t' (f s)) -> - infer_sort Q Σ' Γ' t'. - Proof. - intros ? (? & Hs) s HPQ. eexists. now apply HPQ. - Qed. + Lemma All_local_env_prod_inv : + forall P Q Γ, + All_local_env (fun Δ j => P Δ j × Q Δ j) Γ -> + All_local_env P Γ × All_local_env Q Γ. + Proof using Type. + intros P Q Γ h. + split; apply All_local_env_impl with (1 := h). + all: now intros ??[]. + Qed. - Lemma infer_typing_sort_impl {P Q} {Σ Σ' : global_env_ext} {Γ Γ' : context} {t t' : term} : - forall f, forall tu: infer_sort (typing_sort P) Σ Γ t, - let s := tu.π1 in - (P Σ Γ t (tSort s) -> Q Σ' Γ' t' (tSort (f s))) -> - infer_sort (typing_sort Q) Σ' Γ' t'. - Proof. - apply (infer_sort_impl (P := typing_sort P) (Q := typing_sort Q)). - Qed. + Lemma All_local_env_lift_prod_inv : + forall P Q Δ, + All_local_env (lift_typing1 (Prop_local_conj P Q)) Δ -> + All_local_env (lift_typing1 P) Δ × All_local_env (lift_typing1 Q) Δ. + Proof using Type. + intros P Q Δ h. + split; apply All_local_env_impl with (1 := h); intros ?? H; apply lift_typing_impl with (1 := H). + all: move => ??[] //. + Qed. - Lemma lift_typing_impl {P Q Σ Σ' Γ Γ' t t' T} : - lift_typing P Σ Γ t T -> - (forall T, P Σ Γ t T -> Q Σ' Γ' t' T) -> - lift_typing Q Σ' Γ' t' T. - Proof. - intros HT HPQ. - apply lift_judgment_impl with (1 := HT); tas. - intro Hs; apply infer_typing_sort_impl with id Hs; apply HPQ. - Qed. + End All_local_env_rel. Section TypeLocalOver. - Context (checking : global_env_ext -> context -> term -> term -> Type). - Context (sorting : global_env_ext -> context -> term -> Type). - Context (cproperty : forall (Σ : global_env_ext) (Γ : context), - All_local_env (lift_judgment checking sorting Σ) Γ -> - forall (t T : term), checking Σ Γ t T -> Type). - Context (sproperty : forall (Σ : global_env_ext) (Γ : context), - All_local_env (lift_judgment checking sorting Σ) Γ -> - forall (t : term), sorting Σ Γ t -> Type). - - Inductive All_local_env_over_gen (Σ : global_env_ext) : - forall (Γ : context), All_local_env (lift_judgment checking sorting Σ) Γ -> Type := + Context (checking : context -> term -> term -> Type). + Context (sorting : context -> term -> sort -> Type). + Context (cproperty : forall (Γ : context), + All_local_env (lift_sorting1 checking sorting) Γ -> + forall (t T : term), checking Γ t T -> Type). + Context (sproperty : forall (Γ : context), + All_local_env (lift_sorting1 checking sorting) Γ -> + forall (t : term) (u : sort), sorting Γ t u -> Type). + + Inductive All_local_env_over_sorting : + forall (Γ : context), All_local_env (lift_sorting1 checking sorting) Γ -> Type := | localenv_over_nil : - All_local_env_over_gen Σ [] localenv_nil + All_local_env_over_sorting [] localenv_nil | localenv_over_cons_abs Γ na t - (all : All_local_env (lift_judgment checking sorting Σ) Γ) : - All_local_env_over_gen Σ Γ all -> - forall (tu : lift_judgment checking sorting Σ Γ t Sort) - (Hs: sproperty Σ Γ all _ tu), - All_local_env_over_gen Σ (Γ ,, vass na t) + (all : All_local_env (lift_sorting1 checking sorting) Γ) : + All_local_env_over_sorting Γ all -> + forall (tu : lift_sorting1 checking sorting Γ (j_vass na t)) + (Hs: sproperty Γ all _ _ tu.2.π2.1), + All_local_env_over_sorting (Γ ,, vass na t) (localenv_cons_abs all tu) | localenv_over_cons_def Γ na b t - (all : All_local_env (lift_judgment checking sorting Σ) Γ) (tb : checking Σ Γ b t) : - All_local_env_over_gen Σ Γ all -> - forall (Hc: cproperty Σ Γ all _ _ tb), - forall (tu : lift_judgment checking sorting Σ Γ t Sort) - (Hs: sproperty Σ Γ all _ tu), - All_local_env_over_gen Σ (Γ ,, vdef na b t) - (localenv_cons_def all tu tb). + (all : All_local_env (lift_sorting1 checking sorting) Γ) : + All_local_env_over_sorting Γ all -> + forall (tu : lift_sorting1 checking sorting Γ (j_vdef na b t)) + (Hc: cproperty Γ all _ _ tu.1) + (Hs: sproperty Γ all _ _ tu.2.π2.1), + All_local_env_over_sorting (Γ ,, vdef na b t) + (localenv_cons_def all tu). End TypeLocalOver. - Derive Signature for All_local_env_over_gen. + Derive Signature for All_local_env_over_sorting. Definition All_local_env_over typing property := - (All_local_env_over_gen typing (infer_sort (typing_sort typing)) property (fun Σ Γ H t tu => property _ _ H _ _ tu.π2)). + (All_local_env_over_sorting typing (typing_sort1 typing) property (fun Γ H t u tu => property _ H _ _ tu)). + + Lemma All_local_env_over_sorting_2 c s Pc Ps Γ (H : All_local_env (lift_sorting1 c s) Γ) : + All_local_env_over_sorting _ _ (fun Γ _ t T _ => Pc Γ t T) (fun Γ _ t s _ => Ps Γ t s) _ H -> + All_local_env (lift_sorting1 (Prop_local_conj c Pc) (Prop_local_conj s Ps)) Γ. + Proof. + induction 1; constructor; eauto. + - destruct tu as (Htm & u & Hty & e). + repeat (eexists; tea). + - destruct tu as (Htm & u & Hty & e). + repeat (eexists; tea). + Defined. - Definition All_local_env_over_sorting checking sorting cproperty (sproperty : forall Σ Γ _ t s, sorting Σ Γ t s -> Type) := - (All_local_env_over_gen checking (infer_sort sorting) cproperty (fun Σ Γ H t tu => sproperty Σ Γ H t tu.π1 tu.π2)). + Definition on_wf_local_decl {typing Γ} + (P : forall Γ (wfΓ : All_local_env (lift_typing1 typing) Γ) t T, typing Γ t T -> Type) + (wfΓ : All_local_env (lift_typing1 typing) Γ) {d} + (H : on_local_decl (lift_typing1 typing) Γ d) := + match d return (on_local_decl (lift_typing1 typing) Γ d) -> Type with + | {| decl_name := na; decl_body := Some b; decl_type := ty |} + => fun H => P Γ wfΓ b ty H.1 × P Γ wfΓ ty _ H.2.π2.1 + | {| decl_name := na; decl_body := None; decl_type := ty |} + => fun H => P Γ wfΓ ty _ H.2.π2.1 + end H. + + + Lemma nth_error_All_local_env_over {typing} {P Γ n decl} (eq : nth_error Γ n = Some decl) {wfΓ : All_local_env (lift_typing1 typing) Γ} : + All_local_env_over typing P Γ wfΓ -> + let Γ' := skipn (S n) Γ in + let wfΓ' := All_local_env_skipn _ wfΓ in + let p := All_local_env_nth_error wfΓ eq in + (All_local_env_over typing P Γ' wfΓ' * on_wf_local_decl P wfΓ' p)%type. + Proof. + induction 1 in n, decl, eq |- *. + - exfalso. destruct n => //. + - destruct n; simpl. + + simpl in *. split; tas. clear -Hs. + destruct f_equal; cbn. + assumption. + + apply IHX. + - destruct n; simpl. + + simpl in *. split; tas. clear -Hc Hs. + destruct f_equal; cbn. + split; assumption. + + apply IHX. + Defined. + + Lemma All_local_env_over_2 typing P Γ (H : All_local_env (lift_typing1 typing) Γ) : + All_local_env_over _ (fun Γ _ t T _ => P Γ t T) _ H -> + All_local_env (lift_typing_conj typing P) Γ. + Proof. + apply All_local_env_over_sorting_2 with (Ps := fun Γ t u => P Γ t (tSort u)). + Defined. Section TypeCtxInst. - Context (typing : forall (Σ : global_env_ext) (Γ : context), term -> term -> Type). + Context (typing : forall (Γ : context), term -> term -> Type). (* Γ |- s : Δ, where Δ is a telescope (reverse context) *) - Inductive ctx_inst Σ (Γ : context) : list term -> context -> Type := - | ctx_inst_nil : ctx_inst Σ Γ [] [] + Inductive ctx_inst (Γ : context) : list term -> context -> Type := + | ctx_inst_nil : ctx_inst Γ [] [] | ctx_inst_ass na t i inst Δ : - typing Σ Γ i t -> - ctx_inst Σ Γ inst (subst_telescope [i] 0 Δ) -> - ctx_inst Σ Γ (i :: inst) (vass na t :: Δ) + typing Γ i t -> + ctx_inst Γ inst (subst_telescope [i] 0 Δ) -> + ctx_inst Γ (i :: inst) (vass na t :: Δ) | ctx_inst_def na b t inst Δ : - ctx_inst Σ Γ inst (subst_telescope [b] 0 Δ) -> - ctx_inst Σ Γ inst (vdef na b t :: Δ). + ctx_inst Γ inst (subst_telescope [b] 0 Δ) -> + ctx_inst Γ inst (vdef na b t :: Δ). Derive Signature NoConfusion for ctx_inst. End TypeCtxInst. - Lemma ctx_inst_impl_gen Σ Γ inst Δ args P : - { P' & ctx_inst P' Σ Γ inst Δ } -> + Lemma ctx_inst_impl_gen Γ inst Δ args P : + { P' & ctx_inst P' Γ inst Δ } -> (forall t T, - All (fun P' => P' Σ Γ t T) args -> - P Σ Γ t T) -> - All (fun P' => ctx_inst P' Σ Γ inst Δ) args -> - ctx_inst P Σ Γ inst Δ. + All (fun P' => P' Γ t T) args -> + P Γ t T) -> + All (fun P' => ctx_inst P' Γ inst Δ) args -> + ctx_inst P Γ inst Δ. Proof. intros [? Hexists] HPQ H. induction Hexists; constructor; tea. @@ -592,37 +961,113 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). all: eapply All_impl; tea; cbn; intros *; inversion 1; subst; eauto. Qed. - Lemma ctx_inst_impl P Q Σ Γ inst Δ : - ctx_inst P Σ Γ inst Δ -> - (forall t T, P Σ Γ t T -> Q Σ Γ t T) -> - ctx_inst Q Σ Γ inst Δ. + Lemma ctx_inst_impl P Q Γ inst Δ : + ctx_inst P Γ inst Δ -> + (forall t T, P Γ t T -> Q Γ t T) -> + ctx_inst Q Γ inst Δ. Proof. intros H HPQ. induction H; econstructor; auto. Qed. + Definition option_default_size {A f} (fsize : forall (a : A), f a -> size) o (w : option_default f o (unit : Type)) : size := + match o as tm return option_default _ tm (unit : Type) -> size with + | Some tm => fun w => fsize _ w + | None => fun w => 0 + end w. + + Section lift_sorting_size_gen. + Context {checking : term -> term -> Type}. + Context {sorting : term -> sort -> Type}. + Context (csize : forall (t T : term), checking t T -> size). + Context (ssize : forall (t : term) (u : sort), sorting t u -> size). + + Definition lift_sorting_size_gen base j (w : lift_sorting checking sorting j) : size := + base + option_default_size (fun tm => csize tm _) (j_term j) w.1 + ssize _ _ w.2.π2.1. + + + Lemma lift_sorting_size_gen_impl {Qc Qs j} : + forall tu: lift_sorting checking sorting j, + (forall t T, forall Hty: checking t T, csize _ _ Hty <= lift_sorting_size_gen 0 _ tu -> Qc t T) -> + (forall t u, forall Hty: sorting t u, ssize _ _ Hty <= lift_sorting_size_gen 0 _ tu -> Qs t u) -> + lift_sorting Qc Qs j. + Proof. + intros (Htm & s & Hty & es) HPQc HPQs. + unfold lift_sorting_size_gen in *; cbn in *. + repeat (eexists; tea). + - destruct (j_term j) => //=. + eapply HPQc with (Hty := Htm); cbn. + lia. + - eapply HPQs with (Hty := Hty); cbn. + lia. + Qed. + + End lift_sorting_size_gen. + + Definition on_def_type_size_gen {c s} (ssize : forall Γ t u, s Γ t u -> size) base + Γ d (w : on_def_type (lift_sorting1 c s) Γ d) : size := + base + ssize _ _ _ w.2.π2.1. + Definition on_def_body_size_gen {c s} (csize : forall Γ t u, c Γ t u -> size) (ssize : forall Γ t u, s Γ t u -> size) base + types Γ d (w : on_def_body (lift_sorting1 c s) types Γ d) : size := + base + csize _ _ _ w.1 + ssize _ _ _ w.2.π2.1. + + Notation lift_sorting_size csize ssize := (lift_sorting_size_gen csize ssize 1). + Notation typing_sort_size typing_size := (fun t s (tu: typing_sort _ t s) => typing_size t (tSort s) tu). + Notation lift_typing_size typing_size := (lift_sorting_size_gen typing_size (typing_sort_size typing_size) 0). + Notation typing_sort_size1 typing_size := (fun Γ t s (tu: typing_sort1 _ Γ t s) => typing_size Γ t (tSort s) tu). + Notation on_def_type_sorting_size ssize := (on_def_type_size_gen ssize 1). + Notation on_def_type_size typing_size := (on_def_type_size_gen (typing_sort_size1 typing_size) 0). + Notation on_def_body_sorting_size csize ssize := (on_def_body_size_gen csize ssize 1). + Notation on_def_body_size typing_size := (on_def_body_size_gen typing_size (typing_sort_size1 typing_size) 0). + (* Will probably not pass the guard checker if in a list, must be unrolled like in on_def_* *) + + Lemma lift_sorting_size_impl {checking sorting Qc Qs j} csize ssize : + forall tu: lift_sorting checking sorting j, + (forall t T, forall Hty: checking t T, csize _ _ Hty < lift_sorting_size csize ssize _ tu -> Qc t T) -> + (forall t u, forall Hty: sorting t u, ssize _ _ Hty < lift_sorting_size csize ssize _ tu -> Qs t u) -> + lift_sorting Qc Qs j. + Proof. + intros tu Xc Xs. + eapply lift_sorting_size_gen_impl with (tu := tu). + all: intros. + 1: eapply Xc. 2: eapply Xs. + all: apply le_n_S, H. + Qed. + + Lemma lift_typing_size_impl {P Q j} Psize : + forall tu: lift_typing0 P j, + (forall t T, forall Hty: P t T, Psize _ _ Hty <= lift_typing_size Psize _ tu -> Q t T) -> + lift_typing0 Q j. + Proof. + intros. + eapply lift_sorting_size_gen_impl with (csize := Psize). + all: intros t T; apply X. + Qed. + Section All_local_env_size. - Context {P : forall (Σ : global_env_ext) (Γ : context), term -> typ_or_sort -> Type}. - Context (Psize : forall (Σ : global_env_ext) Γ t T, P Σ Γ t T -> size). + Context {checking : forall (Γ : context), term -> term -> Type}. + Context {sorting : forall (Γ : context), term -> sort -> Type}. + Context (csize : forall Γ t T, checking Γ t T -> size). + Context (ssize : forall Γ t u, sorting Γ t u -> size). - Fixpoint All_local_env_size_gen base Σ Γ (w : All_local_env (P Σ) Γ) : size := + Fixpoint All_local_env_size_gen base Γ (w : All_local_env (lift_sorting1 checking sorting) Γ) : size := match w with | localenv_nil => base - | localenv_cons_abs Γ' na t w' p => Psize _ _ _ _ p + All_local_env_size_gen base _ _ w' - | localenv_cons_def Γ' na b t w' pt pb => Psize _ _ _ _ pt + Psize _ _ _ _ pb + All_local_env_size_gen base _ _ w' + | localenv_cons_abs Γ' na t w' p => ssize _ _ _ p.2.π2.1 + All_local_env_size_gen base _ w' + | localenv_cons_def Γ' na b t w' p => csize _ _ _ p.1 + ssize _ _ _ p.2.π2.1 + All_local_env_size_gen base _ w' end. - Lemma All_local_env_size_pos base Σ Γ w : base <= All_local_env_size_gen base Σ Γ w. + Lemma All_local_env_size_pos base Γ w : base <= All_local_env_size_gen base Γ w. Proof using Type. induction w. all: simpl ; lia. Qed. End All_local_env_size. - Notation All_local_rel_size_gen Psize base := (fun Σ Γ Δ (w : All_local_rel _ Γ Δ) => - All_local_env_size_gen (fun Σ Δ => Psize Σ (Γ ,,, Δ)) base Σ Δ w). + Notation All_local_rel_size_gen c s csize ssize base := (fun Γ Δ (w : All_local_rel (lift_sorting1 c s) Γ Δ) => + All_local_env_size_gen (fun Δ => csize (Γ ,,, Δ)) (fun Δ => ssize (Γ ,,, Δ)) base Δ w). - Lemma All_local_env_size_app P Psize base Σ Γ Γ' (wfΓ : All_local_env (P Σ) Γ) (wfΓ' : All_local_rel (P Σ) Γ Γ') : - All_local_env_size_gen Psize base _ _ (All_local_rel_app wfΓ wfΓ') + base = All_local_env_size_gen Psize base _ _ wfΓ + All_local_rel_size_gen Psize base _ _ _ wfΓ'. + Lemma All_local_env_size_app c s csize ssize base Γ Γ' (wfΓ : All_local_env (lift_sorting1 c s) Γ) (wfΓ' : All_local_rel (lift_sorting1 c s) Γ Γ') : + All_local_env_size_gen csize ssize base _ (All_local_env_app wfΓ wfΓ') + base = All_local_env_size_gen csize ssize base _ wfΓ + All_local_rel_size_gen c s csize ssize base _ _ wfΓ'. Proof. induction Γ'. - dependent inversion wfΓ'. @@ -639,41 +1084,23 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). reflexivity. Qed. - Section lift_judgment_size. - Context {checking : global_env_ext -> context -> term -> term -> Type}. - Context {sorting : global_env_ext -> context -> term -> Type}. - Context (csize : forall (Σ : global_env_ext) (Γ : context) (t T : term), checking Σ Γ t T -> size). - Context (ssize : forall (Σ : global_env_ext) (Γ : context) (t : term), sorting Σ Γ t -> size). - - Definition lift_judgment_size Σ Γ t T (w : lift_judgment checking sorting Σ Γ t T) : size := - match T return lift_judgment checking sorting Σ Γ t T -> size with - | Typ T => csize _ _ _ _ - | Sort => ssize _ _ _ - end w. - End lift_judgment_size. - Implicit Types (Σ : global_env_ext) (Γ : context) (t : term). - Notation infer_sort_size typing_size := (fun Σ Γ t (tu: infer_sort _ Σ Γ t) => let 'existT s d := tu in typing_size Σ Γ t s d). - Notation typing_sort_size typing_size := (fun Σ Γ t s (tu: typing_sort _ Σ Γ t s) => typing_size Σ Γ t (tSort s) tu). - Section Regular. - Context {typing : global_env_ext -> context -> term -> term -> Type}. - Context (typing_size : forall Σ Γ t T, typing Σ Γ t T -> size). + Context {typing : context -> term -> term -> Type}. + Context (typing_size : forall Γ t T, typing Γ t T -> size). - Definition lift_typing_size := lift_judgment_size typing_size (infer_sort_size (typing_sort_size typing_size)). - Definition All_local_env_size := All_local_env_size_gen lift_typing_size 0. - Definition All_local_rel_size := All_local_rel_size_gen lift_typing_size 0. + Definition All_local_env_size := All_local_env_size_gen typing_size (typing_sort1 typing_size) 0. + Definition All_local_rel_size Γ Γ' (wfΓ : All_local_rel (lift_typing1 typing) Γ Γ') := All_local_rel_size_gen typing (typing_sort1 typing) typing_size (typing_sort_size1 typing_size) 0 Γ Γ' wfΓ. End Regular. Section Bidirectional. - Context {checking : global_env_ext -> context -> term -> term -> Type} {sorting : global_env_ext -> context -> term -> Universe.t -> Type}. - Context (checking_size : forall Σ Γ t T, checking Σ Γ t T -> size). - Context (sorting_size : forall Σ Γ t s, sorting Σ Γ t s -> size). + Context {checking : context -> term -> term -> Type} {sorting : context -> term -> sort -> Type}. + Context (checking_size : forall Γ t T, checking Γ t T -> size). + Context (sorting_size : forall Γ t s, sorting Γ t s -> size). - Definition lift_sorting_size := lift_judgment_size checking_size (infer_sort_size sorting_size). - Definition All_local_env_sorting_size := All_local_env_size_gen lift_sorting_size 1. - Definition All_local_rel_sorting_size := All_local_rel_size_gen lift_sorting_size 1. + Definition All_local_env_sorting_size := All_local_env_size_gen checking_size sorting_size 1. + Definition All_local_rel_sorting_size := All_local_rel_size_gen _ _ checking_size sorting_size 1. End Bidirectional. End EnvTyping. @@ -744,33 +1171,35 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT Context {cf: checker_flags}. Context (Pcmp: global_env_ext -> context -> conv_pb -> term -> term -> Type). - Context (P : global_env_ext -> context -> term -> typ_or_sort -> Type). + Context (P : global_env_ext -> context -> judgment -> Type). Definition on_context Σ ctx := All_local_env (P Σ) ctx. (** 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]. We also force well-typing of the let-ins in any universe to imply wf_local. *) - Fixpoint type_local_ctx Σ (Γ Δ : context) (u : Universe.t) : Type := + Fixpoint type_local_ctx Σ (Γ Δ : context) (u : sort) : Type := match Δ with - | [] => wf_universe Σ u - | {| decl_body := None; decl_type := t |} :: Δ => type_local_ctx Σ Γ Δ u × P Σ (Γ ,,, Δ) t (Typ (tSort u)) - | {| decl_body := Some b; decl_type := t |} :: Δ => type_local_ctx Σ Γ Δ u × P Σ (Γ ,,, Δ) t Sort × P Σ (Γ ,,, Δ) b (Typ t) + | [] => wf_sort Σ u + | {| decl_name := na; decl_body := None; decl_type := t |} :: Δ => + type_local_ctx Σ Γ Δ u × P Σ (Γ ,,, Δ) (TypUniv t u (* na.(binder_relevance) *)) + | {| decl_body := Some _; |} as d :: Δ => + type_local_ctx Σ Γ Δ u × P Σ (Γ ,,, Δ) (j_decl d) end. - Fixpoint sorts_local_ctx Σ (Γ Δ : context) (us : list Universe.t) : Type := + Fixpoint sorts_local_ctx Σ (Γ Δ : context) (us : list sort) : Type := match Δ, us with | [], [] => unit - | {| decl_body := None; decl_type := t |} :: Δ, u :: us => - sorts_local_ctx Σ Γ Δ us × P Σ (Γ ,,, Δ) t (Typ (tSort u)) - | {| decl_body := Some b; decl_type := t |} :: Δ, us => - sorts_local_ctx Σ Γ Δ us × P Σ (Γ ,,, Δ) t Sort × P Σ (Γ ,,, Δ) b (Typ t) + | {| decl_name := na; decl_body := None; decl_type := t |} :: Δ, u :: us => + sorts_local_ctx Σ Γ Δ us × P Σ (Γ ,,, Δ) (TypUniv t u (* na.(binder_relevance) *)) + | {| decl_body := Some _ |} as d :: Δ, us => + sorts_local_ctx Σ Γ Δ us × P Σ (Γ ,,, Δ) (j_decl d) | _, _ => False end. Implicit Types (mdecl : mutual_inductive_body) (idecl : one_inductive_body) (cdecl : constructor_body). - Definition on_type Σ Γ T := P Σ Γ T Sort. + Definition on_type Σ Γ T := P Σ Γ (Typ T). Open Scope type_scope. @@ -1017,7 +1446,7 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT sorts_local_ctx Σ (arities_context mdecl.(ind_bodies) ,,, mdecl.(ind_params)) cdecl.(cstr_args) cunivs; on_cindices : - ctx_inst (fun Σ Γ t T => P Σ Γ t (Typ T)) Σ (arities_context mdecl.(ind_bodies) ,,, mdecl.(ind_params) ,,, cdecl.(cstr_args)) + ctx_inst (fun Γ t T => P Σ Γ (TermTyp t T)) (arities_context mdecl.(ind_bodies) ,,, mdecl.(ind_params) ,,, cdecl.(cstr_args)) cdecl.(cstr_indices) (List.rev (lift_context #|cdecl.(cstr_args)| 0 ind_indices)); @@ -1093,20 +1522,20 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT Definition check_constructors_smaller φ cunivss ind_sort := Forall (fun cunivs => - Forall (fun argsort => leq_universe φ argsort ind_sort) cunivs) cunivss. + Forall (fun argsort => leq_sort φ argsort ind_sort) cunivs) cunivss. (** This ensures that all sorts in kelim are lower or equal to the top elimination sort, if set. For inductives in Type we do not check [kelim] currently. *) - Definition constructor_univs := list Universe.t. + Definition constructor_univs := list sort. (* The sorts of the arguments context (without lets) *) Definition elim_sort_prop_ind (ind_ctors_sort : list constructor_univs) := match ind_ctors_sort with | [] => (* Empty inductive proposition: *) IntoAny | [ s ] => - if forallb Universes.is_propositional s then + if forallb Sort.is_propositional s then IntoAny (* Singleton elimination *) else IntoPropSProp (* Squashed: some arguments are higher than Prop, restrict to Prop *) @@ -1121,23 +1550,25 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT Definition check_ind_sorts (Σ : global_env_ext) params kelim ind_indices cdecls ind_sort : Type := - if Universe.is_prop ind_sort then + match Sort.to_family ind_sort with + | Sort.fProp => (** The inductive is declared in the impredicative sort Prop *) (** No universe-checking to do: any size of constructor argument is allowed, however elimination restrictions apply. *) (allowed_eliminations_subset kelim (elim_sort_prop_ind cdecls) : Type) - else if Universe.is_sprop ind_sort then + | Sort.fSProp => (** The inductive is declared in the impredicative sort SProp *) (** No universe-checking to do: any size of constructor argument is allowed, however elimination restrictions apply. *) (allowed_eliminations_subset kelim (elim_sort_sprop_ind cdecls) : Type) - else + | _ => (** The inductive is predicative: check that all constructors arguments are smaller than the declared universe. *) check_constructors_smaller Σ cdecls ind_sort × if indices_matter then type_local_ctx Σ params ind_indices ind_sort - else True. + else True + end. Record on_ind_body Σ mind mdecl i idecl := { (** The type of the inductive must be an arity, sharing the same params @@ -1212,10 +1643,7 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT (** *** Typing of constant declarations *) Definition on_constant_decl Σ d := - match d.(cst_body) with - | Some trm => P Σ [] trm (Typ d.(cst_type)) - | None => on_type Σ [] d.(cst_type) - end. + P Σ [] (TermoptTyp d.(cst_body) d.(cst_type)). Definition on_global_decl Σ kn decl := match decl with @@ -1288,6 +1716,23 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT Definition on_global_env_ext (Σ : global_env_ext) := on_global_env Σ.1 × on_udecl Σ.(universes) Σ.2. + Lemma on_global_env_ext_empty_ext g : + on_global_env g -> on_global_env_ext (empty_ext g). + Proof. + intro H; split => //. + unfold empty_ext, snd. repeat split. + - unfold levels_of_udecl. intros x e. lsets. + - unfold constraints_of_udecl. intros x e. csets. + - unfold satisfiable_udecl, univs_ext_constraints, constraints_of_udecl, fst_ctx, fst => //. + destruct H as ((cstrs & _ & consistent) & decls). + destruct consistent; eexists. + intros v e. specialize (H v e); tea. + - unfold valid_on_mono_udecl, constraints_of_udecl, consistent_extension_on. + intros v sat; exists v; split. + + intros x e. csets. + + intros x e => //. + Qed. + End GlobalMaps. Arguments cstr_args_length {_ Pcmp P Σ mdecl i idecl ind_indices cdecl cunivs}. @@ -1316,9 +1761,9 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT repeat match goal with H : _ × _ |- _ => destruct H end. Lemma type_local_ctx_impl_gen Σ Γ Δ u args P : { P' & type_local_ctx P' Σ Γ Δ u } -> - (forall Γ t T, - All (fun P' => P' Σ Γ t T) args -> - P Σ Γ t T) -> + (forall Γ j, + All (fun P' => P' Σ Γ j) args -> + P Σ Γ j) -> All (fun P' => type_local_ctx P' Σ Γ Δ u) args -> type_local_ctx P Σ Γ Δ u. Proof. @@ -1337,21 +1782,22 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT all: repeat match goal with H : _ × _ |- _ => destruct H end => //=. } Qed. - Lemma type_local_ctx_impl (P Q : global_env_ext -> context -> term -> typ_or_sort -> Type) Σ Γ Δ u : + Lemma type_local_ctx_impl (P Q : global_env_ext -> context -> judgment -> Type) Σ Σ' Γ Δ u : type_local_ctx P Σ Γ Δ u -> - (forall Γ t T, P Σ Γ t T -> Q Σ Γ t T) -> - type_local_ctx Q Σ Γ Δ u. + (forall u, wf_sort Σ u -> wf_sort Σ' u) -> + (forall Γ j, P Σ Γ j -> Q Σ' Γ j) -> + type_local_ctx Q Σ' Γ Δ u. Proof. intros HP HPQ. revert HP; induction Δ in Γ, HPQ |- *; simpl; auto. - destruct a as [na [b|] ty]; simpl; auto. - intros. intuition auto. intuition auto. + destruct a as [na [b|] ty]. + all: intros (? & ?); now split. Qed. Lemma sorts_local_ctx_impl_gen Σ Γ Δ u args P : { P' & sorts_local_ctx P' Σ Γ Δ u } -> - (forall Γ t T, - All (fun P' => P' Σ Γ t T) args -> - P Σ Γ t T) -> + (forall Γ j, + All (fun P' => P' Σ Γ j) args -> + P Σ Γ j) -> All (fun P' => sorts_local_ctx P' Σ Γ Δ u) args -> sorts_local_ctx P Σ Γ Δ u. Proof. @@ -1370,15 +1816,15 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT all: repeat match goal with H : _ × _ |- _ => destruct H end => //=. } Qed. - Lemma sorts_local_ctx_impl (P Q : global_env_ext -> context -> term -> typ_or_sort -> Type) Σ Γ Δ u : + Lemma sorts_local_ctx_impl (P Q : global_env_ext -> context -> judgment -> Type) Σ Σ' Γ Δ u : sorts_local_ctx P Σ Γ Δ u -> - (forall Γ t T, P Σ Γ t T -> Q Σ Γ t T) -> - sorts_local_ctx Q Σ Γ Δ u. + (forall Γ j, P Σ Γ j -> Q Σ' Γ j) -> + sorts_local_ctx Q Σ' Γ Δ u. Proof. intros HP HPQ. revert HP; induction Δ in Γ, HPQ, u |- *; simpl; auto. destruct a as [na [b|] ty]; simpl; auto. - intros. intuition auto. intuition auto. - destruct u; auto. intuition eauto. + 2: destruct u => //. + all: intros (? & ?); now split. Qed. Lemma cstr_respects_variance_impl_gen Σ mdecl v cs args Pcmp : @@ -1415,16 +1861,16 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT all: inversion 1; subst => //=. } Qed. - Lemma cstr_respects_variance_impl Σ mdecl v cs Pcmp Pcmp' : + Lemma cstr_respects_variance_impl Σ Σ' mdecl v cs Pcmp Pcmp' : (match variance_universes (ind_universes mdecl) v with | Some (univs, u, u') => forall Γ t T pb, - Pcmp' (Σ, univs) Γ pb t T -> - Pcmp (Σ, univs) Γ pb t T + Pcmp (Σ, univs) Γ pb t T -> + Pcmp' (Σ', univs) Γ pb t T | None => True end) -> - @cstr_respects_variance Pcmp' Σ mdecl v cs -> - @cstr_respects_variance Pcmp Σ mdecl v cs. + @cstr_respects_variance Pcmp Σ mdecl v cs -> + @cstr_respects_variance Pcmp' Σ' mdecl v cs. Proof. rewrite /cstr_respects_variance/cumul_ctx_rel/cumul_pb_decls. destruct variance_universes; destr_prod; try now case. @@ -1435,9 +1881,9 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT Lemma on_constructor_impl_config_gen Σ mdecl i idecl ind_indices cdecl cunivs args cf Pcmp P : { '(cf', Pcmp', P') & config.impl cf' cf × @on_constructor cf' Pcmp' P' Σ mdecl i idecl ind_indices cdecl cunivs } -> - (forall Γ t T, - All (fun '(cf', Pcmp', P') => P' Σ Γ t T) args -> - P Σ Γ t T) -> + (forall Γ j, + All (fun '(cf', Pcmp', P') => P' Σ Γ j) args -> + P Σ Γ j) -> (forall u Γ t T pb, All (fun '(cf', Pcmp', P') => Pcmp' (Σ.1, u) Γ pb t T) args -> Pcmp (Σ.1, u) Γ pb t T) -> @@ -1452,11 +1898,11 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT eapply All_impl; tea; intros *; destr_prod; destruct 1; tea. } { eapply sorts_local_ctx_impl_gen; tea. { eexists; tea. } - { intros; eapply H1, All_eta3; cbn; apply All_map_inv with (P:=fun P => P _ _ _ _) (f:=snd); tea. } + { intros; eapply H1, All_eta3; cbn; apply All_map_inv with (P:=fun P => P _ _ _) (f:=snd); tea. } { eapply All_map, All_impl; tea; intros *; destr_prod; destruct 1; cbn; tea. } } { eapply ctx_inst_impl_gen; tea. { eexists; tea. } - { intros; eapply H1, All_eta3; cbn; apply All_map_inv with (P:=fun P => P _ _ _ _) (f:=fun P Σ Γ t T => snd P Σ Γ t (Typ T)); tea. } + { intros; eapply H1, All_eta3; cbn. apply All_map_inv with (P:=fun P => P _ t T1) (f:=fun P Γ t T => snd P Σ Γ (TermTyp t T)); tea. } { eapply All_map, All_impl; tea; intros *; destr_prod; destruct 1; cbn; tea. } } { move => ? H'. match goal with H : _ |- _ => specialize (H _ H'); revert H end => H''. @@ -1475,9 +1921,9 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT Lemma on_constructors_impl_config_gen Σ mdecl i idecl ind_indices cdecl cunivs args cf Pcmp P : { '(cf', Pcmp', P') & config.impl cf' cf × @on_constructors cf' Pcmp' P' Σ mdecl i idecl ind_indices cdecl cunivs } -> All (fun '(cf', Pcmp', P') => config.impl cf' cf) args -> - (forall Γ t T, - All (fun '(cf', Pcmp', P') => P' Σ Γ t T) args -> - P Σ Γ t T) -> + (forall Γ j, + All (fun '(cf', Pcmp', P') => P' Σ Γ j) args -> + P Σ Γ j) -> (forall u Γ t T pb, All (fun '(cf', Pcmp', P') => Pcmp' (Σ.1, u) Γ pb t T) args -> Pcmp (Σ.1, u) Γ pb t T) -> @@ -1503,16 +1949,16 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT now eapply All_impl; [ multimatch goal with H : _ |- _ => exact H end | ]; intros; destr_prod; eauto. Qed. - Lemma ind_respects_variance_impl Σ mdecl v cs Pcmp' Pcmp : + Lemma ind_respects_variance_impl Σ Σ' mdecl v cs Pcmp Pcmp' : match variance_universes (ind_universes mdecl) v with | Some (univs, u, u') => forall Γ t T pb, - Pcmp' (Σ, univs) Γ pb t T -> - Pcmp (Σ, univs) Γ pb t T + Pcmp (Σ, univs) Γ pb t T -> + Pcmp' (Σ', univs) Γ pb t T | None => True end -> - @ind_respects_variance Pcmp' Σ mdecl v cs -> - @ind_respects_variance Pcmp Σ mdecl v cs. + @ind_respects_variance Pcmp Σ mdecl v cs -> + @ind_respects_variance Pcmp' Σ' mdecl v cs. Proof. rewrite /ind_respects_variance/cumul_ctx_rel/cumul_pb_decls. destruct variance_universes; destr_prod; try now case. @@ -1546,89 +1992,159 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT end ]. Qed. + Lemma check_constructors_smaller_impl Σ cdecls ind_sort cf' cf : + config.impl cf' cf -> + @check_constructors_smaller cf' Σ cdecls ind_sort -> + @check_constructors_smaller cf Σ cdecls ind_sort. + Proof. + intro Xcf. + unfold check_constructors_smaller. + intro H; apply Forall_impl with (1 := H). + intros l Hl; apply Forall_impl with (1 := Hl). + intro u. now apply leq_sort_config_impl. + Qed. + + Lemma on_global_decl_impl_full {cf1 cf2 : checker_flags} Pcmp1 Pcmp2 P1 P2 Σ Σ' kn d : + config.impl cf1 cf2 -> + (forall Γ j, P1 Σ Γ j -> P2 Σ' Γ j) -> + (forall u Γ pb t t', Pcmp1 (Σ.1, u) Γ pb t t' -> Pcmp2 (Σ'.1, u) Γ pb t t') -> + (forall u, wf_sort Σ u -> wf_sort Σ' u) -> + (forall l s, @check_constructors_smaller cf1 (global_ext_constraints Σ) l s -> + @check_constructors_smaller cf2 (global_ext_constraints Σ') l s) -> + (forall u l, @on_variance cf1 Σ.1 u l -> @on_variance cf2 Σ'.1 u l) -> + @on_global_decl cf1 Pcmp1 P1 Σ kn d -> @on_global_decl cf2 Pcmp2 P2 Σ' kn d. + Proof. + intros Xcf XP Xcmp Xu Xc Xv. + destruct d; simpl. + - apply XP => //. + - intros [onI onP onNP onV]. + constructor; auto. + 2: { eapply All_local_env_impl with (1 := onP). intros; apply XP => //. } + + eapply Alli_impl; tea. intros. + refine {| ind_arity_eq := X.(ind_arity_eq); + ind_cunivs := X.(ind_cunivs) |}. + + eapply XP => //. now apply onArity in X. + + pose proof X.(onConstructors) as X1. + apply All2_impl with (1 := X1). + intros ? ? [? ? ? ?]; unshelve econstructor; eauto. + * apply XP; eauto. + * eapply sorts_local_ctx_impl with (1 := on_cargs0). + intros; apply XP => //. + * eapply ctx_inst_impl with (1 := on_cindices0). + intros; apply XP => //. + * intros v e. + eapply cstr_respects_variance_impl. + 2: eauto. + destruct variance_universes as [[[]]|] => //=. + intros; now eapply Xcmp. + * move: on_lets_in_type0. + move: Xcf. + rewrite /config.impl. + do 3 destruct lets_in_constructor_types => //=. + all: rewrite ?andb_false_r //=. + + exact (onProjections X). + + pose proof (ind_sorts X) as X1. unfold check_ind_sorts in *. + destruct Sort.to_family; auto. + destruct X1 as [constr_smaller type_local_ctx]. + split. + * apply Xc, constr_smaller. + * unfold config.impl in Xcf. + revert type_local_ctx. + do 2 destruct indices_matter => //=. + 2: now rewrite ?andb_false_r //= in Xcf. + intro. eapply type_local_ctx_impl; eauto. + + generalize (X.(onIndices)). destruct ind_variance => //. + apply ind_respects_variance_impl. + destruct variance_universes as [[[]]|] => //=. + intros; now eapply Xcmp. + Qed. + + Lemma on_global_decl_impl_only_config {cf cf1 cf2 : checker_flags} Pcmp Pcmp' P Q Σ kn d : + config.impl cf1 cf2 -> + (forall Γ j, + @on_global_env cf Pcmp P Σ.1 -> + P Σ Γ j -> Q Σ Γ j) -> + (forall u Γ pb t t', + @on_global_env cf Pcmp P Σ.1 -> + Pcmp (Σ.1, u) Γ pb t t' -> Pcmp' (Σ.1, u) Γ pb t t') -> + @on_global_env cf Pcmp P Σ.1 -> + @on_global_decl cf1 Pcmp P Σ kn d -> @on_global_decl cf2 Pcmp' Q Σ kn d. + Proof. + destruct Σ as [Σ u]. cbn. + intros ??? H. + apply on_global_decl_impl_full => //. + - intros ??. now apply X. + - intros ?????. now apply X0. + - intros ??; now apply check_constructors_smaller_impl. + - intros ??; now apply on_variance_impl. + Qed. + + Lemma on_global_decl_impl_simple {cf : checker_flags} Pcmp P Q Σ kn d : + (forall Γ j, on_global_env Pcmp P Σ.1 -> P Σ Γ j -> Q Σ Γ j) -> + on_global_env Pcmp P Σ.1 -> + on_global_decl Pcmp P Σ kn d -> on_global_decl Pcmp Q Σ kn d. + Proof. + intro X. + apply on_global_decl_impl_only_config; tas. + 1: reflexivity. + easy. + Qed. + + Lemma on_global_env_impl_config {cf1 cf2 : checker_flags} Pcmp Pcmp' P Q : config.impl cf1 cf2 -> - (forall Σ Γ t T, + (forall Σ Γ j, @on_global_env cf1 Pcmp P Σ.1 -> @on_global_env cf2 Pcmp' Q Σ.1 -> - P Σ Γ t T -> Q Σ Γ t T) -> + P Σ Γ j -> Q Σ Γ j) -> (forall Σ Γ t T pb, @on_global_env cf1 Pcmp P Σ.1 -> @on_global_env cf2 Pcmp' Q Σ.1 -> Pcmp Σ Γ pb t T -> Pcmp' Σ Γ pb t T) -> forall Σ, @on_global_env cf1 Pcmp P Σ -> @on_global_env cf2 Pcmp' Q Σ. Proof. - intros Hcf X Y Σ [cu IH]. split; auto. + intros Xcf X Xcmp Σ [cu IH]. split; auto. revert cu IH; generalize (universes Σ) as univs, (retroknowledge Σ) as retro, (declarations Σ). clear Σ. induction g; intros; auto. constructor; auto. depelim IH. specialize (IHg cu IH). constructor; auto. - pose proof (globenv_decl _ _ _ _ _ _ _ IH o). destruct o. constructor; auto. - assert (X' := fun Γ t T => X ({| universes := univs; declarations := _ |}, udecl0) Γ t T - (cu, IH) (cu, IHg)); clear X. - rename X' into X. - assert (Y' := fun udecl0 Γ t T pb => Y ({| universes := univs; declarations := _ |}, udecl0) Γ t T pb - (cu, IH) (cu, IHg)); clear Y. - rename Y' into Y. - clear IH IHg. destruct d; simpl. - - destruct c; simpl. destruct cst_body0; cbn in *; now eapply X. - - destruct on_global_decl_d0 as [onI onP onNP]. - constructor; auto. - -- eapply Alli_impl; tea. intros. - refine {| ind_arity_eq := X1.(ind_arity_eq); - ind_cunivs := X1.(ind_cunivs) |}. - --- apply onArity in X1. unfold on_type in *; simpl in *. - now eapply X. - --- pose proof X1.(onConstructors) as X11. red in X11. - eapply All2_impl; eauto. - simpl. intros. destruct X2 as [? ? ? ?]; unshelve econstructor; eauto. - * apply X; eauto. - * clear -X0 X on_cargs0. revert on_cargs0. - generalize (cstr_args x0). - induction c in y |- *; destruct y; simpl; auto; - destruct a as [na [b|] ty]; simpl in *; auto; - split; intuition eauto. - * clear -X0 X on_cindices0. - revert on_cindices0. - generalize (List.rev (lift_context #|cstr_args x0| 0 (ind_indices x))). - generalize (cstr_indices x0). - induction 1; simpl; constructor; auto. - * intros; eapply cstr_respects_variance_impl; [ | now eauto ]. - repeat destruct ?; subst; eauto. - * move: Hcf; unfold config.impl. - do 2 destruct lets_in_constructor_types => //=; rewrite ?andb_false_r => //. - --- simpl; intros. pose (onProjections X1) as X2. simpl in *; auto. - --- destruct X1. simpl. unfold check_ind_sorts in *. - destruct Universe.is_prop; auto. - destruct Universe.is_sprop; auto. - split. - * unfold check_constructors_smaller in *. - eapply Forall_impl; [ apply ind_sorts0 | ]; cbv beta; intros *. - move => H'; eapply Forall_impl; [ exact H' | ]; cbv beta; intros *. - unfold leq_universe, leq_universe_n, leq_universe_n_, leq_levelalg_n, config.impl in *. - repeat match goal with |- context[match ?x with _ => _ end] => destruct x eqn:?; subst end => //=. - move: Hcf; do 2 destruct prop_sub_type => //=; rewrite ?andb_false_r => //=. - * move: Hcf; unfold config.impl. - destr_prod. - do 2 destruct indices_matter => //=; rewrite ?andb_false_r => //= Hcf; auto. - eapply type_local_ctx_impl; eauto. - --- generalize X1.(onIndices). - destruct ind_variance eqn:? => //. - eapply ind_respects_variance_impl; eauto. - repeat destruct ?; eauto. - -- red in onP. red. - eapply All_local_env_impl; tea. - -- move: onVariance0; eapply on_variance_impl; eauto. + eapply @on_global_decl_impl_only_config with (cf := cf1) (1 := Xcf) (5 := on_global_decl_d0). + { intros. eapply X; [tea| split; tea |tea]. } + { cbn. intros. eapply Xcmp. 1,3: eassumption. split; cbn. all: eassumption. } + split; eauto. Qed. Lemma on_global_env_impl {cf : checker_flags} Pcmp P Q : - (forall Σ Γ t T, + (forall Σ Γ j, on_global_env Pcmp P Σ.1 -> on_global_env Pcmp Q Σ.1 -> - P Σ Γ t T -> Q Σ Γ t T) -> + P Σ Γ j -> Q Σ Γ j) -> forall Σ, on_global_env Pcmp P Σ -> on_global_env Pcmp Q Σ. Proof. intros; eapply on_global_env_impl_config; eauto; reflexivity. Qed. + Lemma lookup_on_global_env `{checker_flags} {Pcmp P} {Σ : global_env} {c decl} : + on_global_env Pcmp P Σ -> + lookup_env Σ c = Some decl -> + ∑ Σ' : global_env_ext, on_global_env Pcmp P Σ' × strictly_extends_decls Σ' Σ × on_global_decl Pcmp P Σ' c decl. + Proof. + unfold on_global_env. + destruct Σ as [univs Σ retro]; cbn. intros [cu ond]. + induction ond; cbn in * => //. + destruct o. rename udecl0 into udecl. + case: eqb_specT => [-> [= <-]| ne]. + - exists ({| universes := univs; declarations := Σ; retroknowledge := retro |}, udecl); cbn. + split; try constructor; tas. + split => //=. now exists [(kn, d)]. + - intros hl. + specialize (IHond hl) as [[Σ' udecl'] [ong [[equ ext extretro] ond']]]. + exists (Σ', udecl'). cbn in equ |- *. subst univs. split; cbn; auto; try apply ong. + split; cbn; auto. split; cbn; auto. + cbn in ext. destruct ext as [Σ'' ->]. cbn in *. + now exists ((kn, d) :: Σ''). + Qed. + + End GlobalMaps. Module Type GlobalMapsSig (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvTypingSig T E TU) (C: ConversionSig T E TU ET) (L: LookupSig T E). @@ -1665,7 +2181,7 @@ Module DeclarationTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E) Import T E L TU ET CT GM CS Ty. Definition isType `{checker_flags} (Σ : global_env_ext) (Γ : context) (t : term) := - infer_sort (typing_sort typing) Σ Γ t. + on_type (lift_typing typing) Σ Γ t. (** This predicate enforces that there exists typing derivations for every typable term in env. *) @@ -1686,99 +2202,36 @@ Module DeclarationTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E) Lemma refine_type `{checker_flags} Σ Γ t T U : Σ ;;; Γ |- t : T -> T = U -> Σ ;;; Γ |- t : U. Proof. now intros Ht ->. Qed. - Definition wf_local_rel `{checker_flags} Σ := All_local_rel (lift_typing typing Σ). + Definition wf_local_rel `{checker_flags} Σ := All_local_rel (lift_typing1 (typing Σ)). (** Functoriality of global environment typing derivations + folding of the well-formed environment assumption. *) Lemma on_wf_global_env_impl_config {cf1 cf2 cf3 : checker_flags} {Σ : global_env} {wfΣ : @on_global_env cf1 (@cumul_gen cf1) (lift_typing (@typing cf1)) Σ} P Q : config.impl cf2 cf3 -> (forall Σ Γ pb t T, @cumul_gen cf2 Σ Γ pb t T -> @cumul_gen cf3 Σ Γ pb t T) -> - (forall Σ Γ t T, @on_global_env cf1 (@cumul_gen cf1) (lift_typing (@typing cf1)) Σ.1 -> + (forall Σ Γ j, @on_global_env cf1 (@cumul_gen cf1) (lift_typing (@typing cf1)) Σ.1 -> @on_global_env cf2 (@cumul_gen cf2) P Σ.1 -> @on_global_env cf3 (@cumul_gen cf3) Q Σ.1 -> - P Σ Γ t T -> Q Σ Γ t T) -> + P Σ Γ j -> Q Σ Γ j) -> @on_global_env cf2 (@cumul_gen cf2) P Σ -> @on_global_env cf3 (@cumul_gen cf3) Q Σ. Proof. - unfold on_global_env in *. - intros Hcf Hcumul X [hu X0]. split; auto. - simpl in *. destruct wfΣ as [cu wfΣ]. revert cu wfΣ. - revert X0. generalize (universes Σ) as univs, (retroknowledge Σ) as retro, (declarations Σ). clear hu Σ. - induction 1; constructor; try destruct o; try constructor; auto. - { depelim wfΣ. eauto. } - depelim wfΣ. specialize (IHX0 cu wfΣ). destruct o. - assert (X' := fun Γ t T => X ({| universes := univs; declarations := Σ |}, udecl0) Γ t T - (cu, wfΣ) (cu, X0) (cu, IHX0)); clear X. - rename X' into X. - clear IHX0. destruct d; simpl. - - destruct c; simpl. destruct cst_body0; simpl in *; now eapply X. - - simpl in *. destruct on_global_decl_d0 as [onI onP onNP]. - constructor; auto. - -- eapply Alli_impl; tea. intros. - refine {| ind_arity_eq := X1.(ind_arity_eq); - ind_cunivs := X1.(ind_cunivs) |}. - --- apply onArity in X1. unfold on_type in *; simpl in *. - now eapply X. - --- pose proof X1.(onConstructors) as X11. red in X11. - eapply All2_impl; eauto. - simpl. intros. destruct X2 as [? ? ? ?]; unshelve econstructor; eauto. - * apply X; eauto. - * clear -X0 X on_cargs0. revert on_cargs0. - generalize (cstr_args x0). - induction c in y |- *; destruct y; simpl; auto; - destruct a as [na [b|] ty]; simpl in *; auto; - split; intuition eauto. - * clear -X0 X on_cindices0. - revert on_cindices0. - generalize (List.rev (lift_context #|cstr_args x0| 0 (ind_indices x))). - generalize (cstr_indices x0). - induction 1; simpl; constructor; auto. - * move => v H. - move: Hcf (on_ctype_variance0 v H). - rewrite /config.impl/cstr_respects_variance/cumul_ctx_rel/cumul_pb_decls. - repeat match goal with - | [ |- context[match ?x with _ => _ end] ] => destruct x eqn:?; subst - end => //=. - move => Hcf [H1 H2]; split. - all: [ > eapply All2_fold_impl; try eapply H1 | eapply All2_impl; try eapply H2 ]; intros *; cbn; try eapply All_decls_alpha_pb_impl; intros *. - all: auto. - * move: on_lets_in_type0. - move: Hcf. - rewrite /config.impl. - do 3 destruct lets_in_constructor_types => //=. - all: rewrite ?andb_false_r //=. - --- simpl; intros. pose (onProjections X1). simpl in *; auto. - --- destruct X1. simpl. unfold check_ind_sorts in *. - destruct Universe.is_prop; auto. - destruct Universe.is_sprop; auto. - split. - * eapply Forall_impl; [ apply ind_sorts0 | cbn ]; intros. - eapply Forall_impl; [ eassumption | cbn ] => ?. - move: Hcf. - rewrite /leq_universe/leq_universe_n/leq_universe_n_/leq_levelalg_n/config.impl. - do 2 destruct check_univs => //=. - all: do 2 destruct prop_sub_type => //=. - all: repeat match goal with - | [ |- context[match ?x with _ => _ end] ] => destruct x eqn:?; subst - end => //=. - * move: Hcf; unfold config.impl. - repeat match goal with H : _ × _ |- _ => destruct H end. - do 2 destruct indices_matter; cbn [implb andb]; rewrite ?andb_false_r => //= Hcf. - eapply type_local_ctx_impl; eauto. - --- move: X1.(onIndices). - destruct ind_variance eqn:Heq => //. - eapply ind_respects_variance_impl. - repeat match goal with |- context[match ?x with _ => _ end] => destruct x eqn:?; subst end; eauto. - -- red in onP. red. - eapply All_local_env_impl; tea. - -- move: onVariance0. - eapply on_variance_impl; eauto. + intros Xcf Xcmp X [cu IH]. destruct wfΣ as [_ wfΣ]. split; auto. + revert cu IH wfΣ; generalize (universes Σ) as univs, (retroknowledge Σ) as retro, (declarations Σ). clear Σ. + induction g; intros; auto. constructor; auto. + depelim IH. depelim wfΣ. + specialize (IHg cu IH wfΣ). + constructor; auto. + destruct o. constructor; auto. + eapply @on_global_decl_impl_only_config with (cf := cf2) (5 := on_global_decl_d0) => //. + { intros. eapply X; cycle -1; tea; split; tea. } + { intros. now eapply Xcmp. } Qed. Lemma on_wf_global_env_impl `{checker_flags} {Σ : global_env} {wfΣ : on_global_env cumul_gen (lift_typing typing) Σ} P Q : - (forall Σ Γ t T, on_global_env cumul_gen (lift_typing typing) Σ.1 -> + (forall Σ Γ j, on_global_env cumul_gen (lift_typing typing) Σ.1 -> on_global_env cumul_gen P Σ.1 -> on_global_env cumul_gen Q Σ.1 -> - P Σ Γ t T -> Q Σ Γ t T) -> + P Σ Γ j -> Q Σ Γ j) -> on_global_env cumul_gen P Σ -> on_global_env cumul_gen Q Σ. Proof. unshelve eapply on_wf_global_env_impl_config; eauto; reflexivity. @@ -1787,7 +2240,7 @@ Module DeclarationTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E) Section Properties. Context {cf : checker_flags}. Context {Pcmp: global_env_ext -> context -> conv_pb -> term -> term -> Type}. - Context {P: global_env_ext -> context -> term -> typ_or_sort -> Type}. + Context {P: global_env_ext -> context -> judgment -> Type}. Let wf := on_global_env Pcmp P. diff --git a/common/theories/MonadBasicAst.v b/common/theories/MonadBasicAst.v index cafdeca2d..b1cb40fbb 100644 --- a/common/theories/MonadBasicAst.v +++ b/common/theories/MonadBasicAst.v @@ -22,10 +22,9 @@ Section with_monad. dbody <- bodyf dbody;; ret {| dname := dname; dtype := dtype; dbody := dbody; rarg := rarg |}. - Definition monad_typ_or_sort_map {T' T''} (f: T' -> T T'') t := + Definition monad_typ_or_sort_map {univ T' T''} (f: T' -> T T'') (t : judgment_ univ T') := match t with - | Typ ty => fty <- f ty;; ret (Typ fty) - | Sort => ret Sort + | Judge tm ty u => ftm <- monad_option_map f tm;; fty <- f ty;; ret (Judge ftm fty u) end. Definition monad_map_decl {term term'} (f : term -> T term') (d : context_decl term) : T (context_decl term') := diff --git a/common/theories/Reflect.v b/common/theories/Reflect.v index 431c56c98..923b0e564 100644 --- a/common/theories/Reflect.v +++ b/common/theories/Reflect.v @@ -397,7 +397,7 @@ Ltac finish_reflect := end); constructor; trivial; congruence. -Definition eqb_universes_decl x y := +Definition eqb_sorts_decl x y := match x, y with | Monomorphic_ctx, Monomorphic_ctx => true | Polymorphic_ctx cx, Polymorphic_ctx cy => eqb cx cy @@ -405,9 +405,9 @@ Definition eqb_universes_decl x y := end. #[global,program] Instance reflect_universes_decl : ReflectEq universes_decl := - {| eqb := eqb_universes_decl |}. + {| eqb := eqb_sorts_decl |}. Next Obligation. - unfold eqb_universes_decl. + unfold eqb_sorts_decl. intros [] []; finish_reflect. Qed. diff --git a/common/theories/Universes.v b/common/theories/Universes.v index 71e71ac40..eca654b27 100644 --- a/common/theories/Universes.v +++ b/common/theories/Universes.v @@ -26,7 +26,7 @@ Hint Extern 10 => absurd : core. (** A valuation is a universe level (nat) given for each universe lvariable (Level.t). - It is >= for polymorphic concreteUniverses and > 0 for monomorphic concreteUniverses. *) + It is >= for polymorphic concrete_sort and > 0 for monomorphic concrete_sort. *) Record valuation := { valuation_mono : string -> positive ; valuation_poly : nat -> nat }. @@ -247,7 +247,6 @@ Module PropLevel. End PropLevel. Module LevelExpr. - (* npe = no prop expression *) Definition t := (Level.t * nat)%type. Global Instance Evaluable : Evaluable t @@ -270,15 +269,6 @@ Module LevelExpr. Definition set : t := (Level.lzero, 0%nat). Definition type1 : t := (Level.lzero, 1%nat). - (* Used for (un)quoting. *) - Definition from_kernel_repr (e : Level.t * bool) : t - := (e.1, if e.2 then 1%nat else 0%nat). - - Definition to_kernel_repr (e : t) : Level.t * bool - := match e with - | (l, b) => (l, match b with 0%nat => false | _ => true end) - end. - Definition eq : t -> t -> Prop := eq. Definition eq_equiv : Equivalence eq := _. @@ -349,7 +339,7 @@ Module LevelExprSetDecide := LevelExprSetProp.Dec. Module LevelExprSetExtraOrdProp := MSets.ExtraOrdProperties LevelExprSet LevelExprSetOrdProp. (* We have decidable equality w.r.t leibniz equality for sets of levels. - This means concreteUniverses also have a decidable equality. *) + This means concrete_sort also has a decidable equality. *) #[global, program] Instance levelexprset_reflect : ReflectEq LevelExprSet.t := { eqb := LevelExprSet.equal }. Next Obligation. @@ -596,7 +586,7 @@ End NonEmptySetFacts. Import NonEmptySetFacts. -Module LevelAlgExpr. +Module Universe. (** A universe / an algebraic expression is a list of universe expressions which is: - sorted - without duplicate @@ -618,15 +608,20 @@ Module LevelAlgExpr. #[global] Instance eq_dec_univ0 : EqDec t := eq_dec. Definition make (e: LevelExpr.t) : t := singleton e. - Definition make' (l: Level.t) := singleton (LevelExpr.make l). + Definition make' (l: Level.t) : t := singleton (LevelExpr.make l). + + Lemma make'_inj l l' : make' l = make' l' -> l = l'. + Proof. + destruct l, l' => //=; now inversion 1. + Qed. (** 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 *) Definition exprs : t -> LevelExpr.t * list LevelExpr.t := to_nonempty_list. - Global Instance Evaluable : Evaluable LevelAlgExpr.t + Global Instance Evaluable : Evaluable Universe.t := fun v u => - let '(e, u) := LevelAlgExpr.exprs u in + let '(e, u) := Universe.exprs u in List.fold_left (fun n e => Nat.max (val v e) n) u (val v e). (** Test if the universe is a lub of levels or contains +n's. *) @@ -637,6 +632,9 @@ Module LevelAlgExpr. Definition is_level (u : t) : bool := (LevelExprSet.cardinal u =? 1)%nat && is_levels u. + (* Used for quoting. *) + Definition from_kernel_repr (e : Level.t * nat) (es : list (Level.t * nat)) : t + := add_list es (Universe.make e). Definition succ : t -> t := map LevelExpr.succ. @@ -665,7 +663,7 @@ Module LevelAlgExpr. { intros ? H. apply irreflexivity in H; assumption. } { intros ??? H1 H2; etransitivity; tea. } Qed. -End LevelAlgExpr. +End Universe. Ltac u := change LevelSet.elt with Level.t in *; @@ -673,21 +671,21 @@ Ltac u := (* change ConstraintSet.elt with UnivConstraint.t in *. *) -Lemma val_fold_right (u : LevelAlgExpr.t) v : - val v u = fold_right (fun e x => Nat.max (val v e) x) (val v (LevelAlgExpr.exprs u).1) - (List.rev (LevelAlgExpr.exprs u).2). +Lemma val_fold_right (u : Universe.t) v : + val v u = fold_right (fun e x => Nat.max (val v e) x) (val v (Universe.exprs u).1) + (List.rev (Universe.exprs u).2). Proof. - unfold val at 1, LevelAlgExpr.Evaluable. - destruct (LevelAlgExpr.exprs u). + unfold val at 1, Universe.Evaluable. + destruct (Universe.exprs u). now rewrite fold_left_rev_right. Qed. -Lemma val_In_le (u : LevelAlgExpr.t) v e : +Lemma val_In_le (u : Universe.t) v e : LevelExprSet.In e u -> val v e <= val v u. Proof. intro H. rewrite val_fold_right. apply In_to_nonempty_list_rev in H. - fold LevelAlgExpr.exprs in H; destruct (LevelAlgExpr.exprs u); cbn in *. + fold Universe.exprs in H; destruct (Universe.exprs u); cbn in *. destruct H as [H|H]. - subst. induction (List.rev l); cbnr. lia. - induction (List.rev l); cbn; invs H. @@ -695,12 +693,12 @@ Proof. + specialize (IHl0 H0). lia. Qed. -Lemma val_In_max (u : LevelAlgExpr.t) v : +Lemma val_In_max (u : Universe.t) v : exists e, LevelExprSet.In e u /\ val v e = val v u. Proof. eapply iff_ex. { intro. eapply and_iff_compat_r. apply In_to_nonempty_list_rev. } - rewrite val_fold_right. fold LevelAlgExpr.exprs; destruct (LevelAlgExpr.exprs u) as [e l]; cbn in *. + rewrite val_fold_right. fold Universe.exprs; 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]]. @@ -712,7 +710,7 @@ Proof. destruct H1 as [H1|H1]; [now left|right]. now constructor 2. Qed. -Lemma val_ge_caract (u : LevelAlgExpr.t) v k : +Lemma val_ge_caract (u : Universe.t) v k : (forall e, LevelExprSet.In e u -> val v e <= k) <-> val v u <= k. Proof. split. @@ -720,7 +718,7 @@ Proof. eapply iff_forall; intro. eapply imp_iff_compat_r. apply In_to_nonempty_list_rev. } rewrite val_fold_right. - fold LevelAlgExpr.exprs; destruct (LevelAlgExpr.exprs u) as [e l]; cbn; clear. + fold Universe.exprs; destruct (Universe.exprs u) as [e l]; cbn; clear. induction (List.rev l); cbn. + intros H. apply H. left; reflexivity. + intros H. @@ -733,14 +731,14 @@ Proof. - intros H e He. eapply val_In_le in He. etransitivity; eassumption. Qed. -Lemma val_le_caract (u : LevelAlgExpr.t) v k : +Lemma val_le_caract (u : Universe.t) v k : (exists e, LevelExprSet.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 In_to_nonempty_list_rev. } rewrite val_fold_right. - fold LevelAlgExpr.exprs; destruct (LevelAlgExpr.exprs u) as [e l]; cbn; clear. + fold Universe.exprs; 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. @@ -756,7 +754,7 @@ Qed. -Lemma val_caract (u : LevelAlgExpr.t) v k : +Lemma val_caract (u : Universe.t) v k : val v u = k <-> (forall e, LevelExprSet.In e u -> val v e <= k) /\ exists e, LevelExprSet.In e u /\ val v e = k. @@ -772,7 +770,7 @@ Proof. exists e. split; tas. lia. Qed. -Lemma val_add v e (s: LevelAlgExpr.t) +Lemma val_add v e (s: Universe.t) : val v (add e s) = Nat.max (val v e) (val v s). Proof. apply val_caract. split. @@ -786,7 +784,7 @@ Proof. Qed. Lemma val_sup v s1 s2 : - val v (LevelAlgExpr.sup s1 s2) = Nat.max (val v s1) (val v s2). + val v (Universe.sup s1 s2) = Nat.max (val v s1) (val v s2). Proof. eapply val_caract. cbn. split. - intros e' H. eapply LevelExprSet.union_spec in H. destruct H as [H|H]. @@ -808,11 +806,11 @@ Proof. Qed. -Lemma levelalg_get_is_level_correct u l : - LevelAlgExpr.get_is_level u = Some l -> u = LevelAlgExpr.make' l. +Lemma universe_get_is_level_correct u l : + Universe.get_is_level u = Some l -> u = Universe.make' l. Proof. intro H. - unfold LevelAlgExpr.get_is_level in *. + unfold Universe.get_is_level in *. destruct (LevelExprSet.elements u) as [|l0 L] eqn:Hu1; [discriminate |]. destruct l0, L; try discriminate. * destruct n; inversion H; subst. @@ -822,1288 +820,1286 @@ Qed. Lemma sup0_comm x1 x2 : - LevelAlgExpr.sup x1 x2 = LevelAlgExpr.sup x2 x1. + Universe.sup x1 x2 = Universe.sup x2 x1. Proof. apply eq_univ'; simpl. unfold LevelExprSet.Equal. intros H. rewrite !LevelExprSet.union_spec. intuition. Qed. +(* +Lemma val_zero_exprs v (l : Universe.t) : 0 <= val v l. +Proof. + rewrite val_fold_right. + destruct (Universe.exprs l) as [e u']; clear l; cbn. + induction (List.rev u'); simpl. + - destruct e as [npl_expr]. + destruct npl_expr as [t b]. + cbn. + assert (0 <= val v t) by apply Level.val_zero. + destruct b;lia. + - pose proof (LevelExpr.val_zero a v); lia. +Qed. *) -Declare Scope univ_scope. -Delimit Scope univ_scope with u. - -(** Universes with linear hierarchy *) -Section ConcreteUniverses. - Context {cf}. - Inductive concreteUniverses := - | UProp - | USProp - | UType (i : nat). - Derive NoConfusion EqDec for concreteUniverses. +Module ConstraintType. + Inductive t_ : Set := Le (z : Z) | Eq. + Derive NoConfusion EqDec for t_. - (** u + n <= u' *) - Definition leq_cuniverse_n n u u' := - match u, u' with - | UProp, UProp - | USProp, USProp => (n = 0)%Z - | UType u, UType u' => (Z.of_nat u <= Z.of_nat u' - n)%Z - | UProp, UType u => - if prop_sub_type then True else False - | _, _ => False - end. + Definition t := t_. + Definition eq : t -> t -> Prop := eq. + Definition eq_equiv : Equivalence eq := _. - Definition leq_cuniverse := leq_cuniverse_n 0. - Definition lt_cuniverse := leq_cuniverse_n 1. + Definition Le0 := Le 0. + Definition Lt := Le 1. - Notation "x <_ n y" := (leq_cuniverse_n n x y) (at level 10, n name) : univ_scope. - Notation "x < y" := (lt_cuniverse x y) : univ_scope. - Notation "x <= y" := (leq_cuniverse x y) : univ_scope. + Inductive lt_ : t -> t -> Prop := + | LeLe n m : (n < m)%Z -> lt_ (Le n) (Le m) + | LeEq n : lt_ (Le n) Eq. + Derive Signature for lt_. + Definition lt := lt_. - Definition cuniv_sup u1 u2 := - match u1, u2 with - | UProp, UProp => UProp - | USProp, USProp => USProp - | UType v, UType v' => UType (Nat.max v v') - | _, UType _ => u2 - | UType _, _ => u1 - | UProp, USProp => UProp - | USProp, UProp => UProp - end. + Global Instance lt_strorder : StrictOrder lt. + Proof. + constructor. + - intros []; intro X; inversion X. lia. + - intros ? ? ? X Y; invs X; invs Y; constructor. lia. + Qed. - Definition is_uprop u := match u with UProp => True | _ => False end. - Definition is_usprop u := match u with USProp => True | _ => False end. - Definition is_uproplevel u := match u with USProp | UProp => True | _ => False end. - Definition is_uproplevel_or_set u := match u with USProp | UProp | UType 0 => True | _ => False end. - Definition is_utype u := match u with UType _ => True | _ => False end. + Global Instance lt_compat : Proper (eq ==> eq ==> iff) lt. + Proof. + intros ? ? X ? ? Y; invs X; invs Y. reflexivity. + Qed. - (** Type of a product *) - Definition cuniv_of_product domsort rangsort := - match rangsort with - (* Prop and SProp impredicativity *) - | UProp | USProp => rangsort - | _ => cuniv_sup domsort rangsort + Definition compare (x y : t) : comparison := + match x, y with + | Le n, Le m => Z.compare n m + | Le _, Eq => Datatypes.Lt + | Eq, Eq => Datatypes.Eq + | Eq, _ => Datatypes.Gt end. - Lemma cuniv_sup_comm u u' : cuniv_sup u u' = cuniv_sup u' u. - Proof using Type. destruct u, u'; cbn; try congruence. f_equal; lia. Qed. - - Lemma cuniv_sup_not_uproplevel u u' : - ~ is_uproplevel u -> ∑ n, cuniv_sup u u' = UType n. - Proof using Type. - destruct u, u'; cbn; intros Hp; try absurd; - now eexists. + Lemma compare_spec x y : CompareSpec (eq x y) (lt x y) (lt y x) (compare x y). + Proof. + destruct x, y; repeat constructor. simpl. + destruct (Z.compare_spec z z0); simpl; constructor. + subst; constructor. now constructor. now constructor. Qed. - Lemma cuniv_le_uprop_inv u : (u <= UProp)%u -> u = UProp. - Proof using Type. destruct u; simpl; intros Hle; try absurd; now reflexivity. Qed. - - Lemma cuniv_le_usprop_inv u : (u <= USProp)%u -> u = USProp. - Proof using Type. destruct u; simpl; intros Hle; try absurd; now reflexivity. Qed. - - Lemma cuniv_uprop_le_inv u : (UProp <= u)%u -> - (u = UProp \/ (prop_sub_type /\ exists n, u = UType n)). - Proof using Type. - destruct u; simpl; intros Hle; [ left; reflexivity | absurd | right ]. - destruct prop_sub_type; [| absurd]. - split; [ trivial | now eexists ]. + Lemma eq_dec x y : {eq x y} + {~ eq x y}. + Proof. + unfold eq. decide equality. apply Z.eq_dec. Qed. +End ConstraintType. - Lemma cuniv_sup_mon u u' v v' : (u <= u')%u -> (UType v <= UType v')%u -> - (cuniv_sup u (UType v) <= cuniv_sup u' (UType v'))%u. - Proof using Type. - destruct u, u'; simpl; intros Hle Hle'; try absurd; - lia. - Qed. +Module UnivConstraint. + Definition t : Set := Level.t * ConstraintType.t * Level.t. - Lemma leq_cuniv_of_product_mon u u' v v' : - (u <= u')%u -> - (v <= v')%u -> - (cuniv_of_product u v <= cuniv_of_product u' v')%u. - Proof using Type. - intros Hle1 Hle2. - destruct v, v'; cbn in Hle2 |- *; auto. - - destruct u'; cbn; assumption. - - apply cuniv_sup_mon; assumption. - Qed. + Definition eq : t -> t -> Prop := eq. + Definition eq_equiv : Equivalence eq := _. - Lemma impredicative_cuniv_product {l u} : - is_uproplevel u -> - (cuniv_of_product l u <= u)%u. - Proof using Type. now destruct u. Qed. + 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_. - Global Instance leq_cuniverse_refl : Reflexive leq_cuniverse. - Proof using Type. - intros []; cbnr; lia. + 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. - Global Instance leq_cuniverse_n_trans n : Transitive (leq_cuniverse_n (Z.of_nat n)). - Proof using Type. - intros [] [] []; cbnr; trivial; try absurd; lia. + Lemma lt_compat : Proper (eq ==> eq ==> iff) lt. + Proof. + intros ? ? X ? ? Y; invs X; invs Y. reflexivity. Qed. - Global Instance leq_cuniverse_trans : Transitive leq_cuniverse. - Proof using Type. apply (leq_cuniverse_n_trans 0). Qed. - - Global Instance lt_cuniverse_trans : Transitive lt_cuniverse. - Proof using Type. apply (leq_cuniverse_n_trans 1). Qed. - - Global Instance leq_cuniverse_preorder : PreOrder leq_cuniverse := - Build_PreOrder _ _ _. + Definition compare : t -> t -> comparison := + fun '(l1, t, l2) '(l1', t', l2') => + compare_cont (Level.compare l1 l1') + (compare_cont (ConstraintType.compare t t') + (Level.compare l2 l2')). - Global Instance lt_cuniverse_str_order : StrictOrder lt_cuniverse. - Proof using Type. - split. - - intros []; unfold complement; cbnr; lia. - - exact _. + 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. - Global Instance leq_cuniverse_antisym : Antisymmetric _ eq leq_cuniverse. - Proof using Type. - intros [] []; cbnr; try absurd. - intros. f_equal; lia. - Qed. + Lemma eq_dec x y : {eq x y} + {~ eq x y}. + Proof. + unfold eq. decide equality; apply eq_dec. + Defined. -End ConcreteUniverses. + Definition eq_leibniz (x y : t) : eq x y -> x = y := id. +End UnivConstraint. +Module ConstraintSet := MSetAVL.Make UnivConstraint. +Module ConstraintSetFact := WFactsOn UnivConstraint ConstraintSet. +Module ConstraintSetOrdProp := MSetProperties.OrdProperties ConstraintSet. +Module ConstraintSetProp := ConstraintSetOrdProp.P. +Module CS := ConstraintSet. +Module ConstraintSetDecide := ConstraintSetProp.Dec. +Module ConstraintSetExtraOrdProp := MSets.ExtraOrdProperties ConstraintSet ConstraintSetOrdProp. +Module ConstraintSetExtraDecide := MSetAVL.Decide UnivConstraint ConstraintSet. +Ltac csets := ConstraintSetDecide.fsetdec. -(** This inductive classifies which eliminations are allowed for inductive types - in lvarious sorts. *) -Inductive allowed_eliminations : Set := - | IntoSProp - | IntoPropSProp - | IntoSetPropSProp - | IntoAny. -Derive NoConfusion EqDec for allowed_eliminations. +Notation "(=_cset)" := ConstraintSet.Equal (at level 0). +Infix "=_cset" := ConstraintSet.Equal (at level 30). +Notation "(==_cset)" := ConstraintSet.equal (at level 0). +Infix "==_cset" := ConstraintSet.equal (at level 30). +Definition declared_cstr_levels levels (cstr : UnivConstraint.t) := + let '(l1,_,l2) := cstr in + LevelSet.In l1 levels /\ LevelSet.In l2 levels. -Definition is_allowed_elimination_cuniv (allowed : allowed_eliminations) : concreteUniverses -> Prop := - match allowed with - | IntoSProp => is_usprop - | IntoPropSProp => is_uproplevel - | IntoSetPropSProp => is_uproplevel_or_set - | IntoAny => fun _ => True - end. +Definition is_declared_cstr_levels levels (cstr : UnivConstraint.t) : bool := + let '(l1,_,l2) := cstr in + LevelSet.mem l1 levels && LevelSet.mem l2 levels. +Lemma CS_union_empty s : ConstraintSet.union ConstraintSet.empty s =_cset s. +Proof. + intros x; rewrite ConstraintSet.union_spec. lsets. +Qed. -Module Universe. - Inductive t_ := - lProp | lSProp | lType (_ : LevelAlgExpr.t). - Derive NoConfusion for t_. +Lemma CS_For_all_union f cst cst' : ConstraintSet.For_all f (ConstraintSet.union cst cst') -> + ConstraintSet.For_all f cst. +Proof. + unfold CS.For_all. + intros IH x inx. apply (IH x). + now eapply CS.union_spec; left. +Qed. - Definition t := t_. +Lemma CS_For_all_add P x s : CS.For_all P (CS.add x s) -> P x /\ CS.For_all P s. +Proof. + intros. + split. + * apply (H x), CS.add_spec; left => //. + * intros y iny. apply (H y), CS.add_spec; right => //. +Qed. - Definition eqb (u1 u2 : t) : bool := - match u1, u2 with - | lSProp, lSProp => true - | lProp, lProp => true - | lType e1, lType e2 => eqb e1 e2 - | _, _ => false - end. +#[global] Instance CS_For_all_proper P : Morphisms.Proper ((=_cset) ==> iff)%signature (ConstraintSet.For_all P). +Proof. + intros s s' eqs. + unfold CS.For_all. split; intros IH x inxs; apply (IH x); + now apply eqs. +Qed. - #[global, program] Instance reflect_eq_universe : ReflectEq t := - { eqb := eqb }. - Next Obligation. - destruct x, y; cbn; try constructor; auto; try congruence. - destruct (eqb_spec t0 t1); constructor. now f_equal. - congruence. - Qed. +(** {6 Sort instances} *) - #[global] Instance eq_dec_univ : EqDec t := eq_dec. +Module Instance. - Definition on_sort (P: LevelAlgExpr.t -> Prop) (def: Prop) (s : t) := - match s with - | lProp | lSProp => def - | lType l => P l + (** A universe instance represents a vector of argument concrete_sort + to a polymorphic definition (constant, inductive or constructor). *) + Definition t : Set := list Level.t. + + Definition empty : t := []. + Definition is_empty (i : t) : bool := + match i with + | [] => true + | _ => false end. - (** Create a universe representing the given level. *) - Definition make (l : Level.t) : t := - lType (LevelAlgExpr.make (LevelExpr.make l)). + Definition eqb (i j : t) := + forallb2 Level.eqb i j. - Definition of_expr e := (lType (LevelAlgExpr.make e)). + Definition equal_upto (f : Level.t -> Level.t -> bool) (i j : t) := + forallb2 f i j. +End Instance. - Definition add_to_exprs (e : LevelExpr.t) (u : t) : t := - match u with - | lSProp | lProp => u - | lType l => lType (add e l) - end. +Module UContext. + Definition t := list name × (Instance.t × ConstraintSet.t). - Definition add_list_to_exprs (es : list LevelExpr.t) (u : t) : t := - match u with - | lSProp | lProp => u - | lType l => lType (add_list es l) - end. + Definition make' : Instance.t -> ConstraintSet.t -> Instance.t × ConstraintSet.t := pair. + Definition make (ids : list name) (inst_ctrs : Instance.t × ConstraintSet.t) : t := (ids, inst_ctrs). - (** Test if the universe is a lub of levels or contains +n's. *) - Definition is_levels (u : t) : bool := - match u with - | lSProp | lProp => true - | lType l => LevelAlgExpr.is_levels l - end. + Definition empty : t := ([], (Instance.empty, ConstraintSet.empty)). - (** Test if the universe is a level or an algebraic universe. *) - Definition is_level (u : t) : bool := - match u with - | lSProp | lProp => true - | lType l => LevelAlgExpr.is_level l - end. + Definition instance : t -> Instance.t := fun x => fst (snd x). + Definition constraints : t -> ConstraintSet.t := fun x => snd (snd x). - Definition is_sprop (u : t) : bool := - match u with - | lSProp => true - | _ => false - end. + Definition dest : t -> list name * (Instance.t * ConstraintSet.t) := fun x => x. +End UContext. - Definition is_prop (u : t) : bool := - match u with - | lProp => true - | _ => false - end. +Module AUContext. + Definition t := list name × ConstraintSet.t. - Definition is_type_sort (u : t) : bool := - match u with - | lType _ => true - | _ => false - end. + Definition make (ids : list name) (ctrs : ConstraintSet.t) : t := (ids, ctrs). + Definition repr (x : t) : UContext.t := + let (u, cst) := x in + (u, (mapi (fun i _ => Level.lvar i) u, cst)). - Definition type0 : t := make Level.lzero. - Definition type1 : t := lType (LevelAlgExpr.make LevelExpr.type1). + Definition levels (uctx : t) : LevelSet.t := + LevelSetProp.of_list (fst (snd (repr uctx))). - Definition of_levels (l : PropLevel.t + Level.t) : t := - match l with - | inl PropLevel.lSProp => lSProp - | inl PropLevel.lProp => lProp - | inr l => lType (LevelAlgExpr.make' l) - end. + #[local] + Existing Instance EqDec_ReflectEq. + Definition inter (au av : AUContext.t) : AUContext.t := + let prefix := (split_prefix au.1 av.1).1.1 in + let lvls := fold_left_i (fun s i _ => LevelSet.add (Level.lvar i) s) prefix LevelSet.empty in + let filter := ConstraintSet.filter (is_declared_cstr_levels lvls) in + make prefix (ConstraintSet.union (filter au.2) (filter av.2)). +End AUContext. - (* Used for quoting. *) - Definition from_kernel_repr (e : Level.t * bool) (es : list (Level.t * bool)) : t - := lType (add_list (List.map LevelExpr.from_kernel_repr es) - (LevelAlgExpr.make (LevelExpr.from_kernel_repr e))). - - (* Definition uex_to_kernel_repr (e : LevelExpr.t) : Level.t * bool := *) - (* match e with *) - (* | LevelExpr.npe (l, b) => (NoPropLevel.to_level l, b) *) - (* end. *) - - (* Definition to_kernel_repr (u : t) : list (Level.t * bool) *) - (* := map (LevelExpr.to_kernel_repr) (LevelExprSet.elements u). *) - (* match u with *) - (* | lProp => [(Level.lProp, false)] *) - (* | lSProp => [(Level.lSProp, false)] *) - (* | lType l => map uex_to_kernel_repr (LevelExprSet.elements l) *) - (* end. *) +Module ContextSet. + Definition t := LevelSet.t × ConstraintSet.t. - (** The universe strictly above FOR TYPING (not cumulativity) *) + Definition levels : t -> LevelSet.t := fst. + Definition constraints : t -> ConstraintSet.t := snd. - Definition super (l : t) : t := - match l with - | lSProp => type1 - | lProp => type1 - | lType l => lType (LevelAlgExpr.succ l) - end. + Definition empty : t := (LevelSet.empty, ConstraintSet.empty). - Definition sup (u v : t) : t := - match u,v with - | lSProp, lProp => lProp - | lProp, lSProp => lProp - | (lSProp | lProp), u' => u' - | u', (lSProp | lProp) => u' - | lType l1, lType l2 => lType (LevelAlgExpr.sup l1 l2) - end. + Definition is_empty (uctx : t) + := LevelSet.is_empty (fst uctx) && ConstraintSet.is_empty (snd uctx). - Definition get_univ_exprs (u : t) (H1 : is_prop u = false) (H2 : is_sprop u = false) : LevelAlgExpr.t. - destruct u; try discriminate; easy. Defined. + Definition Equal (x y : t) : Prop := + x.1 =_lset y.1 /\ x.2 =_cset y.2. - (** Type of a product *) - Definition sort_of_product (domsort rangsort : t) := - (* Prop and SProp impredicativity *) - if Universe.is_prop rangsort || Universe.is_sprop rangsort then rangsort - else Universe.sup domsort rangsort. + Definition equal (x y : t) : bool := + x.1 ==_lset y.1 && x.2 ==_cset y.2. - Lemma eqb_refl u : eqb u u. - Proof. - destruct u; auto. - now apply LevelExprSet.equal_spec. - Qed. + Definition Subset (x y : t) : Prop := + LevelSet.Subset (levels x) (levels y) /\ + ConstraintSet.Subset (constraints x) (constraints y). - Definition get_is_level (u : t) : option Level.t := - match u with - | lSProp => None - | lProp => None - | lType l => LevelAlgExpr.get_is_level l - end. + Definition subset (x y : t) : bool := + LevelSet.subset (levels x) (levels y) && + ConstraintSet.subset (constraints x) (constraints y). - Definition to_cuniv v u := - match u with - | lSProp => USProp - | lProp => UProp - | lType l => UType (val v l) - end. + Definition inter (x y : t) : t := + (LevelSet.inter (levels x) (levels y), + ConstraintSet.inter (constraints x) (constraints y)). - Lemma val_make v l - : to_cuniv v (make l) = UType (val v l). + Definition inter_spec (x y : t) : + Subset (inter x y) x /\ + Subset (inter x y) y /\ + forall z, Subset z x -> Subset z y -> Subset z (inter x y). Proof. - destruct l; cbnr. + split; last split. + 1,2: split=> ?; [move=> /LevelSet.inter_spec [//]|move=> /ConstraintSet.inter_spec [//]]. + move=> ? [??] [??]; split=> ??; + [apply/LevelSet.inter_spec|apply/ConstraintSet.inter_spec]; split; auto. Qed. - Lemma make_inj x y : - make x = make y -> x = y. + Definition union (x y : t) : t := + (LevelSet.union (levels x) (levels y), ConstraintSet.union (constraints x) (constraints y)). + + Definition union_spec (x y : t) : + Subset x (union x y) /\ + Subset y (union x y) /\ + forall z, Subset x z -> Subset y z -> Subset (union x y) z. Proof. - destruct x, y; try now inversion 1. + split; last split. + 1,2: split=> ??; [apply/LevelSet.union_spec|apply/ConstraintSet.union_spec ]; by constructor. + move=> ? [??] [??]; split=> ?; + [move=>/LevelSet.union_spec|move=>/ConstraintSet.union_spec]=>-[]; auto. Qed. - Inductive lt_ : t -> t -> Prop := - | ltPropSProp : lt_ lProp lSProp - | ltPropType s : lt_ lProp (lType s) - | ltSPropType s : lt_ lSProp (lType s) - | ltTypeType s1 s2 : LevelAlgExpr.lt s1 s2 -> lt_ (lType s1) (lType s2). - Derive Signature for lt_. - - Definition lt := lt_. - - Module OT <: OrderedType. - Definition t := t. - #[local] Definition eq : t -> t -> Prop := eq. - #[local] Definition eq_equiv : Equivalence eq := _. - Definition lt := lt. - #[local] Instance lt_strorder : StrictOrder lt. - Proof. - constructor. - - intros [| |] X; inversion X. - now eapply irreflexivity in H1. - - intros [| |] [| |] [| |] X1 X2; - inversion X1; inversion X2; constructor. - subst. - etransitivity; tea. - Qed. - - Definition lt_compat : Proper (eq ==> eq ==> iff) lt. - Proof. - intros x y e z t e'. hnf in * |- ; subst. reflexivity. - Qed. - Definition compare (x y : t) : comparison - := match x, y with - | lProp, lProp => Eq - | lProp, _ => Lt - | _, lProp => Gt - | lSProp, lSProp => Eq - | lSProp, _ => Lt - | _, lSProp => Gt - | lType x, lType y => LevelExprSet.compare x y - end. - Lemma compare_spec x y : CompareSpec (eq x y) (lt x y) (lt y x) (compare x y). - Proof. - cbv [compare eq]. - destruct x, y. - all: lazymatch goal with - | [ |- context[LevelExprSet.compare ?x ?y] ] - => destruct (LevelExprSet.compare_spec x y) - | _ => idtac - end. - all: lazymatch goal with - | [ H : LevelExprSet.eq (?f ?x) (?f ?y) |- _ ] - => apply LevelExprSet.eq_leibniz in H; - is_var x; is_var y; destruct x, y; cbn in H; subst - | _ => idtac - end. - all: repeat constructor; try apply f_equal; try assumption. - f_equal; apply Eqdep_dec.UIP_dec; decide equality. - Qed. - Definition eq_dec (x y : t) : {x = y} + {x <> y}. - Proof. repeat decide equality. apply LevelAlgExpr.eq_dec_univ0. Defined. - End OT. - Module OTOrig <: OrderedTypeOrig := Backport_OT OT. -End Universe. + Lemma equal_spec s s' : equal s s' <-> Equal s s'. + Proof. + rewrite /equal/Equal/is_true Bool.andb_true_iff LevelSet.equal_spec ConstraintSet.equal_spec. + reflexivity. + Qed. -Module UniverseMap := FMapAVL.Make Universe.OTOrig. -Module UniverseMapFact := FMapFacts.WProperties UniverseMap. -Module UniverseMapExtraFact := FSets.WFactsExtra_fun Universe.OTOrig UniverseMap UniverseMapFact.F. -Module UniverseMapDecide := FMapAVL.Decide Universe.OTOrig UniverseMap. + Lemma subset_spec s s' : subset s s' <-> Subset s s'. + Proof. + rewrite /subset/Subset/is_true Bool.andb_true_iff LevelSet.subset_spec ConstraintSet.subset_spec. + reflexivity. + Qed. -Definition is_propositional u := - Universe.is_prop u || Universe.is_sprop u. + Lemma subsetP s s' : reflect (Subset s s') (subset s s'). + Proof. + generalize (subset_spec s s'). + destruct subset; case; constructor; intuition. + Qed. +End ContextSet. -Notation "⟦ u ⟧_ v" := (Universe.to_cuniv v u) (at level 0, format "⟦ u ⟧_ v", v name) : univ_scope. +Export (hints) ContextSet. +Notation "(=_cs)" := ContextSet.Equal (at level 0). +Notation "(⊂_cs)" := ContextSet.Subset (at level 0). +Infix "=_cs" := ContextSet.Equal (at level 30). +Infix "⊂_cs" := ContextSet.Subset (at level 30). +Notation "(==_cs)" := ContextSet.equal (at level 0). +Notation "(⊂?_cs)" := ContextSet.subset (at level 0). +Infix "==_cs" := ContextSet.equal (at level 30). +Infix "⊂?_cs" := ContextSet.subset (at level 30). -Lemma val_universe_sup v u1 u2 : - Universe.to_cuniv v (Universe.sup u1 u2) = cuniv_sup (Universe.to_cuniv v u1) (Universe.to_cuniv v u2). +Lemma incl_cs_refl cs : cs ⊂_cs cs. Proof. - destruct u1 as [ | | l1]; destruct u2 as [ | | l2];cbn;try lia; auto. - f_equal. apply val_sup. + split; [lsets|csets]. Qed. -Lemma is_prop_val u : - Universe.is_prop u -> forall v, Universe.to_cuniv v u = UProp. +Lemma incl_cs_trans cs1 cs2 cs3 : cs1 ⊂_cs cs2 -> cs2 ⊂_cs cs3 -> cs1 ⊂_cs cs3. Proof. - destruct u as [| | u];try discriminate;auto. + intros [? ?] [? ?]; split; [lsets|csets]. Qed. -Lemma is_sprop_val u : - Universe.is_sprop u -> forall v, Universe.to_cuniv v u = USProp. +Lemma empty_contextset_subset u : ContextSet.empty ⊂_cs u. Proof. - destruct u as [| | u];try discriminate;auto. + red. split; cbn; [lsets|csets]. Qed. -(* -Lemma val_zero_exprs v (l : LevelAlgExpr.t) : 0 <= val v l. -Proof. - rewrite val_fold_right. - destruct (LevelAlgExpr.exprs l) as [e u']; clear l; cbn. - induction (List.rev u'); simpl. - - destruct e as [npl_expr]. - destruct npl_expr as [t b]. - cbn. - assert (0 <= val v t) by apply Level.val_zero. - destruct b;lia. - - pose proof (LevelExpr.val_zero a v); lia. -Qed. *) +(* Variance info is needed to do full universe polymorphism *) +Module Variance. + (** A universe position in the instance given to a cumulative + inductive can be the following. Note there is no Contralvariant + case because [forall x : A, B <= forall x : A', B'] requires [A = + A'] as opposed to [A' <= A]. *) + Inductive t := + | Irrelevant : t + | Covariant : t + | Invariant : t. + Derive NoConfusion EqDec for t. + (* val check_subtype : t -> t -> bool *) + (* val sup : t -> t -> t *) +End Variance. -Lemma val_is_prop u v : - Universe.to_cuniv v u = UProp <-> Universe.is_prop u. -Proof. - destruct u; auto;cbn in *; intuition congruence. -Qed. +(** What to check of two instances *) +Variant opt_variance := + AllEqual | AllIrrelevant | Variance of list Variance.t. -Lemma val_is_sprop u v : - Universe.to_cuniv v u = USProp <-> Universe.is_sprop u. -Proof. - destruct u;auto;cbn in *; intuition congruence. -Qed. +Inductive universes_decl : Type := +| Monomorphic_ctx +| Polymorphic_ctx (cst : AUContext.t). -Lemma is_prop_and_is_sprop_val_false u : - Universe.is_prop u = false -> Universe.is_sprop u = false -> - forall v, ∑ n, Universe.to_cuniv v u = UType n. -Proof. - intros Hp Hsp v. - destruct u; try discriminate. simpl. eexists; eauto. -Qed. +Derive NoConfusion for universes_decl. -Lemma val_is_prop_false u v n : - Universe.to_cuniv v u = UType n -> 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). discriminate. -Qed. +Definition levels_of_udecl u := + match u with + | Monomorphic_ctx => LevelSet.empty + | Polymorphic_ctx ctx => AUContext.levels ctx + end. -Lemma get_is_level_correct u l : - Universe.get_is_level u = Some l -> u = Universe.lType (LevelAlgExpr.make' l). -Proof. - intro H; destruct u; cbnr; try discriminate. - f_equal; now apply levelalg_get_is_level_correct. +Definition constraints_of_udecl u := + match u with + | Monomorphic_ctx => ConstraintSet.empty + | Polymorphic_ctx ctx => snd (snd (AUContext.repr ctx)) + end. + +Declare Scope univ_scope. +Delimit Scope univ_scope with u. + +Inductive satisfies0 (v : valuation) : UnivConstraint.t -> Prop := +| satisfies0_Lt (l l' : Level.t) (z : Z) : (Z.of_nat (val v l) <= Z.of_nat (val v l') - z)%Z + -> satisfies0 v (l, ConstraintType.Le z, l') +| satisfies0_Eq (l l' : Level.t) : val v l = val v l' + -> satisfies0 v (l, ConstraintType.Eq, l'). + +Definition satisfies v : ConstraintSet.t -> Prop := + ConstraintSet.For_all (satisfies0 v). + +Lemma satisfies_union v φ1 φ2 : + satisfies v (CS.union φ1 φ2) + <-> (satisfies v φ1 /\ satisfies v φ2). +Proof using Type. + unfold satisfies. split. + - intros H; split; intros c Hc; apply H; now apply CS.union_spec. + - intros [H1 H2] c Hc; apply CS.union_spec in Hc; destruct Hc; auto. Qed. -Lemma eqb_true_iff u v : - Universe.eqb u v <-> u = v. -Proof. - split. 2: intros ->; apply Universe.eqb_refl. - intro H. - destruct u, v;auto;try discriminate. - apply f_equal. now apply univ_expr_eqb_true_iff. +Lemma satisfies_subset φ φ' val : + ConstraintSet.Subset φ φ' -> + satisfies val φ' -> + satisfies val φ. +Proof using Type. + intros sub sat ? isin. + apply sat, sub; auto. Qed. -Lemma sup_comm x1 x2 : - Universe.sup x1 x2 = Universe.sup x2 x1. -Proof. - destruct x1,x2;auto. - cbn;apply f_equal;apply sup0_comm. +Definition consistent ctrs := exists v, satisfies v ctrs. + +Definition consistent_extension_on cs cstr := + forall v, satisfies v (ContextSet.constraints cs) -> exists v', + satisfies v' cstr /\ + LevelSet.For_all (fun l => val v l = val v' l) (ContextSet.levels cs). + +Lemma consistent_extension_on_empty Σ : + consistent_extension_on Σ CS.empty. +Proof using Type. + move=> v hv; exists v; split; [move=> ? /CS.empty_spec[]| move=> ??//]. Qed. -Lemma is_not_prop_and_is_not_sprop u : - Universe.is_prop u = false -> Universe.is_sprop u = false -> - ∑ u', u = Universe.lType u'. -Proof. - intros Hp Hsp. - destruct u; try discriminate. now eexists. +Lemma consistent_extension_on_union X cstrs + (wfX : forall c, CS.In c X.2 -> LS.In c.1.1 X.1 /\ LS.In c.2 X.1) : + consistent_extension_on X cstrs <-> + consistent_extension_on X (CS.union cstrs X.2). +Proof using Type. + split. + 2: move=> h v /h [v' [/satisfies_union [??] eqv']]; exists v'; split=> //. + move=> hext v /[dup] vsat /hext [v' [v'sat v'eq]]. + exists v'; split=> //. + apply/satisfies_union; split=> //. + move=> c hc. destruct (wfX c hc). + destruct (vsat c hc); constructor; rewrite -!v'eq //. Qed. -Lemma is_prop_sort_sup x1 x2 : - Universe.is_prop (Universe.sup x1 x2) - -> Universe.is_prop x2 \/ Universe.is_sprop x2 . -Proof. destruct x1,x2;auto. Qed. -Lemma is_prop_sort_sup' x1 x2 : - Universe.is_prop (Universe.sup x1 x2) - -> Universe.is_prop x1 \/ Universe.is_sprop x1 . -Proof. destruct x1,x2;auto. Qed. +Definition leq0_universe_n n φ (u u' : Universe.t) := + forall v, satisfies v φ -> (Z.of_nat (val v u) <= Z.of_nat (val v u') - n)%Z. -Lemma is_prop_or_sprop_sort_sup x1 x2 : - Universe.is_sprop (Universe.sup x1 x2) -> Universe.is_sprop x2. -Proof. destruct x1,x2;auto. Qed. +Definition leq_universe_n {cf} n φ (u u' : Universe.t) := + if check_univs then leq0_universe_n n φ u u' else True. -Lemma is_prop_sort_sup_prop x1 x2 : - Universe.is_prop x1 && Universe.is_prop x2 -> - Universe.is_prop (Universe.sup x1 x2). -Proof. destruct x1,x2;auto. Qed. +Definition lt_universe {cf} := leq_universe_n 1. +Definition leq_universe {cf} := leq_universe_n 0. -Lemma is_prop_or_sprop_sort_sup_prop x1 x2 : - Universe.is_sprop x1 && Universe.is_sprop x2 -> - Universe.is_sprop (Universe.sup x1 x2). -Proof. destruct x1,x2;auto. Qed. +Definition eq0_universe φ (u u' : Universe.t) := + forall v, satisfies v φ -> val v u = val v u'. -Lemma is_prop_sup u s : - Universe.is_prop (Universe.sup u s) -> - (Universe.is_prop u \/ Universe.is_sprop u) /\ - (Universe.is_prop s \/ Universe.is_sprop s). -Proof. destruct u,s;auto. Qed. +Definition eq_universe {cf} φ (u u' : Universe.t) := + if check_univs then eq0_universe φ u u' else True. -Lemma is_sprop_sup_iff u s : - Universe.is_sprop (Universe.sup u s) <-> - (Universe.is_sprop u /\ Universe.is_sprop s). -Proof. split;destruct u,s;intuition. Qed. +(* ctrs are "enforced" by φ *) -Lemma is_type_sup_r s1 s2 : - Universe.is_type_sort s2 -> - Universe.is_type_sort (Universe.sup s1 s2). -Proof. destruct s2; try absurd; destruct s1; cbnr; intros; absurd. Qed. +Definition valid_constraints0 φ ctrs + := forall v, satisfies v φ -> satisfies v ctrs. -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. - destruct x3;cbn;auto. - intros;simpl in *;destruct x2;auto. -Qed. +Definition valid_constraints {cf} φ ctrs + := if check_univs then valid_constraints0 φ ctrs else True. -Lemma is_sprop_sort_prod x2 x3 : - Universe.is_sprop (Universe.sort_of_product x2 x3) - -> Universe.is_sprop x3. -Proof. - unfold Universe.sort_of_product. - destruct x3;cbn;auto. - intros;simpl in *;destruct x2;auto. -Qed. +Definition compare_universe {cf} φ (pb : conv_pb) := + match pb with + | Conv => eq_universe φ + | Cumul => leq_universe φ + end. -Module ConstraintType. - Inductive t_ : Set := Le (z : Z) | Eq. - Derive NoConfusion EqDec for t_. - Definition t := t_. - Definition eq : t -> t -> Prop := eq. - Definition eq_equiv : Equivalence eq := _. +Ltac unfold_univ_rel0 := + unfold eq0_universe, leq0_universe_n, valid_constraints0 in *; + try ( + match goal with |- forall v : valuation, _ -> _ => idtac end; + intros v Hv; + repeat match goal with H : forall v : valuation, _ -> _ |- _ => specialize (H v Hv) end; + cbnr + ). - Definition Le0 := Le 0. - Definition Lt := Le 1. +Ltac unfold_univ_rel := + unfold eq_universe, leq_universe, lt_universe, leq_universe_n, valid_constraints in *; + destruct check_univs; [unfold_univ_rel0 | trivial]. - Inductive lt_ : t -> t -> Prop := - | LeLe n m : (n < m)%Z -> lt_ (Le n) (Le m) - | LeEq n : lt_ (Le n) Eq. - Derive Signature for lt_. - Definition lt := lt_. +Section Univ. + Context {cf}. - Global Instance lt_strorder : StrictOrder lt. - Proof. - constructor. - - intros []; intro X; inversion X. lia. - - intros ? ? ? X Y; invs X; invs Y; constructor. lia. + Lemma valid_subset φ φ' ctrs + : ConstraintSet.Subset φ φ' -> valid_constraints φ ctrs + -> valid_constraints φ' ctrs. + Proof using Type. + unfold_univ_rel. + intros Hφ H v Hv. apply H. + intros ctr Hc. apply Hv. now apply Hφ. Qed. - Global Instance lt_compat : Proper (eq ==> eq ==> iff) lt. - Proof. - intros ? ? X ? ? Y; invs X; invs Y. reflexivity. + Lemma switch_minus (x y z : Z) : (x <= y - z <-> x + z <= y)%Z. + Proof using Type. split; lia. Qed. + + (** **** Lemmas about eq and leq **** *) + + Global Instance eq_universe_refl φ : Reflexive (eq_universe φ). + Proof using Type. + intros u; unfold_univ_rel. Qed. - Definition compare (x y : t) : comparison := - match x, y with - | Le n, Le m => Z.compare n m - | Le _, Eq => Datatypes.Lt - | Eq, Eq => Datatypes.Eq - | Eq, _ => Datatypes.Gt - end. + Global Instance leq_universe_refl φ : Reflexive (leq_universe φ). + Proof using Type. + intros u; unfold_univ_rel. lia. + Qed. - Lemma compare_spec x y : CompareSpec (eq x y) (lt x y) (lt y x) (compare x y). - Proof. - destruct x, y; repeat constructor. simpl. - destruct (Z.compare_spec z z0); simpl; constructor. - subst; constructor. now constructor. now constructor. + Global Instance eq_universe_sym φ : Symmetric (eq_universe φ). + Proof using Type. + intros u u' H; unfold_univ_rel. + lia. Qed. - Lemma eq_dec x y : {eq x y} + {~ eq x y}. - Proof. - unfold eq. decide equality. apply Z.eq_dec. + Global Instance eq_universe_trans φ : Transitive (eq_universe φ). + Proof using Type. + intros u u' u'' H1 H2; unfold_univ_rel. + lia. Qed. -End ConstraintType. -Module UnivConstraint. - Definition t : Set := Level.t * ConstraintType.t * Level.t. + Global Instance leq_universe_n_trans n φ : Transitive (leq_universe_n (Z.of_nat n) φ). + Proof using Type. + intros u u' u'' H1 H2; unfold_univ_rel. + lia. + Qed. - Definition eq : t -> t -> Prop := eq. - Definition eq_equiv : Equivalence eq := _. + Global Instance leq_universe_trans φ : Transitive (leq_universe φ). + Proof using Type. apply (leq_universe_n_trans 0). Qed. - Definition make l1 ct l2 : t := (l1, ct, l2). + Global Instance lt_universe_trans φ : Transitive (lt_universe φ). + Proof using Type. apply (leq_universe_n_trans 1). Qed. - 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 eq0_leq0_universe φ u u' : + eq0_universe φ u u' <-> leq0_universe_n 0 φ u u' /\ leq0_universe_n 0 φ u' u. + Proof using Type. + split. + - intros H. split; unfold_univ_rel0; lia. + - intros [H1 H2]. unfold_univ_rel0; lia. + Qed. - 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. + Lemma eq_universe_leq_universe φ u u' : + eq_universe φ u u' <-> leq_universe φ u u' /\ leq_universe φ u' u. + Proof using Type. + unfold_univ_rel => //. + apply eq0_leq0_universe. Qed. - Lemma lt_compat : Proper (eq ==> eq ==> iff) lt. - Proof. - intros ? ? X ? ? Y; invs X; invs Y. reflexivity. + Lemma leq_universe_sup_l φ u1 u2 : leq_universe φ u1 (Universe.sup u1 u2). + Proof using Type. unfold_univ_rel. rewrite val_sup; lia. Qed. + + Lemma leq_universe_sup_r φ u1 u2 : leq_universe φ u2 (Universe.sup u1 u2). + Proof using Type. unfold_univ_rel. rewrite val_sup; lia. Qed. + + Lemma leq_universe_sup_mon φ u1 u1' u2 u2' : leq_universe φ u1 u1' -> leq_universe φ u2 u2' -> + leq_universe φ (Universe.sup u1 u2) (Universe.sup u1' u2'). + Proof using Type. + intros H1 H2; unfold_univ_rel. + rewrite !val_sup. lia. Qed. - Definition compare : t -> t -> comparison := - fun '(l1, t, l2) '(l1', t', l2') => - compare_cont (Level.compare l1 l1') - (compare_cont (ConstraintType.compare t t') - (Level.compare l2 l2')). + Global Instance eq_leq_universe φ : subrelation (eq_universe φ) (leq_universe φ). + Proof using Type. + intros u u'. apply eq_universe_leq_universe. + Qed. - Lemma compare_spec x y - : CompareSpec (eq x y) (lt x y) (lt y x) (compare x y). + Global Instance eq_universe_equivalence φ : Equivalence (eq_universe φ) := Build_Equivalence _ _ _ _. + + Global Instance leq_universe_preorder φ : PreOrder (leq_universe φ) := Build_PreOrder _ _ _. + + Global Instance lt_universe_irrefl {c: check_univs} φ (H: consistent φ) : Irreflexive (lt_universe φ). + Proof using Type. + intro u. unfold complement. + unfold_univ_rel => //. + destruct H as [v Hv]; intros nH; specialize (nH v Hv); lia. + Qed. + + Global Instance lt_universe_str_order {c: check_univs} φ (H: consistent φ) : StrictOrder (lt_universe φ). 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. + refine (Build_StrictOrder _ _ _). + now unshelve eapply lt_universe_irrefl. Qed. - Lemma eq_dec x y : {eq x y} + {~ eq x y}. + Global Instance leq_universe_antisym φ : Antisymmetric _ (eq_universe φ) (leq_universe φ). + Proof using Type. intros t u tu ut. now apply eq_universe_leq_universe. Qed. + + Global Instance leq_universe_partial_order φ + : PartialOrder (eq_universe φ) (leq_universe φ). Proof. - unfold eq. decide equality; apply eq_dec. + intros x y; split; apply eq_universe_leq_universe. Defined. - Definition eq_leibniz (x y : t) : eq x y -> x = y := id. -End UnivConstraint. -Module ConstraintSet := MSetAVL.Make UnivConstraint. -Module ConstraintSetFact := WFactsOn UnivConstraint ConstraintSet. -Module ConstraintSetOrdProp := MSetProperties.OrdProperties ConstraintSet. -Module ConstraintSetProp := ConstraintSetOrdProp.P. -Module CS := ConstraintSet. -Module ConstraintSetDecide := ConstraintSetProp.Dec. -Module ConstraintSetExtraOrdProp := MSets.ExtraOrdProperties ConstraintSet ConstraintSetOrdProp. -Module ConstraintSetExtraDecide := MSetAVL.Decide UnivConstraint ConstraintSet. -Ltac csets := ConstraintSetDecide.fsetdec. + Global Instance compare_universe_subrel φ pb : subrelation (eq_universe φ) (compare_universe φ pb). + Proof using Type. + destruct pb; tc. + Qed. -Notation "(=_cset)" := ConstraintSet.Equal (at level 0). -Infix "=_cset" := ConstraintSet.Equal (at level 30). -Notation "(==_cset)" := ConstraintSet.equal (at level 0). -Infix "==_cset" := ConstraintSet.equal (at level 30). + Global Instance compare_universe_refl φ pb : Reflexive (compare_universe φ pb). + Proof using Type. + destruct pb; tc. + Qed. -Definition declared_cstr_levels levels (cstr : UnivConstraint.t) := - let '(l1,_,l2) := cstr in - LevelSet.In l1 levels /\ LevelSet.In l2 levels. + Global Instance compare_universe_trans φ pb : Transitive (compare_universe φ pb). + Proof using Type. + destruct pb; tc. + Qed. -Definition is_declared_cstr_levels levels (cstr : UnivConstraint.t) : bool := - let '(l1,_,l2) := cstr in - LevelSet.mem l1 levels && LevelSet.mem l2 levels. + Global Instance compare_universe_preorder φ pb : PreOrder (compare_universe φ pb). + Proof using Type. + destruct pb; tc. + Qed. -Lemma CS_union_empty s : ConstraintSet.union ConstraintSet.empty s =_cset s. -Proof. - intros x; rewrite ConstraintSet.union_spec. lsets. -Qed. + Definition eq_leq_universe' φ u u' + := @eq_leq_universe φ u u'. + Definition leq_universe_refl' φ u + := @leq_universe_refl φ u. -Lemma CS_For_all_union f cst cst' : ConstraintSet.For_all f (ConstraintSet.union cst cst') -> - ConstraintSet.For_all f cst. -Proof. - unfold CS.For_all. - intros IH x inx. apply (IH x). - now eapply CS.union_spec; left. -Qed. + Hint Resolve eq_leq_universe' leq_universe_refl' : core. -Lemma CS_For_all_add P x s : CS.For_all P (CS.add x s) -> P x /\ CS.For_all P s. -Proof. - intros. - split. - * apply (H x), CS.add_spec; left => //. - * intros y iny. apply (H y), CS.add_spec; right => //. -Qed. -#[global] Instance CS_For_all_proper P : Morphisms.Proper ((=_cset) ==> iff)%signature (ConstraintSet.For_all P). -Proof. - intros s s' eqs. - unfold CS.For_all. split; intros IH x inxs; apply (IH x); - now apply eqs. -Qed. + Lemma cmp_universe_subset φ φ' pb t u : + ConstraintSet.Subset φ φ' -> compare_universe φ pb t u -> compare_universe φ' pb t u. + Proof using Type. + intros Hctrs. + destruct pb, t, u; cbnr; trivial. + all: intros H; unfold_univ_rel; + apply H. + all: eapply satisfies_subset; eauto. + Qed. -(** {6 Universe instances} *) + Lemma eq_universe_subset φ φ' t u + : ConstraintSet.Subset φ φ' + -> eq_universe φ t u -> eq_universe φ' t u. + Proof using Type. apply cmp_universe_subset with (pb := Conv). Qed. -Module Instance. + Lemma leq_universe_subset φ φ' t u + : ConstraintSet.Subset φ φ' + -> leq_universe φ t u -> leq_universe φ' t u. + Proof using Type. apply cmp_universe_subset with (pb := Cumul). Qed. - (** A universe instance represents a vector of argument concreteUniverses - to a polymorphic definition (constant, inductive or constructor). *) - Definition t : Set := list Level.t. +End Univ. - Definition empty : t := []. - Definition is_empty (i : t) : bool := - match i with - | [] => true - | _ => false +Module Sort. + Inductive t_ {univ} := + sProp | sSProp | sType (_ : univ). + Derive NoConfusion for t_. + Arguments t_ : clear implicits. + + Definition t := t_ Universe.t. + + Inductive family : Set := + | fSProp + | fProp + | fType. + Derive NoConfusion EqDec for family. + + Definition eqb {univ} `{ReflectEq univ} (u1 u2 : t_ univ) : bool := + match u1, u2 with + | sSProp, sSProp => true + | sProp, sProp => true + | sType e1, sType e2 => eqb e1 e2 + | _, _ => false end. - Definition eqb (i j : t) := - forallb2 Level.eqb i j. + #[global, program] Instance reflect_eq_sort {univ} `{ReflectEq univ} : ReflectEq (t_ univ) := + { eqb := eqb }. + Next Obligation. + destruct x, y; cbn; try constructor; auto; try congruence. + destruct (eqb_spec u u0); constructor. now f_equal. + congruence. + Qed. - Definition equal_upto (f : Level.t -> Level.t -> bool) (i j : t) := - forallb2 f i j. -End Instance. + #[global] Instance eq_dec_sort {univ} `{EqDec univ} : EqDec (t_ univ) := ltac:(intros s s'; decide equality). -Module UContext. - Definition t := list name × (Instance.t × ConstraintSet.t). + Definition map {u u'} (f : u -> u') s := + match s with + | sType u => sType (f u) + | sProp => sProp + | sSProp => sSProp + end. - Definition make' : Instance.t -> ConstraintSet.t -> Instance.t × ConstraintSet.t := pair. - Definition make (ids : list name) (inst_ctrs : Instance.t × ConstraintSet.t) : t := (ids, inst_ctrs). + Definition on_sort {univ} {T} (P: univ -> T) (def: T) (s : t_ univ) := + match s with + | sProp | sSProp => def + | sType l => P l + end. - Definition empty : t := ([], (Instance.empty, ConstraintSet.empty)). + (** Test if the universe is a lub of levels or contains +n's. *) + Definition is_levels (s : t) : bool := + match s with + | sSProp | sProp => true + | sType l => Universe.is_levels l + end. - Definition instance : t -> Instance.t := fun x => fst (snd x). - Definition constraints : t -> ConstraintSet.t := fun x => snd (snd x). + (** Test if the universe is a level or an algebraic universe. *) + Definition is_level (s : t) : bool := + match s with + | sSProp | sProp => true + | sType l => Universe.is_level l + end. - Definition dest : t -> list name * (Instance.t * ConstraintSet.t) := fun x => x. -End UContext. + Definition is_sprop {univ} (s : t_ univ) : bool := + match s with + | sSProp => true + | _ => false + end. -Module AUContext. - Definition t := list name × ConstraintSet.t. + Definition is_prop {univ} (s : t_ univ) : bool := + match s with + | sProp => true + | _ => false + end. - Definition make (ids : list name) (ctrs : ConstraintSet.t) : t := (ids, ctrs). - Definition repr (x : t) : UContext.t := - let (u, cst) := x in - (u, (mapi (fun i _ => Level.lvar i) u, cst)). + Definition is_propositional {univ} (s : t_ univ) : bool := + match s with + | sProp | sSProp => true + | _ => false + end. - Definition levels (uctx : t) : LevelSet.t := - LevelSetProp.of_list (fst (snd (repr uctx))). + Lemma is_prop_propositional {univ} (s : t_ univ) : + is_prop s -> is_propositional s. + Proof. now destruct s. Qed. + Lemma is_sprop_propositional {univ} (s : t_ univ) : + is_sprop s -> is_propositional s. + Proof. now destruct s. Qed. - #[local] - Existing Instance EqDec_ReflectEq. - Definition inter (au av : AUContext.t) : AUContext.t := - let prefix := (split_prefix au.1 av.1).1.1 in - let lvls := fold_left_i (fun s i _ => LevelSet.add (Level.lvar i) s) prefix LevelSet.empty in - let filter := ConstraintSet.filter (is_declared_cstr_levels lvls) in - make prefix (ConstraintSet.union (filter au.2) (filter av.2)). -End AUContext. + Definition is_type_sort {univ} (s : t_ univ) : bool := + match s with + | sType _ => true + | _ => false + end. -Module ContextSet. - Definition t := LevelSet.t × ConstraintSet.t. + Definition type0 : t := sType (Universe.make LevelExpr.set). + Definition type1 : t := sType (Universe.make LevelExpr.type1). - Definition levels : t -> LevelSet.t := fst. - Definition constraints : t -> ConstraintSet.t := snd. + Definition of_levels (l : PropLevel.t + Level.t) : t := + match l with + | inl PropLevel.lSProp => sSProp + | inl PropLevel.lProp => sProp + | inr l => sType (Universe.make' l) + end. - Definition empty : t := (LevelSet.empty, ConstraintSet.empty). + (** The universe strictly above FOR TYPING (not cumulativity) *) - Definition is_empty (uctx : t) - := LevelSet.is_empty (fst uctx) && ConstraintSet.is_empty (snd uctx). + Definition super_ {univ} type1 univ_succ (s : t_ univ) : t_ univ := + match s with + | sSProp | sProp => sType type1 + | sType l => sType (univ_succ l) + end. + Definition super : t -> t := super_ (Universe.make LevelExpr.type1) Universe.succ. + Definition csuper := super_ 1 S. - Definition Equal (x y : t) : Prop := - x.1 =_lset y.1 /\ x.2 =_cset y.2. + Definition sup_ {univ} univ_sup (s s' : t_ univ) : t_ univ := + match s, s' with + | sSProp, s' => s' + | sProp, sSProp => sProp + | sProp, s' => s' + | sType u1, sType u2 => sType (univ_sup u1 u2) + | sType _ as s, _ => s + end. + Definition sup : t -> t -> t := sup_ Universe.sup. + Definition csup := sup_ Nat.max. - Definition equal (x y : t) : bool := - x.1 ==_lset y.1 && x.2 ==_cset y.2. + (** Type of a product *) + Definition sort_of_product_ {univ} univ_sup (domsort rangsort : t_ univ) : t_ univ := + match rangsort with + | sSProp | sProp => rangsort + (* Prop and SProp impredicativity *) + | _ => Sort.sup_ univ_sup domsort rangsort + end. + Definition sort_of_product : t -> t -> t := sort_of_product_ Universe.sup. + Definition csort_of_product := sort_of_product_ Nat.max. - Definition Subset (x y : t) : Prop := - LevelSet.Subset (levels x) (levels y) /\ - ConstraintSet.Subset (constraints x) (constraints y). + Definition get_is_level (s : t) : option Level.t := + match s with + | sSProp => None + | sProp => None + | sType l => Universe.get_is_level l + end. - Definition subset (x y : t) : bool := - LevelSet.subset (levels x) (levels y) && - ConstraintSet.subset (constraints x) (constraints y). + Definition to_family {univ} (s : t_ univ) := + match s with + | sSProp => fSProp + | sProp => fProp + | sType _ => fType + end. - Definition inter (x y : t) : t := - (LevelSet.inter (levels x) (levels y), - ConstraintSet.inter (constraints x) (constraints y)). + Definition to_csort v s := + match s with + | sSProp => sSProp + | sProp => sProp + | sType u => sType (val v u) + end. - Definition inter_spec (x y : t) : - Subset (inter x y) x /\ - Subset (inter x y) y /\ - forall z, Subset z x -> Subset z y -> Subset z (inter x y). + Lemma to_family_to_csort s v : + to_family (to_csort v s) = to_family s. Proof. - split; last split. - 1,2: split=> ?; [move=> /LevelSet.inter_spec [//]|move=> /ConstraintSet.inter_spec [//]]. - move=> ? [??] [??]; split=> ??; - [apply/LevelSet.inter_spec|apply/ConstraintSet.inter_spec]; split; auto. + destruct s; cbnr. Qed. - Definition union (x y : t) : t := - (LevelSet.union (levels x) (levels y), ConstraintSet.union (constraints x) (constraints y)). + Lemma sType_super_ {univ type1 univ_succ} (s : t_ univ) : + to_family (super_ type1 univ_succ s) = fType. + Proof. now destruct s. Qed. - Definition union_spec (x y : t) : - Subset x (union x y) /\ - Subset y (union x y) /\ - forall z, Subset x z -> Subset y z -> Subset (union x y) z. - Proof. - split; last split. - 1,2: split=> ??; [apply/LevelSet.union_spec|apply/ConstraintSet.union_spec ]; by constructor. - move=> ? [??] [??]; split=> ?; - [move=>/LevelSet.union_spec|move=>/ConstraintSet.union_spec]=>-[]; auto. - Qed. + Lemma sType_super (s : t) : + to_family (super s) = fType. + Proof. apply sType_super_. Qed. - Lemma equal_spec s s' : equal s s' <-> Equal s s'. - Proof. - rewrite /equal/Equal/is_true Bool.andb_true_iff LevelSet.equal_spec ConstraintSet.equal_spec. - reflexivity. - Qed. + Inductive lt_ {univ univ_lt} : t_ univ -> t_ univ -> Prop := + | ltPropSProp : lt_ sProp sSProp + | ltPropType s : lt_ sProp (sType s) + | ltSPropType s : lt_ sSProp (sType s) + | ltTypeType s1 s2 : univ_lt s1 s2 -> lt_ (sType s1) (sType s2). + Derive Signature for lt_. + Arguments lt_ {univ} univ_lt. - Lemma subset_spec s s' : subset s s' <-> Subset s s'. - Proof. - rewrite /subset/Subset/is_true Bool.andb_true_iff LevelSet.subset_spec ConstraintSet.subset_spec. - reflexivity. - Qed. + Definition lt := lt_ Universe.lt. + Definition clt := lt_ Nat.lt. - Lemma subsetP s s' : reflect (Subset s s') (subset s s'). - Proof. - generalize (subset_spec s s'). - destruct subset; case; constructor; intuition. - Qed. -End ContextSet. + Module OT <: OrderedType. + Definition t := t. + #[local] Definition eq : t -> t -> Prop := eq. + #[local] Definition eq_equiv : Equivalence eq := _. + Definition lt := lt. + #[local] Instance lt_strorder : StrictOrder lt. + Proof. + constructor. + - intros [| |] X; inversion X. + now eapply irreflexivity in H1. + - intros [| |] [| |] [| |] X1 X2; + inversion X1; inversion X2; constructor. + subst. + etransitivity; tea. + Qed. -Export (hints) ContextSet. + Definition lt_compat : Proper (eq ==> eq ==> iff) lt. + Proof. + intros x y e z t e'. hnf in * |- ; subst. reflexivity. + Qed. + Definition compare (x y : t) : comparison + := match x, y with + | sProp, sProp => Eq + | sProp, _ => Lt + | _, sProp => Gt + | sSProp, sSProp => Eq + | sSProp, _ => Lt + | _, sSProp => Gt + | sType x, sType y => LevelExprSet.compare x y + end. + Lemma compare_spec x y : CompareSpec (eq x y) (lt x y) (lt y x) (compare x y). + Proof. + cbv [compare eq]. + destruct x, y. + all: lazymatch goal with + | [ |- context[LevelExprSet.compare ?x ?y] ] + => destruct (LevelExprSet.compare_spec x y) + | _ => idtac + end. + all: lazymatch goal with + | [ H : LevelExprSet.eq (?f ?x) (?f ?y) |- _ ] + => apply LevelExprSet.eq_leibniz in H; + is_var x; is_var y; destruct x, y; cbn in H; subst + | _ => idtac + end. + all: repeat constructor; try apply f_equal; try assumption. + f_equal; apply Eqdep_dec.UIP_dec; decide equality. + Qed. + Definition eq_dec (x y : t) : {x = y} + {x <> y}. + Proof. repeat decide equality. apply Universe.eq_dec_univ0. Defined. + End OT. + Module OTOrig <: OrderedTypeOrig := Backport_OT OT. +End Sort. -Notation "(=_cs)" := ContextSet.Equal (at level 0). -Notation "(⊂_cs)" := ContextSet.Subset (at level 0). -Infix "=_cs" := ContextSet.Equal (at level 30). -Infix "⊂_cs" := ContextSet.Subset (at level 30). -Notation "(==_cs)" := ContextSet.equal (at level 0). -Notation "(⊂?_cs)" := ContextSet.subset (at level 0). -Infix "==_cs" := ContextSet.equal (at level 30). -Infix "⊂?_cs" := ContextSet.subset (at level 30). -Lemma incl_cs_refl cs : cs ⊂_cs cs. +Module SortMap := FMapAVL.Make Sort.OTOrig. +Module SortMapFact := FMapFacts.WProperties SortMap. +Module SortMapExtraFact := FSets.WFactsExtra_fun Sort.OTOrig SortMap SortMapFact.F. +Module SortMapDecide := FMapAVL.Decide Sort.OTOrig SortMap. + +Notation sProp := Sort.sProp. +Notation sSProp := Sort.sSProp. +Notation sType := Sort.sType. +Notation sort := Sort.t. + +Notation "⟦ u ⟧_ v" := (Sort.to_csort v u) (at level 0, format "⟦ u ⟧_ v", v name) : univ_scope. + + +Lemma val_sort_sup v s1 s2 : + Sort.to_csort v (Sort.sup s1 s2) = Sort.csup (Sort.to_csort v s1) (Sort.to_csort v s2). Proof. - split; [lsets|csets]. + destruct s1 as [ | | u1]; destruct s2 as [ | | u2]; cbnr. + f_equal. apply val_sup. +Qed. + +Lemma is_prop_val s : + Sort.is_prop s -> forall v, Sort.to_csort v s = Sort.sProp. +Proof. + destruct s => //. +Qed. + +Lemma is_sprop_val s : + Sort.is_sprop s -> forall v, Sort.to_csort v s = Sort.sSProp. +Proof. + destruct s => //. +Qed. + + +Lemma val_is_prop s v : + Sort.to_csort v s = sProp <-> Sort.is_prop s. +Proof. + destruct s => //. +Qed. + +Lemma val_is_sprop s v : + Sort.to_csort v s = sSProp <-> Sort.is_sprop s. +Proof. + destruct s => //. +Qed. + +Lemma is_prop_and_is_sprop_val_false s : + Sort.is_prop s = false -> Sort.is_sprop s = false -> + forall v, ∑ n, Sort.to_csort v s = sType n. +Proof. + intros Hp Hsp v. + destruct s => //. simpl. eexists; eauto. +Qed. + +Lemma val_is_prop_false s v n : + Sort.to_csort v s = sType n -> Sort.is_prop s = false. +Proof. + destruct s => //. Qed. -Lemma incl_cs_trans cs1 cs2 cs3 : cs1 ⊂_cs cs2 -> cs2 ⊂_cs cs3 -> cs1 ⊂_cs cs3. +Lemma get_is_level_correct s l : + Sort.get_is_level s = Some l -> s = sType (Universe.make' l). Proof. - intros [? ?] [? ?]; split; [lsets|csets]. + intro H; destruct s => //=. + f_equal; now apply universe_get_is_level_correct. Qed. -Lemma empty_contextset_subset u : ContextSet.empty ⊂_cs u. +Lemma eqb_true_iff (s s' : sort) : + eqb s s' <-> s = s'. Proof. - red. split; cbn; [lsets|csets]. + split. apply /eqb_spec. eapply introP. apply /eqb_spec. Qed. -(* Variance info is needed to do full universe polymorphism *) -Module Variance. - (** A universe position in the instance given to a cumulative - inductive can be the following. Note there is no Contralvariant - case because [forall x : A, B <= forall x : A', B'] requires [A = - A'] as opposed to [A' <= A]. *) - Inductive t := - | Irrelevant : t - | Covariant : t - | Invariant : t. - Derive NoConfusion EqDec for t. - - (* val check_subtype : t -> t -> bool *) - (* val sup : t -> t -> t *) -End Variance. - -Inductive universes_decl : Type := -| Monomorphic_ctx -| Polymorphic_ctx (cst : AUContext.t). - -Derive NoConfusion for universes_decl. +Lemma sup_comm x1 x2 : + Sort.sup x1 x2 = Sort.sup x2 x1. +Proof. + destruct x1,x2;auto. + cbn;apply f_equal;apply sup0_comm. +Qed. -Definition levels_of_udecl u := - match u with - | Monomorphic_ctx => LevelSet.empty - | Polymorphic_ctx ctx => AUContext.levels ctx - end. +Lemma is_not_prop_and_is_not_sprop {univ} (s : Sort.t_ univ) : + Sort.is_prop s = false -> Sort.is_sprop s = false -> + ∑ s', s = sType s'. +Proof. + intros Hp Hsp. + destruct s => //. now eexists. +Qed. -Definition constraints_of_udecl u := - match u with - | Monomorphic_ctx => ConstraintSet.empty - | Polymorphic_ctx ctx => snd (snd (AUContext.repr ctx)) - end. +Lemma is_prop_sort_sup x1 x2 : + Sort.is_prop (Sort.sup x1 x2) + -> Sort.is_prop x2 \/ Sort.is_sprop x2 . +Proof. destruct x1,x2;auto. Qed. -Section Univ. - Context {cf: checker_flags}. +Lemma is_prop_sort_sup' x1 x2 : + Sort.is_prop (Sort.sup x1 x2) + -> Sort.is_prop x1 \/ Sort.is_sprop x1 . +Proof. destruct x1,x2;auto. Qed. - Inductive satisfies0 (v : valuation) : UnivConstraint.t -> Prop := - | satisfies0_Lt (l l' : Level.t) (z : Z) : (Z.of_nat (val v l) <= Z.of_nat (val v l') - z)%Z - -> satisfies0 v (l, ConstraintType.Le z, l') - | satisfies0_Eq (l l' : Level.t) : val v l = val v l' - -> satisfies0 v (l, ConstraintType.Eq, l'). +Lemma is_prop_or_sprop_sort_sup x1 x2 : + Sort.is_sprop (Sort.sup x1 x2) -> Sort.is_sprop x2. +Proof. destruct x1,x2;auto. Qed. - Definition satisfies v : ConstraintSet.t -> Prop := - ConstraintSet.For_all (satisfies0 v). +Lemma is_prop_sort_sup_prop x1 x2 : + Sort.is_prop x1 && Sort.is_prop x2 -> + Sort.is_prop (Sort.sup x1 x2). +Proof. destruct x1,x2;auto. Qed. - Lemma satisfies_union v φ1 φ2 : - satisfies v (CS.union φ1 φ2) - <-> (satisfies v φ1 /\ satisfies v φ2). - Proof using Type. - unfold satisfies. split. - - intros H; split; intros c Hc; apply H; now apply CS.union_spec. - - intros [H1 H2] c Hc; apply CS.union_spec in Hc; destruct Hc; auto. - Qed. +Lemma is_prop_or_sprop_sort_sup_prop x1 x2 : + Sort.is_sprop x1 && Sort.is_sprop x2 -> + Sort.is_sprop (Sort.sup x1 x2). +Proof. destruct x1,x2;auto. Qed. - Lemma satisfies_subset φ φ' val : - ConstraintSet.Subset φ φ' -> - satisfies val φ' -> - satisfies val φ. - Proof using Type. - intros sub sat ? isin. - apply sat, sub; auto. - Qed. +Lemma is_prop_sup s s' : + Sort.is_prop (Sort.sup s s') -> + Sort.is_propositional s /\ Sort.is_propositional s'. +Proof. destruct s, s'; auto. Qed. - Definition consistent ctrs := exists v, satisfies v ctrs. +Lemma is_sprop_sup_iff s s' : + Sort.is_sprop (Sort.sup s s') <-> + (Sort.is_sprop s /\ Sort.is_sprop s'). +Proof. split; destruct s, s' => //=; intuition. Qed. - Definition consistent_extension_on cs cstr := - forall v, satisfies v (ContextSet.constraints cs) -> exists v', - satisfies v' cstr /\ - LevelSet.For_all (fun l => val v l = val v' l) (ContextSet.levels cs). +Lemma is_type_sup_r s1 s2 : + Sort.is_type_sort s2 -> + Sort.is_type_sort (Sort.sup s1 s2). +Proof. destruct s2; try absurd; destruct s1; cbnr; intros; absurd. Qed. - Lemma consistent_extension_on_empty Σ : - consistent_extension_on Σ CS.empty. - Proof using Type. - move=> v hv; exists v; split; [move=> ? /CS.empty_spec[]| move=> ??//]. - Qed. +Lemma is_prop_sort_prod x2 x3 : + Sort.is_prop (Sort.sort_of_product x2 x3) + -> Sort.is_prop x3. +Proof. + unfold Sort.sort_of_product. + destruct x3;cbn;auto. + intros;simpl in *;destruct x2;auto. +Qed. - Lemma consistent_extension_on_union X cstrs - (wfX : forall c, CS.In c X.2 -> LS.In c.1.1 X.1 /\ LS.In c.2 X.1) : - consistent_extension_on X cstrs <-> - consistent_extension_on X (CS.union cstrs X.2). - Proof using Type. - split. - 2: move=> h v /h [v' [/satisfies_union [??] eqv']]; exists v'; split=> //. - move=> hext v /[dup] vsat /hext [v' [v'sat v'eq]]. - exists v'; split=> //. - apply/satisfies_union; split=> //. - move=> c hc. destruct (wfX c hc). - destruct (vsat c hc); constructor; rewrite -!v'eq //. - Qed. +Lemma is_sprop_sort_prod x2 x3 : + Sort.is_sprop (Sort.sort_of_product x2 x3) + -> Sort.is_sprop x3. +Proof. + unfold Sort.sort_of_product. + destruct x3;cbn;auto. + intros;simpl in *;destruct x2;auto. +Qed. - Definition leq0_levelalg_n n φ (u u' : LevelAlgExpr.t) := - forall v, satisfies v φ -> (Z.of_nat (val v u) <= Z.of_nat (val v u') - n)%Z. - Definition leq_levelalg_n n φ (u u' : LevelAlgExpr.t) := - if check_univs then leq0_levelalg_n n φ u u' else True. - Definition leq_universe_n_ {CS} leq_levelalg_n n (φ: CS) s s' := +Section SortCompare. + Context {cf}. + Definition leq_sort_n_ {univ} (leq_universe_n : Z -> univ -> univ -> Prop) n s s' : Prop := match s, s' with - | Universe.lProp, Universe.lProp - | Universe.lSProp, Universe.lSProp => (n = 0)%Z - | Universe.lType u, Universe.lType u' => leq_levelalg_n n φ u u' - | Universe.lProp, Universe.lType u => prop_sub_type + | sProp, sProp + | sSProp, sSProp => (n = 0)%Z + | sType u, sType u' => leq_universe_n n u u' + | sProp, sType u => prop_sub_type | _, _ => False end. - Definition leq_universe_n := leq_universe_n_ leq_levelalg_n. + Definition leq_sort_n n φ := leq_sort_n_ (fun n => leq_universe_n n φ) n. + Definition lt_sort := leq_sort_n 1. + Definition leq_sort := leq_sort_n 0. - Definition leqb_universe_n_ leqb_levelalg_n b s s' := + Definition leqb_sort_n_ {univ} (leqb_universe_n : bool -> univ -> univ -> bool) b s s' : bool := match s, s' with - | Universe.lProp, Universe.lProp - | Universe.lSProp, Universe.lSProp => negb b - | Universe.lType u, Universe.lType u' => leqb_levelalg_n b u u' - | Universe.lProp, Universe.lType u => prop_sub_type + | sProp, sProp + | sSProp, sSProp => negb b + | sType u, sType u' => leqb_universe_n b u u' + | sProp, sType u => prop_sub_type | _, _ => false end. - Definition eq0_levelalg φ (u u' : LevelAlgExpr.t) := - forall v, satisfies v φ -> val v u = val v u'. - - Definition eq_levelalg φ (u u' : LevelAlgExpr.t) := - if check_univs then eq0_levelalg φ u u' else True. - - Definition eq_universe_ {CS} eq_levelalg (φ: CS) s s' := + Definition eq_sort_ {univ} (eq_universe : univ -> univ -> Prop) s s' : Prop := match s, s' with - | Universe.lProp, Universe.lProp - | Universe.lSProp, Universe.lSProp => True - | Universe.lType u, Universe.lType u' => eq_levelalg φ u u' + | sProp, sProp + | sSProp, sSProp => True + | sType u, sType u' => eq_universe u u' | _, _ => False end. - Definition eq_universe := eq_universe_ eq_levelalg. + Definition eq_sort φ := eq_sort_ (eq_universe φ). - Definition lt_levelalg := leq_levelalg_n 1. - Definition leq_levelalg := leq_levelalg_n 0. - Definition lt_universe := leq_universe_n 1. - Definition leq_universe := leq_universe_n 0. + Definition eqb_sort_ {univ} (eqb_universe : univ -> univ -> bool) s s' : bool := + match s, s' with + | sProp, sProp + | sSProp, sSProp => true + | sType u, sType u' => eqb_universe u u' + | _, _ => false + end. - Definition compare_universe (pb : conv_pb) := + Definition compare_sort φ (pb : conv_pb) := match pb with - | Conv => eq_universe - | Cumul => leq_universe + | Conv => eq_sort φ + | Cumul => leq_sort φ end. - Lemma leq_levelalg_leq_levelalg_n (φ : ConstraintSet.t) u u' : - leq_levelalg φ u u' <-> leq_levelalg_n 0 φ u u'. + Lemma leq_sort_leq_sort_n (φ : ConstraintSet.t) s s' : + leq_sort φ s s' <-> leq_sort_n 0 φ s s'. Proof using Type. intros. reflexivity. Qed. - Lemma leq_universe_leq_universe_n (φ : ConstraintSet.t) u u' : - leq_universe φ u u' <-> leq_universe_n 0 φ u u'. - Proof using Type. intros. reflexivity. Qed. + Lemma compare_sort_type φ pb u u' : + compare_sort φ pb (sType u) (sType u') = compare_universe φ pb u u'. + Proof. now destruct pb. Qed. - (* ctrs are "enforced" by φ *) + Section GeneralLemmas. + Context {univ} {leq_universe_n : Z -> univ -> univ -> Prop} {eq_universe : univ -> univ -> Prop}. - Definition valid_constraints0 φ ctrs - := forall v, satisfies v φ -> satisfies v ctrs. + Let leq_sort_n := leq_sort_n_ leq_universe_n. + Let lt_sort := leq_sort_n_ leq_universe_n 1. + Let leq_sort := leq_sort_n_ leq_universe_n 0. + Let eq_sort := eq_sort_ eq_universe. + Notation "x <_ n y" := (leq_sort_n n x y) (at level 10, n name). + Notation "x < y" := (lt_sort x y). + Notation "x <= y" := (leq_sort x y). - Definition valid_constraints φ ctrs - := if check_univs then valid_constraints0 φ ctrs else True. - Lemma valid_subset φ φ' ctrs - : ConstraintSet.Subset φ φ' -> valid_constraints φ ctrs - -> valid_constraints φ' ctrs. - Proof using Type. - unfold valid_constraints. - destruct check_univs; [|trivial]. - intros Hφ H v Hv. apply H. - intros ctr Hc. apply Hv. now apply Hφ. - Qed. + Lemma sort_le_prop_inv s : s <= sProp -> s = sProp. + Proof using Type. destruct s => //. Qed. - Lemma switch_minus (x y z : Z) : (x <= y - z <-> x + z <= y)%Z. - Proof using Type. split; lia. Qed. + Lemma sort_le_sprop_inv s : s <= sSProp -> s = sSProp. + Proof using Type. destruct s => //. Qed. - (* Lemma llt_lt n m : (n < m)%u -> (n < m)%Z. - Proof. lled; lia. Qed. + Lemma sort_prop_le_inv s : sProp <= s -> + (s = sProp \/ (prop_sub_type /\ exists n, s = sType n)). + Proof using Type. + destruct s => //= Hle. + - now left. + - right; split => //; now eexists. + Qed. - Lemma lle_le n m : (n <= m)%u -> (n <= m)%Z. - Proof. lled; lia. Qed. + Lemma sort_sprop_le_inv s : sSProp <= s -> s = sSProp. + Proof using Type. destruct s => //. Qed. - Lemma lt_llt n m : prop_sub_type -> (n < m)%Z -> (n < m)%u. - Proof. unfold llt. now intros ->. Qed. + Global Instance leq_sort_refl `{Reflexive univ (leq_universe_n 0)} : Reflexive leq_sort. + Proof using Type. intros []; cbnr. Qed. - Lemma le_lle n m : prop_sub_type -> (n <= m)%Z -> (n <= m)%u. - Proof. lled; [lia|discriminate]. Qed. + Global Instance eq_sort_refl `{Reflexive univ eq_universe} : Reflexive eq_sort. + Proof using Type. intros []; cbnr. Qed. - Lemma lt_llt' n m : (0 <= n)%Z -> (n < m)%Z -> (n < m)%u. - Proof. lled; lia. Qed. + Global Instance eq_sort_sym `{Symmetric univ eq_universe} : Symmetric eq_sort. + Proof using Type. intros [] [] => //=. apply H. Qed. - Lemma le_lle' n m : (0 <= n)%Z -> (n <= m)%Z -> (n <= m)%u. - Proof. lled; lia. Qed. *) + Global Instance leq_sort_n_trans n `{Transitive univ (leq_universe_n n)} : Transitive (leq_sort_n n). + Proof using Type. + intros [] [] [] => //=. apply H. + Qed. + Global Instance leq_sort_trans `{Transitive univ (leq_universe_n 0)} : Transitive leq_sort. + Proof using Type. apply (leq_sort_n_trans 0). Qed. - (** **** Lemmas about eq and leq **** *) + Global Instance lt_sort_trans `{Transitive univ (leq_universe_n 1)} : Transitive lt_sort. + Proof using Type. apply (leq_sort_n_trans 1). Qed. - Ltac unfold_univ_rel0 := - unfold eq0_levelalg, leq0_levelalg_n in *; - intros v Hv; cbnr. + Global Instance eq_sort_trans `{Transitive univ eq_universe} : Transitive eq_sort. + Proof using Type. + intros [] [] [] => //=. apply H. + Qed. - Ltac unfold_univ_rel := - unfold eq_levelalg, leq_levelalg, lt_levelalg, leq_levelalg_n in *; - destruct check_univs; [unfold_univ_rel0 | trivial]. + Global Instance leq_sort_preorder `{PreOrder univ (leq_universe_n 0)} : PreOrder leq_sort := + Build_PreOrder _ _ _. + (* Can't be a global instance since it can lead to infinite search *) + Lemma lt_sort_irrefl : Irreflexive (leq_universe_n 1) -> Irreflexive lt_sort. + Proof using Type. + intros H []; unfold complement; cbnr. 1,2: lia. apply H. + Qed. - Global Instance eq_levelalg_refl φ : Reflexive (eq_levelalg φ). - Proof using Type. - intros u; unfold_univ_rel. - Qed. + Global Instance lt_sort_str_order `{StrictOrder univ (leq_universe_n 1)} : StrictOrder lt_sort := + Build_StrictOrder _ (lt_sort_irrefl _) _. - Global Instance eq_universe_refl φ : Reflexive (eq_universe φ). - Proof using Type. - intros []; cbnr. - Qed. + Global Instance eq_leq_sort `{subrelation univ eq_universe (leq_universe_n 0)}: subrelation eq_sort leq_sort. + Proof using Type. + intros [] [] => //=. apply H. + Qed. - Global Instance leq_levelalg_n_refl φ : Reflexive (leq_levelalg φ). - Proof using Type. - intros u; unfold_univ_rel. lia. - Qed. + Global Instance eq_sort_equivalence `{Equivalence univ eq_universe} : Equivalence eq_sort := Build_Equivalence _ _ _ _. - Global Instance leq_universe_refl φ : Reflexive (leq_universe φ). - Proof using Type. - intros []; cbnr. - Qed. + Global Instance leq_sort_antisym `{Antisymmetric _ eq_universe (leq_universe_n 0)} : Antisymmetric _ eq_sort leq_sort. + Proof using Type. + intros [] [] => //=. apply H. + Qed. - Global Instance eq_levelalg_sym φ : Symmetric (eq_levelalg φ). - Proof using Type. - intros u u' H; unfold_univ_rel. - specialize (H v Hv); lia. - Qed. + Global Instance leq_sort_partial_order `{PartialOrder _ eq_universe (leq_universe_n 0)}: PartialOrder eq_sort leq_sort. + Proof. + assert (subrelation eq_universe (leq_universe_n 0)). + { intros u u' Hu. specialize (H u u'); cbn in H. apply H in Hu. apply Hu. } + assert (subrelation eq_universe (flip (leq_universe_n 0))). + { intros u u' Hu. specialize (H u u'); cbn in H. apply H in Hu. apply Hu. } + intros s s'. split. + - intro Heq. split. + + now eapply eq_leq_sort. + + now eapply eq_leq_sort. + - intros [Hle Hge]. now eapply leq_sort_antisym. + Qed. - Global Instance eq_universe_sym φ : Symmetric (eq_universe φ). - Proof using Type. - intros [] []; cbnr; auto. - now symmetry. - Qed. + End GeneralLemmas. - Global Instance eq_levelalg_trans φ : Transitive (eq_levelalg φ). - Proof using Type. - intros u u' u'' H1 H2; unfold_univ_rel. - specialize (H1 v Hv); specialize (H2 v Hv); lia. - Qed. - Global Instance eq_universe_trans φ : Transitive (eq_universe φ). - Proof using Type. - intros [] [] []; cbnr; trivial; try absurd. - etransitivity; eauto. - Qed. + (** Universes with linear hierarchy *) + Definition concrete_sort := Sort.t_ nat. - Global Instance leq_levelalg_n_trans n φ : Transitive (leq_levelalg_n (Z.of_nat n) φ). - Proof using Type. - intros u u' u'' H1 H2; unfold_univ_rel. - specialize (H1 v Hv); specialize (H2 v Hv); lia. - Qed. + (** u + n <= u' *) + Definition leq_csort_n : Z -> concrete_sort -> concrete_sort -> Prop := + leq_sort_n_ (fun n u u' => (Z.of_nat u <= Z.of_nat u' - n)%Z). - Global Instance leq_universe_n_trans n φ : Transitive (leq_universe_n (Z.of_nat n) φ). - Proof using Type. - intros [] [] []; cbnr; trivial; try absurd. - now etransitivity. - Qed. + Definition leq_csort := leq_csort_n 0. + Definition lt_csort := leq_csort_n 1. - Global Instance leq_universe_trans φ : Transitive (leq_universe φ). - Proof using Type. apply (leq_universe_n_trans 0). Qed. + Notation "x <_ n y" := (leq_csort_n n x y) (at level 10, n name) : univ_scope. + Notation "x < y" := (lt_csort x y) : univ_scope. + Notation "x <= y" := (leq_csort x y) : univ_scope. - Global Instance lt_universe_trans φ : Transitive (lt_universe φ). - Proof using Type. apply (leq_universe_n_trans 1). Qed. + Definition is_propositional_or_set s := match s with sSProp | sProp | sType 0 => true | _ => false end. - Lemma eq0_leq0_levelalg φ u u' : - eq0_levelalg φ u u' <-> leq0_levelalg_n 0 φ u u' /\ leq0_levelalg_n 0 φ u' u. - Proof using Type. - split. - - intros H. split; unfold eq0_levelalg, leq_levelalg_n in *; - intros v Hv; specialize (H v Hv); lia. - - intros [H1 H2]. unfold eq0_levelalg, leq_levelalg_n in *. - intros v Hv. specialize (H1 v Hv); specialize (H2 v Hv); lia. - Qed. + Lemma csort_sup_comm s s' : Sort.csup s s' = Sort.csup s' s. + Proof using Type. destruct s, s' => //=. f_equal; lia. Qed. - Lemma eq_leq_levelalg φ u u' : - eq_levelalg φ u u' <-> leq_levelalg φ u u' /\ leq_levelalg φ u' u. + Lemma csort_sup_not_uproplevel s s' : + ~ Sort.is_propositional s -> ∑ n, Sort.csup s s' = sType n. Proof using Type. - split. - - intros H. split; unfold_univ_rel; specialize (H v Hv); lia. - - intros [H1 H2]. unfold_univ_rel. specialize (H1 v Hv); specialize (H2 v Hv); lia. + destruct s => //=. + destruct s'; now eexists. Qed. - Lemma eq_leq_universe φ u u' : - eq_universe φ u u' <-> leq_universe φ u u' /\ leq_universe φ u' u. + Lemma csort_sup_mon s s' v v' : (s <= s')%u -> (sType v <= sType v')%u -> + (Sort.csup s (sType v) <= Sort.csup s' (sType v'))%u. Proof using Type. - destruct u, u'; cbnr; intuition auto. - all: now apply eq_leq_levelalg. + destruct s, s' => //=; intros Hle Hle'; + lia. Qed. - Lemma leq_levelalg_sup0_l φ u1 u2 : leq_levelalg φ u1 (LevelAlgExpr.sup u1 u2). - Proof using Type. unfold_univ_rel. rewrite val_sup; lia. Qed. - - Lemma leq_levelalg_sup0_r φ u1 u2 : leq_levelalg φ u2 (LevelAlgExpr.sup u1 u2). - Proof using Type. unfold_univ_rel. rewrite val_sup; lia. Qed. - - Lemma leq_levelalg_sup0_mon φ u1 u1' u2 u2' : leq_levelalg φ u1 u1' -> leq_levelalg φ u2 u2' -> - leq_levelalg φ (LevelAlgExpr.sup u1 u2) (LevelAlgExpr.sup u1' u2'). + Lemma leq_csort_of_product_mon u u' v v' : + (u <= u')%u -> + (v <= v')%u -> + (Sort.csort_of_product u v <= Sort.csort_of_product u' v')%u. Proof using Type. - intros H1 H2; unfold_univ_rel. - specialize (H1 v Hv); specialize (H2 v Hv). - rewrite !val_sup. lia. + intros Hle1 Hle2. + destruct v, v'; cbn in Hle2 |- *; auto. + - destruct u'; cbn; assumption. + - apply csort_sup_mon; assumption. Qed. - Lemma leq_universe_sup_l φ u1 s2 : - let s1 := Universe.lType u1 in - leq_universe φ s1 (Universe.sup s1 s2). + Lemma impredicative_csort_product {univ} {univ_sup} {l u} : + Sort.is_propositional u -> + Sort.sort_of_product_ (univ := univ) univ_sup l u = u. + Proof using Type. now destruct u. Qed. + + Lemma leq_sort_sup_l φ u1 s2 : + let s1 := sType u1 in + leq_sort φ s1 (Sort.sup s1 s2). Proof using Type. destruct s2 as [| | u2]; cbnr. - apply leq_levelalg_sup0_l. + apply leq_universe_sup_l. Qed. - Lemma leq_universe_sup_r φ s1 u2 : - let s2 := Universe.lType u2 in - leq_universe φ s2 (Universe.sup s1 s2). + Lemma leq_sort_sup_r φ s1 u2 : + let s2 := sType u2 in + leq_sort φ s2 (Sort.sup s1 s2). Proof using Type. destruct s1 as [| | u1]; cbnr. - apply leq_levelalg_sup0_r. + apply leq_universe_sup_r. Qed. - Lemma leq_universe_product φ (s1 s2 : Universe.t) - : leq_universe φ s2 (Universe.sort_of_product s1 s2). + Lemma leq_sort_product φ (s1 s2 : Sort.t) + : leq_sort φ s2 (Sort.sort_of_product s1 s2). Proof using Type. - destruct s2 as [| | u2]. - - apply leq_universe_refl. - - apply leq_universe_refl. - - apply leq_universe_sup_r. + destruct s2 as [| | u2] => //. + apply leq_sort_sup_r. Qed. (* 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 using Type. - intros u u'. apply eq_leq_universe. - Qed. - - Global Instance eq_levelalg_equivalence φ : Equivalence (eq_levelalg φ) := Build_Equivalence _ _ _ _. - - Global Instance leq_levelalg_preorder φ : PreOrder (leq_levelalg φ) := - {| PreOrder_Transitive := leq_levelalg_n_trans 0 _ |}. + impredicativity. *) - Global Instance lt_levelalg_str_order {c: check_univs} φ (H: consistent φ) : StrictOrder (lt_levelalg φ). - Proof using Type. - split. - - intros u; unfold complement, lt_levelalg, leq_levelalg_n; cbn. - rewrite c; destruct H as [v Hv]; intros nH; specialize (nH v Hv); lia. - - apply (leq_levelalg_n_trans 1). - Qed. - - Global Instance leq_levelalg_antisym φ - : Antisymmetric _ (eq_levelalg φ) (leq_levelalg φ). - Proof using Type. intros t u tu ut. now apply eq_leq_levelalg. Qed. - Global Instance leq_levelalg_partial_order φ - : PartialOrder (eq_levelalg φ) (leq_levelalg φ). + Global Instance lt_sort_irrefl' {c: check_univs} φ (H: consistent φ) : Irreflexive (lt_sort φ). Proof. - intros x y; split; apply eq_leq_levelalg. - Defined. - - - Global Instance eq_universe_equivalence φ : Equivalence (eq_universe φ) := Build_Equivalence _ _ _ _. - - Global Instance leq_universe_preorder φ : PreOrder (leq_universe φ) := Build_PreOrder _ _ _. + unshelve eapply lt_sort_irrefl. + now unshelve eapply lt_universe_irrefl. + Qed. - Global Instance lt_universe_str_order {c: check_univs} φ (H: consistent φ) : StrictOrder (lt_universe φ). + Global Instance lt_sort_str_order' {c: check_univs} φ (H: consistent φ) : StrictOrder (lt_sort φ). Proof using Type. - split. - - intros []; unfold complement; cbn; [lia|lia|]. - apply @StrictOrder_Irreflexive; apply @lt_levelalg_str_order; assumption. - - exact _. + unshelve eapply lt_sort_str_order. + now unshelve eapply lt_universe_str_order. Qed. - Global Instance leq_universe_antisym φ - : Antisymmetric _ (eq_universe φ) (leq_universe φ). - Proof using Type. intros t u tu ut. now apply eq_leq_universe. Qed. - - Global Instance leq_universe_partial_order φ - : PartialOrder (eq_universe φ) (leq_universe φ). - Proof. - intros x y; split; apply eq_leq_universe. - Defined. - - Global Instance compare_universe_subrel pb Σ : subrelation (eq_universe Σ) (compare_universe pb Σ). + Global Instance compare_sort_subrel φ pb : subrelation (eq_sort φ) (compare_sort φ pb). Proof using Type. destruct pb; tc. Qed. - Global Instance compare_universe_refl pb Σ : Reflexive (compare_universe pb Σ). + Global Instance compare_sort_refl φ pb : Reflexive (compare_sort φ pb). Proof using Type. destruct pb; tc. Qed. - Global Instance compare_universe_trans pb Σ : Transitive (compare_universe pb Σ). + Global Instance compare_sort_trans φ pb : Transitive (compare_sort φ pb). Proof using Type. destruct pb; tc. Qed. - Global Instance compare_universe_preorder pb Σ : PreOrder (compare_universe pb Σ). + Global Instance compare_sort_preorder φ pb : PreOrder (compare_sort φ pb). Proof using Type. destruct pb; tc. Qed. - Definition eq_universe_leq_universe' φ u u' - := @eq_universe_leq_universe φ u u'. - Definition leq_universe_refl' φ u - := @leq_universe_refl φ u. + Definition eq_leq_sort' φ leq_universe eq_universe Hsub u u' + := @eq_leq_sort φ leq_universe eq_universe Hsub u u'. + Definition leq_sort_refl' φ leq_universe leq_refl u + := @leq_sort_refl φ leq_universe leq_refl u. - Hint Resolve eq_universe_leq_universe' leq_universe_refl' : core. + Hint Resolve eq_leq_universe' leq_universe_refl' : core. - Lemma cmp_universe_subset ctrs ctrs' pb t u - : ConstraintSet.Subset ctrs ctrs' - -> compare_universe pb ctrs t u -> compare_universe pb ctrs' t u. + + Lemma cmp_sort_subset φ φ' pb t u + : ConstraintSet.Subset φ φ' + -> compare_sort φ pb t u -> compare_sort φ' pb t u. Proof using Type. intros Hctrs. destruct pb, t, u; cbnr; trivial. @@ -2112,54 +2108,82 @@ Section Univ. all: eapply satisfies_subset; eauto. Qed. - Lemma eq_universe_subset ctrs ctrs' t u + Lemma eq_sort_subset ctrs ctrs' t u : ConstraintSet.Subset ctrs ctrs' - -> eq_universe ctrs t u -> eq_universe ctrs' t u. - Proof using Type. apply cmp_universe_subset with (pb := Conv). Qed. + -> eq_sort ctrs t u -> eq_sort ctrs' t u. + Proof using Type. apply cmp_sort_subset with (pb := Conv). Qed. - Lemma leq_universe_subset ctrs ctrs' t u + Lemma leq_sort_subset ctrs ctrs' t u : ConstraintSet.Subset ctrs ctrs' - -> leq_universe ctrs t u -> leq_universe ctrs' t u. - Proof using Type. apply cmp_universe_subset with (pb := Cumul). Qed. + -> leq_sort ctrs t u -> leq_sort ctrs' t u. + Proof using Type. apply cmp_sort_subset with (pb := Cumul). Qed. +End SortCompare. - (** Elimination restriction *) +(** Elimination restriction *) - Definition is_lSet φ s := eq_universe φ s Universe.type0. - (* Unfolded definition : - match s with - | Universe.lType u => - if check_univs then forall v, satisfies v φ -> val v u = 0 else True - | _ => False - end. *) - - Definition is_allowed_elimination φ allowed : Universe.t -> Prop := - match allowed with - | IntoSProp => Universe.is_sprop - | IntoPropSProp => is_propositional - | IntoSetPropSProp => fun s => is_propositional s \/ is_lSet φ s - | IntoAny => fun s => True - end. - (* Is [a] a subset of [a']? *) - Definition allowed_eliminations_subset (a a' : allowed_eliminations) : bool := - match a, a' with - | IntoSProp, _ - | IntoPropSProp, (IntoPropSProp | IntoSetPropSProp | IntoAny) - | IntoSetPropSProp, (IntoSetPropSProp | IntoAny) - | IntoAny, IntoAny => true - | _, _ => false - end. +(** This inductive classifies which eliminations are allowed for inductive types + in various sorts. *) +Inductive allowed_eliminations : Set := + | IntoSProp + | IntoPropSProp + | IntoSetPropSProp + | IntoAny. +Derive NoConfusion EqDec for allowed_eliminations. - Lemma allowed_eliminations_subset_impl φ a a' s - : allowed_eliminations_subset a a' -> - is_allowed_elimination φ a s -> is_allowed_elimination φ a' s. - Proof using Type. - destruct a, a'; cbnr; trivial; - destruct s; cbnr; trivial; - intros H1 H2; try absurd; constructor; trivial. - Qed. -End Univ. +Definition is_allowed_elimination_cuniv (allowed : allowed_eliminations) : concrete_sort -> bool := + match allowed with + | IntoSProp => Sort.is_sprop + | IntoPropSProp => Sort.is_propositional + | IntoSetPropSProp => is_propositional_or_set + | IntoAny => fun _ => true + end. + +Definition is_lSet {cf} φ s := eq_sort φ s Sort.type0. + (* Unfolded definition : + match s with + | Sort.sType u => + if check_univs then forall v, satisfies v φ -> val v u = 0 else true + | _ => false + end. *) + +Definition is_allowed_elimination {cf} φ allowed : Sort.t -> Prop := + match allowed with + | IntoSProp => Sort.is_sprop + | IntoPropSProp => Sort.is_propositional + | IntoSetPropSProp => fun s => Sort.is_propositional s \/ is_lSet φ s + | IntoAny => fun s => true + end. + +(* Is [a] a subset of [a']? *) +Definition allowed_eliminations_subset (a a' : allowed_eliminations) : bool := + match a, a' with + | IntoSProp, _ + | IntoPropSProp, (IntoPropSProp | IntoSetPropSProp | IntoAny) + | IntoSetPropSProp, (IntoSetPropSProp | IntoAny) + | IntoAny, IntoAny => true + | _, _ => false + end. + +Lemma allowed_eliminations_subset_impl {cf} φ a a' s + : allowed_eliminations_subset a a' -> + is_allowed_elimination φ a s -> is_allowed_elimination φ a' s. +Proof using Type. + destruct a, a'; cbnr; trivial; + destruct s; cbnr; trivial; + intros H1 H2; try absurd; constructor; trivial. +Qed. + +Lemma is_allowed_elimination_monotone `{cf : checker_flags} Σ s1 s2 a : + leq_sort Σ s1 s2 -> is_allowed_elimination Σ a s2 -> is_allowed_elimination Σ a s1. +Proof. + destruct a, s2 as [| |u2], s1 as [| |u1] => //=. 1: now left. + intros Hle [H|]; right => //. + unfold_univ_rel. + cbn in H. + lia. +Qed. Section UnivCF2. Context {cf1 cf2 : checker_flags}. @@ -2174,11 +2198,10 @@ Section UnivCF2. Lemma cmp_universe_config_impl ctrs pb t u : config.impl cf1 cf2 - -> @compare_universe cf1 pb ctrs t u -> @compare_universe cf2 pb ctrs t u. + -> @compare_universe cf1 ctrs pb t u -> @compare_universe cf2 ctrs pb t u. Proof using Type. - destruct pb, t, u; cbnr; trivial. - all: unfold config.impl, eq_levelalg, leq_levelalg_n, is_true. - all: do 2 destruct check_univs, prop_sub_type; cbn => //. + unfold config.impl, compare_universe, leq_universe, eq_universe, leq_universe_n, is_true. + destruct pb; do 2 destruct check_univs => //=. Qed. Lemma eq_universe_config_impl ctrs t u @@ -2191,6 +2214,27 @@ Section UnivCF2. -> @leq_universe cf1 ctrs t u -> @leq_universe cf2 ctrs t u. Proof using Type. apply cmp_universe_config_impl with (pb := Cumul). Qed. + Lemma cmp_sort_config_impl ctrs pb t u + : config.impl cf1 cf2 + -> @compare_sort cf1 ctrs pb t u -> @compare_sort cf2 ctrs pb t u. + Proof using Type. + unfold compare_sort, leq_sort, eq_sort, eq_sort_, leq_sort_n, leq_sort_n_, is_true. + destruct pb, t, u => //=. + - apply eq_universe_config_impl. + - unfold config.impl. do 2 destruct check_univs, prop_sub_type; cbn => //=. + - apply leq_universe_config_impl. + Qed. + + Lemma eq_sort_config_impl ctrs t u + : config.impl cf1 cf2 + -> @eq_sort cf1 ctrs t u -> @eq_sort cf2 ctrs t u. + Proof using Type. apply cmp_sort_config_impl with (pb := Conv). Qed. + + Lemma leq_sort_config_impl ctrs t u + : config.impl cf1 cf2 + -> @leq_sort cf1 ctrs t u -> @leq_sort cf2 ctrs t u. + Proof using Type. apply cmp_sort_config_impl with (pb := Cumul). Qed. + (** Elimination restriction *) Lemma allowed_eliminations_config_impl φ a s @@ -2198,109 +2242,103 @@ Section UnivCF2. @is_allowed_elimination cf1 φ a s -> @is_allowed_elimination cf2 φ a s. Proof using Type. destruct a, s; cbnr; trivial. - unfold eq_levelalg, config.impl, is_true. + unfold eq_universe, config.impl, is_true. do 2 destruct check_univs; cbnr; auto => //. Qed. End UnivCF2. -Ltac unfold_univ_rel0 := - unfold eq0_levelalg, leq0_levelalg_n in *; - intros v Hv; cbnr. -Ltac unfold_univ_rel := - unfold is_allowed_elimination, is_lSet, - leq_universe, eq_universe, leq_universe_n, leq_universe_n_, eq_universe_, - eq_levelalg, leq_levelalg, lt_levelalg, leq_levelalg_n in *; +Ltac unfold_univ_rel ::= + unfold is_allowed_elimination, is_lSet, valid_constraints, + compare_sort, eq_sort, leq_sort, lt_sort, leq_sort_n, leq_sort_n_, eq_sort_, leqb_sort_n_, eqb_sort_, + compare_universe, leq_universe, eq_universe, leq_universe_n in *; destruct check_univs; [unfold_univ_rel0 | trivial]. Tactic Notation "unfold_univ_rel" "eqn" ":"ident(H) := - unfold is_allowed_elimination, is_lSet, - leq_universe, eq_universe, leq_universe_n, leq_universe_n_, eq_universe_, - eq_levelalg, leq_levelalg, lt_levelalg, leq_levelalg_n in *; + unfold is_allowed_elimination, is_lSet, valid_constraints, + compare_sort, eq_sort, leq_sort, lt_sort, leq_sort_n, leq_sort_n_, eq_sort_, leqb_sort_n_, eqb_sort_, + compare_universe, leq_universe, eq_universe, leq_universe_n in *; destruct check_univs eqn:H; [unfold_univ_rel0 | trivial]. (* Ltac prop_non_prop := match goal with - | |- context[ Universe.is_prop ?u || Universe.is_sprop ?u] => - destruct (Universe.is_prop u || Universe.is_sprop u) - | H : context[ Universe.is_prop ?u || Universe.is_sprop ?u] |- _ => - destruct (Universe.is_prop u || Universe.is_sprop u) + | |- context[ Sort.is_prop ?u || Sort.is_sprop ?u] => + destruct (Sort.is_prop u || Sort.is_sprop u) + | H : context[ Sort.is_prop ?u || Sort.is_sprop ?u] |- _ => + destruct (Sort.is_prop u || Sort.is_sprop u) end. *) Ltac cong := intuition congruence. - -Lemma leq_universe_product_mon {cf} ϕ s1 s1' s2 s2' : - leq_universe ϕ s1 s1' -> - leq_universe ϕ s2 s2' -> - leq_universe ϕ (Universe.sort_of_product s1 s2) (Universe.sort_of_product s1' s2'). +Lemma leq_sort_product_mon {cf} ϕ s1 s1' s2 s2' : + leq_sort ϕ s1 s1' -> + leq_sort ϕ s2 s2' -> + leq_sort ϕ (Sort.sort_of_product s1 s2) (Sort.sort_of_product s1' s2'). Proof. destruct s2 as [| | u2], s2' as [| | u2']; cbnr; try absurd; destruct s1 as [| | u1], s1' as [| | u1']; cbnr; try absurd; trivial. - - intros _ H2; etransitivity; [apply H2 | apply leq_levelalg_sup0_r]. - - apply leq_levelalg_sup0_mon. + - intros _ H2; etransitivity; [apply H2 | apply leq_universe_sup_r]. + - apply leq_universe_sup_mon. Qed. Lemma impredicative_product {cf} {ϕ l u} : - Universe.is_prop u -> - leq_universe ϕ (Universe.sort_of_product l u) u. + Sort.is_propositional u -> + leq_sort ϕ (Sort.sort_of_product l u) u. Proof. - unfold Universe.sort_of_product. - intros ->. reflexivity. + destruct u => //; reflexivity. Qed. Section UniverseLemmas. Context {cf: checker_flags}. - Lemma sup0_idem s : LevelAlgExpr.sup s s = s. + Lemma univ_sup_idem s : Universe.sup s s = s. Proof using Type. apply eq_univ'; cbn. intro; rewrite !LevelExprSet.union_spec. intuition. Qed. - Lemma sup_idem s : Universe.sup s s = s. + Lemma sup_idem s : Sort.sup s s = s. Proof using Type. destruct s; cbn; auto. apply f_equal. - apply sup0_idem. + apply univ_sup_idem. Qed. Lemma sort_of_product_idem s - : Universe.sort_of_product s s = s. + : Sort.sort_of_product s s = s. Proof using Type. - unfold Universe.sort_of_product; destruct s; try reflexivity. + unfold Sort.sort_of_product; destruct s; try reflexivity. apply sup_idem. Qed. - Lemma sup0_assoc s1 s2 s3 : - LevelAlgExpr.sup s1 (LevelAlgExpr.sup s2 s3) = LevelAlgExpr.sup (LevelAlgExpr.sup s1 s2) s3. + Lemma univ_sup_assoc s1 s2 s3 : + Universe.sup s1 (Universe.sup s2 s3) = Universe.sup (Universe.sup s1 s2) s3. Proof using Type. apply eq_univ'; cbn. symmetry; apply LevelExprSetProp.union_assoc. Qed. - Instance proper_sup0_eq_levelalg φ : - Proper (eq_levelalg φ ==> eq_levelalg φ ==> eq_levelalg φ) LevelAlgExpr.sup. + Instance proper_univ_sup_eq_univ φ : + Proper (eq_universe φ ==> eq_universe φ ==> eq_universe φ) Universe.sup. Proof using Type. intros u1 u1' H1 u2 u2' H2. unfold_univ_rel. - specialize (H1 v Hv); specialize (H2 v Hv). rewrite !val_sup. lia. Qed. - Instance universe_sup_eq_universe φ : - Proper (eq_universe φ ==> eq_universe φ ==> eq_universe φ) Universe.sup. + Instance proper_sort_sup_eq_sort φ : + Proper (eq_sort φ ==> eq_sort φ ==> eq_sort φ) Sort.sup. Proof using Type. intros [| | u1] [| |u1'] H1 [| |u2] [| |u2'] H2; cbn in *; try absurd; auto. - now apply proper_sup0_eq_levelalg. + now apply proper_univ_sup_eq_univ. Qed. Lemma sort_of_product_twice u s : - Universe.sort_of_product u (Universe.sort_of_product u s) - = Universe.sort_of_product u s. + Sort.sort_of_product u (Sort.sort_of_product u s) + = Sort.sort_of_product u s. Proof using Type. destruct u,s; cbnr. - now rewrite sup0_assoc sup0_idem. + now rewrite univ_sup_assoc univ_sup_idem. Qed. End UniverseLemmas. @@ -2316,7 +2354,7 @@ Section no_prop_leq_type. Qed. Lemma spec_map_succ l x : - LevelExprSet.In x (LevelAlgExpr.succ l) <-> + LevelExprSet.In x (Universe.succ l) <-> exists x', LevelExprSet.In x' l /\ x = LevelExpr.succ x'. Proof using Type. rewrite map_spec. reflexivity. @@ -2327,10 +2365,10 @@ Section no_prop_leq_type. destruct l as []; simpl. cbn. lia. Qed. - Lemma val_map_succ v l : val v (LevelAlgExpr.succ l) = val v l + 1. + Lemma val_map_succ v l : val v (Universe.succ l) = val v l + 1. Proof using Type. pose proof (spec_map_succ l). - set (n := LevelAlgExpr.succ l) in *. + set (n := Universe.succ l) in *. destruct (val_In_max l v) as [max [inmax eqv]]. rewrite <-eqv. rewrite val_caract. split. intros. @@ -2342,138 +2380,45 @@ Section no_prop_leq_type. now rewrite val_succ. Qed. - Lemma leq_universe_super s s' : - leq_universe ϕ s s' -> - leq_universe ϕ (Universe.super s) (Universe.super s'). + Lemma leq_sort_super s s' : + leq_sort ϕ s s' -> + leq_sort ϕ (Sort.super s) (Sort.super s'). Proof using Type. destruct s as [| | u1], s' as [| | u1']; cbnr; try absurd; intros H; unfold_univ_rel; - rewrite !val_map_succ. lia. - specialize (H v Hv). lia. - Qed. - - Lemma leq_universe_props s1 s2 : - check_univs -> - consistent ϕ -> - leq_universe ϕ s1 s2 -> - match s1, s2 with - | Universe.lProp, Universe.lProp => True - | Universe.lSProp, Universe.lSProp => True - | Universe.lProp, Universe.lSProp => False - | Universe.lSProp, Universe.lProp => False - | Universe.lProp, Universe.lType _ => prop_sub_type - | Universe.lSProp, Universe.lType _ => False - | Universe.lType l, Universe.lType l' => True - | Universe.lType _, _ => False - end. - Proof using Type. - destruct s1, s2; cbnr; trivial. - Qed. - - Lemma leq_universe_prop_r s1 s2 : - check_univs -> - consistent ϕ -> - leq_universe ϕ s1 s2 -> - Universe.is_prop s2 -> Universe.is_prop s1. - Proof using Type. - intros Hcf cu. - destruct s2; cbn; [ | absurd | absurd]. - destruct s1; cbn; [ auto | absurd | absurd]. - Qed. - - Lemma leq_universe_sprop_r s1 s2 : - check_univs -> - consistent ϕ -> - leq_universe ϕ s1 s2 -> - Universe.is_sprop s2 -> Universe.is_sprop s1. - Proof using Type. - intros Hcf cu. - destruct s2; cbn; [ absurd | | absurd]. - destruct s1; cbn; [ absurd | auto | absurd]. + rewrite !val_map_succ; lia. Qed. - Lemma leq_universe_prop_no_prop_sub_type s1 s2 : - check_univs -> + Lemma leq_sort_prop_no_prop_sub_type s1 s2 : prop_sub_type = false -> - consistent ϕ -> - leq_universe ϕ s1 s2 -> - Universe.is_prop s1 -> Universe.is_prop s2. + leq_sort ϕ s1 s2 -> + Sort.is_prop s1 -> Sort.is_prop s2. Proof using Type. - intros Hcf ps cu. + intros ps. destruct s1; cbn; [ | absurd | absurd]. rewrite ps. destruct s2; cbn; [ auto | absurd | absurd]. Qed. - Lemma leq_universe_sprop_l s1 s2 : - check_univs -> - consistent ϕ -> - leq_universe ϕ s1 s2 -> - Universe.is_sprop s1 -> Universe.is_sprop s2. - Proof using Type. - intros Hcf cu. - destruct s1; cbn; [ absurd | | absurd]. - destruct s2; cbn; [ absurd | auto | absurd]. - Qed. - - Hint Resolve leq_universe_sprop_l leq_universe_sprop_r - leq_universe_prop_r - leq_universe_prop_no_prop_sub_type - : univ_lemmas. - - Lemma leq_prop_prop {l l'} : - Universe.is_prop l -> Universe.is_prop l' -> - leq_universe ϕ l l'. - Proof using Type. - destruct l, l'; cbnr; absurd. - Qed. - - Lemma leq_sprop_sprop {l l'} : - Universe.is_sprop l -> Universe.is_sprop l' -> - leq_universe ϕ l l'. - Proof using Type. - destruct l, l'; cbnr; absurd. - Qed. + Hint Resolve leq_sort_prop_no_prop_sub_type : univ_lemmas. - Lemma leq_prop_is_prop_sprop {s1 s2} : - check_univs -> + Lemma leq_prop_is_propositonal {s1 s2} : prop_sub_type = false -> - consistent ϕ -> - leq_universe ϕ s1 s2 -> - is_propositional s1 <-> is_propositional s2. + leq_sort ϕ s1 s2 -> + Sort.is_propositional s1 <-> Sort.is_propositional s2. Proof using Type. - intros Hcf ps cu. + intros ps. destruct s1, s2; cbn; try absurd; intros H; split; trivial. now rewrite ps in H. Qed. - Lemma is_prop_gt s1 s2 : - check_univs -> - consistent ϕ -> - leq_universe ϕ (Universe.super s1) s2 -> Universe.is_prop s2 -> False. - Proof using Type. - intros Hcf cu Hleq Hprop. - apply leq_universe_prop_r in Hleq; tas. - now destruct s1. - Qed. - - Lemma is_sprop_gt s1 s2 : - check_univs -> - consistent ϕ -> - leq_universe ϕ (Universe.super s1) s2 -> Universe.is_sprop s2 -> False. - Proof using Type. - intros Hcf cu Hleq Hprop. - apply leq_universe_sprop_r in Hleq; tas. - now destruct s1. - Qed. - End no_prop_leq_type. (* This level is a hack used in plugings to generate fresh levels *) Definition fresh_level : Level.t := Level.level "__metacoq_fresh_level__". -(* This universe is a hack used in plugins to generate fresh concreteUniverses *) -Definition fresh_universe : Universe.t := Universe.make fresh_level. +(* This universe is a hack used in plugins to generate fresh universes *) +Definition fresh_universe : Universe.t := Universe.make' fresh_level. (** * Universe substitution @@ -2512,19 +2457,25 @@ Notation "x @[ u ]" := (subst_instance u x) (at level 3, end end. -#[global] Instance subst_instance_univ0 : UnivSubst LevelAlgExpr.t := +#[global] Instance subst_instance_universe : UnivSubst Universe.t := fun u => map (subst_instance_level_expr u). -#[global] Instance subst_instance_univ : UnivSubst Universe.t := +#[global] Instance subst_instance_sort : UnivSubst Sort.t := fun u e => match e with - | Universe.lProp | Universe.lSProp => e - | Universe.lType l => Universe.lType (subst_instance u l) + | sProp | sSProp => e + | sType u' => sType (subst_instance u u') end. +Lemma subst_instance_to_family s u : + Sort.to_family s@[u] = Sort.to_family s. +Proof. + destruct s => //. +Qed. + #[global] 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 lvariables *) +(** Tests that the term is closed over [k] universe variables *) Section Closedu. Context (k : nat). @@ -2537,13 +2488,13 @@ Section Closedu. Definition closedu_level_expr (s : LevelExpr.t) := closedu_level (LevelExpr.get_level s). - Definition closedu_universe_levels (u : LevelAlgExpr.t) := + Definition closedu_universe (u : Universe.t) := LevelExprSet.for_all closedu_level_expr u. - Definition closedu_universe (u : Universe.t) := + Definition closedu_sort (u : Sort.t) := match u with - | Universe.lSProp | Universe.lProp => true - | Universe.lType l => closedu_universe_levels l + | sSProp | sProp => true + | sType l => closedu_universe l end. Definition closedu_instance (u : Instance.t) := @@ -2566,13 +2517,13 @@ Section UniverseClosedSubst. Qed. Lemma closedu_subst_instance_univ u s - : closedu_universe 0 s -> subst_instance_univ u s = s. + : closedu_sort 0 s -> subst_instance_sort u s = s. Proof. intro H. destruct s as [| | t]; cbnr. apply f_equal. apply eq_univ'. destruct t as [ts H1]. - unfold closedu_universe_levels in *;cbn in *. + unfold closedu_universe in *;cbn in *. intro e; split; intro He. - apply map_spec in He. destruct He as [e' [He' X]]. rewrite closedu_subst_instance_level_expr in X. @@ -2625,7 +2576,7 @@ Section SubstInstanceClosed. Qed. Lemma subst_instance_univ_closedu s - : closedu_universe #|u| s -> closedu_universe 0 (subst_instance_univ u s). + : closedu_sort #|u| s -> closedu_sort 0 (subst_instance_sort u s). Proof using Hcl. intro H. destruct s as [| |t]; cbnr. @@ -2662,11 +2613,11 @@ Definition string_of_level (l : Level.t) : string := Definition string_of_level_expr (e : LevelExpr.t) : string := let '(l, n) := e in string_of_level l ^ (if n is 0 then "" else "+" ^ string_of_nat n). -Definition string_of_sort (u : Universe.t) := +Definition string_of_sort (u : Sort.t) := match u with - | Universe.lSProp => "SProp" - | Universe.lProp => "Prop" - | Universe.lType l => "Type(" ^ string_of_list string_of_level_expr (LevelExprSet.elements l) ^ ")" + | sSProp => "SProp" + | sProp => "Prop" + | sType l => "Type(" ^ string_of_list string_of_level_expr (LevelExprSet.elements l) ^ ")" end. Definition string_of_universe_instance u := diff --git a/common/theories/UniversesDec.v b/common/theories/UniversesDec.v index 5b7e5f46e..f73f495c7 100644 --- a/common/theories/UniversesDec.v +++ b/common/theories/UniversesDec.v @@ -77,7 +77,7 @@ Proof. split; apply global_uctx_invariants_union_or; constructor; apply levels_of_cs2_spec. Qed. -Definition levels_of_algexpr (u : LevelAlgExpr.t) : VSet.t +Definition levels_of_universe (u : Universe.t) : VSet.t := LevelExprSet.fold (fun gc acc => match LevelExpr.get_noprop gc with | Some l => VSet.add l acc @@ -85,10 +85,10 @@ Definition levels_of_algexpr (u : LevelAlgExpr.t) : VSet.t end) u VSet.empty. -Lemma levels_of_algexpr_spec u cstr (lvls := levels_of_algexpr u) +Lemma levels_of_universe_spec u cstr (lvls := levels_of_universe u) : gc_levels_declared (lvls, cstr) u. Proof. - subst lvls; cbv [levels_of_algexpr gc_levels_declared gc_expr_declared on_Some_or_None LevelExpr.get_noprop]; cbn [fst snd]. + subst lvls; cbv [levels_of_universe gc_levels_declared gc_expr_declared on_Some_or_None LevelExpr.get_noprop]; cbn [fst snd]. cbv [LevelExprSet.For_all]; cbn [fst snd]. repeat first [ apply conj | progress intros @@ -751,14 +751,14 @@ Proof. destruct b; [ left; apply H; reflexivity | right; intro H'; apply H in H'; auto ]. Defined. -Definition leq0_levelalg_n_dec n ϕ u u' : {@leq0_levelalg_n (uGraph.Z_of_bool n) ϕ u u'} + {~@leq0_levelalg_n (uGraph.Z_of_bool n) ϕ u u'}. +Definition leq0_universe_n_dec n ϕ u u' : {@leq0_universe_n (uGraph.Z_of_bool n) ϕ u u'} + {~@leq0_universe_n (uGraph.Z_of_bool n) ϕ u u'}. Proof. - pose proof (@uGraph.gc_leq0_levelalg_n_iff config.default_checker_flags (uGraph.Z_of_bool n) ϕ u u') as H. + pose proof (@uGraph.gc_leq0_universe_n_iff config.default_checker_flags (uGraph.Z_of_bool n) ϕ u u') as H. pose proof (@uGraph.gc_consistent_iff config.default_checker_flags ϕ). cbv [on_Some on_Some_or_None] in *. destruct gc_of_constraints eqn:?. all: try solve [ left; cbv [consistent] in *; hnf; intros; exfalso; intuition eauto ]. - pose proof (fun G cstr => @uGraph.leqb_levelalg_n_spec G (LevelSet.union (levels_of_cs ϕ) (LevelSet.union (levels_of_algexpr u) (levels_of_algexpr u')), cstr)). + pose proof (fun G cstr => @uGraph.leqb_universe_n_spec G (LevelSet.union (levels_of_cs ϕ) (LevelSet.union (levels_of_universe u) (levels_of_universe u')), cstr)). pose proof (fun x y => @gc_of_constraints_of_uctx config.default_checker_flags (x, y)) as H'. pose proof (@is_consistent_spec config.default_checker_flags (levels_of_cs ϕ, ϕ)). specialize_under_binders_by eapply gc_levels_declared_union_or. @@ -770,40 +770,40 @@ Proof. specialize_under_binders_by eapply levels_of_cs_spec. specialize_under_binders_by reflexivity. destruct is_consistent; - [ | left; now cbv [leq0_levelalg_n consistent] in *; intros; exfalso; intuition eauto ]. + [ | left; now cbv [leq0_universe_n consistent] in *; intros; exfalso; intuition eauto ]. specialize_by intuition eauto. let H := match goal with H : forall (b : bool), _ |- _ => H end in specialize (H n u u'). - specialize_under_binders_by (constructor; eapply gc_levels_declared_union_or; constructor; eapply levels_of_algexpr_spec). + specialize_under_binders_by (constructor; eapply gc_levels_declared_union_or; constructor; eapply levels_of_universe_spec). match goal with H : is_true ?b <-> ?x, H' : ?y <-> ?x |- {?y} + {_} => destruct b eqn:?; [ left | right ] end. all: intuition. Defined. -Definition leq_levelalg_n_dec cf n ϕ u u' : {@leq_levelalg_n cf (uGraph.Z_of_bool n) ϕ u u'} + {~@leq_levelalg_n cf (uGraph.Z_of_bool n) ϕ u u'}. +Definition leq_universe_n_dec cf n ϕ u u' : {@leq_universe_n cf (uGraph.Z_of_bool n) ϕ u u'} + {~@leq_universe_n cf (uGraph.Z_of_bool n) ϕ u u'}. Proof. - cbv [leq_levelalg_n]; destruct (@leq0_levelalg_n_dec n ϕ u u'); destruct ?; auto. + cbv [leq_universe_n]; destruct (@leq0_universe_n_dec n ϕ u u'); destruct ?; auto. Defined. -Definition eq0_levelalg_dec ϕ u u' : {@eq0_levelalg ϕ u u'} + {~@eq0_levelalg ϕ u u'}. +Definition eq0_universe_dec ϕ u u' : {@eq0_universe ϕ u u'} + {~@eq0_universe ϕ u u'}. Proof. - pose proof (@eq0_leq0_levelalg ϕ u u') as H. - destruct (@leq0_levelalg_n_dec false ϕ u u'), (@leq0_levelalg_n_dec false ϕ u' u); constructor; destruct H; split_and; now auto. + pose proof (@eq0_leq0_universe ϕ u u') as H. + destruct (@leq0_universe_n_dec false ϕ u u'), (@leq0_universe_n_dec false ϕ u' u); constructor; destruct H; split_and; now auto. Defined. -Definition eq_levelalg_dec {cf ϕ} u u' : {@eq_levelalg cf ϕ u u'} + {~@eq_levelalg cf ϕ u u'}. +Definition eq_universe_dec {cf ϕ} u u' : {@eq_universe cf ϕ u u'} + {~@eq_universe cf ϕ u u'}. Proof. - cbv [eq_levelalg]; destruct ?; auto using eq0_levelalg_dec. + cbv [eq_universe]; destruct ?; auto using eq0_universe_dec. Defined. -Definition eq_universe__dec {CS eq_levelalg ϕ} - (eq_levelalg_dec : forall u u', {@eq_levelalg ϕ u u'} + {~@eq_levelalg ϕ u u'}) +Definition eq_sort__dec {univ eq_universe} + (eq_universe_dec : forall u u', {@eq_universe u u'} + {~@eq_universe u u'}) s s' - : {@eq_universe_ CS eq_levelalg ϕ s s'} + {~@eq_universe_ CS eq_levelalg ϕ s s'}. + : {@eq_sort_ univ eq_universe s s'} + {~@eq_sort_ univ eq_universe s s'}. Proof. - cbv [eq_universe_]; repeat destruct ?; auto. + cbv [eq_sort_]; repeat destruct ?; auto. all: destruct pst; auto. Defined. -Definition eq_universe_dec {cf ϕ} s s' : {@eq_universe cf ϕ s s'} + {~@eq_universe cf ϕ s s'} := eq_universe__dec eq_levelalg_dec _ _. +Definition eq_sort_dec {cf ϕ} s s' : {@eq_sort cf ϕ s s'} + {~@eq_sort cf ϕ s s'} := eq_sort__dec eq_universe_dec _ _. Definition valid_constraints_dec cf ϕ cstrs : {@valid_constraints cf ϕ cstrs} + {~@valid_constraints cf ϕ cstrs}. Proof. @@ -830,12 +830,12 @@ Definition valid_constraints0_dec ϕ ctrs : {@valid_constraints0 ϕ ctrs} + {~@v Definition is_lSet_dec cf ϕ s : {@is_lSet cf ϕ s} + {~@is_lSet cf ϕ s}. Proof. - apply eq_universe_dec. + apply eq_sort_dec. Defined. Definition is_allowed_elimination_dec cf ϕ allowed u : {@is_allowed_elimination cf ϕ allowed u} + {~@is_allowed_elimination cf ϕ allowed u}. Proof. cbv [is_allowed_elimination is_true]; destruct ?; auto; try solve [ repeat decide equality ]. - destruct (@is_lSet_dec cf ϕ u), is_propositional; intuition auto. + destruct (@is_lSet_dec cf ϕ u), Sort.is_propositional; intuition auto. Defined. diff --git a/common/theories/uGraph.v b/common/theories/uGraph.v index ca51cda54..3f0fdc162 100644 --- a/common/theories/uGraph.v +++ b/common/theories/uGraph.v @@ -544,36 +544,36 @@ Qed. Local Open Scope univ_scope. -Definition gc_leq0_levelalg_n n ctrs (u u' : LevelAlgExpr.t) := +Definition gc_leq0_universe_n n ctrs (u u' : Universe.t) := forall v, gc_satisfies v ctrs -> (Z.of_nat (val v u) <= Z.of_nat (val v u') - n)%Z. -Definition gc_leq_levelalg_n n ctrs (u u' : LevelAlgExpr.t) := - if check_univs then gc_leq0_levelalg_n n ctrs u u' else True. +Definition gc_leq_universe_n n ctrs (u u' : Universe.t) := + if check_univs then gc_leq0_universe_n n ctrs u u' else True. -Definition gc_eq0_levelalg φ (u u' : LevelAlgExpr.t) := +Definition gc_eq0_universe φ (u u' : Universe.t) := forall v, gc_satisfies v φ -> val v u = val v u'. -Definition gc_eq_levelalg φ (u u' : LevelAlgExpr.t) := - if check_univs then gc_eq0_levelalg φ u u' else True. +Definition gc_eq_universe φ (u u' : Universe.t) := + if check_univs then gc_eq0_universe φ u u' else True. -Definition gc_leq0_levelalg := gc_leq0_levelalg_n 0. -Definition gc_lt0_levelalg := gc_leq0_levelalg_n 1. -Definition gc_leq_levelalg := gc_leq_levelalg_n 0. -Definition gc_lt_levelalg := gc_leq_levelalg_n 1. +Definition gc_leq0_universe := gc_leq0_universe_n 0. +Definition gc_lt0_universe := gc_leq0_universe_n 1. +Definition gc_leq_universe := gc_leq_universe_n 0. +Definition gc_lt_universe := gc_leq_universe_n 1. Ltac unfold_univ_rel0 := - unfold eq0_levelalg, leq0_levelalg_n, - gc_eq0_levelalg, gc_leq0_levelalg, gc_lt0_levelalg, gc_leq0_levelalg_n in *; + unfold eq0_universe, leq0_universe_n, + gc_eq0_universe, gc_leq0_universe, gc_lt0_universe, gc_leq0_universe_n in *; intros v Hv; cbnr. Ltac unfold_univ_rel := - unfold eq_levelalg, leq_levelalg, lt_levelalg, leq_levelalg_n, - gc_eq_levelalg, gc_leq_levelalg, gc_lt_levelalg, gc_leq_levelalg_n in *; + unfold eq_universe, leq_universe, lt_universe, leq_universe_n, + gc_eq_universe, gc_leq_universe, gc_lt_universe, gc_leq_universe_n in *; destruct check_univs; [| trivial]. -Lemma gc_leq0_levelalg_n_iff (n: Z) ctrs u u' : - leq0_levelalg_n n ctrs u u' - <-> on_Some_or_None (fun ctrs => gc_leq0_levelalg_n n ctrs u u') +Lemma gc_leq0_universe_n_iff (n: Z) ctrs u u' : + leq0_universe_n n ctrs u u' + <-> on_Some_or_None (fun ctrs => gc_leq0_universe_n n ctrs u u') (gc_of_constraints ctrs). Proof. split. @@ -593,18 +593,18 @@ Proof. rewrite e in Hv; contradiction. Defined. -Lemma gc_leq0_levelalg_iff ctrs u u': - leq0_levelalg_n 0 ctrs u u' - <-> on_Some_or_None (fun ctrs => gc_leq0_levelalg_n 0 ctrs u u') +Lemma gc_leq0_universe_iff ctrs u u': + leq0_universe_n 0 ctrs u u' + <-> on_Some_or_None (fun ctrs => gc_leq0_universe_n 0 ctrs u u') (gc_of_constraints ctrs). Proof using Type. - apply gc_leq0_levelalg_n_iff. + apply gc_leq0_universe_n_iff. Qed. -Lemma gc_eq0_levelalg_iff ctrs u u' : - eq0_levelalg ctrs u u' - <-> on_Some_or_None (fun ctrs => gc_eq0_levelalg ctrs u u') +Lemma gc_eq0_universe_iff ctrs u u' : + eq0_universe ctrs u u' + <-> on_Some_or_None (fun ctrs => gc_eq0_universe ctrs u u') (gc_of_constraints ctrs). Proof. split. @@ -623,33 +623,33 @@ Proof. rewrite e in Hv; contradiction. Defined. -Lemma gc_leq_levelalg_n_iff n ctrs u u' : - leq_levelalg_n n ctrs u u' - <-> on_Some_or_None (fun ctrs => gc_leq_levelalg_n n ctrs u u') +Lemma gc_leq_universe_n_iff n ctrs u u' : + leq_universe_n n ctrs u u' + <-> on_Some_or_None (fun ctrs => gc_leq_universe_n n ctrs u u') (gc_of_constraints ctrs). Proof using Type. unfold_univ_rel. - apply gc_leq0_levelalg_n_iff. + apply gc_leq0_universe_n_iff. destruct (gc_of_constraints ctrs); reflexivity. Qed. -Lemma gc_leq_levelalg_iff ctrs u u' : - leq_levelalg ctrs u u' - <-> on_Some_or_None (fun ctrs => gc_leq_levelalg ctrs u u') +Lemma gc_leq_universe_iff ctrs u u' : + leq_universe ctrs u u' + <-> on_Some_or_None (fun ctrs => gc_leq_universe ctrs u u') (gc_of_constraints ctrs). Proof using Type. unfold_univ_rel. - apply gc_leq0_levelalg_iff. + apply gc_leq0_universe_iff. destruct (gc_of_constraints ctrs); reflexivity. Qed. -Lemma gc_eq_levelalg_iff ctrs u u' : - eq_levelalg ctrs u u' - <-> on_Some_or_None (fun ctrs => gc_eq_levelalg ctrs u u') +Lemma gc_eq_universe_iff ctrs u u' : + eq_universe ctrs u u' + <-> on_Some_or_None (fun ctrs => gc_eq_universe ctrs u u') (gc_of_constraints ctrs). Proof using Type. unfold_univ_rel. - apply gc_eq0_levelalg_iff. + apply gc_eq0_universe_iff. destruct (gc_of_constraints ctrs); reflexivity. Qed. @@ -855,7 +855,7 @@ Proof. rewrite EdgeSet.add_spec. split. * intros [[c [eq inl]]|?]. - subst e. firstorder auto. + subst e. left. repeat eexists; tea. now right. destruct H as [->|ing]; [left|right]; auto. exists a; firstorder auto. * intros [[c [eq [->|inl]]]|?]; auto. @@ -1206,30 +1206,30 @@ Section CheckLeqProcedure. end. (* this is function [exists_bigger] of kernel/uGraph.ml *) - Definition leqb_expr_univ_n_gen lt (e1 : LevelExpr.t) (u : LevelAlgExpr.t) := + Definition leqb_expr_univ_n_gen lt (e1 : LevelExpr.t) (u : Universe.t) := (* CHECKME:SPROP: should we use [prop_sub_type] here somehow? *) (* if LevelExpr.is_prop e1 && (n =? 0) then *) - (* prop_sub_type || Universe.is_prop u *) + (* prop_sub_type || Sort.is_prop u *) (* else *) - let '(e2, u) := LevelAlgExpr.exprs u in + let '(e2, u) := Universe.exprs u in List.fold_left (fun b e2 => leqb_expr_n_gen lt e1 e2 || b) u (leqb_expr_n_gen lt e1 e2). (* this is function [real_check_leq] of kernel/uGraph.ml *) - Definition leqb_levelalg_n_gen lt (l1 l2 : LevelAlgExpr.t) := - let '(e1, u1) := LevelAlgExpr.exprs l1 in + Definition leqb_universe_n_gen lt (l1 l2 : Universe.t) := + let '(e1, u1) := Universe.exprs l1 in List.fold_left (fun b e1 => leqb_expr_univ_n_gen ⎩ lt ⎭ e1 l2 && b) u1 (leqb_expr_univ_n_gen ⎩ lt ⎭ e1 l2). - Definition check_leqb_levelalg_gen (u1 u2 : LevelAlgExpr.t) := + Definition check_leqb_universe_gen (u1 u2 : Universe.t) := ~~ check_univs || (u1 == u2) - || leqb_levelalg_n_gen false u1 u2. + || leqb_universe_n_gen false u1 u2. - Definition check_eqb_levelalg_gen (u1 u2 : LevelAlgExpr.t) := + Definition check_eqb_universe_gen (u1 u2 : Universe.t) := ~~ check_univs || (u1 == u2) - || (leqb_levelalg_n_gen false u1 u2 && leqb_levelalg_n_gen false u2 u1). + || (leqb_universe_n_gen false u1 u2 && leqb_universe_n_gen false u2 u1). Definition check_gc_constraint_gen (gc : GoodConstraint.t) := ~~ check_univs || @@ -1251,19 +1251,19 @@ Section CheckLeqProcedure. end. Definition eqb_univ_instance_gen (u1 u2 : Instance.t) : bool := - forallb2 (fun l1 l2 => check_eqb_levelalg_gen - (LevelAlgExpr.make' l1) (LevelAlgExpr.make' l2)) u1 u2. + forallb2 (fun l1 l2 => check_eqb_universe_gen + (Universe.make' l1) (Universe.make' l2)) u1 u2. - Definition leqb_universe_gen (s1 s2 : Universe.t) := - leqb_universe_n_ (fun _ => check_leqb_levelalg_gen) false s1 s2. + Definition leqb_sort_gen (s1 s2 : Sort.t) := + leqb_sort_n_ (fun _ => check_leqb_universe_gen) false s1 s2. - Definition check_leqb_universe_gen (u1 u2 : Universe.t) := - Universe.eqb u1 u2 - || leqb_universe_gen u1 u2. + Definition check_leqb_sort_gen (s1 s2 : Sort.t) := + (s1 == s2) + || leqb_sort_gen s1 s2. - Definition check_eqb_universe_gen (u1 u2 : Universe.t) := - Universe.eqb u1 u2 - || (leqb_universe_gen u1 u2 && leqb_universe_gen u2 u1). + Definition check_eqb_sort_gen (s1 s2 : Sort.t) := + (s1 == s2) + || (leqb_sort_gen s1 s2 && leqb_sort_gen s2 s1). End CheckLeqProcedure. @@ -1294,14 +1294,11 @@ Section CheckLeq. Definition gc_expr_declared e := on_Some_or_None (fun l => VSet.In l uctx.1) (LevelExpr.get_noprop e). - Definition gc_levels_declared (u : LevelAlgExpr.t) + Definition gc_levels_declared (u : Universe.t) := LevelExprSet.For_all gc_expr_declared u. - Definition gc_levels_declared_univ (u : Universe.t) - := match u with - | Universe.lSProp | Universe.lProp => True - | Universe.lType l => gc_levels_declared l - end. + Definition gc_levels_declared_sort (s : Sort.t) + := Sort.on_sort gc_levels_declared True s. Lemma val_level_of_variable_level v (l : VariableLevel.t) : val v (l : Level.t) = val v l. @@ -1318,7 +1315,7 @@ Section CheckLeq. Qed. Lemma val_labelling_of_valuation' v (l : Level.t) n : - val v (LevelAlgExpr.make (l, n)) + val v (Universe.make (l, n)) = n + labelling_of_valuation v l. Proof using Type. reflexivity. @@ -1361,13 +1358,13 @@ Section CheckLeq. (** ** Check of leq ** *) Ltac unfold_univ_rel0 := - unfold eq0_levelalg, leq0_levelalg_n, leq_vertices, - gc_eq0_levelalg, gc_leq0_levelalg, gc_lt0_levelalg, gc_leq0_levelalg_n in *; + unfold eq0_universe, leq0_universe_n, leq_vertices, + gc_eq0_universe, gc_leq0_universe, gc_lt0_universe, gc_leq0_universe_n in *; intros v Hv; cbnr. - Lemma leq_levelalg_vertices0 n (l l' : Level.t) + Lemma leq_universe_vertices0 n (l l' : Level.t) : leq_vertices G n l l' - -> gc_leq0_levelalg_n n uctx.2 (LevelAlgExpr.make' l) (LevelAlgExpr.make' l'). + -> gc_leq0_universe_n n uctx.2 (Universe.make' l) (Universe.make' l'). Proof using HG. intros H. unfold_univ_rel0. apply make_graph_spec in Hv; tas. @@ -1377,16 +1374,16 @@ Section CheckLeq. rewrite !val_labelling_of_valuation; lia. Qed. - Lemma leq_levelalg_vertices1 n (l l' : Level.t) + Lemma leq_universe_vertices1 n (l l' : Level.t) (Hl : VSet.In l (wGraph.V G)) (Hl' : VSet.In l' (wGraph.V G)) - : gc_leq0_levelalg_n n uctx.2 (LevelAlgExpr.make' l) (LevelAlgExpr.make' l') + : gc_leq0_universe_n n uctx.2 (Universe.make' l) (Universe.make' l') -> leq_vertices G n l l'. Proof using HG Huctx. intros H. unfold_univ_rel0. eapply correct_labelling_proper in Hv. 2:symmetry; tea. 2:reflexivity. specialize (H _ (make_graph_spec' _ Huctx _ Hv)) as HH. eapply HG in Hl, Hl'. - rewrite !LevelAlgExpr.val_make' in HH. + rewrite !Universe.val_make' in HH. rewrite <- (valuation_labelling_eq _ _ Hv l Hl). rewrite <- (valuation_labelling_eq _ _ Hv l' Hl'). pose proof (val_labelling_of_valuation (valuation_of_labelling v) l). @@ -1394,14 +1391,14 @@ Section CheckLeq. rewrite H0 H1 in HH. lia. Qed. - Lemma leq_levelalg_vertices n (l l' : Level.t) + Lemma leq_universe_vertices n (l l' : Level.t) (Hl : VSet.In l (wGraph.V G)) (Hl' : VSet.In l' (wGraph.V G)) - : gc_leq0_levelalg_n n uctx.2 (LevelAlgExpr.make' l) (LevelAlgExpr.make' l') + : gc_leq0_universe_n n uctx.2 (Universe.make' l) (Universe.make' l') <-> leq_vertices G n l l'. Proof using HG Huctx. split. - - intros H. unfold_univ_rel0. apply leq_levelalg_vertices1; tas. - - apply leq_levelalg_vertices0. + - intros H. unfold_univ_rel0. apply leq_universe_vertices1; tas. + - apply leq_universe_vertices0. Qed. Definition leqb_level_n n (l l' : Level.t) @@ -1410,12 +1407,12 @@ Section CheckLeq. Definition leqb_level_n_spec_gen (leqb_level_n : Z -> Level.t -> Level.t -> bool) := forall n (l l' : Level.t) (Hl : VSet.In l uctx.1) (Hl' : VSet.In l' uctx.1), leqb_level_n n l l' - <-> gc_leq0_levelalg_n n uctx.2 (LevelAlgExpr.make' l) (LevelAlgExpr.make' l'). + <-> gc_leq0_universe_n n uctx.2 (Universe.make' l) (Universe.make' l'). Lemma leqb_level_n_spec : leqb_level_n_spec_gen leqb_level_n. Proof using HC HG Huctx. unfold leqb_level_n_spec_gen; intros; - symmetry. etransitivity. apply leq_levelalg_vertices; now apply HG. + symmetry. etransitivity. apply leq_universe_vertices; now apply HG. etransitivity. apply leqb_vertices_correct; try exact _. 1-2:now rewrite HG; exact _. now unfold leqb_level_n. Qed. @@ -1428,7 +1425,7 @@ Section CheckLeq. : gc_expr_declared e -> gc_expr_declared e' -> leqb_expr_n_gen leqb_level_n_gen lt e e' -> - gc_leq0_levelalg_n lt uctx.2 (LevelAlgExpr.make e) (LevelAlgExpr.make e'). + gc_leq0_universe_n lt uctx.2 (Universe.make e) (Universe.make e'). Proof using Type. unfold leqb_expr_n. destruct e as [l k], e' as [l' k']; @@ -1448,7 +1445,7 @@ Section CheckLeq. (HHl : gc_expr_declared e) (HHl' : gc_expr_declared e') : leqb_expr_n_gen leqb_level_n_gen ⎩ n ⎭ e e' - <-> gc_leq0_levelalg_n ⎩ n ⎭ uctx.2 (LevelAlgExpr.make e) (LevelAlgExpr.make e'). + <-> gc_leq0_universe_n ⎩ n ⎭ uctx.2 (Universe.make e) (Universe.make e'). Proof using HC. split; [apply (leqb_expr_n_spec0_gen _ leqb_correct)|]; try assumption. destruct e as [l k] eqn:eqe, e' as [l' k'] eqn:eqe'; cbn; intro H; @@ -1470,18 +1467,18 @@ Section CheckLeq. (leqb_correct : leqb_level_n_spec_gen leqb_level_n_gen) n e1 u : gc_expr_declared e1 -> gc_levels_declared u -> leqb_expr_univ_n_gen leqb_level_n_gen n e1 u - -> gc_leq0_levelalg_n n uctx.2 (LevelAlgExpr.make e1) u. + -> gc_leq0_universe_n n uctx.2 (Universe.make e1) u. Proof using Type. unfold leqb_expr_univ_n_gen; intros He1 Hu H. unfold_univ_rel0. rewrite val_fold_right. - destruct (LevelAlgExpr.exprs u) as [e u'] eqn:ee;cbn in *. + destruct (Universe.exprs u) as [e u'] eqn:ee;cbn in *. rewrite <- !fold_left_rev_right in H; cbn in *. red in Hu. assert (Hu': gc_expr_declared e /\ Forall gc_expr_declared u'). { - split. apply Hu. apply In_to_nonempty_list. fold LevelAlgExpr.exprs. left. now rewrite ee. + split. apply Hu. apply In_to_nonempty_list. fold Universe.exprs. left. now rewrite ee. apply Forall_forall. intros e' He'. apply Hu. - apply In_to_nonempty_list. fold LevelAlgExpr.exprs. right. now rewrite ee. } + apply In_to_nonempty_list. fold Universe.exprs. right. now rewrite ee. } destruct Hu' as [He Hu']. apply Forall_rev in Hu'. revert Hu'. induction (List.rev u'); cbn in *; intros. @@ -1496,7 +1493,7 @@ Section CheckLeq. Import Nbar Datatypes. - Lemma val_le_caract' (u : LevelAlgExpr.t) v k : + Lemma val_le_caract' (u : Universe.t) v k : (exists e, LevelExprSet.In e u /\ Z.of_nat k <= Z.of_nat (val v e))%Z <-> (Z.of_nat k <= Z.of_nat (val v u))%Z. Proof using Type. epose proof (val_le_caract u v k). @@ -1509,7 +1506,7 @@ Section CheckLeq. lia. Qed. - Lemma val_ge_caract' (u : LevelAlgExpr.t) v k : + Lemma val_ge_caract' (u : Universe.t) v k : (forall e, LevelExprSet.In e u -> (Z.of_nat (val v e) <= Z.of_nat k)%Z) <-> (Z.of_nat (val v u) <= Z.of_nat k)%Z. Proof using Type. epose proof (val_ge_caract u v k). @@ -1558,13 +1555,13 @@ Section CheckLeq. (* Non trivial lemma *) (* l + n <= max (l1, ... ln) -> exists i, l+n <= li *) - Lemma gc_leq0_levelalg_n_sup lt (l : Level.t) b (u : LevelAlgExpr.t) + Lemma gc_leq0_universe_n_sup lt (l : Level.t) b (u : Universe.t) (e := (l, b)) : gc_level_declared l -> gc_levels_declared u -> - gc_leq0_levelalg_n ⎩ lt ⎭ uctx.2 (LevelAlgExpr.make e) u -> + gc_leq0_universe_n ⎩ lt ⎭ uctx.2 (Universe.make e) u -> exists (e' : LevelExpr.t), LevelExprSet.In e' u - /\ gc_leq0_levelalg_n ⎩ lt ⎭ uctx.2 (LevelAlgExpr.make e) (LevelAlgExpr.make e'). + /\ gc_leq0_universe_n ⎩ lt ⎭ uctx.2 (Universe.make e) (Universe.make e'). Proof using HC HG Huctx. intros Hl Hu H. assert (HG1 : invariants G) by (rewrite HG; exact _). @@ -1612,7 +1609,7 @@ Section CheckLeq. forward Hv; [now rewrite <- HG|]. specialize (H _ Hv). specialize (Hl' _ Hv). specialize (Hl'' _ Hv). - rewrite LevelAlgExpr.val_make in H. + rewrite Universe.val_make in H. rewrite (val_valuation_of_labelling' _ l b) in H; tas. apply switch_minus in H. subst e. @@ -1768,7 +1765,7 @@ Section CheckLeq. pose proof (make_graph_spec' _ Huctx lab) as Hv. forward Hv; [now rewrite <- HG|]. specialize (H _ Hv); clear Hv. - rewrite LevelAlgExpr.val_make in H. + rewrite Universe.val_make in H. rewrite val_valuation_of_labelling' in H; tas. apply switch_minus in H. @@ -1850,23 +1847,23 @@ Section CheckLeq. Lemma leqb_expr_univ_n_spec_gen leqb_level_n_gen (leqb_correct : leqb_level_n_spec_gen leqb_level_n_gen) - lt e1 (u : LevelAlgExpr.t) + lt e1 (u : Universe.t) (He1 : gc_expr_declared e1) (Hu : gc_levels_declared u) : leqb_expr_univ_n_gen leqb_level_n_gen ⎩ lt ⎭ e1 u - <-> gc_leq0_levelalg_n ⎩ lt ⎭ uctx.2 (LevelAlgExpr.make e1) u. + <-> gc_leq0_universe_n ⎩ lt ⎭ uctx.2 (Universe.make e1) u. Proof using HC HG Huctx. split; [eapply leqb_expr_univ_n_spec0_gen; eauto|]. unfold leqb_expr_univ_n_gen; intro HH. - case_eq (LevelAlgExpr.exprs u). intros e u' ee. + case_eq (Universe.exprs u). intros e u' ee. assert (Hu': gc_expr_declared e /\ Forall gc_expr_declared u'). { - split. apply Hu. apply In_to_nonempty_list. fold LevelAlgExpr.exprs. left. now rewrite ee. + split. apply Hu. apply In_to_nonempty_list. fold Universe.exprs. left. now rewrite ee. apply Forall_forall. intros e' He'. apply Hu. - apply In_to_nonempty_list. fold LevelAlgExpr.exprs. right. now rewrite ee. } + apply In_to_nonempty_list. fold Universe.exprs. right. now rewrite ee. } destruct e1 as [l1 b1]. - apply gc_leq0_levelalg_n_sup in HH; tas. + apply gc_leq0_universe_n_sup in HH; tas. destruct HH as [e' [He' HH]]. eapply leqb_expr_n_spec_gen in HH; eauto; tas. - apply In_to_nonempty_list in He'. fold LevelAlgExpr.exprs in He'; rewrite ee in He'; cbn in He'. + apply In_to_nonempty_list in He'. fold Universe.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]. @@ -1878,27 +1875,27 @@ Section CheckLeq. Definition leqb_expr_univ_n_spec := leqb_expr_univ_n_spec_gen _ leqb_level_n_spec. - Definition leqb_levelalg_n := (leqb_levelalg_n_gen leqb_level_n). + Definition leqb_universe_n := (leqb_universe_n_gen leqb_level_n). Lemma fold_right_xpred0 {A} (l : list A) : fold_right (fun _ => xpred0) false l = false. Proof using Type. induction l; simpl; auto. Qed. - Lemma leqb_levelalg_n_spec0_gen leqb_level_n_gen + Lemma leqb_universe_n_spec0_gen leqb_level_n_gen (leqb_correct : leqb_level_n_spec_gen leqb_level_n_gen) - lt (u1 u2 : LevelAlgExpr.t) + lt (u1 u2 : Universe.t) (Hu1 : gc_levels_declared u1) (Hu2 : gc_levels_declared u2) -: leqb_levelalg_n_gen leqb_level_n_gen lt u1 u2 -> gc_leq0_levelalg_n ⎩ lt ⎭ uctx.2 u1 u2. +: leqb_universe_n_gen leqb_level_n_gen lt u1 u2 -> gc_leq0_universe_n ⎩ lt ⎭ uctx.2 u1 u2. Proof using Type. - unfold leqb_levelalg_n_gen. intros H. + unfold leqb_universe_n_gen. intros H. unfold_univ_rel0. - unfold val, LevelAlgExpr.Evaluable. - destruct (LevelAlgExpr.exprs u1) as [e1 u1'] eqn:Hu1'. + unfold val, Universe.Evaluable. + destruct (Universe.exprs u1) as [e1 u1'] eqn:Hu1'. rewrite <- fold_left_rev_right in *; cbn in *. assert (Hu': gc_expr_declared e1 /\ Forall gc_expr_declared u1'). { - split. apply Hu1. apply In_to_nonempty_list. fold LevelAlgExpr.exprs. left. now rewrite Hu1'. + split. apply Hu1. apply In_to_nonempty_list. fold Universe.exprs. left. now rewrite Hu1'. apply Forall_forall. intros e' He'. apply Hu1. - apply In_to_nonempty_list. fold LevelAlgExpr.exprs. right. now rewrite Hu1'. } + apply In_to_nonempty_list. fold Universe.exprs. right. now rewrite Hu1'. } destruct Hu' as [? Hu']. apply Forall_rev in Hu'. revert Hu'. induction (List.rev u1'); cbn in *; intros. + eapply leqb_expr_univ_n_spec0_gen in H; eauto. @@ -1912,19 +1909,19 @@ Section CheckLeq. * now inversion Hu'. Qed. - Definition leqb_levelalg_n_spec0 := leqb_levelalg_n_spec0_gen _ leqb_level_n_spec. + Definition leqb_universe_n_spec0 := leqb_universe_n_spec0_gen _ leqb_level_n_spec. - Lemma leqb_levelalg_n_spec_gen leqb_level_n_gen + Lemma leqb_universe_n_spec_gen leqb_level_n_gen (leqb_correct : leqb_level_n_spec_gen leqb_level_n_gen) - lt (l1 l2 : LevelAlgExpr.t) + lt (l1 l2 : Universe.t) (Hu1 : gc_levels_declared l1) (Hu2 : gc_levels_declared l2) - : leqb_levelalg_n_gen leqb_level_n_gen lt l1 l2 - <-> gc_leq0_levelalg_n ⎩ lt ⎭ uctx.2 l1 l2. + : leqb_universe_n_gen leqb_level_n_gen lt l1 l2 + <-> gc_leq0_universe_n ⎩ lt ⎭ uctx.2 l1 l2. Proof using HC HG Huctx. - split; [eapply leqb_levelalg_n_spec0_gen; eauto |]. - unfold leqb_levelalg_n_gen; intro HH. - unfold LevelAlgExpr.exprs. + split; [eapply leqb_universe_n_spec0_gen; eauto |]. + unfold leqb_universe_n_gen; intro HH. + unfold Universe.exprs. case_eq (to_nonempty_list l1); intros e1 uu1 Huu1. rewrite (fold_left_andb_forallb (fun e => _)). pose proof (to_nonempty_list_spec' l1) as X; rewrite Huu1 in X; cbn in X. @@ -1939,59 +1936,59 @@ Section CheckLeq. eapply (val_ge_caract' l1 v (val v l1)).p2. lia. auto. Qed. - Definition leqb_levelalg_n_spec := leqb_levelalg_n_spec_gen _ leqb_level_n_spec. + Definition leqb_universe_n_spec := leqb_universe_n_spec_gen _ leqb_level_n_spec. - Definition check_leqb_levelalg := (check_leqb_levelalg_gen leqb_level_n). + Definition check_leqb_universe := (check_leqb_universe_gen leqb_level_n). - Lemma check_leqb_levelalg_spec_gen leqb_level_n_gen + Lemma check_leqb_universe_spec_gen leqb_level_n_gen (leqb_correct : leqb_level_n_spec_gen leqb_level_n_gen) - (u1 u2 : LevelAlgExpr.t) + (u1 u2 : Universe.t) (Hu1 : gc_levels_declared u1) (Hu2 : gc_levels_declared u2) - : check_leqb_levelalg_gen leqb_level_n_gen u1 u2 <-> gc_leq_levelalg uctx.2 u1 u2. + : check_leqb_universe_gen leqb_level_n_gen u1 u2 <-> gc_leq_universe uctx.2 u1 u2. Proof using HC HG Huctx. - unfold check_leqb_levelalg_gen, - gc_leq_levelalg, gc_leq_levelalg_n, - leqb_levelalg_n_gen, gc_leq0_levelalg_n. + unfold check_leqb_universe_gen, + gc_leq_universe, gc_leq_universe_n, + leqb_universe_n_gen, gc_leq0_universe_n. destruct check_univs; [|split; trivial]. split; cbn. - move/orP => [|]. + rewrite univ_expr_eqb_true_iff. intros <- v Hv. lia. - + now eapply (leqb_levelalg_n_spec0_gen _ _ false). - - intros H; eapply (leqb_levelalg_n_spec_gen _ _ false) in H; tas. - unfold leqb_levelalg_n_gen in H. rewrite H. + + now eapply (leqb_universe_n_spec0_gen _ _ false). + - intros H; eapply (leqb_universe_n_spec_gen _ _ false) in H; tas. + unfold leqb_universe_n_gen in H. rewrite H. now rewrite orb_true_r. Unshelve. all:eauto. Qed. - Definition check_leqb_levelalg_spec := check_leqb_levelalg_spec_gen _ leqb_level_n_spec. + Definition check_leqb_universe_spec := check_leqb_universe_spec_gen _ leqb_level_n_spec. - Definition check_eqb_levelalg := (check_eqb_levelalg_gen leqb_level_n). + Definition check_eqb_universe := (check_eqb_universe_gen leqb_level_n). - Lemma check_eqb_levelalg_spec_gen leqb_level_n_gen + Lemma check_eqb_universe_spec_gen leqb_level_n_gen (leqb_correct : leqb_level_n_spec_gen leqb_level_n_gen) - (l1 l2 : LevelAlgExpr.t) + (l1 l2 : Universe.t) (Hu1 : gc_levels_declared l1) (Hu2 : gc_levels_declared l2) - : check_eqb_levelalg_gen leqb_level_n_gen l1 l2 <-> gc_eq_levelalg uctx.2 l1 l2. + : check_eqb_universe_gen leqb_level_n_gen l1 l2 <-> gc_eq_universe uctx.2 l1 l2. Proof using HC HG Huctx. - unfold check_eqb_levelalg_gen, gc_eq_levelalg. + unfold check_eqb_universe_gen, gc_eq_universe. destruct check_univs; [|split; trivial]. split; cbn. - move/orP => [ | /andP [Hle Hge]]. + rewrite univ_expr_eqb_true_iff. now intros <- v Hv. - + eapply leqb_levelalg_n_spec0_gen in Hle, Hge; eauto. + + eapply leqb_universe_n_spec0_gen in Hle, Hge; eauto. unfold_univ_rel0. specialize (Hle v Hv); specialize (Hge v Hv). simpl in *. lia. - intros H. toProp; right. - toProp; eapply leqb_levelalg_n_spec_gen; tas; intros v Hv; specialize (H v Hv). + toProp; eapply leqb_universe_n_spec_gen; tas; intros v Hv; specialize (H v Hv). rewrite H. cbn; lia. rewrite H. cbn; lia. Qed. - Definition check_eqb_levelalg_spec := check_eqb_levelalg_spec_gen _ leqb_level_n_spec. + Definition check_eqb_universe_spec := check_eqb_universe_spec_gen _ leqb_level_n_spec. Lemma fold_left_false {A} l : fold_left (B:=A) (fun _ : bool => xpred0) l false = false. Proof using Type. @@ -2070,56 +2067,54 @@ Section CheckLeq. Definition eqb_univ_instance := (eqb_univ_instance_gen leqb_level_n). - Definition leqb_universe := (leqb_universe_gen leqb_level_n). + Definition leqb_sort := (leqb_sort_gen leqb_level_n). - Definition check_leqb_universe := (check_leqb_universe_gen leqb_level_n). + Definition check_leqb_sort := (check_leqb_sort_gen leqb_level_n). - Definition check_eqb_universe := (check_eqb_universe_gen leqb_level_n). + Definition check_eqb_sort := (check_eqb_sort_gen leqb_level_n). - Lemma check_eqb_universe_refl_gen leqb_level_n_gen + Lemma check_eqb_sort_refl_gen leqb_level_n_gen (leqb_correct : leqb_level_n_spec_gen leqb_level_n_gen) u : - check_eqb_universe_gen leqb_level_n_gen u u. + check_eqb_sort_gen leqb_level_n_gen u u. Proof using Type. - unfold check_eqb_universe_gen; toProp; left. - apply Universe.eqb_refl. + unfold check_eqb_sort_gen; toProp; left. + apply eqb_refl. Qed. - Definition check_eqb_universe_refl := check_eqb_universe_refl_gen _ leqb_level_n_spec. + Definition check_eqb_sort_refl := check_eqb_sort_refl_gen _ leqb_level_n_spec. - Definition gc_leq_universe := - leq_universe_n_ (fun n φ u u' => if check_univs then gc_leq0_levelalg_n n φ u u' else True) 0. + Definition gc_leq_sort φ := + leq_sort_n_ (fun n u u' => if check_univs then gc_leq0_universe_n n φ u u' else True) 0. - Definition gc_eq_universe := - eq_universe_ (fun φ u u' => if check_univs then gc_eq0_levelalg φ u u' else True). + Definition gc_eq_sort φ := + eq_sort_ (fun u u' => if check_univs then gc_eq0_universe φ u u' else True). - Let levels_declared_univ (u : Universe.t) := - match u with - | Universe.lSProp | Universe.lProp => True - | Universe.lType l => gc_levels_declared l - end. + Let levels_declared_sort (s : Sort.t) := + Sort.on_sort gc_levels_declared True s. - Lemma check_eqb_universe_spec_gen leqb_level_n_gen + Lemma check_eqb_sort_spec_gen leqb_level_n_gen (leqb_correct : leqb_level_n_spec_gen leqb_level_n_gen) - (u1 u2 : Universe.t) - (Hu1 : levels_declared_univ u1) - (Hu2 : levels_declared_univ u2) - : check_eqb_universe_gen leqb_level_n_gen u1 u2 <-> gc_eq_universe uctx.2 u1 u2. + (u1 u2 : Sort.t) + (Hu1 : levels_declared_sort u1) + (Hu2 : levels_declared_sort u2) + : check_eqb_sort_gen leqb_level_n_gen u1 u2 <-> gc_eq_sort uctx.2 u1 u2. Proof. - unfold check_eqb_universe_gen, gc_eq_universe. + unfold check_eqb_sort_gen, gc_eq_sort. destruct u1, u2; cbnr; split; intuition auto. - now destruct prop_sub_type. - - eapply check_eqb_levelalg_spec_gen; eauto; tas. - unfold check_eqb_universe_gen, check_eqb_levelalg_gen in *; cbn in *. - unfold check_leqb_levelalg_gen in *. + - eapply check_eqb_universe_spec_gen; eauto; tas. + unfold check_eqb_sort_gen, check_eqb_universe_gen in *; cbn in *. + unfold check_leqb_universe_gen in *. destruct check_univs; cbnr. - move/orP: H => [-> | /andP [/orP [/orP [Hf | ->] | H1] /orP [/orP [Hf' | e] | H2]]] //. + unfold eqb at 1, Sort.reflect_eq_sort, Sort.eqb in H. cbn in H. + move/orP : H => /= [-> //|] /andP[] /orP[-> //|] H1 /orP[e | H2]. 1: apply NonEmptySetFacts.univ_expr_eqb_true_iff in e as ->. 1: toProp; left; now apply NonEmptySetFacts.univ_expr_eqb_true_iff. toProp; right; now toProp. - toProp; right. - eapply check_eqb_levelalg_spec_gen in H; eauto; tas. - unfold check_eqb_levelalg_gen in *; cbn in *. - unfold check_leqb_levelalg_gen in *. + eapply check_eqb_universe_spec_gen in H; eauto; tas. + unfold check_eqb_universe_gen in *; cbn in *. + unfold check_leqb_universe_gen in *. destruct check_univs; [cbn in * | trivial]. move/orP : H => [H | /andP [H1 H2]]. + apply NonEmptySetFacts.univ_expr_eqb_true_iff in H as ->. @@ -2127,7 +2122,7 @@ Section CheckLeq. + toProp; toProp; right; assumption. Defined. - Definition check_eqb_universe_spec := check_eqb_universe_spec_gen _ leqb_level_n_spec. + Definition check_eqb_sort_spec := check_eqb_sort_spec_gen _ leqb_level_n_spec. End CheckLeq. @@ -2178,7 +2173,7 @@ Section CheckLeq2. := on_Some_or_None (fun l : Level.t => level_declared l) (LevelExpr.get_noprop e). - Let levels_declared (u : LevelAlgExpr.t) := + Let levels_declared (u : Universe.t) := LevelExprSet.For_all expr_declared u. Lemma level_gc_declared_declared l @@ -2198,7 +2193,7 @@ Section CheckLeq2. intro; now apply (level_gc_declared_declared l) in H. Qed. - Lemma levels_gc_declared_declared (u : LevelAlgExpr.t) + Lemma levels_gc_declared_declared (u : Universe.t) : levels_declared u -> gc_levels_declared uctx' u. Proof using HG expr_declared. unfold levels_declared, gc_levels_declared. @@ -2212,13 +2207,13 @@ Section CheckLeq2. (He1 : expr_declared e1) (Hu : levels_declared u) : leqb_expr_univ_n_gen leqb_level_n_gen ⎩ lt ⎭ e1 u - <-> leq0_levelalg_n ⎩ lt ⎭ uctx.2 (LevelAlgExpr.make e1) u. + <-> leq0_universe_n ⎩ lt ⎭ uctx.2 (Universe.make e1) u. Proof using HG' Huctx'. etransitivity. eapply (leqb_expr_univ_n_spec_gen G uctx' Huctx' HC' HG'); eauto; tas. - apply expr_gc_declared_declared; tas. - apply levels_gc_declared_declared; tas. - - symmetry. etransitivity. apply gc_leq0_levelalg_n_iff. + - symmetry. etransitivity. apply gc_leq0_universe_n_iff. unfold uctx'; cbn; clear -HG. unfold is_graph_of_uctx, gc_of_uctx in *. destruct (gc_of_constraints uctx.2) as [ctrs|]. @@ -2228,19 +2223,19 @@ Section CheckLeq2. Definition leqb_univ_expr_n_spec' := leqb_univ_expr_n_spec_gen' _ (leqb_level_n_spec _ _ Huctx' HC' HG'). - Lemma check_leqb_levelalg_spec_gen' leqb_level_n_gen + Lemma check_leqb_universe_spec_gen' leqb_level_n_gen (leqb_correct : leqb_level_n_spec_gen uctx' leqb_level_n_gen) u1 u2 : levels_declared u1 -> levels_declared u2 -> - check_leqb_levelalg_gen leqb_level_n_gen u1 u2 -> leq_levelalg uctx.2 u1 u2. + check_leqb_universe_gen leqb_level_n_gen u1 u2 -> leq_universe uctx.2 u1 u2. Proof using HG' Huctx'. - unfold check_leqb_levelalg_gen; intros Hu1 Hu2 H. + unfold check_leqb_universe_gen; intros Hu1 Hu2 H. unfold_univ_rel. cbn in H; toProp H; destruct H as [e | ]. { apply NonEmptySetFacts.univ_expr_eqb_true_iff in e. destruct e; lia. } - eapply leqb_levelalg_n_spec0_gen in H; eauto. - eapply gc_leq0_levelalg_iff; tea. + eapply leqb_universe_n_spec0_gen in H; eauto. + eapply gc_leq0_universe_iff; tea. unfold uctx' in *. unfold is_graph_of_uctx, gc_of_uctx in HG. destruct (gc_of_constraints uctx.2). cbn in *. exact H. @@ -2248,80 +2243,80 @@ Section CheckLeq2. Unshelve. all: try eapply levels_gc_declared_declared; eauto. Qed. - Definition check_leqb_levelalg_spec' := - check_leqb_levelalg_spec_gen' _ (leqb_level_n_spec _ _ Huctx' HC' HG'). + Definition check_leqb_universe_spec' := + check_leqb_universe_spec_gen' _ (leqb_level_n_spec _ _ Huctx' HC' HG'). - Lemma check_leqb_levelalg_complete_gen leqb_level_n_gen + Lemma check_leqb_universe_complete_gen leqb_level_n_gen (leqb_correct : leqb_level_n_spec_gen uctx' leqb_level_n_gen) u1 u2 : levels_declared u1 -> levels_declared u2 -> - leq_levelalg uctx.2 u1 u2 -> - check_leqb_levelalg_gen leqb_level_n_gen u1 u2. + leq_universe uctx.2 u1 u2 -> + check_leqb_universe_gen leqb_level_n_gen u1 u2. Proof using HG' Huctx'. intros decl1 decl2. apply levels_gc_declared_declared in decl1. apply levels_gc_declared_declared in decl2. - rewrite gc_leq_levelalg_iff. + rewrite gc_leq_universe_iff. unfold is_graph_of_uctx, gc_of_uctx in HG. unfold uctx' in *. destruct gc_of_constraints; [cbn in *|contradiction HG]. intros eq. - apply <- check_leqb_levelalg_spec_gen; eauto. + apply <- check_leqb_universe_spec_gen; eauto. exact eq. Qed. - Definition check_leqb_levelalg_complete := - check_leqb_levelalg_complete_gen _ (leqb_level_n_spec _ _ Huctx' HC' HG'). + Definition check_leqb_universe_complete := + check_leqb_universe_complete_gen _ (leqb_level_n_spec _ _ Huctx' HC' HG'). - Lemma check_eqb_levelalg_spec_gen' leqb_level_n_gen + Lemma check_eqb_universe_spec_gen' leqb_level_n_gen (leqb_correct : leqb_level_n_spec_gen uctx' leqb_level_n_gen) u1 u2 : levels_declared u1 -> levels_declared u2 -> - check_eqb_levelalg_gen leqb_level_n_gen u1 u2 -> eq_levelalg uctx.2 u1 u2. + check_eqb_universe_gen leqb_level_n_gen u1 u2 -> eq_universe uctx.2 u1 u2. Proof using HG' Huctx'. - unfold check_eqb_levelalg_gen; intros Hu1 Hu2 H. + unfold check_eqb_universe_gen; intros Hu1 Hu2 H. unfold_univ_rel. cbn in H; toProp H; destruct H as [e | ]. { apply NonEmptySetFacts.univ_expr_eqb_true_iff in e. destruct e; lia. } apply andb_prop in H. destruct H as [H1 H2]. - unshelve eapply leqb_levelalg_n_spec0_gen in H1; eauto. - unshelve eapply leqb_levelalg_n_spec0_gen in H2; eauto. + unshelve eapply leqb_universe_n_spec0_gen in H1; eauto. + unshelve eapply leqb_universe_n_spec0_gen in H2; eauto. unfold uctx' in *. unfold is_graph_of_uctx, gc_of_uctx in HG. - apply <- eq0_leq0_levelalg; tea. - split; eapply gc_leq0_levelalg_iff; + apply <- eq0_leq0_universe; tea. + split; eapply gc_leq0_universe_iff; (destruct (gc_of_constraints uctx.2); [cbn in *|contradiction HG]); tas. all: now eapply levels_gc_declared_declared. Qed. - Definition check_eqb_levelalg_spec' := - check_eqb_levelalg_spec_gen' _ (leqb_level_n_spec _ _ Huctx' HC' HG'). + Definition check_eqb_universe_spec' := + check_eqb_universe_spec_gen' _ (leqb_level_n_spec _ _ Huctx' HC' HG'). - Lemma check_eqb_levelalg_complete_gen leqb_level_n_gen + Lemma check_eqb_universe_complete_gen leqb_level_n_gen (leqb_correct : leqb_level_n_spec_gen uctx' leqb_level_n_gen) u1 u2 : levels_declared u1 -> levels_declared u2 -> - eq_levelalg uctx.2 u1 u2 -> - check_eqb_levelalg_gen leqb_level_n_gen u1 u2. + eq_universe uctx.2 u1 u2 -> + check_eqb_universe_gen leqb_level_n_gen u1 u2. Proof using HG' Huctx'. intros decl1 decl2. apply levels_gc_declared_declared in decl1. apply levels_gc_declared_declared in decl2. - rewrite gc_eq_levelalg_iff. + rewrite gc_eq_universe_iff. unfold is_graph_of_uctx, gc_of_uctx in HG. unfold uctx' in *. destruct gc_of_constraints; [cbn in *|contradiction HG]. intros eq. - apply <- check_eqb_levelalg_spec_gen; eauto. + apply <- check_eqb_universe_spec_gen; eauto. exact eq. Qed. - Definition check_eqb_levelalg_complete := - check_eqb_levelalg_complete_gen _ (leqb_level_n_spec _ _ Huctx' HC' HG'). + Definition check_eqb_universe_complete := + check_eqb_universe_complete_gen _ (leqb_level_n_spec _ _ Huctx' HC' HG'). Definition leq0_level_n z l l' := - leq0_levelalg_n z uctx.2 (LevelAlgExpr.make' l) (LevelAlgExpr.make' l'). + leq0_universe_n z uctx.2 (Universe.make' l) (Universe.make' l'). Definition valid_gc_constraint (gc : GoodConstraint.t) := match gc with @@ -2344,7 +2339,7 @@ Section CheckLeq2. Proof using HG' Huctx'. intros decll decll'. unfold leq0_level_n. - intros le; eapply gc_leq0_levelalg_n_iff in le. + intros le; eapply gc_leq0_universe_n_iff in le. unfold is_graph_of_uctx, gc_of_uctx in HG. unfold uctx' in *. destruct gc_of_constraints; [cbn in *|contradiction HG]. @@ -2482,65 +2477,62 @@ Section CheckLeq2. Definition check_constraints_complete := check_constraints_complete_gen _ (leqb_level_n_spec _ _ Huctx' HC' HG'). - Let levels_declared_univ (u : Universe.t) - := match u with - | Universe.lSProp | Universe.lProp => True - | Universe.lType l => levels_declared l - end. + Let levels_declared_sort (s : Sort.t) + := Sort.on_sort levels_declared True s. - Lemma levels_univ_gc_declared_declared (u : Universe.t) - : levels_declared_univ u -> gc_levels_declared_univ uctx' u. + Lemma levels_univ_gc_declared_declared (s : Sort.t) + : levels_declared_sort s -> gc_levels_declared_sort uctx' s. Proof using HG levels_declared. - destruct u; cbnr. + destruct s; cbnr. apply levels_gc_declared_declared. Qed. - Lemma check_leqb_universe_spec_gen' leqb_level_n_gen - (leqb_correct : leqb_level_n_spec_gen uctx' leqb_level_n_gen) u1 u2 - : levels_declared_univ u1 -> - levels_declared_univ u2 -> - check_leqb_universe_gen leqb_level_n_gen u1 u2 -> leq_universe uctx.2 u1 u2. + Lemma check_leqb_sort_spec_gen' leqb_level_n_gen + (leqb_correct : leqb_level_n_spec_gen uctx' leqb_level_n_gen) s1 s2 + : levels_declared_sort s1 -> + levels_declared_sort s2 -> + check_leqb_sort_gen leqb_level_n_gen s1 s2 -> leq_sort uctx.2 s1 s2. Proof using HG' Huctx'. intros Hu1 Hu2. move => /orP [H | H]. - apply eqb_true_iff in H as ->. reflexivity. - - destruct u1, u2; cbn in *; trivial; try discriminate H. - now eapply check_leqb_levelalg_spec_gen'. + - destruct s1, s2; cbn in *; trivial; try discriminate H. + now eapply check_leqb_universe_spec_gen'. Qed. - Definition check_leqb_universe_spec' := - check_leqb_universe_spec_gen' _ (leqb_level_n_spec _ _ Huctx' HC' HG'). + Definition check_leqb_sort_spec' := + check_leqb_sort_spec_gen' _ (leqb_level_n_spec _ _ Huctx' HC' HG'). - Lemma check_leqb_universe_complete_gen leqb_level_n_gen - (leqb_correct : leqb_level_n_spec_gen uctx' leqb_level_n_gen) u1 u2 : - levels_declared_univ u1 -> - levels_declared_univ u2 -> - leq_universe uctx.2 u1 u2 -> - check_leqb_universe_gen leqb_level_n_gen u1 u2. + Lemma check_leqb_sort_complete_gen leqb_level_n_gen + (leqb_correct : leqb_level_n_spec_gen uctx' leqb_level_n_gen) s1 s2 : + levels_declared_sort s1 -> + levels_declared_sort s2 -> + leq_sort uctx.2 s1 s2 -> + check_leqb_sort_gen leqb_level_n_gen s1 s2. Proof using HG' Huctx'. - move : u1 u2 => [| | u1] [| | u2] //. cbn. + move : s1 s2 => [| | u1] [| | u2] //. cbn. intros decl1 decl2 Hle. - unfold check_leqb_universe_gen. + unfold check_leqb_sort_gen. toProp; right. - apply check_leqb_levelalg_complete_gen => //. + apply check_leqb_universe_complete_gen => //. Qed. - Definition check_leqb_universe_complete := - check_leqb_universe_complete_gen _ (leqb_level_n_spec _ _ Huctx' HC' HG'). + Definition check_leqb_sort_complete := + check_leqb_sort_complete_gen _ (leqb_level_n_spec _ _ Huctx' HC' HG'). - Lemma check_eqb_universe_spec_gen' leqb_level_n_gen - (leqb_correct : leqb_level_n_spec_gen uctx' leqb_level_n_gen) u1 u2 - : levels_declared_univ u1 -> - levels_declared_univ u2 -> - check_eqb_universe_gen leqb_level_n_gen u1 u2 -> eq_universe uctx.2 u1 u2. + Lemma check_eqb_sort_spec_gen' leqb_level_n_gen + (leqb_correct : leqb_level_n_spec_gen uctx' leqb_level_n_gen) s1 s2 + : levels_declared_sort s1 -> + levels_declared_sort s2 -> + check_eqb_sort_gen leqb_level_n_gen s1 s2 -> eq_sort uctx.2 s1 s2. Proof using HG' Huctx'. - move : u1 u2 => [| | u1] [| | u2] //; intros Hu1 Hu2. + move : s1 s2 => [| | u1] [| | u2] //; intros Hu1 Hu2. { move/andP => [H HH] //. } move/orP => [H | H]. - apply eqb_true_iff in H as ->. reflexivity. - - eapply check_eqb_levelalg_spec_gen'; eauto. - cbn in H. unfold check_eqb_levelalg_gen in *. + - eapply check_eqb_universe_spec_gen'; eauto. + cbn in H. unfold check_eqb_universe_gen in *. move/andP: H => [/orP [/orP [-> | ->] | ->] /orP [/orP [He | HH] | ->]] //. all: try now rewrite orb_true_r. now rewrite He. @@ -2548,28 +2540,29 @@ Section CheckLeq2. toProp; left; toProp; right; now apply NonEmptySetFacts.univ_expr_eqb_true_iff. Qed. - Definition check_eqb_universe_spec' := - check_eqb_universe_spec_gen' _ (leqb_level_n_spec _ _ Huctx' HC' HG'). + Definition check_eqb_sort_spec' := + check_eqb_sort_spec_gen' _ (leqb_level_n_spec _ _ Huctx' HC' HG'). - Lemma check_eqb_universe_complete_gen leqb_level_n_gen + Lemma check_eqb_sort_complete_gen leqb_level_n_gen (leqb_correct : leqb_level_n_spec_gen uctx' leqb_level_n_gen) u1 u2 : - levels_declared_univ u1 -> - levels_declared_univ u2 -> - eq_universe uctx.2 u1 u2 -> - check_eqb_universe_gen leqb_level_n_gen u1 u2. + levels_declared_sort u1 -> + levels_declared_sort u2 -> + eq_sort uctx.2 u1 u2 -> + check_eqb_sort_gen leqb_level_n_gen u1 u2. Proof using HG' Huctx'. move : u1 u2 => [| | u1] [| | u2] //. cbn. intros decl1 decl2 Hle. - eapply check_eqb_levelalg_complete_gen in Hle => //; eauto. - unfold check_eqb_universe_gen, leqb_universe_gen, check_leqb_levelalg_gen; cbn. - unfold check_eqb_levelalg_gen in Hle. - move/orP: Hle => [/orP [-> | ->] | /andP [H1 H2]] //. + eapply check_eqb_universe_complete_gen in Hle => //; eauto. + unfold check_eqb_sort_gen, leqb_sort_gen, check_leqb_universe_gen; cbn. + unfold check_eqb_universe_gen in Hle. + move/orP: Hle => [/orP [-> | e] | /andP [H1 H2]] //=. now rewrite orb_true_r. + apply eqb_eq in e as ->; rewrite eqb_refl //. toProp; right; toProp; toProp; right; assumption. Qed. - Definition check_eqb_universe_complete := - check_eqb_universe_complete_gen _ (leqb_level_n_spec _ _ Huctx' HC' HG'). + Definition check_eqb_sort_complete := + check_eqb_sort_complete_gen _ (leqb_level_n_spec _ _ Huctx' HC' HG'). End CheckLeq2. diff --git a/erasure/theories/EArities.v b/erasure/theories/EArities.v index cc7c33887..f167c829a 100644 --- a/erasure/theories/EArities.v +++ b/erasure/theories/EArities.v @@ -35,16 +35,17 @@ Qed. Lemma isType_isErasable Σ Γ T : isType Σ Γ T -> isErasable Σ Γ T. Proof. - intros [s Hs]. + intros (_ & s & Hs & _). exists (tSort s). intuition auto. Qed. Lemma isType_red: forall (Σ : global_env_ext) (Γ : context) (T : term), wf Σ -> wf_local Σ Γ -> - isType Σ Γ T -> forall x5 : term, red Σ Γ T x5 -> isType Σ Γ x5. + isType Σ Γ T -> forall T' : term, red Σ Γ T T' -> isType Σ Γ T'. Proof. - intros. destruct X1 as []. - eexists. eapply subject_reduction ; eauto. + intros. + apply lift_sorting_it_impl_gen with X1 => // HT. + eapply subject_reduction ; eauto. Qed. Lemma it_mkProd_isArity: @@ -262,7 +263,7 @@ Lemma nIs_conv_to_Arity_nArity {Σ : global_env_ext} {wfΣ : wf Σ} {Γ T} : Proof. intros isty nisc isa. apply nisc. exists T. split => //. sq. - destruct isty as [s Hs]. + destruct isty as (_ & s & Hs & _). eapply wt_closed_red_refl; tea. Qed. @@ -328,48 +329,47 @@ Section Elim'. Context `{cf : checker_flags}. Context {Σ : global_env_ext} {wfΣ : wf_ext Σ}. Variable Hcf : prop_sub_type = false. -Variable Hcf' : check_univs. Lemma cumul_prop1 Γ A B u : - Universe.is_prop u -> + Sort.is_prop u -> isType Σ Γ A -> Σ ;;; Γ |- B : tSort u -> Σ ;;; Γ ⊢ A ≤ B -> Σ ;;; Γ |- A : tSort u. -Proof using Hcf Hcf' wfΣ. +Proof using Hcf wfΣ. intros; eapply cumul_prop1; tea. now apply ws_cumul_pb_forget in X1. Qed. Lemma cumul_prop2 Γ A B u : - Universe.is_prop u -> + Sort.is_prop u -> isType Σ Γ B -> Σ ;;; Γ ⊢ A ≤ B -> Σ ;;; Γ |- A : tSort u -> Σ ;;; Γ |- B : tSort u. -Proof using Hcf Hcf' wfΣ. +Proof using Hcf wfΣ. intros. eapply cumul_prop2; tea. now apply ws_cumul_pb_forget in X0. Qed. Lemma cumul_sprop1 Γ A B u : - Universe.is_sprop u -> + Sort.is_sprop u -> isType Σ Γ A -> Σ ;;; Γ |- B : tSort u -> Σ ;;; Γ ⊢ A ≤ B -> Σ ;;; Γ |- A : tSort u. -Proof using Hcf Hcf' wfΣ. +Proof using Hcf wfΣ. intros. eapply cumul_sprop1; tea. now apply ws_cumul_pb_forget in X1. Qed. Lemma cumul_sprop2 Γ A B u : - Universe.is_sprop u -> + Sort.is_sprop u -> isType Σ Γ B -> Σ ;;; Γ ⊢ A ≤ B -> Σ ;;; Γ |- A : tSort u -> Σ ;;; Γ |- B : tSort u. -Proof using Hcf Hcf' wfΣ. +Proof using Hcf wfΣ. intros. eapply cumul_sprop2; tea. now apply ws_cumul_pb_forget in X0. Qed. @@ -377,27 +377,26 @@ End Elim'. Lemma cumul_propositional (Σ : global_env_ext) Γ A B u : wf_ext Σ -> - is_propositional u -> + Sort.is_propositional u -> isType Σ Γ B -> Σ ;;; Γ ⊢ A ≤ B -> Σ ;;; Γ |- A : tSort u -> Σ ;;; Γ |- B : tSort u. Proof. intros wf. + intros pu isTy cum Ha. destruct u => //. - intros _ [s Hs] cum Ha. - eapply cumul_prop2; eauto. now exists s. - intros _ [s Hs] cum Ha. - eapply cumul_sprop2; eauto. now exists s. + eapply cumul_prop2; eauto. + eapply cumul_sprop2; eauto. Qed. Lemma sort_typing_spine: - forall (Σ : global_env_ext) (Γ : context) (L : list term) (u : Universe.t) (x x0 : term), + forall (Σ : global_env_ext) (Γ : context) (L : list term) (u : sort) (x x0 : term), wf_ext Σ -> - is_propositional u -> + Sort.is_propositional u -> typing_spine Σ Γ x L x0 -> Σ;;; Γ |- x : tSort u -> - ∑ u', Σ;;; Γ |- x0 : tSort u' × is_propositional u'. + ∑ u', Σ;;; Γ |- x0 : tSort u' × Sort.is_propositional u'. Proof. intros Σ Γ L u x x0 HΣ ? t1 c0. assert (X : wf Σ) by apply HΣ. @@ -408,16 +407,10 @@ Proof. - eapply cumul_propositional in c0; auto. 2-3: tea. eapply inversion_Prod in c0 as (? & ? & ? & ? & e0) ; auto. eapply ws_cumul_pb_Sort_inv in e0. - unfold is_propositional in H. - destruct (Universe.is_prop u) eqn:isp => //. - eapply leq_universe_prop_r in e0 as H0; cbn; eauto. - eapply is_prop_sort_prod in H0. eapply IHt1; [unfold is_propositional; now rewrite -> H0|]. - change (tSort x0) with ((tSort x0) {0 := hd}). - eapply substitution0; eauto. - eapply leq_universe_sprop_r in e0 as H0; cbn; eauto. - eapply is_sprop_sort_prod in H0. eapply IHt1; [unfold is_propositional; now rewrite -> H0, orb_true_r|]. - change (tSort x0) with ((tSort x0) {0 := hd}). - eapply substitution0; eauto. + eapply IHt1. + 2: eapply @substitution0 with (T := tSort _); tea. + unfold compare_sort, leq_sort in *. + destruct u, x0, x => //. Qed. Lemma arity_type_inv (Σ : global_env_ext) Γ t T1 T2 : wf_ext Σ -> wf_local Σ Γ -> @@ -428,68 +421,58 @@ Proof. eapply invert_cumul_arity_l_gen; tea. eapply invert_cumul_arity_r_gen. 2:exact e. exists T1. split; auto. sq. - eapply PCUICValidity.validity in X as [s Hs]. + eapply PCUICValidity.validity in X as (_ & s & Hs & _). eapply wt_closed_red_refl; eauto. Qed. Lemma cumul_prop1' (Σ : global_env_ext) Γ A B u : - check_univs -> wf_ext Σ -> isType Σ Γ A -> - is_propositional u -> + Sort.is_propositional u -> Σ ;;; Γ |- B : tSort u -> Σ ;;; Γ ⊢ A ≤ B -> Σ ;;; Γ |- A : tSort u. Proof. intros. - destruct X0 as [s Hs]. destruct u => //. - eapply cumul_prop1 in X2; eauto. now exists s. - eapply cumul_sprop1 in X2; eauto. now exists s. + eapply cumul_prop1 in X2; eauto. + eapply cumul_sprop1 in X2; eauto. Qed. Lemma cumul_prop2' (Σ : global_env_ext) Γ A B u : - check_univs -> wf_ext Σ -> isType Σ Γ A -> - is_propositional u -> + Sort.is_propositional u -> Σ ;;; Γ |- B : tSort u -> Σ ;;; Γ ⊢ B ≤ A -> Σ ;;; Γ |- A : tSort u. Proof. intros. - destruct X0 as [s Hs]. destruct u => //. - eapply cumul_prop2 in X2; eauto. now exists s. - eapply cumul_sprop2 in X2; eauto. now exists s. + eapply cumul_prop2 in X2; eauto. + eapply cumul_sprop2 in X2; eauto. Qed. Lemma leq_term_propositional_sorted_l {Σ Γ v v' u u'} : wf_ext Σ -> PCUICEquality.leq_term Σ (global_ext_constraints Σ) v v' -> Σ;;; Γ |- v : tSort u -> - Σ;;; Γ |- v' : tSort u' -> is_propositional u -> - leq_universe (global_ext_constraints Σ) u' u. + Σ;;; Γ |- v' : tSort u' -> Sort.is_propositional u -> + leq_sort (global_ext_constraints Σ) u' u. Proof. intros wf leq Hv Hv' isp. - unfold is_propositional in isp. - destruct u => //. eapply leq_term_prop_sorted_l; eauto. - eapply leq_term_sprop_sorted_l; eauto. Qed. Lemma leq_term_propopositional_sorted_r {Σ Γ v v' u u'} : wf_ext Σ -> PCUICEquality.leq_term Σ (global_ext_constraints Σ) v v' -> Σ;;; Γ |- v : tSort u -> - Σ;;; Γ |- v' : tSort u' -> is_propositional u' -> - leq_universe (global_ext_constraints Σ) u u'. + Σ;;; Γ |- v' : tSort u' -> Sort.is_propositional u' -> + leq_sort (global_ext_constraints Σ) u u'. Proof. intros wfΣ leq hv hv' isp. - unfold is_propositional in isp. - destruct u' => //. eapply leq_term_prop_sorted_r; eauto. - eapply leq_term_sprop_sorted_r; eauto. Qed. Lemma Is_type_app (Σ : global_env_ext) Γ t L T : @@ -528,45 +511,23 @@ Proof. now apply PCUICValidity.validity in t2. Qed. -Lemma leq_universe_propositional_r {cf : checker_flags} (ϕ : ConstraintSet.t) (u1 u2 : Universe.t_) : - check_univs -> - consistent ϕ -> - leq_universe ϕ u1 u2 -> is_propositional u2 -> is_propositional u1. +Lemma leq_sort_propositional_r {cf : checker_flags} (ϕ : ConstraintSet.t) (u1 u2 : sort) : + leq_sort ϕ u1 u2 -> Sort.is_propositional u2 -> Sort.is_propositional u1. Proof. - intros cu cons leq; unfold is_propositional. - destruct u2 => //. - apply leq_universe_prop_r in leq => //. - now rewrite leq. - intros _. - apply leq_universe_sprop_r in leq => //. - now rewrite leq orb_true_r. + destruct u1, u2 => //. Qed. -Lemma leq_universe_propositional_l {cf : checker_flags} (ϕ : ConstraintSet.t) (u1 u2 : Universe.t_) : - check_univs -> +Lemma leq_sort_propositional_l {cf : checker_flags} (ϕ : ConstraintSet.t) (u1 u2 : sort) : prop_sub_type = false -> - consistent ϕ -> - leq_universe ϕ u1 u2 -> is_propositional u1 -> is_propositional u2. + leq_sort ϕ u1 u2 -> Sort.is_propositional u1 -> Sort.is_propositional u2. Proof. - intros cu ps cons leq; unfold is_propositional. - destruct u1 => //. - eapply leq_universe_prop_no_prop_sub_type in leq => //. - now rewrite leq. - intros _. - apply leq_universe_sprop_l in leq => //. - now rewrite leq orb_true_r. + destruct u1, u2 => //= -> //. Qed. Lemma is_propositional_sort_prod x2 x3 : - is_propositional (Universe.sort_of_product x2 x3) -> is_propositional x3. + Sort.is_propositional (Sort.sort_of_product x2 x3) -> Sort.is_propositional x3. Proof. - unfold is_propositional. - destruct (Universe.is_prop (Universe.sort_of_product x2 x3)) eqn:eq => //. - simpl. - intros _. - apply is_prop_sort_prod in eq. now rewrite eq. - destruct (Universe.is_sprop (Universe.sort_of_product x2 x3)) eqn:eq' => //. - apply is_sprop_sort_prod in eq'. now rewrite eq' !orb_true_r. + destruct x2, x3 => //. Qed. Lemma Is_type_lambda (Σ : global_env_ext) Γ na T1 t : @@ -576,20 +537,20 @@ Lemma Is_type_lambda (Σ : global_env_ext) Γ na T1 t : ∥isErasable Σ (vass na T1 :: Γ) t∥. Proof. intros ? ? (T & ? & ?). - eapply inversion_Lambda in t0 as (? & ? & ? & ? & e); auto. + eapply inversion_Lambda in t0 as (? & h1 & ? & e); auto. destruct s as [ | (u & ? & ?)]. - eapply invert_cumul_arity_r in e; eauto. destruct e as (? & [] & ?). eapply invert_red_prod in X1 as (? & ? & []); eauto; subst. cbn in H. - econstructor. exists x3. econstructor. + econstructor. exists x2. econstructor. eapply type_reduction_closed; eauto. econstructor; eauto. - sq. eapply cumul_prop1' in e; eauto. eapply inversion_Prod in e as (? & ? & ? & ? & e) ; auto. eapply ws_cumul_pb_Sort_inv in e. - eapply leq_universe_propositional_r in e as H0; cbn; eauto. + eapply leq_sort_propositional_r in e as H0; cbn; eauto. eexists. split. eassumption. right. eexists. split. eassumption. eapply is_propositional_sort_prod in H0; eauto. - eapply type_Lambda in t1; eauto. - now apply PCUICValidity.validity in t1. + eapply type_Lambda in h1; eauto. + now apply PCUICValidity.validity in h1. Qed. Lemma Is_type_red (Σ : global_env_ext) Γ t v: @@ -644,7 +605,7 @@ Qed. Lemma isType_closed_red_refl {Σ} {wfΣ : wf Σ} {Γ T} : isType Σ Γ T -> Σ ;;; Γ ⊢ T ⇝ T. Proof. - intros [s hs]; eapply wt_closed_red_refl; tea. + intros (_ & s & hs & _); eapply wt_closed_red_refl; tea. Qed. Lemma nIs_conv_to_Arity_isWfArity_elim {Σ} {wfΣ : wf Σ} {Γ x} : @@ -662,7 +623,7 @@ Qed. Definition isErasable_Type (Σ : global_env_ext) Γ T := (Is_conv_to_Arity Σ Γ T + - (∑ u : Universe.t, Σ;;; Γ |- T : tSort u × is_propositional u))%type. + (∑ u : sort, Σ;;; Γ |- T : tSort u × Sort.is_propositional u))%type. Lemma isErasable_any_type {Σ} {wfΣ : wf_ext Σ} {Γ t T} : isErasable Σ Γ t -> @@ -702,68 +663,54 @@ Proof. Qed. -Lemma is_propositional_bottom {Σ Γ T s s'} : - wf_ext Σ -> - check_univs -> - prop_sub_type = false -> - Σ ;;; Γ ⊢ T ≤ tSort s -> - Σ ;;; Γ ⊢ T ≤ tSort s' -> - PCUICCumulProp.eq_univ_prop s s'. -Proof. - intros wf cu pst h h'; rewrite /PCUICCumulProp.eq_univ_prop. - split. split; eapply PCUICCumulProp.is_prop_bottom; tea. - split; eapply PCUICCumulProp.is_sprop_bottom; tea. -Qed. - Import PCUICGlobalEnv PCUICUnivSubst PCUICValidity PCUICCumulProp. Notation " Σ ;;; Γ |- t ~~ u " := (cumul_prop Σ Γ t u) (at level 50, Γ, t, u at next level) : type_scope. Lemma is_propositional_bottom' {Σ Γ T s s'} : wf_ext Σ -> - check_univs -> prop_sub_type = false -> Σ ;;; Γ |- T ~~ tSort s -> Σ ;;; Γ |- T ~~ tSort s' -> PCUICCumulProp.eq_univ_prop s s'. Proof. - intros wf cu pst h h'; rewrite /PCUICCumulProp.eq_univ_prop. - pose proof (cumul_prop_trans _ _ _ _ _ _ (cumul_prop_sym _ _ _ _ _ h') h). - split. split; intros; eapply PCUICCumulProp.cumul_prop_props; tea. now symmetry. - split; intros; eapply PCUICCumulProp.cumul_sprop_props; tea. now symmetry. + intros. + eapply PCUICCumulProp.cumul_prop_sort. + etransitivity; tea. + now symmetry. +Qed. + +Lemma is_propositional_bottom {Σ Γ T s s'} : + wf_ext Σ -> + prop_sub_type = false -> + Σ ;;; Γ ⊢ T ≤ tSort s -> + Σ ;;; Γ ⊢ T ≤ tSort s' -> + PCUICCumulProp.eq_univ_prop s s'. +Proof. + move => wf pst /cumul_pb_cumul_prop h /cumul_pb_cumul_prop h'. + now eapply is_propositional_bottom'. Qed. Lemma is_propositional_lower {Σ s u u'} : - consistent Σ -> - leq_universe Σ s u -> - leq_universe Σ s u' -> + leq_sort Σ s u -> + leq_sort Σ s u' -> PCUICCumulProp.eq_univ_prop u u'. Proof. - intros wf leu leu'. - unfold eq_univ_prop; split. - - split. intros pu. eapply leq_universe_prop_r in leu; tea => //. - eapply leq_universe_prop_no_prop_sub_type in leu'; trea => //. - intros pu'. eapply leq_universe_prop_r in leu'; tea => //. - eapply leq_universe_prop_no_prop_sub_type in leu; tea => //. - - split. intros pu. eapply leq_universe_sprop_r in leu; tea => //. - eapply leq_universe_sprop_l in leu'; tea => //. - intros pu'. eapply leq_universe_sprop_r in leu'; tea => //. - eapply leq_universe_sprop_l in leu; tea => //. + destruct s, u, u' => //. Qed. Lemma typing_spine_inj {Σ Γ Δ s args args' u u'} : wf_ext Σ -> - check_univs -> prop_sub_type = false -> let T := it_mkProd_or_LetIn Δ (tSort s) in typing_spine Σ Γ T args (tSort u) -> typing_spine Σ Γ T args' (tSort u') -> PCUICCumulProp.eq_univ_prop u u'. Proof. - intros wf cu ips T. + intros wf ips T. move/typing_spine_it_mkProd_or_LetIn_full_inv => su. move/typing_spine_it_mkProd_or_LetIn_full_inv => su'. - eapply is_propositional_lower; tea. apply wf. + eapply is_propositional_lower; tea. Qed. Lemma Is_proof_ind Σ Γ t : @@ -784,7 +731,7 @@ Proof. eapply cumul_prop2'; tea => //. now eapply validity. eapply inversion_mkApps in X0 as x1. destruct x1 as [? []]. eapply inversion_Ind in t1 as [mdecl [idecl [wf [decli ?]]]]; eauto. - destruct (validity Hty'') as [u'' tyargs']. + destruct (validity Hty'') as (_ & u'' & tyargs' & _). eapply inversion_mkApps in X0 as x1. destruct x1 as [? []]. eapply invert_type_mkApps_ind in X0 as [sp cum]; eauto. eapply invert_type_mkApps_ind in tyargs' as f; tea. destruct f as [sp' cum']; eauto. @@ -792,8 +739,8 @@ Proof. split => //. rewrite (declared_inductive_type decli) in sp, sp'. rewrite subst_instance_it_mkProd_or_LetIn /= in sp, sp'. - eapply typing_spine_inj in sp. 5:exact sp'. all:eauto. - destruct sp as [H H0]. apply/orP. rewrite H H0. now apply/orP. + eapply typing_spine_inj in sp. 4:exact sp'. all:eauto. + destruct u, u'' => //. Qed. @@ -827,7 +774,7 @@ Proof. epose proof (PCUICPrincipality.common_typing _ wfΣ Hty Ht) as [C [Cty [Cty' Ht'']]]. eapply PCUICSpine.typing_spine_strengthen in sp. 3:tea. edestruct (sort_typing_spine _ _ _ u _ _ _ pu sp) as [u' [Hty' isp']]. - eapply cumul_prop1'. 5:tea. all:eauto. + eapply cumul_prop1'. 4:tea. all:eauto. eapply validity; eauto. exists ty, u'; split; auto. eapply PCUICSpine.type_mkApps; tea; eauto. @@ -887,7 +834,7 @@ Proof. rewrite (declared_inductive_lookup_gen decli_). rewrite oib.(ind_arity_eq). rewrite /isPropositionalArity !destArity_it_mkProd_or_LetIn /=. - destruct (is_propositional (ind_sort x0)) eqn:isp; auto. + destruct (Sort.is_propositional (ind_sort x0)) eqn:isp; auto. elimtype False; eapply ise. red. eexists; intuition eauto. right. unfold type_of_constructor in e, X. @@ -898,20 +845,20 @@ Proof. rewrite subst_cstr_concl_head in e, X. destruct decli. eapply nth_error_Some_length in H1; eauto. rewrite -it_mkProd_or_LetIn_app in e, X. - exists (subst_instance_univ u (ind_sort x0)). + exists (subst_instance_sort u (ind_sort x0)). rewrite is_propositional_subst_instance => //. split; auto. eapply cumul_propositional; eauto. rewrite is_propositional_subst_instance => //. eapply PCUICValidity.validity; eauto. - destruct X as [cty ty]. + destruct X as (_ & cty & ty & _). eapply type_Cumul_alt; eauto. eapply isType_Sort. 2:eauto. destruct (ind_sort x0) => //. eapply PCUICSpine.inversion_it_mkProd_or_LetIn in ty; eauto. - epose proof (typing_spine_proofs _ _ [] _ _ _ [] _ _ eq_refl wfΣ ty). - forward H0 by constructor. eexists; eauto. - simpl. now exists cty. eapply PCUICConversion.ws_cumul_pb_eq_le_gen, PCUICSR.wt_cumul_pb_refl; eauto. + epose proof (typing_spine_proofs _ _ [] _ _ _ [] _ _ wfΣ ty). + forward H0 by constructor. eapply has_sort_isType; eauto. + simpl. now eapply has_sort_isType. eapply PCUICConversion.ws_cumul_pb_eq_le_gen, PCUICSR.wt_cumul_pb_refl; eauto. destruct H0 as [_ sorts]. specialize (sorts _ _ decli c) as [sorts sorts']. forward sorts' by constructor. @@ -1038,7 +985,7 @@ Lemma not_isErasable Σ Γ f A u : (forall B, ∥Σ ;;; Γ |- f : B∥ -> ∥Σ ;;; Γ ⊢ A ≤ B∥) -> ~ ∥ isArity A ∥ -> ∥ Σ;;; Γ |- A : tSort u ∥ -> - ~ is_propositional u -> + ~ Sort.is_propositional u -> ~ ∥ Extract.isErasable Σ Γ f ∥. Proof. intros wfΣ Hlocal Hf Hnf Hprinc Harity Hfu Hu [[T [HT []]]]; sq. diff --git a/erasure/theories/EDeps.v b/erasure/theories/EDeps.v index 9bdf6047c..88a9d7716 100644 --- a/erasure/theories/EDeps.v +++ b/erasure/theories/EDeps.v @@ -627,8 +627,9 @@ Proof. cbn in *; try solve [constructor]. - now apply inversion_Evar in wt. - constructor. - now apply inversion_Lambda in wt as (? & ? & ? & ? & ?); eauto. - - apply inversion_LetIn in wt as (? & ? & ? & ? & ? & ?); eauto. + now apply inversion_Lambda in wt as (? & ? & ? & ?); eauto. + - apply inversion_LetIn in wt as (? & h1 & ? & ?); eauto. + apply unlift_TermTyp in h1. now constructor; eauto. - apply inversion_App in wt as (? & ? & ? & ? & ? & ?); eauto. now constructor; eauto. @@ -673,7 +674,7 @@ Proof. apply inversion_Fix in wt as (?&?&?&?&?&?&?); eauto. clear -wf a0 X H Σer. revert a0 X H Σer. - generalize mfix at 1 2 4 6. + generalize mfix at 1 3 5. intros mfix_gen. revert mfix'. induction mfix; cbn in *; intros mfix' typ er all_deps deps. @@ -682,12 +683,13 @@ Proof. depelim er. depelim all_deps. destruct a0 as [? ? ? ?]. + apply unlift_TermTyp in o. now constructor; eauto. - constructor. apply inversion_CoFix in wt as (?&?&?&?&?&?&?); eauto. clear -wf a0 X H Σer. revert a0 X H Σer. - generalize mfix at 1 2 4 6. + generalize mfix at 1 3 5. intros mfix_gen. revert mfix'. induction mfix; cbn in *; intros mfix' typ er all_deps deps. @@ -696,6 +698,7 @@ Proof. depelim er. depelim all_deps. destruct p as (?&?&?). + apply unlift_TermTyp in o. now constructor; eauto. - eapply inversion_Prim in wt as [prim_ty [decl []]]; eauto. depelim H0; depelim p1; depelim X; cbn in *; try constructor; cbn; intuition eauto. solve_all. @@ -743,7 +746,7 @@ Proof. depelim wf. depelim o0. destruct o1. cbn in *. eapply (erases_extends ({| universes := univs; declarations := Σ |}, cst_universes cst')); eauto. cbn. 4:{ split; eauto; cbn; try reflexivity. eexists [_]; cbn; reflexivity. } - constructor; auto. cbn. red in on_global_decl_d. rewrite E in on_global_decl_d. exact on_global_decl_d. + constructor; auto. cbn. red in on_global_decl_d. rewrite E in on_global_decl_d. apply unlift_TermTyp in on_global_decl_d. exact on_global_decl_d. split; auto. * intros. eapply (erases_deps_cons {| universes := univs; declarations := Σ |} _ kn (PCUICEnvironment.ConstantDecl cst')); auto. @@ -755,7 +758,8 @@ Proof. eapply (erases_deps_single (_, _)). 3:eauto. depelim wf. depelim o0. destruct o1. now split; cbn; eauto. - depelim wf. depelim o0. destruct o1. do 2 red in on_global_decl_d. now rewrite E in on_global_decl_d. + depelim wf. depelim o0. destruct o1. do 2 red in on_global_decl_d. rewrite E in on_global_decl_d. + now apply unlift_TermTyp in on_global_decl_d. apply IH; eauto. depelim wf. now depelim o0. + set (Σu := {| universes := univs; declarations := Σ; retroknowledge := retro |}). assert (wfΣu : PCUICTyping.wf Σu). @@ -781,6 +785,7 @@ Proof. unfold on_constant_decl, erases_constant_body in *. destruct ?; [|easy]. destruct ?; [|easy]. + apply unlift_TermTyp in decl_ext. eapply (erases_extends (Σu, cst_universes cst')). 4:{ split; cbn; auto. eexists [_]; cbn; reflexivity. } all: cbn; eauto. diff --git a/erasure/theories/EOptimizePropDiscr.v b/erasure/theories/EOptimizePropDiscr.v index 0fb577261..7257c75df 100644 --- a/erasure/theories/EOptimizePropDiscr.v +++ b/erasure/theories/EOptimizePropDiscr.v @@ -297,11 +297,11 @@ Proof. - destruct t => //. Qed. -Lemma isType_tSort {cf:checker_flags} {Σ : global_env_ext} {Γ l A} {wfΣ : wf Σ} : Σ ;;; Γ |- tSort (Universe.make l) : A -> isType Σ Γ (tSort (Universe.make l)). +Lemma isType_tSort {cf:checker_flags} {Σ : global_env_ext} {Γ s A} {wfΣ : wf Σ} : Σ ;;; Γ |- tSort s : A -> isType Σ Γ (tSort s). Proof. intros HT. eapply inversion_Sort in HT as [l' [wfΓ Hs]]; auto. - eexists; econstructor; eauto. + eapply has_sort_isType. econstructor; eauto. Qed. Lemma isType_it_mkProd {cf:checker_flags} {Σ : global_env_ext} {Γ na dom codom A} {wfΣ : wf Σ} : @@ -310,7 +310,7 @@ Lemma isType_it_mkProd {cf:checker_flags} {Σ : global_env_ext} {Γ na dom codom Proof. intros HT. eapply inversion_Prod in HT as (? & ? & ? & ? & ?); auto. - eexists; econstructor; eauto. + eapply has_sort_isType. econstructor; eauto. Qed. Definition remove_match_on_box_constant_decl Σ cb := diff --git a/erasure/theories/ESubstitution.v b/erasure/theories/ESubstitution.v index f693f836a..9813cb182 100644 --- a/erasure/theories/ESubstitution.v +++ b/erasure/theories/ESubstitution.v @@ -73,12 +73,18 @@ Require Import ssrbool. Lemma erases_extends : env_prop (fun Σ Γ t T => forall Σ', wf Σ' -> strictly_extends_decls Σ Σ' -> forall t', erases Σ Γ t t' -> erases (Σ', Σ.2) Γ t t') + (fun Σ Γ j => + forall Σ', wf Σ' -> strictly_extends_decls Σ Σ' -> + lift_wf_term (fun t => forall t', erases Σ Γ t t' -> erases (Σ', Σ.2) Γ t t') j) (fun Σ Γ => wf_local Σ Γ). Proof. apply typing_ind_env; intros; rename_all_hyps; auto. + { destruct X as (Htm & s & (_ & Hty) & _). split; eauto. destruct j_term => //. cbn in *. now eapply Htm. } all: match goal with [ H : erases _ _ ?a _ |- _ ] => tryif is_var a then idtac else inv H end. all: try now (econstructor; eauto). all: try now (econstructor; eapply Is_type_extends; eauto; tc). + - edestruct forall_Σ' as [Htm _]; tea. cbn in Htm. + econstructor; eauto. - econstructor. move: H4; apply contraNN. unfold isPropositional. @@ -102,14 +108,14 @@ Proof. - econstructor. destruct isdecl. 2:eauto. eapply Subsingleton_extends; eauto. exact H.p1. - econstructor. - eapply All2_All_mix_left in X1; eauto. - eapply All2_impl. exact X1. - intros ? ? [[?] [? []]]. + eapply All2_All_mix_left in X3; eauto. + eapply All2_impl. exact X3. + intros ? ? [(? & _) [? []]]; tea. split; eauto. - econstructor. - eapply All2_All_mix_left in X1; eauto. - eapply All2_impl. exact X1. - intros ? ? [[] [? []]]. + eapply All2_All_mix_left in X3; eauto. + eapply All2_impl. exact X3. + intros ? ? [(? & _) [? []]]; tea. split; eauto. - econstructor. induction H3; constructor. @@ -186,7 +192,7 @@ Lemma erases_weakening' (Σ : global_env_ext) (Γ Γ' Γ'' : context) (t T : ter Σ ;;; Γ ,,, Γ'' ,,, lift_context #|Γ''| 0 Γ' |- (lift #|Γ''| #|Γ'| t) ⇝ℇ (ELiftSubst.lift #|Γ''| #|Γ'| t'). Proof. intros HΣ HΓ'' * H He. - generalize_eqs H. intros eqw. rewrite <- eqw in *. + remember (Γ ,,, Γ') as Γ0 eqn:eqw. revert Γ Γ' Γ'' HΓ'' eqw t' He. revert Σ HΣ Γ0 t T H . apply (typing_ind_env (fun Σ Γ0 t T => forall Γ Γ' Γ'', @@ -195,9 +201,15 @@ Proof. forall t', Σ;;; Γ0 |- t ⇝ℇ t' -> _) + (fun Σ Γ0 j => forall Γ Γ' Γ'', + wf_local Σ (Γ ,,, Γ'' ,,, lift_context #|Γ''| 0 Γ') -> + Γ0 = Γ ,,, Γ' -> + lift_wf_term (fun t => forall t', + Σ;;; Γ0 |- t ⇝ℇ t' -> Σ;;; Γ,,, Γ'',,, lift_context #|Γ''| 0 Γ' |- lift #|Γ''| #|Γ'| t ⇝ℇ ELiftSubst.lift #|Γ''| #|Γ'| t')j ) (fun Σ Γ => wf_local Σ Γ) ); intros Σ wfΣ Γ0; intros; try subst Γ0; auto. + { destruct X as (Htm & s & (_ & Hty) & _). split; eauto. destruct j_term => //. cbn in *. now eapply Htm. } all: match goal with [ H : erases _ _ ?a _ |- _ ] => tryif is_var a then idtac else inv H end. all: try now (cbn; econstructor; eauto). all: try now (econstructor; eapply Is_type_weakening; eauto). @@ -205,19 +217,19 @@ Proof. - destruct ?; econstructor. - econstructor. unfold app_context, snoc in *. - pose proof (H0 Γ (vass n t :: Γ') Γ''). - rewrite lift_context_snoc0 - plus_n_O in H1. - eapply H1; eauto. cbn. econstructor. - eauto. cbn. exists s1. eapply (weakening_typing (T := tSort s1)); eauto. - now apply All_local_env_app_inv in X2. + pose proof (H Γ (vass na t :: Γ') Γ''). + rewrite lift_context_snoc0 - plus_n_O in H0. + eapply H0; eauto. cbn. econstructor. + eauto. cbn. eapply lift_typing_f_impl with (1 := X0) => // ??. eapply weakening_typing. + now apply All_local_env_app_inv in X3. - econstructor. - + eapply H0; eauto. - + pose proof (H1 Γ (vdef n b b_ty :: Γ') Γ''). - rewrite lift_context_snoc0 -plus_n_O in H2. - eapply H2; eauto. cbn. econstructor. - eauto. hnf. 2: cbn; eapply weakening_typing; eauto. - eapply weakening_typing in X0; eauto. - now apply All_local_env_app_inv in X3. + + edestruct X1 as [Htm _]; tea; eauto. + + pose proof (H Γ (vdef na b b_ty :: Γ') Γ''). + rewrite lift_context_snoc0 -plus_n_O in H0. + eapply H0; eauto. cbn. econstructor. + eauto. + eapply lift_typing_f_impl with (1 := X0) => // ??. + eapply weakening_typing. now apply All_local_env_app_inv in X3. - econstructor. + eauto. @@ -247,17 +259,14 @@ Proof. eapply e. eapply weakening_wf_local => //. rewrite app_context_assoc //. - now eapply wf_local_app_inv in X7 as []. + now eapply All_local_env_app_inv in X7 as []. rewrite app_context_assoc. reflexivity. now rewrite [_.1](PCUICCasesContexts.inst_case_branch_context_eq a). - assert (HT : Σ;;; Γ ,,, Γ' |- PCUICAst.tFix mfix n : (decl.(dtype))). - econstructor; eauto. eapply All_impl. eassumption. intros. - destruct X4; cbn in *; pcuicfo. - exists x0; auto. - eapply (All_impl X1). intros d [HT IH]. pcuicfo. + econstructor; eauto. eapply weakening_typing in HT; auto. - 2:{ apply All_local_env_app_inv in X2 as [X2 _]. eapply X2. } + 2:{ apply All_local_env_app_inv in X4 as [X4 _]. eapply X4. } cbn in HT. eapply inversion_Fix in HT as (? & ? & ? & ? & ? & ? & ?) ; auto. @@ -265,8 +274,8 @@ Proof. econstructor. eapply All2_map. eapply All2_impl. eapply All2_All_mix_left. - eapply X1. eassumption. simpl. - intros [] []. simpl. intros [[Hs IH] [<- <- IH']]. + eapply X3. eassumption. simpl. + intros [] []. simpl. intros [IH [<- <- IH']]. repeat split. unfold app_context in *. eapply isLambda_lift => //. eapply ELiftSubst.isLambda_lift => //. @@ -281,17 +290,14 @@ Proof. rewrite lift_fix_context. rewrite lift_context_app - plus_n_O in IH. unfold app_context in IH. rewrite <- !app_assoc in IH. - rewrite (All2_length X3) in IH |- *. + rewrite (All2_length X5) in IH |- *. apply IH. apply e. - assert (HT : Σ;;; Γ ,,, Γ' |- PCUICAst.tCoFix mfix n : (decl.(dtype))). - econstructor; eauto. eapply All_impl. eassumption. intros. - destruct X4; cbn in *; pcuicfo. - now exists x0. - eapply (All_impl X1). intros d [HT IH]. pcuicfo. + econstructor; eauto. eapply weakening_typing in HT; auto. - 2:{ apply All_local_env_app_inv in X2 as [X2 _]. eapply X2. } + 2:{ apply All_local_env_app_inv in X4 as [X4 _]. eapply X4. } cbn in HT. eapply inversion_CoFix in HT as (? & ? & ? & ? & ? & ? & ?) ; auto. clear a0 e. @@ -299,8 +305,8 @@ Proof. econstructor. eapply All2_map. eapply All2_impl. eapply All2_All_mix_left. - eapply X1. eassumption. simpl. - intros [] []. simpl. intros [[Hs IH] [<- [<- IH']]]. + eapply X3. eassumption. simpl. + intros [] []. simpl. intros [IH [<- [<- IH']]]. repeat split. unfold app_context in *. specialize (IH Γ (types ++ Γ') Γ''). subst types. rewrite app_length fix_context_length in IH. @@ -309,11 +315,11 @@ Proof. eapply All_mfix_wf in a; auto. rewrite lift_fix_context in a. now rewrite <- !app_assoc. } - forward IH by now rewrite app_assoc. + forward IH by now rewrite /app_context app_assoc. rewrite lift_fix_context. rewrite lift_context_app -plus_n_O in IH. unfold app_context in IH. rewrite <- !app_assoc in IH. - rewrite (All2_length X3) in IH |- *. + rewrite (All2_length X5) in IH |- *. apply IH. apply IH'. - econstructor. depelim H4. @@ -378,7 +384,7 @@ Qed. Lemma subst_case_branch_context {cf : checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} ind (n : nat) mdecl idecl p br cdecl s k : PCUICAst.declared_constructor Σ (ind, n) mdecl idecl cdecl -> wf_predicate mdecl idecl p -> - All2 (PCUICEquality.compare_decls eq eq) (bcontext br) + PCUICEquality.eq_context_upto_names (bcontext br) (cstr_branch_context ind mdecl cdecl) -> subst_context s k (case_branch_context ind mdecl p (forget_types (bcontext br)) cdecl) = case_branch_context ind mdecl (map_predicate_k id (subst s) k p) (forget_types (bcontext br)) cdecl. @@ -404,7 +410,7 @@ Lemma erases_subst (Σ : global_env_ext) Γ Γ' Δ t s t' s' T : Σ ;;; (Γ ,,, subst_context s 0 Δ) |- (subst s #|Δ| t) ⇝ℇ ELiftSubst.subst s' #|Δ| t'. Proof. intros HΣ HΔ Hs Ht He. - generalize_eqs Ht. intros eqw. + remember (Γ ,,, Γ' ,,, Δ) as Γ0 eqn:eqw in Ht. revert Γ Γ' Δ t' s Hs HΔ He eqw. revert Σ HΣ Γ0 t T Ht. eapply (typing_ind_env (fun Σ Γ0 t T => @@ -416,9 +422,20 @@ Proof. All2 (erases Σ Γ) s s' -> Σ;;; Γ ,,, subst_context s 0 Δ |- subst s #|Δ| t ⇝ℇ ELiftSubst.subst s' #|Δ| t' ) + (fun Σ Γ0 j => + forall (Γ Γ' : context) Δ t' (s : list PCUICAst.term), + wf_local Σ (Γ ,,, subst_context s 0 Δ) -> + subslet Σ Γ s Γ' -> + Γ0 = Γ ,,, Γ' ,,, Δ -> + All2 (erases Σ Γ) s s' -> + lift_wf_term (fun t => + Σ;;; Γ ,,, Γ' ,,, Δ |- t ⇝ℇ t' -> + Σ;;; Γ ,,, subst_context s 0 Δ |- subst s #|Δ| t ⇝ℇ ELiftSubst.subst s' #|Δ| t') j + ) (fun Σ Γ0 => wf_local Σ Γ0) ); intros Σ wfΣ Γ0 wfΓ0; intros; simpl in * |-; auto; subst Γ0. + { destruct X as (Htm & u & (_ & Hty) & _). split; eauto. destruct j_term => // ?. cbn in *. now eapply Htm. } - inv H0. + cbn. destruct ?. destruct ?. * eapply All2_nth_error_Some in X2; eauto. @@ -434,28 +451,26 @@ Proof. + econstructor. eapply is_type_subst; eauto. - inv H0. econstructor. eapply is_type_subst; eauto. - - inv H1. econstructor. + - inv H0. econstructor. eapply is_type_subst; eauto. - - inv H1. + - inv H0. + cbn. econstructor. - specialize (H0 Γ Γ' (vass n t :: Δ) t'0 s). + specialize (H Γ Γ' (vass na t :: Δ) t'0 s). (* unfold app_context, snoc in *. *) - rewrite subst_context_snoc0 in H0. - eapply H0; eauto. + rewrite subst_context_snoc0 in H. + eapply H; eauto. cbn. econstructor. eauto. - cbn. exists s1. eapply (substitution (T := tSort s1)); eauto. + eapply lift_typing_f_impl with (1 := X0) => // ??. now eapply substitution. + econstructor. eapply is_type_subst; eauto. - - inv H2. + - inv H0. + cbn. econstructor. - eauto. - specialize (H1 Γ Γ' (vdef n b b_ty :: Δ) t2' s). - rewrite subst_context_snoc0 in H1. - eapply H1; eauto. + now edestruct X1; tea; eauto. + specialize (H Γ Γ' (vdef na b b_ty :: Δ) t2' s). + rewrite subst_context_snoc0 in H. + eapply H; eauto. cbn. econstructor. eauto. - hnf. - eapply substitution in X0; eauto. - eapply substitution; eauto. + eapply lift_typing_f_impl with (1 := X0) => // ??. now eapply substitution. + econstructor. eapply is_type_subst; eauto. - inv H2. @@ -539,8 +554,8 @@ Proof. cbn. now eapply isLambda_subst. now eapply ELiftSubst.isLambda_subst. eapply In_nth_error in H2 as []. - eapply nth_error_all in X1; eauto. - destruct X1 as [Hs IH]. + eapply nth_error_all in X3 as IH; eauto. + eapply nth_error_all in X2 as Ht; eauto. apply unlift_TermTyp in Ht. specialize (IH Γ Γ' (Δ ,,, fix_context mfix)). rewrite app_context_assoc in IH |- *. eapply IH in e1; eauto. @@ -553,8 +568,8 @@ Proof. * cbn. now rewrite app_context_length fix_context_length. * cbn. now erewrite app_context_length, fix_context_length, All2_length. * pose proof (@substitution _ Σ _ Γ Γ' s (Δ ,,, fix_context mfix)). - rewrite app_context_assoc in X1. - eapply X1 in Hs; eauto. + rewrite app_context_assoc in X9. + eapply X9 in X5; eauto. eapply typing_wf_local. eassumption. + econstructor. eapply is_type_subst; eauto. @@ -566,11 +581,11 @@ Proof. intros. destruct H4 as (? & ? & ?). repeat split; eauto. eapply In_nth_error in H2 as []. - eapply nth_error_all in X1; eauto. - destruct X1. - specialize (e2 Γ Γ' (Δ ,,, fix_context mfix)). - rewrite app_context_assoc in e2. - eapply e2 in e1; eauto. + eapply nth_error_all in X3 as IH; eauto. + eapply nth_error_all in X2 as Ht; eauto. apply unlift_TermTyp in Ht. + specialize (IH Γ Γ' (Δ ,,, fix_context mfix)). + rewrite app_context_assoc in IH. + eapply IH in e1; eauto. eapply erases_eq; eauto. * rewrite subst_context_app. @@ -578,10 +593,10 @@ Proof. rewrite app_context_assoc. f_equal. now rewrite subst_fix_context. * cbn. now rewrite app_context_length fix_context_length. - * cbn. now erewrite app_context_length, fix_context_length, (All2_length X5). + * cbn. now erewrite app_context_length, fix_context_length, All2_length. * pose proof (@substitution _ Σ _ Γ Γ' s (Δ ,,, fix_context mfix)). - rewrite app_context_assoc in X1. - eapply X1 in t; eauto. + rewrite app_context_assoc in X9. + eapply X9 in X5; eauto. eapply typing_wf_local. eassumption. + econstructor. eapply is_type_subst; eauto. diff --git a/erasure/theories/ErasureCorrectness.v b/erasure/theories/ErasureCorrectness.v index f79f46320..fa16aa49f 100644 --- a/erasure/theories/ErasureCorrectness.v +++ b/erasure/theories/ErasureCorrectness.v @@ -60,11 +60,11 @@ Proof. eapply IHeval1 in H4 as (vf' & Hvf' & [He_vf']); eauto. eapply IHeval2 in H6 as (vu' & Hvu' & [He_vu']); eauto. pose proof (subject_reduction_eval t0 H). - eapply inversion_Lambda in X0 as (? & ? & ? & ? & e0). + eapply inversion_Lambda in X0 as (? & ? & ? & e0). assert (Σ ;;; [] |- a' : t). { eapply subject_reduction_eval; eauto. eapply PCUICConversion.ws_cumul_pb_Prod_Prod_inv in e0 as [? e1 e2]. - eapply type_Cumul_alt. eassumption. now exists x2. + eapply type_Cumul_alt; tea. symmetry in e1. eapply ws_cumul_pb_forget in e1. now eapply conv_cumul. } @@ -100,19 +100,22 @@ Proof. + auto. - assert (Hty' := Hty). assert (Σ |-p tLetIn na b0 t b1 ⇓ res) by eauto. - eapply inversion_LetIn in Hty' as (? & ? & ? & ? & ? & ?); auto. + eapply inversion_LetIn in Hty' as (? & h1 & ? & ?); auto. + assert (wf_rel : lift_typing0 (typing Σ []) (j_vdef na b0' t)). + { apply lift_sorting_it_impl_gen with h1 => // HT. eapply subject_reduction_eval; eauto. } + apply unlift_TermTyp in h1 as h1'. invs He. + depelim Hed. eapply IHeval1 in H6 as (vt1' & Hvt2' & [He_vt1']); eauto. assert (Hc : conv_context cumulAlgo_gen Σ ([],, vdef na b0 t) [vdef na b0' t]). { econstructor. econstructor. econstructor. reflexivity. eapply PCUICCumulativity.red_conv. - now eapply wcbveval_red; eauto. + eapply wcbveval_red; eauto. reflexivity. } - assert (Σ;;; [vdef na b0' t] |- b1 : x0). { + assert (Σ;;; [vdef na b0' t] |- b1 : x). { cbn in *. eapply context_conversion. 3:eauto. all:cbn; eauto. - econstructor. all: hnf; eauto. eapply subject_reduction_eval; auto. eauto. eauto. + econstructor. constructor. assumption. } assert (Σ;;; [] |- subst1 b0' 0 b1 ⇝ℇ ELiftSubst.subst1 vt1' 0 t2'). { eapply (erases_subst Σ [] [vdef na b0' t] [] b1 [b0'] t2'); eauto. @@ -123,10 +126,9 @@ Proof. eapply subject_reduction_eval; eauto. eapply erases_context_conversion. 3:eassumption. all: cbn; eauto. - econstructor. all: hnf; eauto. - eapply subject_reduction_eval; eauto. + econstructor. constructor. assumption. } - pose proof (subject_reduction_eval t1 H). + pose proof (subject_reduction_eval h1' H). assert (eqs := type_closed_subst b1 _ X1). rewrite eqs in H1. eapply IHeval2 in H1 as (vres & Hvres & [Hty_vres]); [| |now eauto]. @@ -134,7 +136,7 @@ Proof. exists vres. split. eauto. constructor; econstructor; eauto. enough (ECSubst.csubst vt1' 0 t2' = ELiftSubst.subst10 vt1' t2') as ->; auto. eapply ECSubst.closed_subst. eapply erases_closed in Hvt2'; auto. - eapply eval_closed. eauto. 2:eauto. now eapply subject_closed in t1. + eapply eval_closed. eauto. 2:eauto. now eapply subject_closed in h1'. + exists EAst.tBox. split. 2:constructor; econstructor; eauto. econstructor. eapply Is_type_eval; eauto. @@ -154,13 +156,13 @@ Proof. apply declared_constant_from_gen in isdecl', isdecl''. eapply declared_constant_inv in isdecl'; [| |now eauto|now apply wfΣ]. 2:exact weaken_env_prop_typing. - unfold on_constant_decl in isdecl'. rewrite e in isdecl'. red in isdecl'. + unfold on_constant_decl in isdecl'. rewrite e in isdecl'. eapply unlift_TermTyp in isdecl'. unfold declared_constant in isdecl''. now eapply @typing_subst_instance_decl with (Σ := Σ) (Γ := []); eauto. * assert (isdecl'' := isdecl'). apply declared_constant_from_gen in isdecl', isdecl''. eapply declared_constant_inv in isdecl'; [| |now eauto|now apply wfΣ]. - unfold on_constant_decl in isdecl'. rewrite e in isdecl'. cbn in *. + unfold on_constant_decl in isdecl'. rewrite e in isdecl'. eapply unlift_TermTyp in isdecl'. 2:exact weaken_env_prop_typing. now eapply erases_subst_instance_decl with (Σ := Σ) (Γ := []); eauto. * now eauto. @@ -872,7 +874,7 @@ Proof. move: e0 e. rewrite -closed_unfold_cofix_cunfold_eq //. unfold unfold_cofix. intros hnth; rewrite hnth. intros [=]. subst fn narg. - eapply All_nth_error in a0 as a'; tea. + eapply All_nth_error in a0 as a'; tea. apply unlift_TermTyp in a'. eapply erases_subst0. eauto. 2:eauto. pcuic. all:tea. { rewrite app_context_nil_l. eapply subslet_cofix_subst; eauto. econstructor; eauto. } @@ -992,7 +994,7 @@ Proof. move: hnth e. rewrite -closed_unfold_cofix_cunfold_eq //. unfold unfold_cofix. intros hnth'; rewrite hnth'. intros [=]. subst fn narg. rewrite hnth' in e2. noconf e2. - eapply All_nth_error in a0 as a'; tea. + eapply All_nth_error in a0 as a'; tea. eapply unlift_TermTyp in a'. eapply erases_subst0. eauto. 2:eauto. pcuic. all:tea. { rewrite app_context_nil_l. eapply subslet_cofix_subst; eauto. econstructor; eauto. } @@ -1223,7 +1225,7 @@ Proof. - cbn. destruct cb' as [[]]. cbn in *. depelim wf. destruct o. rewrite [forallb _ _](IHer wf) andb_true_r. - red in H. destruct cb as [ty []]; cbn in *. + red in H. destruct cb as [ty []]; cbn in *. apply unlift_TermTyp in on_global_decl_d. unshelve eapply PCUICClosedTyp.subject_closed in on_global_decl_d. cbn. split; auto. eapply erases_closed in H; tea. elim H. cbn. apply IHer. now depelim wf. @@ -1286,6 +1288,7 @@ Proof. cbn. red in H. do 2 red in on_global_decl_d. destruct (cst_body cb). + eapply unlift_TermTyp in on_global_decl_d. destruct (E.cst_body cb') => //. cbn. set (Σ'' := ({| universes := _ |}, _)) in *. assert (PCUICTyping.wf_ext Σ''). diff --git a/erasure/theories/ErasureFunction.v b/erasure/theories/ErasureFunction.v index f50669ff4..e1b500218 100644 --- a/erasure/theories/ErasureFunction.v +++ b/erasure/theories/ErasureFunction.v @@ -23,7 +23,7 @@ Implicit Types (cf : checker_flags). #[local] Existing Instance extraction_normalizing. -Notation alpha_eq := (All2 (PCUICEquality.compare_decls eq eq)). +Notation alpha_eq := PCUICEquality.eq_context_upto_names. Ltac sq := repeat match goal with @@ -36,8 +36,8 @@ Lemma wf_local_rel_alpha_eq_end {cf} {Σ : global_env_ext} {wfΣ : wf Σ} {Γ Δ wf_local_rel Σ Γ Δ -> wf_local_rel Σ Γ Δ'. Proof. intros wfΓ eqctx wf. - apply wf_local_app_inv. - eapply wf_local_app in wf => //. + apply All_local_env_app_inv. + eapply All_local_env_app in wf => //. eapply PCUICSpine.wf_local_alpha; tea. eapply All2_app => //. reflexivity. Qed. @@ -758,10 +758,9 @@ Section fix_sigma. - intros. destruct H' as []. rewrite <- (abstract_env_ext_irr _ H2) in X0; eauto. rewrite <- (abstract_env_ext_irr _ H2) in wf; eauto. - eapply inversion_Prod in X0 as (? & ? & ? & ? & ?) ; auto. + eapply inversion_Prod in X0 as (? & ? & h1 & ? & ?) ; auto. eapply cored_red in H0 as []. - econstructor. econstructor. econstructor. eauto. - 2:reflexivity. econstructor; pcuic. + econstructor. econstructor; tea. rewrite <- (abstract_env_ext_irr _ H2) in X0; eauto. eapply subject_reduction; eauto. - intros. rewrite <- (abstract_env_ext_irr _ H0) in wf; eauto. @@ -800,8 +799,8 @@ Section fix_sigma. pose proof (abstract_env_ext_wf _ wfΣ') as [wf]. sq. eapply subject_reduction_closed in HT; tea. - eapply inversion_Prod in HT as [? [? [? []]]]. - now eapply typing_wf_local in t1. pcuic. pcuic. + eapply inversion_Prod in HT as [? [? [? [t0 ?]]]]. + now eapply typing_wf_local in t0. pcuic. pcuic. - clear rprod is_arity rsort a0. intros Σ' wfΣ'; specialize (H Σ' wfΣ'). repeat specialize_Σ wfΣ'. @@ -897,7 +896,7 @@ Equations? is_erasableb X_type X {normalization_in : forall Σ, wf_ext Σ -> Σ { | T with is_arity X_type X Γ _ T.π1 _ := { | true => true | false => let s := @sort_of_type extraction_checker_flags _ X_type X _ Γ T.π1 _ in - is_propositional s.π1 } }. + Sort.is_propositional s.π1 } }. Proof. - intros. specialize_Σ H. destruct wt; sq. pcuic. @@ -926,7 +925,7 @@ Proof. 2: eassumption. eauto. - destruct type_of_typing as [x Hx]. cbn -[sort_of_type is_arity] in *. destruct (sort_of_type _ _ _ _). cbn. - destruct (is_propositional x0) eqn:isp; constructor. + destruct (Sort.is_propositional x0) eqn:isp; constructor. * clear Heq. intros. pose proof (abstract_env_ext_wf _ H) as [wf]. specialize_Σ H. @@ -945,7 +944,7 @@ Proof. { apply nisa. intros. rewrite (abstract_env_ext_irr _ H wfΣ). eapply invert_cumul_arity_r; tea. } { destruct s as [Hs]. - unshelve epose proof (H := unique_sorting_equality_propositional _ _ wf Hs Hu' p) => //. reflexivity. reflexivity. congruence. } + unshelve epose proof (H := unique_sorting_equality_propositional _ wf Hs Hu' p) => //. reflexivity. congruence. } Qed. Equations? is_erasable {X_type X} {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} (Γ : context) (t : PCUICAst.term) @@ -1090,12 +1089,13 @@ Section Erase. - now eapply inversion_Evar in Ht. - discriminate. - discriminate. - - eapply inversion_Lambda in Ht as (? & ? & ? & ? & ?); auto. + - eapply inversion_Lambda in Ht as (? & ? & ? & ?); auto. eexists; eauto. - - eapply inversion_LetIn in Ht as (? & ? & ? & ? & ? & ?); auto. + - eapply inversion_LetIn in Ht as (? & h1 & ? & ?); auto. + apply unlift_TermTyp in h1. eexists; eauto. - simpl in *. - eapply inversion_LetIn in Ht as (? & ? & ? & ? & ? & ?); auto. + eapply inversion_LetIn in Ht as (? & ? & ? & ?); auto. eexists; eauto. - eapply inversion_App in Ht as (? & ? & ? & ? & ? & ?); auto. eexists; eauto. @@ -1115,9 +1115,9 @@ Section Erase. eexists; eauto. - eapply inversion_Fix in Ht as (? & ? & ? & ? & ? & ?); auto. sq. destruct p. move/andP: i => [wffix _]. - solve_all. now eexists. + solve_all. destruct a0 as (Hb & _). cbn in Hb. now eexists. - eapply inversion_CoFix in Ht as (? & ? & ? & ? & ? & ?); auto. - sq. eapply All_impl; tea. cbn. intros d Ht; now eexists. + sq. eapply All_impl; tea. cbn. intros d (Hb & _); cbn in Hb; now eexists. - eapply inversion_Prim in Ht as [prim_ty [decl []]]; eauto. depelim p0. now eexists. - eapply inversion_Prim in Ht as [prim_ty [decl []]]; eauto. @@ -1316,7 +1316,7 @@ Program Definition erase_constant_body X_type X {normalization_in : forall Σ, w Next Obligation. Proof. specialize_Σ H. sq. red in Hcb. - rewrite <-Heq_anonymous in Hcb. now eexists. + rewrite <-Heq_anonymous in Hcb. apply unlift_TermTyp in Hcb. now eexists. Qed. Definition erase_one_inductive_body (oib : one_inductive_body) : E.one_inductive_body := @@ -1325,7 +1325,7 @@ Definition erase_one_inductive_body (oib : one_inductive_body) : E.one_inductive let projs := map (fun pdecl => EAst.mkProjection pdecl.(proj_name)) oib.(ind_projs) in let is_propositional := match destArity [] oib.(ind_type) with - | Some (_, u) => is_propositional u + | Some (_, u) => Sort.is_propositional u | None => false (* dummy, impossible case *) end in @@ -1818,7 +1818,7 @@ Lemma erase_constant_body_correct X_type X {normalization_in : forall Σ, wf_ext Proof. red. destruct cb as [name [bod|] univs]; simpl; eauto. intros. set (ecbo := erase_constant_body_obligation_1 X_type X _ _ _ _). clearbody ecbo. - cbn in *. specialize_Σ H. sq. + cbn in *. specialize_Σ H. sq. apply unlift_TermTyp in onc. eapply (erases_weakening_env (Σ := Σ) (Σ' := (Σ', univs))); simpl; eauto. now eapply erases_erase. Qed. diff --git a/erasure/theories/ErasureFunctionProperties.v b/erasure/theories/ErasureFunctionProperties.v index f921232a3..ab2b0542b 100644 --- a/erasure/theories/ErasureFunctionProperties.v +++ b/erasure/theories/ErasureFunctionProperties.v @@ -87,9 +87,10 @@ Proof. | [ H : KernameSet.In _ (KernameSet.union _ _) |- _ ] => apply KernameSet.union_spec in hin as [?|?] end. - - apply inversion_Lambda in wt as (? & ? & ? & ? & ?); eauto. - - apply inversion_LetIn in wt as (? & ? & ? & ? & ? & ?); eauto. - - apply inversion_LetIn in wt as (? & ? & ? & ? & ? & ?); eauto. + - apply inversion_Lambda in wt as (? & ? & ? & ?); eauto. + - apply inversion_LetIn in wt as (? & h1 & ? & ?); eauto. + now apply unlift_TermTyp in h1. + - apply inversion_LetIn in wt as (? & ? & ? & ?); eauto. - apply inversion_App in wt as (? & ? & ? & ? & ? & ?); eauto. - apply inversion_App in wt as (? & ? & ? & ? & ? & ?); eauto. - apply inversion_Const in wt as (? & ? & ? & ? & ?); eauto. @@ -141,6 +142,7 @@ Proof. eapply All2_All_mix_left in H; eauto. eapply All2_In_right in H; eauto. destruct H as [[def [Hty Hdef]]]. + apply unlift_TermTyp in Hty. eapply Hdef; eauto. - apply inversion_CoFix in wt as (?&?&?&?&?&?&?); eauto. @@ -151,6 +153,7 @@ Proof. eapply All2_All_mix_left in H; eauto. eapply All2_In_right in H; eauto. destruct H as [[def [Hty Hdef]]]. + apply unlift_TermTyp in Hty. eapply Hdef; eauto. - eapply inversion_Prim in wt as [prim_ty [decl []]]; eauto. @@ -182,8 +185,9 @@ Proof. end. - now apply inversion_Evar in wt. - constructor. - now apply inversion_Lambda in wt as (? & ? & ? & ? & ?); eauto. - - apply inversion_LetIn in wt as (? & ? & ? & ? & ? & ?); eauto. + now apply inversion_Lambda in wt as (? & ? & ? & ?); eauto. + - apply inversion_LetIn in wt as (? & h1 & ? & ?); eauto. + apply unlift_TermTyp in h1. constructor; eauto. - apply inversion_App in wt as (? & ? & ? & ? & ? & ?); eauto. now constructor; eauto. @@ -292,14 +296,14 @@ Proof. eapply In_Forall in Σer. eapply Forall_All in Σer. eapply Forall2_All2 in H. - ELiftSubst.solve_all. + ELiftSubst.solve_all. apply unlift_TermTyp in b1; eauto. - constructor. apply inversion_CoFix in wt as (?&?&?&?&?&?&?); eauto. eapply All_Forall. eapply includes_deps_fold in Σer as [_ Σer]. eapply In_Forall in Σer. eapply Forall_All in Σer. eapply Forall2_All2 in H. - ELiftSubst.solve_all. + ELiftSubst.solve_all. apply unlift_TermTyp in b1; eauto. - eapply inversion_Prim in wt as [prim_ty [decl []]]; eauto. depelim H; depelim H0. depelim X; depelim X0; depelim p1; constructor; noconf H; cbn; simp prim_global_deps in Σer; simpl in *. @@ -337,7 +341,7 @@ Proof. destruct (E.cst_body cb') eqn:cbe'; auto. specialize (H3 _ eq_refl). eapply on_declared_constant in H_; auto. - red in H_. rewrite cbe in H_. simpl in H_. + red in H_. rewrite cbe in H_. apply unlift_TermTyp in H_. eapply (erases_weakening_env (Σ := (Σ, cst_universes cb)) (Σ' := (add_global_decl Σ (kn, d), cst_universes cb))); eauto. simpl. econstructor; eauto. econstructor; eauto. @@ -719,13 +723,11 @@ Lemma leq_term_propositional_sorted_l {Σ Γ v v' u u'} : wf_ext Σ -> PCUICEquality.leq_term Σ (global_ext_constraints Σ) v v' -> Σ;;; Γ |- v : tSort u -> - Σ;;; Γ |- v' : tSort u' -> is_propositional u -> - leq_universe (global_ext_constraints Σ) u' u. + Σ;;; Γ |- v' : tSort u' -> Sort.is_propositional u -> + leq_sort (global_ext_constraints Σ) u' u. Proof. intros wfΣ leq hv hv' isp. - eapply orb_true_iff in isp as [isp | isp]. - - eapply leq_term_prop_sorted_l; eauto. - - eapply leq_term_sprop_sorted_l; eauto. + eapply leq_term_prop_sorted_l; eauto. Qed. Lemma map_erase_irrel X_type X {normalization_in1 normalization_in2 : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} Γ t H1 H2 : diff --git a/erasure/theories/ErasureProperties.v b/erasure/theories/ErasureProperties.v index 18aa969e4..732f4bf2b 100644 --- a/erasure/theories/ErasureProperties.v +++ b/erasure/theories/ErasureProperties.v @@ -40,29 +40,21 @@ Lemma wf_local_rel_conv: All2_fold (conv_decls cumulAlgo_gen Σ) Γ Γ' -> forall Γ0 : context, wf_local Σ Γ' -> wf_local_rel Σ Γ Γ0 -> wf_local_rel Σ Γ' Γ0. Proof. - intros Σ wfΣ Γ Γ' X1 Γ0 ? w0. induction w0. - - econstructor. - - econstructor; eauto. cbn in *. - destruct t0. exists x. eapply context_conversion with (Γ ,,, Γ0); eauto. - * eapply wf_local_app; eauto. - * eapply conv_context_app_same; eauto. - - econstructor; eauto. - + cbn in *. - destruct t0. exists x. eapply context_conversion with (Γ ,,, Γ0); eauto. - * eapply wf_local_app; eauto. - * eapply conv_context_app_same; eauto. - + cbn in *. eapply context_conversion with (Γ ,,, Γ0); eauto. - * eapply wf_local_app; eauto. - * eapply conv_context_app_same; eauto. + intros Σ wfΣ Γ Γ' X1 Γ0 ? w0. + apply All_local_env_impl_ind with (1 := w0) => Δ j wfΓ' Hj. + apply lift_typing_impl with (1 := Hj) => t T HT. + eapply context_conversion with (Γ ,,, Δ); eauto. + - eapply All_local_env_app; eauto. + - eapply conv_context_app_same; eauto. Qed. Lemma conv_context_wf_local_app {A B A'} {Σ : global_env_ext} {wfΣ : wf Σ} : wf_local Σ (A ,,, B) -> wf_local Σ A' -> conv_context cumulAlgo_gen Σ A A' -> wf_local Σ (A' ,,, B). Proof. intros wfab wfa' cv. - eapply wf_local_app => //. + eapply All_local_env_app => //. eapply wf_local_rel_conv; tea. - now eapply wf_local_app_inv. + now eapply All_local_env_app_inv. Qed. @@ -201,23 +193,30 @@ Lemma erases_context_conversion : conv_context cumulAlgo_gen Σ Γ Γ' -> wf_local Σ Γ' -> forall t', erases Σ Γ t t' -> erases Σ Γ' t t') + (fun (Σ : global_env_ext) (Γ : context) j => + forall Γ' : context, + conv_context cumulAlgo_gen Σ Γ Γ' -> + wf_local Σ Γ' -> + lift_wf_term (fun t => forall t', erases Σ Γ t t' -> erases Σ Γ' t t') j) (fun Σ Γ => wf_local Σ Γ) . Proof. apply typing_ind_env; intros Σ wfΣ Γ wfΓ; intros **; rename_all_hyps; auto. + { destruct X as (Htm & u & (_ & Hty) & _). split; eauto. destruct j_term => // ?. cbn in *. now eapply Htm. } all: match goal with [ H : erases _ _ ?a _ |- ?G ] => tryif is_var a then idtac else invs H end. all: try now (econstructor; eauto). - econstructor. eapply h_forall_Γ'0. econstructor. eauto. now constructor. constructor; auto. - exists s1. + eapply lift_typing_impl with (1 := X0) => ?? HT. eapply context_conversion. 3:eauto. all:eauto. - - econstructor. eauto. eapply h_forall_Γ'1. + - econstructor. + now edestruct forall_Γ'; tea; eauto. + eapply h_forall_Γ'0. econstructor. eauto. now constructor. - constructor; auto. exists s1. - eapply context_conversion with Γ; eauto. - eapply context_conversion with Γ; eauto. - eassumption. + constructor; auto. + eapply lift_typing_impl with (1 := X0) => ?? HT. + eapply context_conversion. 3:eauto. all:eauto. - econstructor. eauto. eauto. eapply (All2i_All2_All2 X6 X5). intros ? ? ? (? & ?) (? & ? & (? & ?) & ? & ?) (? & ?). @@ -234,9 +233,9 @@ Proof. + now rewrite case_branch_type_fst (PCUICCasesContexts.inst_case_branch_context_eq a). - econstructor. - eapply All2_impl. eapply All2_All_mix_left. eassumption. eassumption. + eapply All2_impl. eapply All2_All_mix_left. exact X3. eassumption. intros. cbn in *. - destruct X5 as [[Ht IH] []]. + destruct X7 as [IH []]. split; intuition auto. eapply IH. + subst types. @@ -245,10 +244,11 @@ Proof. + assumption. - econstructor. - eapply All2_impl. eapply All2_All_mix_left. eassumption. eassumption. + eapply All2_impl. eapply All2_All_mix_left. exact X3. eassumption. intros. cbn in *. - decompose [prod] X2. intuition auto. - eapply b0. + destruct X7 as [IH []]. + split; intuition auto. + eapply IH. + subst types. eapply conv_context_app_same; auto. + subst types. eapply conv_context_wf_local_app; eauto. + assumption. @@ -307,25 +307,28 @@ Lemma erases_subst_instance0 consistent_instance_ext (Σ.1,univs) (Σ.2) u -> Σ ;;; Γ |- t ⇝ℇ t' -> (Σ.1,univs) ;;; (subst_instance u Γ) |- subst_instance u t ⇝ℇ t') + (fun Σ Γ j => wf_ext_wk Σ -> + forall u univs, + wf_local (Σ.1, univs) (subst_instance u Γ) -> + consistent_instance_ext (Σ.1,univs) (Σ.2) u -> + lift_wf_term (fun t => forall t', Σ ;;; Γ |- t ⇝ℇ t' -> (Σ.1,univs) ;;; (subst_instance u Γ) |- subst_instance u t ⇝ℇ t') j) (fun Σ Γ => wf_local Σ Γ). Proof. apply typing_ind_env; intros; cbn -[subst_instance] in *; auto. + { destruct X as (Htm & s & (_ & Hty) & _). split; eauto. destruct j_term => // ?. cbn in *. now eapply Htm. } all: match goal with [ H : erases _ _ ?a _ |- ?G ] => tryif is_var a then idtac else invs H end. all: try now (econstructor; eauto 2 using isErasable_subst_instance). - cbn. econstructor. - eapply H0 in X2; eauto. apply X2. - cbn. econstructor. eauto. cbn. econstructor. - eapply typing_subst_instance in X0; eauto. apply snd in X0. - cbn in X0. destruct X0. refine (t0 _ _ _ _); eauto. + eapply H in X3; eauto. apply X3. + cbn. econstructor. eauto. + eapply lift_typing_fu_impl with (1 := X0) => // ?? HT. + now apply typing_subst_instance. - cbn. econstructor. - eapply H0 in X3; eauto. - eapply H1 in X3; eauto. exact X3. - cbn. econstructor. eauto. cbn. econstructor. - eapply typing_subst_instance in X0; eauto. apply snd in X0. - cbn in X0. - eapply X0; eauto. - cbn. eapply typing_subst_instance in X1; eauto. apply snd in X1. - cbn in X1. eapply X1; eauto. + now edestruct X1; tea; eauto. + eapply H in X3; eauto. exact X3. + cbn. econstructor. eauto. + eapply lift_typing_fu_impl with (1 := X0) => // ?? HT. + now apply typing_subst_instance. - unfold subst_instance. cbn [subst_instance_constr]. econstructor; eauto. eapply All2_map_left. @@ -342,34 +345,14 @@ Proof. rewrite subst_instance_app. unfold app_context. f_equal. now rewrite inst_case_branch_context_subst_instance. - assert (Hw : wf_local (Σ.1, univs) (subst_instance u (Γ ,,, types))). - { (* rewrite subst_instance_app. *) - assert(All (fun d => isType Σ Γ (dtype d)) mfix). - eapply (All_impl X0); pcuicfo. - now destruct X5 as [s [Hs ?]]; exists s. - eapply All_mfix_wf in X5; auto. subst types. - - revert X5. clear - wfΣ wfΓ H2 X2 X3. - induction 1. - - eauto. - - cbn. econstructor; eauto. cbn in *. - destruct t0 as (? & ?). eexists. - cbn. eapply typing_subst_instance in t0; eauto. apply snd in t0. cbn in t0. - eapply t0; eauto. - - cbn. econstructor; eauto. cbn in *. - destruct t0 as (? & ?). eexists. - cbn. eapply typing_subst_instance in t0; eauto. apply snd in t0. - eapply t0; eauto. - cbn in *. eapply typing_subst_instance in t1; eauto. - apply snd in t1. eapply t1. all:eauto. - } + { eapply typing_subst_instance_wf_local; eauto. destruct Σ as [Σ Σu]. + now eapply All_mfix_wf. } cbn. econstructor; eauto. eapply All2_map_left. - eapply All2_impl. eapply All2_All_mix_left. eapply X1. - exact X4. - intros; cbn in *. destruct X5. - destruct p as [p IH]. - destruct a. split; auto. + eapply All2_impl. eapply All2_All_mix_left. eapply X3. eassumption. + intros d d' [IH []]; cbn in *. + split; auto. now eapply isLambda_subst_instance. eapply IH in e1. rewrite subst_instance_app in e1. subst types. 2:eauto. @@ -378,34 +361,15 @@ Proof. eapply fix_context_subst_instance. all: eauto. - assert (Hw : wf_local (Σ.1, univs) (subst_instance u (Γ ,,, types))). - { (* rewrite subst_instance_app. *) - assert(All (fun d => isType Σ Γ (dtype d)) mfix). - eapply (All_impl X0); pcuicfo. - destruct X5 as [s [Hs ?]]; now exists s. - eapply All_mfix_wf in X5; auto. subst types. - - revert X5. clear - wfΣ wfΓ H2 X2 X3. - induction 1. - - eauto. - - cbn. econstructor; eauto. cbn in *. - destruct t0 as (? & ?). eexists. - cbn. eapply typing_subst_instance in t0; eauto. apply snd in t0. cbn in t0. - eapply t0; eauto. - - cbn. econstructor; eauto. cbn in *. - destruct t0 as (? & ?). eexists. - cbn. eapply typing_subst_instance in t0; eauto. apply snd in t0. - eapply t0; eauto. - cbn in *. eapply typing_subst_instance in t1; eauto. - apply snd in t1. eapply t1. all:eauto. - } + { eapply typing_subst_instance_wf_local; eauto. destruct Σ as [Σ Σu]. + now eapply All_mfix_wf. } cbn. econstructor; eauto. eapply All2_map_left. - eapply All2_impl. eapply All2_All_mix_left. eapply X1. - exact X4. - intros; cbn in *. destruct X5. destruct p0. destruct p0. - destruct p. repeat split; eauto. - eapply e2 in e1. + eapply All2_impl. eapply All2_All_mix_left. eapply X3. eassumption. + intros d d' [IH (? & ? & ?)]; cbn in *. + repeat split; eauto. + eapply IH in e1. rewrite subst_instance_app in e1; eauto. subst types. 2:eauto. eapply erases_ctx_ext. eassumption. unfold app_context. f_equal. @@ -541,10 +505,12 @@ Section wellscoped. Arguments lookup_projection : simpl never. Lemma typing_wellformed : env_prop (fun Σ Γ a A => wellformed Σ a) + (fun Σ _ j => lift_wfb_term (wellformed Σ) j) (fun Σ Γ => True). Proof. eapply typing_ind_env; cbn; intros => //. - all:try rtoProp; intuition auto. + { destruct X as (Htm & _ & (_ & Hty) & _). unfold lift_wfb_term. rewrite Hty andb_true_r. destruct j_term => //. apply Htm. } + all:try unfold lift_wfb_term in *; rtoProp; intuition auto. - unshelve eapply declared_constant_to_gen in H0; eauto. red in H0. rewrite /lookup_constant /lookup_constant_gen H0 //. - unshelve eapply declared_inductive_to_gen in isdecl; eauto. @@ -560,12 +526,11 @@ Section wellscoped. unfold lookup_projection. now rewrite (declared_projection_lookup_gen isdecl). - now eapply nth_error_Some_length, Nat.ltb_lt in H0. - move/andb_and: H2 => [] hb _. - solve_all. destruct a as [s []], a0. - unfold test_def. len in b0. - rewrite b0. now rewrite i i0. + unfold on_def_body, on_def_type, test_def in *. cbn in *. + solve_all. - now eapply nth_error_Some_length, Nat.ltb_lt in H0. - - solve_all. destruct a as [s []], b. - unfold test_def. len in i0. now rewrite i i0. + - unfold on_def_body, on_def_type, test_def in *. cbn in *. + solve_all. - depelim X0; solve_all; constructor; eauto. Qed. diff --git a/erasure/theories/Extract.v b/erasure/theories/Extract.v index bf09291c7..3a3f72ee9 100644 --- a/erasure/theories/Extract.v +++ b/erasure/theories/Extract.v @@ -17,7 +17,7 @@ Local Existing Instance extraction_checker_flags. - is of propositional type *) Definition isErasable Σ Γ t := ∑ T, Σ ;;; Γ |- t : T × - (isArity T + (∑ u, (Σ ;;; Γ |- T : tSort u) * is_propositional u))%type. + (isArity T + (∑ u, (Σ ;;; Γ |- T : tSort u) * Sort.is_propositional u))%type. (* A more positive notion of relevant terms. Showing that a term is not erasable requires quantification over all its possible typings. @@ -30,7 +30,7 @@ Definition nisErasable Σ Γ t := nf Σ Γ T, ~ isArity T, Σ;;; Γ |- T : tSort u & - ~ is_propositional u]. + ~ Sort.is_propositional u]. Lemma nisErasable_spec Σ Γ t : wf_ext Σ -> wf_local Σ Γ -> @@ -44,12 +44,12 @@ Proof. eapply (proj2 (nf_red _ _ _ _)) in n. 2:eapply hr. subst. contradiction. eapply PCUICClassification.invert_cumul_arity_r_gen. 2:exact w. exists T'. split; auto. sq. - eapply PCUICValidity.validity in t2 as [s Hs]. + eapply PCUICValidity.validity in t2 as (_ & s & Hs & _). eapply PCUICClassification.wt_closed_red_refl; eauto. * destruct (principal_type _ _ t0) as [princ hprinc]. destruct s as [u' [hs isp]]. pose proof (hprinc _ t2) as []. - destruct (PCUICValidity.validity t3). + destruct (PCUICValidity.validity t3) as (_ & s & Hs & _). eapply PCUICElimination.unique_sorting_equality_propositional in hs; tea; eauto. pose proof (hprinc _ t0) as []. eapply PCUICElimination.unique_sorting_equality_propositional in t1; tea; eauto. @@ -386,7 +386,7 @@ Definition computational_ind Σ ind := match List.nth_error decl.(ind_bodies) n with | Some body => match destArity [] body.(ind_type) with - | Some arity => negb (Universe.is_prop (snd arity)) + | Some arity => negb (Sort.is_prop (snd arity)) | None => false end | None => false diff --git a/erasure/theories/Prelim.v b/erasure/theories/Prelim.v index 030afcb95..49d362695 100644 --- a/erasure/theories/Prelim.v +++ b/erasure/theories/Prelim.v @@ -211,7 +211,7 @@ Proof. eapply PCUICSpine.typing_spine_strengthen in t0; eauto. eapply PCUICSpine.type_mkApps; eauto. pose proof a0 as a0'. - eapply nth_error_all in a0'; eauto. simpl in a0'. + eapply nth_error_all in a0'; eauto. apply unlift_TermTyp in a0'. eapply (substitution (Δ := [])) in a0'; eauto. 2:{ eapply subslet_cofix_subst; pcuic. constructor; eauto. } rewrite PCUICLiftSubst.simpl_subst_k in a0'. now autorewrite with len. @@ -261,7 +261,7 @@ Lemma expand_lets_erasure (wfl := default_wcbv_flags) {Σ mdecl idecl cdecl c br declared_constructor Σ c mdecl idecl cdecl -> wf_branches idecl brs -> All2i (fun i cdecl br => - All2 (PCUICEquality.compare_decls eq eq) (bcontext br) + PCUICEquality.eq_context_upto_names (bcontext br) (cstr_branch_context c.1 mdecl cdecl)) 0 idecl.(ind_ctors) brs -> All (fun br => expand_lets (inst_case_branch_context p br) (bbody br) = bbody br) brs. @@ -312,7 +312,7 @@ Lemma subslet_cstr_branch_context {cf : checker_flags} {Σ : global_env_ext} {wf declared_constructor Σ (ind, n) mdecl idecl cdecl -> consistent_instance_ext Σ (ind_universes mdecl) u -> consistent_instance_ext Σ (ind_universes mdecl) (puinst p) -> - PCUICEquality.R_global_instance Σ (eq_universe Σ) (leq_universe Σ) (IndRef ind) napp u (puinst p) -> + PCUICEquality.cmp_ind_universes Σ ind napp u (puinst p) -> spine_subst Σ Γ pars parsubst (ind_params mdecl)@[u] -> spine_subst Σ Γ (pparams p) parsubst' (ind_params mdecl)@[puinst p] -> assumption_context cdecl.(cstr_args) -> diff --git a/erasure/theories/Typed/Erasure.v b/erasure/theories/Typed/Erasure.v index 967f1f175..df57eefa6 100644 --- a/erasure/theories/Typed/Erasure.v +++ b/erasure/theories/Typed/Erasure.v @@ -92,9 +92,10 @@ Lemma isType_red_sq Σ0 (wf : ∥ wf_ext Σ0 ∥) Γ t t' : ∥red Σ0 Γ t t'∥ -> ∥isType Σ0 Γ t'∥. Proof. - intros [(s & typ)] [r]. + intros [(_ & s & typ & _)] [r]. sq. - exists s. eapply subject_reduction; eauto. + eapply has_sort_isType. + eapply subject_reduction; eauto. Qed. Hint Resolve isType_red_sq : erase. @@ -103,10 +104,10 @@ Lemma isType_prod_dom Σ0 (wf : ∥ wf_ext Σ0 ∥) Γ na A B : ∥isType Σ0 Γ (tProd na A B)∥ -> ∥isType Σ0 Γ A∥. Proof. - intros [(s & typ)]. + intros [(_ & s & typ & _)]. sq. apply inversion_Prod in typ as (s' & ? & ? & ? & ?); [|now eauto]. - now exists s'. + now eapply lift_sorting_forget_univ. Qed. Hint Resolve isType_prod_dom : erase. @@ -115,10 +116,10 @@ Lemma isType_prod_cod Σ0 (wf : ∥ wf_ext Σ0 ∥) Γ na A B : ∥isType Σ0 Γ (tProd na A B)∥ -> ∥isType Σ0 (Γ,, vass na A) B∥. Proof. - intros [(s & typ)]. + intros [(_ & s & typ & _)]. sq. apply inversion_Prod in typ as (s' & ? & ? & ? & ?); [|now eauto]. - now exists x. + now eapply has_sort_isType. Qed. Hint Resolve isType_prod_cod : erase. @@ -419,7 +420,7 @@ Proof. destruct isT. now apply isType_welltyped. Qed. (** Definition of normalized arities *) Definition arity_ass := aname * term. -Fixpoint mkNormalArity (l : list arity_ass) (s : Universe.t) : term := +Fixpoint mkNormalArity (l : list arity_ass) (s : sort) : term := match l with | []%list => tSort s | ((na, A) :: l)%list => tProd na A (mkNormalArity l s) @@ -433,7 +434,7 @@ Qed. Record conv_arity {Σ Γ T} : Type := build_conv_arity { conv_ar_context : list arity_ass; - conv_ar_univ : Universe.t; + conv_ar_univ : sort; conv_ar_red : ∥closed_red Σ Γ T (mkNormalArity conv_ar_context conv_ar_univ)∥ }. @@ -593,7 +594,7 @@ flag_of_type Γ T isT with inspect (hnf (X_type := X_type) Γ T (fun Σ h => (ty end |} }; - | fot_view_sort univ := {| is_logical := Universe.is_prop univ; + | fot_view_sort univ := {| is_logical := Sort.is_prop univ; conv_ar := inl {| conv_ar_context := []; conv_ar_univ := univ; |} |} ; | fot_view_other T0 discr @@ -601,7 +602,7 @@ flag_of_type Γ T isT with inspect (hnf (X_type := X_type) Γ T (fun Σ h => (ty { | @existT K princK with inspect (reduce_to_sort Γ K _) := { | exist (Checked_comp (existT _ univ red_univ)) eq := - {| is_logical := Universe.is_prop univ; + {| is_logical := Sort.is_prop univ; conv_ar := inr _ |}; | exist (TypeError_comp t) eq := ! }; @@ -654,7 +655,7 @@ Next Obligation. Qed. Next Obligation. specialize (isT Σ wfΣ) as isT'. - sq. destruct isT' as [u Htype]. + sq. destruct isT' as (_ & u & Htype & _). now eapply typing_wf_local. Qed. Next Obligation. @@ -680,7 +681,7 @@ Next Obligation. { sq. eapply isType_red;eauto. now apply PCUICWellScopedCumulativity.closed_red_red. } sq. - destruct tyT as [u tyT]. + destruct tyT as (_ & u & tyT & _). apply typing_wf_local in tyT. apply BDToPCUIC.infering_typing in HH;sq;eauto. eapply isType_welltyped. @@ -710,8 +711,8 @@ Next Obligation. assert (tyNf : ∥ isType rΣ Γ nf ∥). { sq. eapply isType_red;eauto. } sq. - destruct tyT as [u tyT]. - destruct tyNf as [v tyNf]. + destruct tyT as (_ & u & tyT & _). + destruct tyNf as (_ & v & tyNf & _). apply typing_wf_local in tyT. apply BDToPCUIC.infering_typing in HH;sq;eauto. specialize (PCUICPrincipality.common_typing _ _ HH tyNf) as [x[x_le_K[x_le_sort?]]]. @@ -886,7 +887,7 @@ Next Obligation. Qed. Next Obligation. reduce_term_sound. - destruct (isT _ wfrΣ) as [(? & typ)]. + destruct (isT _ wfrΣ) as [(_ & ? & typ & _)]. assert (∥ wf rΣ ∥) by now apply HΣ. sq. destruct r. eapply subject_reduction with (u := tRel i) in typ; eauto. @@ -894,7 +895,7 @@ Next Obligation. now apply nth_error_Some. Qed. Next Obligation. - destruct (isT _ wfrΣ) as [(? & typ)]. + destruct (isT _ wfrΣ) as [(_ & ? & typ & _)]. assert (∥ wf rΣ ∥) by eauto using HΣ;sq. reduce_term_sound;destruct r. eapply subject_reduction in typ; eauto. @@ -910,11 +911,11 @@ Next Obligation. Qed. Next Obligation. specialize (isT _ wfΣ) as isT'. - sq. destruct isT' as [u Htype]. + sq. destruct isT' as (_ & u & Htype & _). now eapply typing_wf_local. Qed. Next Obligation. - destruct (isT _ wfΣ) as [(? & typ)]. + destruct (isT _ wfΣ) as [(_ & ? & typ & _)]. assert (w : ∥ wf Σ ∥) by eauto using HΣ;sq. reduce_term_sound;destruct r. eapply subject_reduction in typ; eauto. @@ -938,7 +939,7 @@ Next Obligation. destruct (princaT _ wfΣ) as [inf_aT]. assert (HH : ∥ wf_ext Σ0 ∥) by now apply heΣ. destruct HH. - specialize (isT _ wfΣ) as [[? Hty]]. + specialize (isT _ wfΣ) as [(_ & ? & Hty & _)]. apply typing_wf_local in Hty. apply BDToPCUIC.infering_typing in inf_aT;eauto with erase. sq. @@ -952,12 +953,12 @@ Next Obligation. destruct (princaT _ wfΣ) as [inf_aT]. assert (HH : ∥ wf_ext Σ ∥) by now apply heΣ. destruct HH. - specialize (isT _ wfΣ) as [[? Hty]]. + specialize (isT _ wfΣ) as [(_ & ? & Hty & _)]. apply typing_wf_local in Hty. apply BDToPCUIC.infering_typing in inf_aT;eauto with erase. destruct conv_sort as (univ & reduniv). sq. - exists univ. + eapply has_sort_isType. eapply type_reduction;eauto. Qed. Next Obligation. @@ -1004,7 +1005,7 @@ Equations (noeqns) erase_type_scheme_eta (erΓ : Vector.t tRel_kind #|Γ|) (t : term) (ar_ctx : list arity_ass) - (ar_univ : Universe.t) + (ar_univ : sort) (typ : ∥rΣ;;; Γ |- t : mkNormalArity ar_ctx ar_univ∥) (next_tvar : nat) : list type_var_info × box_type := erase_type_scheme_eta Γ erΓ t [] univ typ next_tvar => ([], (erase_type_aux Γ erΓ t _ None).2); @@ -1029,7 +1030,7 @@ Next Obligation. destruct typ. assert (wf_local rΣ Γ) by (eapply typing_wf_local; eauto). assert (∥ wf rΣ ∥) by now apply HΣ . - constructor; eexists;eassumption. + constructor; eapply has_sort_isType;eassumption. Qed. Next Obligation. destruct typ as [typ]. @@ -1051,9 +1052,9 @@ Next Obligation. eapply isType_wf_local; eauto. } rewrite <- (subst_rel0_lift_id 0 (mkNormalArity ar_ctx univ)). eapply validity in typ as typ_valid;auto. - destruct typ_valid as [u Hty]. + destruct typ_valid as (_ & u & Hty & _). eapply type_App. - + eapply validity in typ as typ;auto. + + eapply validity in typ as (_ & ? & typ & _);auto. eapply (PCUICWeakeningTyp.weakening _ _ [_] _ _ _ wflext Hty). + eapply (PCUICWeakeningTyp.weakening _ _ [_] _ _ _ wflext typ). + fold lift. @@ -1065,7 +1066,7 @@ Equations? (noeqns) erase_type_scheme (erΓ : Vector.t tRel_kind #|Γ|) (t : term) (ar_ctx : list arity_ass) - (ar_univ : Universe.t) + (ar_univ : sort) (typ : forall Σ0 (wfΣ : PCUICWfEnv.abstract_env_ext_rel X Σ0), ∥Σ0;;; Γ |- t : mkNormalArity ar_ctx ar_univ∥) (next_tvar : nat) : list type_var_info × box_type := erase_type_scheme Γ erΓ t [] univ typ next_tvar => ([], (erase_type_aux Γ erΓ t _ None).2); @@ -1089,15 +1090,14 @@ erase_type_scheme Γ erΓ t ((na', A') :: ar_ctx) univ typ next_tvar }. Proof. - destruct (typ _ wfΣ). - constructor; eexists; eauto. + constructor; eapply has_sort_isType; eauto. - destruct (typ _ wfΣ) as [typ0]. reduce_term_sound. assert (∥ wf Σ0 ∥) by now apply HΣ. sq. destruct r as [?? r]. eapply subject_reduction in r; eauto. - apply inversion_Lambda in r as (?&?&?&?&?); auto. - eexists; eassumption. + apply inversion_Lambda in r as (?&?&?&?); auto. - clear inf. destruct (typ _ wfΣ) as [typ0]. reduce_term_sound. @@ -1108,11 +1108,11 @@ Proof. { eapply abstract_env_ext_irr;eauto. } subst. eapply subject_reduction in r; eauto. - apply inversion_Lambda in r as (?&?&?&?&c); auto. + apply inversion_Lambda in r as (?&?&?&c); auto. assert (wf_local Σ0 Γ) by (eapply typing_wf_local; eauto). apply ws_cumul_pb_Prod_Prod_inv_l in c as [???]; auto. eapply validity in typ0 as typ0; auto. - apply isType_tProd in typ0 as (_ & (u&?)); auto. + apply isType_tProd in typ0 as (_ & (_&u&?&_)); auto. assert (PCUICCumulativity.conv_context cumulAlgo_gen Σ0 (Γ,, vass na' A') (Γ,, vass na A)). { constructor; [reflexivity|]. constructor. now symmetry. @@ -1121,7 +1121,6 @@ Proof. eapply type_Cumul. + eassumption. + eapply PCUICContextConversionTyp.context_conversion; eauto. - eapply typing_wf_local; eassumption. + now apply cumulAlgo_cumulSpec. Qed. @@ -1147,6 +1146,7 @@ Proof. assert (∥ wf Σ0 ∥) by now apply HΣ. destruct car as [ctx univ r]. sq. + apply unlift_TermTyp in wt. eapply type_reduction in wt; eauto;cbn. now destruct r. Qed. @@ -1169,7 +1169,7 @@ Proof. assert (∥ wf Σ0 ∥) by now apply HΣ. unfold on_constant_decl in wt. destruct (PCUICEnvironment.cst_body cst); cbn in *. - + sq;eapply validity;eauto. + + sq;eapply validity;eauto. now eapply unlift_TermTyp. + destruct wt. eexists; eassumption. - assert (rΣ = Σ). @@ -1182,11 +1182,8 @@ Proof. unfold on_constant_decl in wt. destruct (PCUICEnvironment.cst_body cst). + sq. - now eapply validity in wt. - + sq. - cbn in wt. - destruct wt as (s & ?). - now exists s. + now eapply unlift_TermTyp, validity in wt. + + assumption. Qed. Import P. @@ -1311,7 +1308,7 @@ Proof. unshelve refine ( let is_propositional := match destArity [] (ind_type oib) with - | Some (_, u) => is_propositional u + | Some (_, u) => Sort.is_propositional u | None => false end in let oib_tvars := erase_ind_arity [] (PCUICEnvironment.ind_type oib) _ in @@ -1356,10 +1353,10 @@ Proof. induction on_ctors; [easy|]; destruct is_in as [->|later]; [|easy]; constructor; - destruct (on_ctype r) as (s & typ); + destruct (on_ctype r) as (_ & s & typ & _); rewrite <- (arities_contexts_1 mind) in typ; cbn in *; - now exists s). + now eapply has_sort_isType). Defined. Program Definition erase_ind diff --git a/erasure/theories/Typed/ErasureCorrectness.v b/erasure/theories/Typed/ErasureCorrectness.v index 94a946048..d61b48c2a 100644 --- a/erasure/theories/Typed/ErasureCorrectness.v +++ b/erasure/theories/Typed/ErasureCorrectness.v @@ -219,7 +219,7 @@ Proof. left. apply isArity_mkNormalArity. clear wfΣext. specialize (wfΣ _ hfull). subst Σfull. - depelim wffull; cbn in *. depelim o0. now depelim o1. + depelim wffull; cbn in *. depelim o0. depelim o1. now eapply unlift_TermTyp. * eexists; split;cbn; unfold EGlobalEnv.declared_constant;cbn. now rewrite eq_kername_refl. unfold trans_cst; cbn. @@ -251,7 +251,7 @@ Proof. 2:{ depelim wfΣ. depelim o0. depelim o2. now cbn in on_global_decl_d. } *) eapply (erases_extends (_, _)); eauto. - 1:{ depelim wffull. depelim o0. depelim o1. now cbn in on_global_decl_d. } + 1:{ depelim wffull. depelim o0. depelim o1. eapply unlift_TermTyp. now cbn in on_global_decl_d. } 1:{ eexists. reflexivity. eexists [_]; reflexivity. reflexivity. } now apply erases_erase. -- intros. @@ -275,6 +275,7 @@ Proof. { now depelim wffull. } depelim wffull. cbn in *. depelim o0. depelim o1. cbn in on_global_decl_d. eapply (@erase_global_erases_deps (_, _)); eauto. + now eapply unlift_TermTyp. now apply erases_erase. eapply IH. ++ intros ? isin'. @@ -283,6 +284,7 @@ Proof. ++ intros ? isin'. eapply term_global_deps_spec in isin' as [(? & ?)]; eauto. ** cbn in *. congruence. + ** now eapply unlift_TermTyp. ** now apply erases_erase. + eexists _, _; split; [left; reflexivity|]; split. unfold EGlobalEnv.declared_minductive;cbn. @@ -300,4 +302,4 @@ Proof. intros; eapply erase_global_decls_deps_recursive_correct_gen; eauto. reflexivity. Qed. -End EEnvCorrect. \ No newline at end of file +End EEnvCorrect. diff --git a/erasure/theories/Typed/ExtractionCorrectness.v b/erasure/theories/Typed/ExtractionCorrectness.v index 30683c4bd..1a9db7311 100644 --- a/erasure/theories/Typed/ExtractionCorrectness.v +++ b/erasure/theories/Typed/ExtractionCorrectness.v @@ -205,7 +205,7 @@ Proof. destruct hΣpop as [? []]. destruct Σpop; cbn in *. subst. depelim wf. depelim o0. depelim o1. cbn in on_global_decl_d. sq. exists ar. split; [|now left]. - unshelve eapply PCUICSR.type_reduction. 4:apply X0. constructor; eauto. exact on_global_decl_d. + unshelve eapply PCUICSR.type_reduction. 4:apply X0. constructor; eauto. now eapply unlift_TermTyp. reflexivity. } { cbn [trans_global_decl]. unfold trans_cst, cst_body. unfold erase_constant_body; cbn -[erase]. do 3 f_equal. @@ -701,7 +701,7 @@ Proof. 2:{ cbn; congruence. } 2:exact deps. intros. eapply Erasure.fake_normalization; eauto. * eapply erase_global_erases_deps; tea. subst er. unfold erase_env. - set (obl := fun Σ' => _). + set (obl := fun (Σ' : global_env) => _). destruct Σ; cbn [declarations retroknowledge universes] in *. eapply (erase_global_decls_deps_recursive_correct _ _ _ obl); eauto. intros k hin. diff --git a/erasure/theories/Typed/TypeAnnotations.v b/erasure/theories/Typed/TypeAnnotations.v index f82f22bd7..55c322504 100644 --- a/erasure/theories/Typed/TypeAnnotations.v +++ b/erasure/theories/Typed/TypeAnnotations.v @@ -291,15 +291,14 @@ Proof. all: try subst erΓ1. - depelim er0. now apply inversion_Evar in X. - - apply inversion_Lambda in X as (? & ? & ? & ? & ?); auto. - constructor;econstructor; eauto. - - apply inversion_Lambda in X as (? & ? & ? & ? & ?); auto. + - apply inversion_Lambda in X as (? & ? & ? & ?); auto. + - apply inversion_Lambda in X as (? & ? & ? & ?); auto. econstructor; eauto. - - apply inversion_LetIn in X as (?&?&?&?&?&?); auto. - constructor;econstructor; eauto. - - apply inversion_LetIn in X as (?&?&?&?&?&?); auto. - econstructor; eauto. - - apply inversion_LetIn in X as (?&?&?&?&?&?); auto. + - apply inversion_LetIn in X as (?&?&?&?); auto. + constructor; eapply lift_sorting_forget_body; eauto. + - apply inversion_LetIn in X as (?&?&?&?); auto. + econstructor; eapply unlift_TermTyp; eauto. + - apply inversion_LetIn in X as (?&?&?&?); auto. econstructor; eauto. - apply inversion_App in X as (?&?&?&?&?&?); auto. econstructor; eauto. @@ -339,12 +338,14 @@ Proof. eapply All2_All_mix_left in X; eauto. clear -X. revert X. + unfold on_def_body. generalize (Γ,,, fix_context defs). clear Γ. intros Γ a. induction a; [now constructor|]. constructor; [|now eauto]. destruct r as [?[??]]. + apply unlift_TermTyp in l0. split; [|now auto]. econstructor; eauto. - apply inversion_CoFix in X as (?&?&?&?&?&?&?); auto. @@ -358,12 +359,14 @@ Proof. eapply All2_All_mix_left in X; eauto. clear -X. revert X. + unfold on_def_body. generalize (Γ,,, fix_context defs). clear Γ. intros Γ1 a. induction a; [now constructor|]. constructor; [|now eauto]. destruct r as (? & ? & ? & ?). + apply unlift_TermTyp in l0. split; [|now auto]. econstructor; eauto. Qed. @@ -389,7 +392,7 @@ Proof. cbn. apply (annotate_types [] []%vector t); [|apply ErasureFunction.erases_erase]. cbn in *. - destruct wt. + destruct wt. apply unlift_TermTyp in X0. econstructor; eauto. exact eq. Defined. diff --git a/examples/demo.v b/examples/demo.v index 6b65cf4ad..0b6e0ddf1 100644 --- a/examples/demo.v +++ b/examples/demo.v @@ -113,7 +113,7 @@ MetaCoq Test Quote ((pair' _ _ true 4).(snd')). Definition one_i : one_inductive_entry := {| mind_entry_typename := "demoBool"; - mind_entry_arity := tSort Universe.type0; + mind_entry_arity := tSort Sort.type0; mind_entry_consnames := ["demoTrue"; "demoFalse"]; mind_entry_lc := [tRel 1; tRel 1]; |}. @@ -121,7 +121,7 @@ Definition one_i : one_inductive_entry := Definition one_i2 : one_inductive_entry := {| mind_entry_typename := "demoBool2"; - mind_entry_arity := tSort Universe.type0; + mind_entry_arity := tSort Sort.type0; mind_entry_consnames := ["demoTrue2"; "demoFalse2"]; mind_entry_lc := [tRel 0; tRel 0]; |}. @@ -150,7 +150,7 @@ Definition mkImpl (A B : term) : term := Definition one_list_i : one_inductive_entry := {| mind_entry_typename := "demoList"; - mind_entry_arity := tSort Universe.type0; + mind_entry_arity := tSort Sort.type0; mind_entry_consnames := ["demoNil"; "demoCons"]; mind_entry_lc := [tApp (tRel 1) [tRel 0]; mkImpl (tRel 0) (mkImpl (tApp (tRel 2) [tRel 1]) (tApp (tRel 3) [tRel 2]))]; @@ -161,7 +161,7 @@ Definition mut_list_i : mutual_inductive_entry := mind_entry_record := None; mind_entry_finite := Finite; mind_entry_params := [{| decl_name := bnamed "A"; decl_body := None; - decl_type := (tSort Universe.type0) |}]; + decl_type := (tSort Sort.type0) |}]; mind_entry_inds := [one_list_i]; mind_entry_universes := Monomorphic_entry (LevelSet.empty, ConstraintSet.empty); mind_entry_template := false; @@ -177,7 +177,7 @@ MetaCoq Unquote Inductive mut_list_i. Definition one_pt_i : one_inductive_entry := {| mind_entry_typename := "Point"; - mind_entry_arity := tSort Universe.type0; + mind_entry_arity := tSort Sort.type0; mind_entry_consnames := ["mkPoint"]; mind_entry_lc := [ mkImpl (tRel 0) (mkImpl (tRel 1) (tApp (tRel 3) [tRel 2]))]; @@ -188,7 +188,7 @@ Definition mut_pt_i : mutual_inductive_entry := mind_entry_record := Some (Some "pp"); mind_entry_finite := BiFinite; mind_entry_params := [{| decl_name := bnamed "A"; decl_body := None; - decl_type := (tSort Universe.type0) |}]; + decl_type := (tSort Sort.type0) |}]; mind_entry_inds := [one_pt_i]; mind_entry_universes := Monomorphic_entry ContextSet.empty; mind_entry_template := false; @@ -374,10 +374,10 @@ Inductive T : Type := MetaCoq Quote Recursively Definition TT := T. Unset MetaCoq Strict Unquote Universe Mode. -MetaCoq Unquote Definition t := (tSort (Universe.make (Level.level "Top.20000"))). -MetaCoq Unquote Definition t' := (tSort fresh_universe). -MetaCoq Unquote Definition myProp := (tSort (Universe.lProp)). -MetaCoq Unquote Definition mySet := (tSort (Universe.make Level.lzero)). +MetaCoq Unquote Definition t := (tSort (sType (Universe.make' (Level.level "Top.20000")))). +MetaCoq Unquote Definition t' := (tSort (sType fresh_universe)). +MetaCoq Unquote Definition myProp := (tSort sProp). +MetaCoq Unquote Definition mySet := (tSort (sType (Universe.make' Level.lzero))). (** Cofixpoints *) CoInductive streamn : Set := diff --git a/examples/typing_correctness.v b/examples/typing_correctness.v index 52db52d6b..5637dd0a8 100644 --- a/examples/typing_correctness.v +++ b/examples/typing_correctness.v @@ -145,9 +145,9 @@ Ltac fill_inh t := | [ |- inh _ ?Γ _ ] => fail "Missing local wellformedness assumption for" Γ end. -(* Lemma identity_typing (u := Universe.make univ): inh gctx_wf_env [] (tImpl (tSort u) (tSort u)). +(* Lemma identity_typing (s := sType (Universe.make' univ)): inh gctx_wf_env [] (tImpl (tSort s) (tSort s)). Proof. - set (impl := tLambda (bNamed "s") (tSort u) (tRel 0)). + set (impl := tLambda (bNamed "s") (tSort s) (tRel 0)). assert (wfΓ : forall Σ0 : global_env_ext, abstract_env_ext_rel gctx_wf_env Σ0 -> ∥ wf_local Σ0 [] ∥) by do 2 constructor. @@ -156,7 +156,7 @@ Proof. Time Qed. *) -Lemma identity_typing (u := Universe.make univ): +Lemma identity_typing (s := sType (Universe.make' univ)): (∑ t : term, forall Σ0 : global_env_ext, Σ0 = @@ -168,13 +168,13 @@ Lemma identity_typing (u := Universe.make univ): retroknowledge := Retroknowledge.empty |}, Monomorphic_ctx) -> ∥ Σ0;;; [] |- t - : tProd (bNamed "s") (tSort u) (tImpl (tRel 0) (tRel 0)) ∥). + : tProd (bNamed "s") (tSort s) (tImpl (tRel 0) (tRel 0)) ∥). (* inh gctx_wf_env [] (tProd (bNamed "s") (tSort u) (tImpl (tRel 0) (tRel 0))). *) Proof. - set (impl := tLambda (bNamed "s") (tSort u) (tLambda bAnon (tRel 0) (tRel 0))). + set (impl := tLambda (bNamed "s") (tSort s) (tLambda bAnon (tRel 0) (tRel 0))). assert (wfΓ : forall Σ0 : global_env_ext, abstract_env_ext_rel gctx_wf_env Σ0 -> ∥ wf_local Σ0 [] ∥) by do 2 constructor. - pose (T := tProd (bNamed "s") (tSort u) (tImpl (tRel 0) (tRel 0))). + pose (T := tProd (bNamed "s") (tSort s) (tImpl (tRel 0) (tRel 0))). pose (Σ := gctx_wf_env). let t := uconstr:(check_inh Σ [] wfΓ impl (T:=T)) in let proof := eval hnf in t in diff --git a/pcuic/theories/Bidirectional/BDFromPCUIC.v b/pcuic/theories/Bidirectional/BDFromPCUIC.v index 8f806cd2d..d92a5a6fb 100644 --- a/pcuic/theories/Bidirectional/BDFromPCUIC.v +++ b/pcuic/theories/Bidirectional/BDFromPCUIC.v @@ -10,34 +10,20 @@ Require Import Equations.Prop.DepElim. Implicit Types (cf : checker_flags) (Σ : global_env_ext). -(** Preliminary lemmas missing from MetaCoq *) -Lemma is_allowed_elimination_monotone `{checker_flags} Σ s1 s2 allowed : - leq_universe Σ s1 s2 -> is_allowed_elimination Σ allowed s2 -> is_allowed_elimination Σ allowed s1. +Lemma ctx_inst_length {ty Γ args Δ} : + PCUICTyping.ctx_inst ty Γ args Δ -> + #|args| = context_assumptions Δ. Proof. - destruct allowed, s2; cbnr; trivial; - destruct s1; cbnr; intros H1 H2; trivial; try now destruct H1. - { now left. } - destruct H2 as [|H2]; [now left|right]. - unfold_univ_rel. - specialize (H1 v Hv); specialize (H2 v Hv). - cbn in H2. - lia. + induction 1; simpl; auto. + rewrite /subst_telescope in IHX. + rewrite context_assumptions_mapi in IHX. congruence. + rewrite context_assumptions_mapi in IHX. congruence. Qed. -Lemma ctx_inst_length {ty Σ Γ args Δ} : -PCUICTyping.ctx_inst ty Σ Γ args Δ -> -#|args| = context_assumptions Δ. -Proof. -induction 1; simpl; auto. -rewrite /subst_telescope in IHX. -rewrite context_assumptions_mapi in IHX. congruence. -rewrite context_assumptions_mapi in IHX. congruence. -Qed. - -Lemma ctx_inst_app_impl {P Q Σ Γ} {Δ : context} {Δ' args} (c : PCUICTyping.ctx_inst P Σ Γ args (Δ ++ Δ')) : - (forall Γ' t T, P Σ Γ' t T -> Q Σ Γ' t T) -> - PCUICTyping.ctx_inst Q Σ Γ (firstn (context_assumptions Δ) args) Δ. +Lemma ctx_inst_app_impl {P Q Γ} {Δ : context} {Δ' args} (c : PCUICTyping.ctx_inst P Γ args (Δ ++ Δ')) : + (forall t T, P Γ t T -> Q Γ t T) -> + PCUICTyping.ctx_inst Q Γ (firstn (context_assumptions Δ) args) Δ. Proof. revert args Δ' c. induction Δ using ctx_length_ind; intros. @@ -69,7 +55,7 @@ Qed. Lemma conv_infer_sort `{checker_flags} Σ (wfΣ : wf Σ) Γ t s : Σ ;;; Γ |- t : tSort s -> (∑ T' : term, Σ ;;; Γ |- t ▹ T' × Σ ;;; Γ ⊢ T' ≤ tSort s) -> - {s' & Σ ;;; Γ |- t ▹□ s' × leq_universe Σ s' s}. + {s' & Σ ;;; Γ |- t ▹□ s' × leq_sort Σ s' s}. Proof. intros tyt (T'&?&Cumt). apply ws_cumul_pb_Sort_r_inv in Cumt as (?&?&?) ; auto. @@ -96,7 +82,7 @@ Lemma conv_infer_ind `{checker_flags} Σ (wfΣ : wf Σ) Γ t ind ui args : Σ ;;; Γ |- t : mkApps (tInd ind ui) args -> (∑ T', (Σ ;;; Γ |- t ▹ T') × (Σ ;;; Γ ⊢ T' ≤ mkApps (tInd ind ui) args)) -> ∑ ui' args', [× Σ ;;; Γ |- t ▹{ind} (ui',args'), - R_global_instance Σ (eq_universe Σ) (leq_universe Σ) (IndRef ind) #|args| ui' ui + cmp_ind_universes Σ ind #|args| ui' ui & ws_cumul_pb_terms Σ Γ args' args]. Proof. intros tyt (T'&?&Cumt). @@ -106,33 +92,49 @@ Proof. now apply closed_red_red. Qed. +Lemma conv_lift_judgment `{checker_flags} Σ Γ na b ty : + lift_sorting (checking Σ Γ) (fun T u => ∑ u', Σ;;; Γ |- T ▹□ u' × leq_sort Σ u' u) (j_decl (mkdecl na b ty)) -> + lift_sorting (checking Σ Γ) (infering_sort Σ Γ) (j_decl (mkdecl na b ty)). +Proof. + intro Hj. + apply lift_sorting_ex_it_impl_gen with Hj => //=. + 1: now destruct b. + intros (u' & ? & _); now eexists. +Qed. + +Lemma conv_lift_judgment_univ `{checker_flags} Σ Γ na b ty u : + lift_sorting (checking Σ Γ) (fun T u => ∑ u', Σ;;; Γ |- T ▹□ u' × leq_sort Σ u' u) (j_decl_s (mkdecl na b ty) (Some u)) -> + ∑ u', lift_sorting (checking Σ Γ) (infering_sort Σ Γ) (j_decl_s (mkdecl na b ty) (Some u')) × leq_sort Σ u' u. +Proof. + intros (Htm & u0 & (u' & Hty & Hle) & <-); cbn in *. + exists u'; split; tas. + repeat (eexists; tea). +Qed. + Section BDFromPCUIC. (** The big theorem*) Lemma bidirectional_from_pcuic `{checker_flags} : env_prop (fun Σ Γ t T => {T' & Σ ;;; Γ |- t ▹ T' × Σ ;;; Γ ⊢ T' ≤ T}) + (fun Σ Γ j => lift_sorting (fun t T => Σ ;;; Γ |- t ◃ T) (fun T u => ∑ u', Σ ;;; Γ |- T ▹□ u' × leq_sort Σ u' u) j) (fun Σ Γ => wf_local_bd Σ Γ). Proof. apply typing_ind_env. + { intros Σ wfΣ Γ j Hj. + apply lift_sorting_impl with (1 := Hj) => //. + - intros t T []. now apply conv_check. + - intros t u [HT0 HT1]. + eapply conv_infer_sort. all: done. } + all: intros Σ wfΣ Γ wfΓ. - - intros bdwfΓ. - induction bdwfΓ. - all: constructor ; auto. - + apply conv_infer_sort in Hs ; auto. - 2: by destruct tu. - destruct Hs as (?&?&?). - eexists. - eassumption. - + apply conv_check in Hc ; auto. - apply conv_infer_sort in Hs ; auto. - 2: by destruct tu. - destruct Hs as (?&?&?). - 1: eexists. - all: eassumption. - + by apply conv_check in Hc. + - intros _ bdwfΓ. + clear wfΓ. + apply All_local_env_impl_gen with (1 := bdwfΓ). clear Γ bdwfΓ. + intros ??. + apply conv_lift_judgment with (na := decl.(decl_name)). - intros. eexists. @@ -148,9 +150,8 @@ Proof. constructor. assumption. - - intros n A B ? ? ? ? CumA ? CumB. - apply conv_infer_sort in CumA ; auto. - destruct CumA as (?&?&?). + - intros na A B ? ? ? ? CumA ? CumB. + apply conv_lift_judgment_univ with (na := na) in CumA as (?&?&?). apply conv_infer_sort in CumB ; auto. destruct CumB as (?&?&?). eexists. @@ -159,23 +160,19 @@ Proof. + constructor ; cbn ; auto. 1: by apply wf_local_closed_context. constructor. - apply leq_universe_product_mon. + apply leq_sort_product_mon. all: assumption. - - intros n A t ? ? ? ? CumA ? (?&?&?). - apply conv_infer_sort in CumA ; auto. - destruct CumA as (?&?&?). + - intros na A t ? ? ? CumA ? (?&?&?). + apply conv_lift_judgment with (na := na) in CumA. eexists. split. + econstructor. all: eassumption. + apply ws_cumul_pb_Prod ; auto. - eapply isType_ws_cumul_pb_refl. - by eexists ; eauto. + by eapply isType_ws_cumul_pb_refl. - - intros n t A u ? ? ? ? CumA ? Cumt ? (?&?&?). - apply conv_infer_sort in CumA ; auto. - destruct CumA as (?&?&?). - apply conv_check in Cumt ; auto. + - intros na t A ? ? ? ? CumtA ? (?&?&?). + apply conv_lift_judgment with (na := na) in CumtA. eexists. split. + econstructor. @@ -235,7 +232,7 @@ Proof. * rewrite subst_instance_app_ctx rev_app_distr in Hinst. replace (pparams p) with (firstn (context_assumptions (List.rev (subst_instance (puinst p)(ind_params mdecl)))) (pparams p ++ indices)). eapply ctx_inst_app_impl ; tea. - 1: intros ??? [] ; by apply conv_check. + 1: intros ?? [] ; by apply conv_check. rewrite context_assumptions_rev context_assumptions_subst_instance. erewrite PCUICGlobalMaps.onNpars. 2: eapply on_declared_minductive ; eauto. @@ -306,13 +303,13 @@ Proof. by (apply infering_ind_typing in i ; auto). assert (consistent_instance_ext Σ (ind_universes mdecl) u). { destruct isdecl. - apply validity in X1 as []. + apply validity in X1 as (_ & s1 & Hs1 & _). eapply invert_type_mkApps_ind ; eauto. apply H1.p1. } assert (consistent_instance_ext Σ (ind_universes mdecl) ui'). { destruct isdecl. - apply validity in X2 as [] ; auto. + apply validity in X2 as (_ & s2 & Hs2 & _) ; auto. eapply invert_type_mkApps_ind ; eauto. eapply H2.p1. } @@ -341,39 +338,27 @@ Proof. unshelve eapply declared_projection_to_gen in isdecl; eauto. eapply (weaken_lookup_on_global_env' _ _ (InductiveDecl _)); eauto. apply isdecl. - - intros mfix n decl types ? ? ? Alltypes Allbodies. + - intros mfix n decl types ? ? ? Htypes Alltypes Hbodies Allbodies. eexists. split. 2:{ apply isType_ws_cumul_pb_refl. - eapply nth_error_all in Alltypes as [? []] ; tea. - eexists ; tea. + eapply nth_error_all in Htypes as Hj ; tea. } constructor ; eauto. - + apply (All_impl Alltypes). - intros ? [? [? s]]. - apply conv_infer_sort in s as [? []] ; auto. - eexists ; eauto. - + apply (All_impl Allbodies). - intros ? [? s]. - by apply conv_check in s ; auto. - - - intros mfix n decl types ? ? ? Alltypes Allbodies. + + apply (All_impl Alltypes) => d. apply conv_lift_judgment with (na := d.(dname)). + + apply (All_impl Allbodies) => d. apply conv_lift_judgment with (na := d.(dname)). + + - intros mfix n decl types ? ? ? Htypes Alltypes Hbodies Allbodies. eexists. split. 2:{ apply isType_ws_cumul_pb_refl. - eapply nth_error_all in Alltypes as [? []] ; tea. - eexists ; tea. + eapply nth_error_all in Htypes as Hj ; tea. } constructor ; eauto. - + apply (All_impl Alltypes). - intros ? [? [? s]]. - apply conv_infer_sort in s as [? []] ; auto. - eexists ; eauto. - + apply (All_impl Allbodies). - intros ? [? s]. - by apply conv_check in s ; auto. + + apply (All_impl Alltypes) => d. apply conv_lift_judgment with (na := d.(dname)). + + apply (All_impl Allbodies) => d. apply conv_lift_judgment with (na := d.(dname)). - intros p prim_ty cdecl wfΓ' hp hdecl pinv. eexists. split; [econstructor; tea|]. @@ -415,7 +400,7 @@ Proof. Qed. Lemma typing_infering_sort `{checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} {Γ t u} : - Σ ;;; Γ |- t : tSort u -> ∑ u', Σ ;;; Γ |- t ▹□ u' × leq_universe Σ u' u. + Σ ;;; Γ |- t : tSort u -> ∑ u', Σ ;;; Γ |- t ▹□ u' × leq_sort Σ u' u. Proof. intros. apply conv_infer_sort ; tea. @@ -425,7 +410,7 @@ Qed. Lemma isType_infering_sort `{checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} {Γ t} : isType Σ Γ t -> ∑ u', Σ ;;; Γ |- t ▹□ u'. Proof. - intros [? ty]. + intros (_ & s & ty & _). eapply typing_infering_sort in ty as [? []]; tea. now eexists. Qed. @@ -443,7 +428,7 @@ Qed. Lemma typing_infer_ind `{checker_flags} Σ (wfΣ : wf Σ) Γ t ind ui args : Σ ;;; Γ |- t : mkApps (tInd ind ui) args -> ∑ ui' args', [× Σ ;;; Γ |- t ▹{ind} (ui',args'), - R_global_instance Σ (eq_universe Σ) (leq_universe Σ) (IndRef ind) #|args| ui' ui + cmp_ind_universes Σ ind #|args| ui' ui & ws_cumul_pb_terms Σ Γ args' args]. Proof. intros. @@ -465,17 +450,16 @@ Lemma wf_local_rel_wf_local_bd `{checker_flags} {Σ} (wfΣ : wf Σ) {Γ Γ'} : wf_local_bd_rel Σ Γ Γ'. Proof. intros wfΓ'. - induction Γ' as [|[? [] ?]] in wfΓ' |- *. - all: constructor ; inversion wfΓ' ; subst ; cbn in *. - 1,4: now apply IHΓ'. - - now apply isType_infering_sort. - - now apply typing_checking. - - now apply isType_infering_sort. + apply All_local_env_impl_gen with (1 := wfΓ'); clear Γ' wfΓ'. + intros Γ' ? Hj. + apply lift_sorting_ex_it_impl_gen with Hj => //. + - destruct decl_body => //= Hb. now apply typing_checking. + - intro Hty. eapply typing_infering_sort in Hty as [? []]; tea. now eexists. Qed. Theorem ctx_inst_typing_bd `{checker_flags} (Σ : global_env_ext) Γ l Δ (wfΣ : wf Σ) : - PCUICTyping.ctx_inst typing Σ Γ l Δ -> - PCUICTyping.ctx_inst checking Σ Γ l Δ. + PCUICTyping.ctx_inst (typing Σ) Γ l Δ -> + PCUICTyping.ctx_inst (checking Σ) Γ l Δ. Proof. intros inl. induction inl. diff --git a/pcuic/theories/Bidirectional/BDStrengthening.v b/pcuic/theories/Bidirectional/BDStrengthening.v index 0b141ed31..40c0c72b5 100644 --- a/pcuic/theories/Bidirectional/BDStrengthening.v +++ b/pcuic/theories/Bidirectional/BDStrengthening.v @@ -358,7 +358,7 @@ Section OnFreeVars. on_free_vars P t -> on_free_vars P T. - Let Psort (Γ : context) (t : term) (u : Universe.t) := True. + Let Psort (Γ : context) (t : term) (u : sort) := True. Let Pprod Γ t (na : aname) A B := forall P, @@ -374,12 +374,14 @@ Section OnFreeVars. Let Pcheck (Γ : context) (t : term) (T : term) := True. + Let Pj (Γ : context) (j : judgment) := True. + Let PΓ (Γ : context) := True. Let PΓ_rel (Γ Γ' : context) := True. - Theorem bidirectional_on_free_vars : env_prop_bd Σ Pcheck Pinfer Psort Pprod Pind PΓ PΓ_rel. + Theorem bidirectional_on_free_vars : env_prop_bd Σ Pcheck Pinfer Psort Pprod Pind Pj PΓ PΓ_rel. Proof using wfΣ. apply bidir_ind_env. @@ -387,6 +389,8 @@ Section OnFreeVars. - constructor. + - constructor. + - intros. red. intros P HΓ Hn. eapply alli_Alli, Alli_nth_error in HΓ ; tea. @@ -403,7 +407,7 @@ Section OnFreeVars. apply /andP ; split ; tea. - intros until A. - move => _ _ _ _ _ Ht ? ? /= /andP [] ? /andP [] ? ?. + move => _ _ _ Ht ? ? /= /andP [] ? /andP [] ? ?. repeat (apply /andP ; split ; tea). apply Ht ; tea. rewrite on_ctx_free_vars_snoc. @@ -459,14 +463,14 @@ Section OnFreeVars. eapply declared_projection_closed_type ; tea. - intros until decl. - move => ? ndec ? ? ? ? ? /= Hmfix. + move => ? ndec ? ? ? ? ? ? ? /= Hmfix. eapply forallb_nth_error in Hmfix. erewrite ndec in Hmfix. cbn in Hmfix. by move: Hmfix => /andP []. - intros until decl. - move => ? ndec ? ? ? ? ? /= Hmfix. + move => ? ndec ? ? ? ? ? ? ? /= Hmfix. eapply forallb_nth_error in Hmfix. erewrite ndec in Hmfix. cbn in Hmfix. @@ -499,7 +503,7 @@ Section OnFreeVars. on_free_vars P T. Proof using wfΣ. intros. - edestruct bidirectional_on_free_vars as (_&_&_&p&_). + edestruct bidirectional_on_free_vars as (_&_&_&_&p&_). eapply p ; eauto. Qed. @@ -525,7 +529,7 @@ Proof. assert (wf_local Σ Γ) by (eapply typing_wf_local ; tea). apply typing_infering in ty as [T' []] ; tea. exists T' ; split. - - edestruct bidirectional_on_free_vars as (_&_&_&?&_). + - edestruct bidirectional_on_free_vars as (_&_&_&_&?&_). all: eauto. - by apply infering_typing. Qed. @@ -573,8 +577,15 @@ Let Pcheck Γ t T := on_free_vars P T -> Σ ;;; Δ |- rename f t ◃ rename f T. -Let PΓ := - All_local_env (lift_sorting (fun _ => Pcheck) (fun _ => Psort) Σ). +Let Pj Γ j := + forall P Δ f, + urenaming P Γ Δ f -> + on_ctx_free_vars P Γ -> + lift_sorting + (fun t T => on_free_vars P t -> on_free_vars P T -> Σ ;;; Δ |- rename f t ◃ rename f T) + (fun T u => on_free_vars P T -> Σ ;;; Δ |- rename f T ▹□ u) j. + +Let PΓ := All_local_env Pj. Let PΓ_rel Γ Γ' := forall P Δ f, @@ -588,8 +599,8 @@ Lemma rename_telescope P f Γ Δ tel tys: on_ctx_free_vars P Γ -> forallb (on_free_vars P) tys -> on_free_vars_ctx P (List.rev tel) -> - PCUICTyping.ctx_inst (fun _ => Pcheck) Σ Γ tys tel -> - PCUICTyping.ctx_inst checking Σ Δ (map (rename f) tys) (rename_telescope f tel). + PCUICTyping.ctx_inst Pcheck Γ tys tel -> + PCUICTyping.ctx_inst (checking Σ) Δ (map (rename f) tys) (rename_telescope f tel). Proof using Type. intros ur hΓ htys htel ins. induction ins in Δ, ur, hΓ, htys, htel |- *. @@ -618,51 +629,30 @@ Proof using Type. by rewrite /= andb_true_r. Qed. -Theorem bidirectional_renaming : env_prop_bd Σ Pcheck Pinfer Psort Pprod Pind PΓ PΓ_rel. +Theorem bidirectional_renaming : env_prop_bd Σ Pcheck Pinfer Psort Pprod Pind Pj PΓ PΓ_rel. Proof using wfΣ. apply bidir_ind_env. - - intros Γ wfΓ hΓ. red. - induction hΓ. - + constructor. - + constructor ; tea. - eexists ; eauto. - + constructor ; tea. - eexists ; eauto. - - - intros Γ Γ' wfΓ' allΓ'. red. move => P Δ f hf hΓ hΓ'. - induction allΓ'. - + constructor. - + rewrite rename_context_snoc. - rewrite on_free_vars_ctx_snoc in hΓ'. - move: hΓ' => /andP [] ? ?. - constructor ; eauto. - 1: by eapply IHallΓ' ; eauto. - eexists. - eapply Hs. - * eapply urenaming_context ; tea. - * rewrite on_ctx_free_vars_concat. - apply /andP ; split ; tea. - by rewrite on_free_vars_ctx_on_ctx_free_vars. - * eassumption. - + rewrite rename_context_snoc. - rewrite on_free_vars_ctx_snoc in hΓ'. - move: hΓ' => /andP [] ? /andP /= [] ? ?. - constructor ; eauto. - * by eapply IHallΓ' ; eauto. - * eexists. - eapply Hs. - 1: eapply urenaming_context ; tea. - 2: eauto. - rewrite on_ctx_free_vars_concat. - apply /andP ; split ; tea. - by rewrite on_free_vars_ctx_on_ctx_free_vars. - * eapply Hc. - 1: eapply urenaming_context ; tea. - all: auto. - rewrite on_ctx_free_vars_concat. - apply /andP ; split ; tea. - by rewrite on_free_vars_ctx_on_ctx_free_vars. + - intros Γ j Hj P Δ f hf hΓ. + apply lift_sorting_impl with (1 := Hj) => //; intros ?? [Ht IHt] **. + all: eapply IHt; eauto. + + - intros Γ _ _ hΓ. red. + apply All_local_env_impl with (1 := hΓ); clear Γ hΓ. + intros Γ j Hj P Δ f hf hΓ'. + apply lift_sorting_impl with (1 := Hj) => //; intros ?? IHt **. + all: eapply IHt; eauto. + + - intros Γ Γ' _ _ HΓ P Δ f hf hΓ hΓ'. + apply All_local_env_fold. change (fold_context_k _ ?Γ) with (rename_context f Γ). + induction HΓ in hΓ' |- * using All_local_rel_ind1. 1: constructor. + rewrite on_free_vars_ctx_snoc in hΓ'. move/andP: hΓ' => [hΓ'] /andP [onb ont]. + apply All_local_env_snoc; auto. + eapply urenaming_context in hf ; tea. + apply lift_sorting_f_it_impl with X => //. + 1: destruct decl_body => //; cbn in onb |- *. + all: intros IHt; eapply IHt; eauto. + all: rewrite on_ctx_free_vars_concat hΓ /= on_free_vars_ctx_on_ctx_free_vars //. - intros Γ n decl isdecl P Δ f hf hΓ ht. eapply hf in isdecl as h => //. @@ -674,7 +664,9 @@ Proof using wfΣ. by constructor. - intros. red. move => P Δ f hf hΓ /= /andP [] ? ?. - econstructor ; eauto. + econstructor. + { specialize (X0 _ _ _ hf hΓ). + eapply lift_sorting_f_it_impl with X0 => //= IH. now apply IH. } eapply X2 ; tea. 1: by apply urenaming_vass. rewrite on_ctx_free_vars_snoc /=. @@ -682,6 +674,8 @@ Proof using wfΣ. - intros. red. move => P Δ f hf hΓ /= /andP [] ? ?. econstructor ; eauto. + { specialize (X0 _ _ _ hf hΓ). + eapply lift_sorting_f_it_impl with X0 => //= IH. now apply IH. } eapply X2 ; tea. 1: by apply urenaming_vass. rewrite on_ctx_free_vars_snoc /=. @@ -689,7 +683,9 @@ Proof using wfΣ. - intros. red. move => P Δ f hf hΓ /= /andP [] ? /andP [] ? ?. econstructor ; eauto. - eapply X4 ; tea. + { specialize (X0 _ _ _ hf hΓ). + eapply lift_sorting_f_it_impl with X0 => //= IH. all: now apply IH. } + eapply X2 ; tea. 1: by apply urenaming_vdef. rewrite on_ctx_free_vars_snoc. repeat (apply /andP ; split ; tea). @@ -824,23 +820,22 @@ Proof using wfΣ. econstructor. + eapply fix_guard_rename ; tea. + by rewrite nth_error_map H0 /=. - + eapply All_mix in X ; tea. - eapply All_map, All_impl ; tea. - move => ? [] /andP [] ? ? [] ? [] ? p. - rewrite -map_dtype. - eexists. - eapply p ; tea. - + eapply All_mix in X0 ; tea. - eapply All_map, All_impl ; tea. - move => ? [] /andP [] ? ? [] ? p. - rewrite -map_dbody -map_dtype rename_fix_context rename_context_length -(fix_context_length mfix) -rename_shiftn. - eapply p ; tea. - * rewrite -(fix_context_length mfix). - eapply urenaming_context ; tea. - * by apply on_ctx_free_vars_fix_context. - * rewrite -(Nat.add_0_r (#|mfix|)) fix_context_length. - apply on_free_vars_lift_impl. - by rewrite shiftnP0. + + apply All_mix with (1 := ht) in X0. + apply All_map, All_impl with (1 := X0). + move => d [] /andP [] hT hb p. + specialize (p _ _ _ hf hΓ). + eapply lift_sorting_it_impl_gen with p => // IHp. + eapply IHp ; tea. + + apply All_mix with (1 := ht) in X2. + apply All_map, All_impl with (1 := X2). + move => d [] /andP [] hT hb p. + rewrite /on_def_body !rename_fix_context !rename_context_length -!(fix_context_length mfix) -!rename_shiftn in hb |- *. + eapply urenaming_context in hf ; tea. + eapply on_ctx_free_vars_fix_context in hΓ; tea. rewrite -(fix_context_length mfix) in hΓ. + specialize (p _ _ _ hf hΓ). + rewrite <-(shiftnP0 P) in hT. eapply on_free_vars_lift_impl in hT. rewrite Nat.add_0_r in hT. + apply lift_sorting_f_it_impl with p => //= IHp. + all: eapply IHp ; tea. + by apply rename_wf_fixpoint. - intros. red. move => P Δ f hf hΓ /= /forallb_All ht. @@ -848,23 +843,22 @@ Proof using wfΣ. econstructor. + eapply cofix_guard_rename ; tea. + by rewrite nth_error_map H0 /=. - + eapply All_mix in X ; tea. - eapply All_map, All_impl ; tea. - move => ? [] /andP [] ? ? [] ? [] ? p. - rewrite -map_dtype. - eexists. - eapply p ; tea. - + eapply All_mix in X0 ; tea. - eapply All_map, All_impl ; tea. - move => ? [] /andP [] ? ? [] ? p. - rewrite -map_dbody -map_dtype rename_fix_context rename_context_length -(fix_context_length mfix) -rename_shiftn. - eapply p ; tea. - * rewrite -(fix_context_length mfix). - eapply urenaming_context ; tea. - * by apply on_ctx_free_vars_fix_context. - * rewrite -(Nat.add_0_r (#|mfix|)) fix_context_length. - apply on_free_vars_lift_impl. - by rewrite shiftnP0. + + apply All_mix with (1 := ht) in X0. + apply All_map, All_impl with (1 := X0). + move => d [] /andP [] hT hb p. + specialize (p _ _ _ hf hΓ). + eapply lift_sorting_it_impl_gen with p => // IHp. + eapply IHp ; tea. + + apply All_mix with (1 := ht) in X2. + apply All_map, All_impl with (1 := X2). + move => d [] /andP [] hT hb p. + rewrite /on_def_body !rename_fix_context !rename_context_length -!(fix_context_length mfix) -!rename_shiftn in hb |- *. + eapply urenaming_context in hf ; tea. + eapply on_ctx_free_vars_fix_context in hΓ; tea. rewrite -(fix_context_length mfix) in hΓ. + specialize (p _ _ _ hf hΓ). + rewrite <-(shiftnP0 P) in hT. eapply on_free_vars_lift_impl in hT. rewrite Nat.add_0_r in hT. + apply lift_sorting_f_it_impl with p => //= IHp. + all: eapply IHp ; tea. + by apply rename_wf_cofixpoint. - intros. red. intros P Δ f hf ht. @@ -1016,10 +1010,10 @@ Lemma strengthening `{cf: checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} Γ Proof. intros Hty. assert (wf_local Σ Γ) by - move: Hty => /typing_wf_local /wf_local_app_inv [] /wf_local_app_inv [] //. + move: Hty => /typing_wf_local /All_local_env_app_inv [] /All_local_env_app_inv [] //. assert (wfΓ'' : wf_local Σ (Γ ,,, Γ'')). - { apply wf_local_app => //. + { apply All_local_env_app => //. erewrite <- (lift_unlift_context Γ''). eapply bidirectional_to_pcuic ; tea. rewrite rename_context_lift_context. @@ -1064,7 +1058,7 @@ Qed. Lemma strengthening_type `{cf: checker_flags} {Σ : global_env_ext} {wfΣ : PCUICTyping.wf Σ} Γ Γ' Γ'' t s : Σ ;;; Γ ,,, Γ' ,,, lift_context #|Γ'| 0 Γ'' |- lift #|Γ'| #|Γ''| t : tSort s -> - ∑ s', (Σ ;;; Γ ,,, Γ'' |- t : tSort s') * (compare_universe Cumul Σ s' s). + ∑ s', Σ ;;; Γ ,,, Γ'' |- t : tSort s' × leq_sort Σ s' s. Proof. intros H; pose proof (H' := H); eapply strengthening in H. destruct H as [T' [? HT Hcumul HT']]. pose proof (Hcumul' := Hcumul). @@ -1075,7 +1069,7 @@ Proof. pose proof (wf_local_closed_context (typing_wf_local H')). eapply into_ws_cumul_pb in Hcumul; eauto. eapply PCUICConversion.ws_cumul_pb_Sort_r_inv in Hcumul as [s' [Hred Hcumul]]. - eapply closed_red_red in Hred. pose proof (HT'':=HT'). eapply PCUICValidity.validity in HT' as [? ?]. + eapply closed_red_red in Hred. pose proof (HT'':=HT'). eapply PCUICValidity.validity in HT' as (_ & ? & ? & _). exists s'; split; eauto. assert (PCUICReduction.red Σ (Γ,,, Γ'') (unlift #|Γ'| #|Γ''| T') (tSort s')). { assert (tSort s' = unlift #|Γ'| #|Γ''| (tSort s')) by reflexivity. rewrite H2; clear H2. @@ -1090,4 +1084,4 @@ Proof. assert (liftSort : unlift #|Γ'| #|Γ''| (tSort s') = tSort s') by reflexivity. rewrite <- liftSort; clear liftSort. eapply PCUICConversion.cumulAlgo_cumulSpec. eapply PCUICConversion.ws_cumul_pb_red_r_inv; eauto. eapply ws_cumul_pb_refl; eauto. eapply PCUICInversion.typing_closed_ctx; eauto. -Qed. \ No newline at end of file +Qed. diff --git a/pcuic/theories/Bidirectional/BDToPCUIC.v b/pcuic/theories/Bidirectional/BDToPCUIC.v index 81c98547b..2fd27e6d3 100644 --- a/pcuic/theories/Bidirectional/BDToPCUIC.v +++ b/pcuic/theories/Bidirectional/BDToPCUIC.v @@ -38,35 +38,30 @@ Proof. change (d :: Γ') with (Γ' ,, d). destruct d as [na' [bd|] ty]; rewrite !app_context_cons; intro HH. - rewrite subst_context_snoc0. simpl. - inversion HH; subst; cbn in *. destruct X0 as [s X0]. + inversion HH; subst; cbn in *. change (Γ,, vdef na b t ,,, Γ') with (Γ ,,, [vdef na b t] ,,, Γ') in *. assert (subslet Σ (Δ ,,, Γ) [b] [vdef na b t]). { pose proof (cons_let_def Σ (Δ ,,, Γ) [] [] na b t) as XX. rewrite !subst_empty in XX. apply XX. constructor. - apply All_local_env_app_l in X. inversion X; subst; cbn in * ; assumption. + apply All_local_env_app_l in X. inversion X; subst; cbn in * ; now eapply unlift_TermTyp. } constructor; cbn; auto. 1: apply IHΓ' ; exact X. - 1: exists s. - 1: change (tSort s) with (subst [b] #|Γ'| (tSort s)). - all: rewrite app_context_assoc. - all: eapply substitution; tea. - 1: rewrite !app_context_assoc in X0 ; assumption. - rewrite !app_context_assoc in X1 ; assumption. + apply lift_typing_f_impl with (1 := X0) => // ?? Hs. + rewrite !app_context_assoc in Hs |- *. + eapply substitution; tea. - rewrite subst_context_snoc0. simpl. - inversion HH; subst; cbn in *. destruct X0 as [s X0]. + inversion HH; subst; cbn in *. change (Γ,, vdef na b t ,,, Γ') with (Γ ,,, [vdef na b t] ,,, Γ') in *. assert (subslet Σ (Δ ,,, Γ) [b] [vdef na b t]). { pose proof (cons_let_def Σ (Δ ,,, Γ) [] [] na b t) as XX. rewrite !subst_empty in XX. apply XX. constructor. - apply All_local_env_app_l in X. inversion X; subst; cbn in *; assumption. } + apply All_local_env_app_l in X. inversion X; subst; cbn in *; now eapply unlift_TermTyp. } constructor; cbn; auto. 1: apply IHΓ' ; exact X. - exists s. - change (tSort s) with (subst [b] #|Γ'| (tSort s)). - rewrite app_context_assoc. - all: eapply substitution; tea. - rewrite !app_context_assoc in X0. eassumption. + apply lift_typing_f_impl with (1 := X0) => // ?? Hs. + rewrite !app_context_assoc in Hs |- *. + eapply substitution; tea. Qed. Lemma wf_rel_weak `{checker_flags} (Σ : global_env_ext) (wfΣ : wf Σ) Γ (wfΓ : wf_local Σ Γ) Γ' : @@ -77,25 +72,19 @@ Proof. - intros wfl. inversion_clear wfl. constructor. + apply IHΓ'. assumption. - + apply infer_typing_sort_impl with id X0; intros Hs. + + apply lift_typing_impl with (1 := X0) => // ?? Hs. apply weaken_ctx ; eauto. - + apply weaken_ctx ; auto. - intros wfl. inversion_clear wfl. constructor. + apply IHΓ'. assumption. - + apply infer_typing_sort_impl with id X0; intros Hs. + + apply lift_typing_impl with (1 := X0) => // ?? Hs. apply weaken_ctx ; eauto. Qed. Lemma wf_local_local_rel `{checker_flags} Σ Γ Γ' : wf_local Σ (Γ ,,, Γ') -> wf_local_rel Σ Γ Γ'. Proof. - induction Γ'. - 1: constructor. - cbn. - intros wfΓ. - inversion_clear wfΓ. - all: constructor ; auto. - all: apply IHΓ' ; eassumption. + intro wfΓ. + apply All_local_env_app_inv in wfΓ as [_ wfr] => //. Qed. Section BDToPCUICTyping. @@ -130,43 +119,44 @@ Section BDToPCUICTyping. Let PΓ_rel Γ Γ' := wf_local Σ Γ -> wf_local_rel Σ Γ Γ'. + Let Pj Γ j := + wf_local Σ Γ -> lift_typing typing Σ Γ j. + (** Preliminary lemmas to go from a bidirectional judgement to the corresponding undirected one *) Lemma bd_wf_local Γ (all: wf_local_bd Σ Γ) : - All_local_env_over_sorting checking infering_sort - (fun Σ Γ _ t T _ => Pcheck Γ t T) - (fun Σ Γ _ t s _ => Psort Γ t s) - Σ Γ all -> + All_local_env_over_sorting (checking Σ) (infering_sort Σ) + (fun Γ _ t T _ => Pcheck Γ t T) + (fun Γ _ t s _ => Psort Γ t s) + Γ all -> wf_local Σ Γ. Proof using Type. intros allo ; induction allo. - all: constructor. - 1,3: assumption. - all: do 2 red. - 3: apply Hc; auto. - all: apply infer_sort_impl with id tu; now intros Ht. + all: constructor; tas. + all: apply lift_sorting_it_impl with tu => //= Ht; eauto. + apply Hc; cbn; auto. + now eapply has_sort_isType, Hs. Qed. Lemma bd_wf_local_rel Γ (wfΓ : wf_local Σ Γ) Γ' (all: wf_local_bd_rel Σ Γ Γ') : All_local_env_over_sorting - (fun Σ Δ => checking Σ (Γ,,,Δ)) - (fun Σ Δ => infering_sort Σ (Γ,,,Δ)) - (fun Σ Δ _ t T _ => Pcheck (Γ,,,Δ) t T) - (fun Σ Δ _ t s _ => Psort (Γ,,,Δ) t s) - Σ Γ' all -> + (fun Δ => checking Σ (Γ,,,Δ)) + (fun Δ => infering_sort Σ (Γ,,,Δ)) + (fun Δ _ t T _ => Pcheck (Γ,,,Δ) t T) + (fun Δ _ t s _ => Psort (Γ,,,Δ) t s) + Γ' all -> wf_local_rel Σ Γ Γ'. Proof using Type. intros allo ; induction allo. - all: constructor. - 1,3: assumption. - all: red. - 3: apply Hc; [by apply wf_local_app|]. - all: apply infer_sort_impl with id tu; intros Ht. - all: now apply Hs, wf_local_app. + all: constructor; tas. + all: apply All_local_env_app in IHallo; tas. + all: apply lift_sorting_it_impl with tu => //= Ht; eauto. + apply Hc; cbn; auto. + now eapply has_sort_isType, Hs. Qed. Lemma ctx_inst_impl Γ (wfΓ : wf_local Σ Γ) (Δ : context) (wfΔ : wf_local_rel Σ Γ (List.rev Δ)) : - forall args, PCUICTyping.ctx_inst (fun _ => Pcheck) Σ Γ args Δ -> ctx_inst Σ Γ args Δ. + forall args, PCUICTyping.ctx_inst Pcheck Γ args Δ -> ctx_inst Σ Γ args Δ. Proof using wfΣ. revert wfΔ. induction Δ using ctx_length_ind. @@ -176,7 +166,7 @@ Section BDToPCUICTyping. subst. assert (isType Σ Γ t). { - eapply wf_local_rel_app_inv in wfΔ as [wfd _]. + eapply All_local_rel_app_inv in wfΔ as [wfd _]. inversion_clear wfd. eassumption. } @@ -185,21 +175,21 @@ Section BDToPCUICTyping. 1: by rewrite subst_telescope_length ; reflexivity. rewrite -(rev_involutive Γ0) -subst_context_telescope. cbn in wfΔ. - apply wf_local_rel_app_inv in wfΔ as []. + apply All_local_rel_app_inv in wfΔ as []. apply wf_local_local_rel. eapply substitution_wf_local ; eauto. 1: repeat constructor. 1: rewrite subst_empty ; eauto. - apply wf_local_app. + apply All_local_env_app. 1: constructor ; eauto. eassumption. - subst d. subst. assert (isType Σ Γ t). { - eapply wf_local_rel_app_inv in wfΔ as [wfd _]. + eapply All_local_rel_app_inv in wfΔ as [wfd _]. inversion_clear wfd. - eassumption. + apply lift_sorting_it_impl_gen with X2 => //. } constructor ; auto. apply X ; auto. @@ -212,10 +202,19 @@ Section BDToPCUICTyping. Qed. (** The big theorem, proven by mutual induction using the custom induction principle *) - Theorem bidirectional_to_pcuic : env_prop_bd Σ Pcheck Pinfer Psort Pprod Pind PΓ PΓ_rel. + Theorem bidirectional_to_pcuic : env_prop_bd Σ Pcheck Pinfer Psort Pprod Pind Pj PΓ PΓ_rel. Proof using wfΣ. apply bidir_ind_env. + - intros. intro. + apply lift_sorting_it_impl with X. + 2: intros [_ IH]; now apply IH. + destruct j_term => //=. + intros [_ IH]; apply IH; tas. + apply lift_sorting_forget_univ in X. + apply lift_sorting_it_impl_gen with X => //. + intros [_ IH']; now apply IH'. + - intros. eapply bd_wf_local. eassumption. - intros. intro. eapply bd_wf_local_rel ; eassumption. @@ -227,30 +226,20 @@ Section BDToPCUICTyping. - red ; intros ; econstructor ; eauto. apply X2. constructor. 1: by auto. - eexists. eauto. + eapply lift_sorting_forget_univ. now eapply X0. - red ; intros ; econstructor ; eauto. - apply X2. - constructor. 1: by auto. - eexists. eauto. - red ; intros ; econstructor ; eauto. - + apply X2 ; auto. - eexists. eauto. - - + apply X4. - constructor ; auto. - * eexists. eauto. - * apply X2 ; auto. eexists. eauto. - red ; intros. eapply type_App' ; auto. apply X2 ; auto. specialize (X0 X3). - apply validity in X0 as [? X0]. - apply inversion_Prod in X0 as (? & ? & ? & _). + apply validity in X0 as (_ & s & X0 & _). + apply inversion_Prod in X0 as (s1 & s2 & X0 & _). 2: done. - eexists. eassumption. + eapply lift_sorting_forget_univ; eassumption. - red ; intros ; econstructor ; eauto. @@ -265,9 +254,9 @@ Section BDToPCUICTyping. rewrite rev_involutive. apply wf_rel_weak ; auto. move: (H) => [decl ?]. - unshelve epose proof (decl' := declared_minductive_to_gen decl); eauto. + pose proof (decl' := declared_minductive_to_gen (wfΣ := wfΣ) decl). eapply wf_local_subst_instance_decl ; eauto. - eapply wf_local_app_inv. + eapply All_local_env_app_inv. now eapply on_minductive_wf_params_indices. } @@ -305,9 +294,9 @@ Section BDToPCUICTyping. } assert (isType Σ Γ (mkApps (tInd ci (puinst p)) - (pparams p ++ skipn (ci_npar ci) args))) as [? tyapp]. + (pparams p ++ skipn (ci_npar ci) args))). { - eexists. + eapply has_sort_isType. eapply type_mkApps_arity. 1: econstructor ; eauto. erewrite PCUICGlobalMaps.ind_arity_eq. @@ -324,9 +313,8 @@ Section BDToPCUICTyping. assert (wf_local Σ (Γ,,,case_predicate_context ci mdecl idecl p)). { eapply wf_case_predicate_context ; tea. - eexists ; tea. } - + pose proof X12 as (_ & ? & X12' & _). econstructor ; eauto. 2-3: split ; eauto. 1: now eapply type_Cumul ; eauto ; apply cumulAlgo_cumulSpec in cum. @@ -334,8 +322,6 @@ Section BDToPCUICTyping. eapply All2i_impl. 1:{ apply All2i_prod ; [eassumption|idtac]. eapply wf_case_branches_types' ; eauto. - econstructor. - eauto. } intros * ((?&?&?&?&Hbody)&[]). split ; tea. @@ -344,83 +330,60 @@ Section BDToPCUICTyping. fold ptm in brctxty. repeat split ; auto. apply Hbody ; auto. - eexists. + eapply has_sort_isType. eassumption. - red ; intros ; econstructor ; eauto. - red ; intros ; econstructor ; eauto. - + clear H H0 H1 X0. - induction X. - all: constructor ; auto. - destruct p as (? & ? & ?). - eexists. - apply p. - auto. + + clear H H0 H1 X. + apply All_impl with (1 := X0) => d Hd. + specialize (Hd X3). + apply lift_sorting_it_impl with Hd => //. + have Htypes : All (fun d => isType Σ Γ (dtype d)) mfix. - { clear H H0 H1 X0. - induction X. - all: constructor ; auto. - destruct p as (? & ? & ?). - eexists. - apply p. - auto. + { apply All_impl with (1 := X0) => d Hd. + specialize (Hd X3). + apply lift_sorting_it_impl with Hd => //. } have wfΓ' : wf_local Σ (Γ,,,fix_context mfix) by apply All_mfix_wf. - remember (fix_context mfix) as Γ'. - clear H H0 H1 HeqΓ'. - induction X0. - all: constructor ; auto. - 2:{ inversion_clear X. apply IHX0 ; auto. inversion_clear Htypes. auto. } - destruct p. - apply p ; auto. - inversion_clear Htypes as [| ? ? [u]]. - exists u. - change (tSort u) with (lift0 #|Γ'| (tSort u)). - apply weakening. - all: auto. + remember (fix_context mfix) as Γ' eqn:e. + clear H H0 H1 e. + apply All_mix with (1 := Htypes) in X2. + apply All_impl with (1 := X2) => d [] isTy Hd. + specialize (Hd wfΓ'). + apply lift_sorting_it_impl with Hd => //=. - red ; intros ; econstructor ; eauto. - + clear H H0 H1 X0. - induction X. - all: constructor ; auto. - destruct p as (? & ? & ?). - eexists. - apply p. - auto. + + clear H H0 H1 X. + apply All_impl with (1 := X0) => d Hd. + specialize (Hd X3). + apply lift_sorting_it_impl with Hd => //. + have Htypes : All (fun d => isType Σ Γ (dtype d)) mfix. - { clear H H0 H1 X0. - induction X. - all: constructor ; auto. - destruct p as (? & ? & ?). - eexists. - apply p. - auto. + { apply All_impl with (1 := X0) => d Hd. + specialize (Hd X3). + apply lift_sorting_it_impl with Hd => //. } - have wfΓ' : wf_local Σ (Γ,,,fix_context mfix) by apply All_mfix_wf ; auto. - remember (fix_context mfix) as Γ'. - clear H H0 H1 HeqΓ'. - induction X0. - all: constructor ; auto. - 2:{ inversion_clear X. apply IHX0 ; auto. inversion_clear Htypes. auto. } - destruct p. - apply p ; auto. - inversion_clear Htypes as [| ? ? [u]]. - exists u. - change (tSort u) with (lift0 #|Γ'| (tSort u)). - apply weakening. - all: auto. + have wfΓ' : wf_local Σ (Γ,,,fix_context mfix) by apply All_mfix_wf. + + remember (fix_context mfix) as Γ' eqn:e. + clear H H0 H1 e. + apply All_mix with (1 := Htypes) in X2. + apply All_impl with (1 := X2) => d [] isTy Hd. + specialize (Hd wfΓ'). + apply lift_sorting_it_impl with Hd => //=. - red; intros. econstructor; eauto. - depelim X0; constructor; eauto. - eapply hty; eauto. eexists. econstructor; eauto. - eapply hdef; eauto. eexists; eauto. eapply hty; eauto. eexists; econstructor; eauto. - solve_all. eapply X6; eauto. eexists. eapply hty; eauto. eexists; econstructor; eauto. + depelim X0; try solve [ constructor; eauto ]. + unfold Pcheck in hty, hdef. + do 2 forward hty; tas. 1: eapply has_sort_isType; now econstructor. + do 2 forward hdef; tas. 1: now eapply has_sort_isType. + constructor; eauto. + solve_all. eapply X6; eauto. now eapply has_sort_isType. - red ; intros. now eapply type_reduction. @@ -432,7 +395,7 @@ Section BDToPCUICTyping. now eapply type_reduction. - red ; intros. - destruct X3. + destruct X3 as (_ & ? & ? & _). econstructor ; eauto. eapply (cumulAlgo_cumulSpec _ (pb := Cumul)), into_ws_cumul_pb ; tea. + fvs. @@ -472,7 +435,7 @@ Theorem infering_sort_isType `{checker_flags} (Σ : global_env_ext) Γ t u (wfΣ wf_local Σ Γ -> Σ ;;; Γ |- t ▹□ u -> isType Σ Γ t. Proof. intros wfΓ Ht. - exists u. + repeat (eexists; tea). now apply infering_sort_typing. Qed. @@ -483,6 +446,24 @@ Proof. now eapply infering_sort_isType. Qed. +Theorem isTypebd_isType `{checker_flags} (Σ : global_env_ext) Γ j (wfΣ : wf Σ) : + wf_local Σ Γ -> lift_sorting (checking Σ Γ) (infering_sort Σ Γ) j -> isType Σ Γ (j_typ j). +Proof. + intros wfΓ (_ & s & Hs & _). + now eapply infering_sort_isType. +Qed. + +Theorem lift_sorting_lift_typing `{cf : checker_flags} (Σ : global_env_ext) Γ j (wfΣ : wf Σ) : + wf_local Σ Γ -> lift_sorting (checking Σ Γ) (infering_sort Σ Γ) j -> lift_typing typing Σ Γ j. +Proof. + intros wfΓ Hj. + apply lift_sorting_it_impl with Hj. + 2: intros H; now apply infering_sort_typing, H. + destruct j_term => //=. + intros H; apply checking_typing; tas. + now apply isTypebd_isType in Hj. +Qed. + Theorem infering_prod_typing `{checker_flags} (Σ : global_env_ext) Γ t na A B (wfΣ : wf Σ) : wf_local Σ Γ -> Σ ;;; Γ |- t ▹Π (na,A,B) -> Σ ;;; Γ |- t : tProd na A B. Proof. @@ -515,8 +496,8 @@ Qed. Theorem ctx_inst_bd_typing `{checker_flags} (Σ : global_env_ext) Γ l Δ (wfΣ : wf Σ) : wf_local Σ (Γ,,,Δ) -> - PCUICTyping.ctx_inst checking Σ Γ l (List.rev Δ) -> - PCUICTyping.ctx_inst typing Σ Γ l (List.rev Δ). + PCUICTyping.ctx_inst (checking Σ) Γ l (List.rev Δ) -> + PCUICTyping.ctx_inst (typing Σ) Γ l (List.rev Δ). Proof. intros wfΓ inl. rewrite -(List.rev_involutive Δ) in wfΓ. @@ -527,7 +508,7 @@ Proof. - assert (Σ ;;; Γ |- i : t). { rewrite /= app_context_assoc in wfΓ. - eapply wf_local_app_inv in wfΓ as [wfΓ _]. + eapply All_local_env_app_inv in wfΓ as [wfΓ _]. inversion wfΓ ; subst. now apply checking_typing. } @@ -548,6 +529,7 @@ Proof. constructor. 1: constructor. rewrite !subst_empty. - eapply wf_local_app_inv in wfΓ as [wfΓ _]. - now inversion wfΓ ; subst. + eapply All_local_env_app_inv in wfΓ as [wfΓ _]. + inversion wfΓ ; subst. + now eapply unlift_TermTyp. Qed. diff --git a/pcuic/theories/Bidirectional/BDTyping.v b/pcuic/theories/Bidirectional/BDTyping.v index c36912543..a384ca756 100644 --- a/pcuic/theories/Bidirectional/BDTyping.v +++ b/pcuic/theories/Bidirectional/BDTyping.v @@ -31,22 +31,21 @@ Inductive infering `{checker_flags} (Σ : global_env_ext) (Γ : context) : term Σ ;;; Γ |- tRel n ▹ lift0 (S n) (decl_type decl) | infer_Sort s : - wf_universe Σ s -> - Σ ;;; Γ |- tSort s ▹ tSort (Universe.super s) + wf_sort Σ s -> + Σ ;;; Γ |- tSort s ▹ tSort (Sort.super s) | infer_Prod na A B s1 s2 : - Σ ;;; Γ |- A ▹□ s1 -> + lift_sorting (checking Σ Γ) (infering_sort Σ Γ) (j_vass_s na A s1) -> Σ ;;; Γ ,, vass na A |- B ▹□ s2 -> - Σ ;;; Γ |- tProd na A B ▹ tSort (Universe.sort_of_product s1 s2) + Σ ;;; Γ |- tProd na A B ▹ tSort (Sort.sort_of_product s1 s2) -| infer_Lambda na A t s B : - Σ ;;; Γ |- A ▹□ s -> +| infer_Lambda na A t B : + lift_sorting (checking Σ Γ) (infering_sort Σ Γ) (j_vass na A) -> Σ ;;; Γ ,, vass na A |- t ▹ B -> Σ ;;; Γ |- tLambda na A t ▹ tProd na A B -| infer_LetIn na b B t s A : - Σ ;;; Γ |- B ▹□ s -> - Σ ;;; Γ |- b ◃ B -> +| infer_LetIn na b B t A : + lift_sorting (checking Σ Γ) (infering_sort Σ Γ) (j_vdef na b B) -> Σ ;;; Γ ,, vdef na b B |- t ▹ A -> Σ ;;; Γ |- tLetIn na b B t ▹ tLetIn na b B A @@ -83,11 +82,10 @@ Inductive infering `{checker_flags} (Σ : global_env_ext) (Γ : context) : term consistent_instance_ext Σ (ind_universes mdecl) (puinst p) -> wf_local_bd_rel Σ Γ predctx -> is_allowed_elimination Σ (ind_kelim idecl) ps -> - ctx_inst checking Σ Γ (pparams p) + ctx_inst (checking Σ) Γ (pparams p) (List.rev mdecl.(ind_params)@[p.(puinst)]) -> isCoFinite mdecl.(ind_finite) = false -> - R_global_instance Σ (eq_universe Σ) (leq_universe Σ) - (IndRef ci) #|args| u (puinst p) -> + cmp_ind_universes Σ ci #|args| u (puinst p) -> All2 (convAlgo Σ Γ) (firstn (ci_npar ci) args) (pparams p) -> (* case_branch_typing *) wf_branches idecl brs -> @@ -108,16 +106,16 @@ Inductive infering `{checker_flags} (Σ : global_env_ext) (Γ : context) : term | infer_Fix mfix n decl : fix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> - All (fun d => {s & Σ ;;; Γ |- d.(dtype) ▹□ s}) mfix -> - All (fun d => Σ ;;; Γ ,,, fix_context mfix |- d.(dbody) ◃ lift0 #|fix_context mfix| d.(dtype)) mfix -> + All (on_def_type (lift_sorting1 (checking Σ) (infering_sort Σ)) Γ) mfix -> + All (on_def_body (lift_sorting1 (checking Σ) (infering_sort Σ)) (fix_context mfix) Γ) mfix -> wf_fixpoint Σ mfix -> Σ ;;; Γ |- tFix mfix n ▹ dtype decl | infer_CoFix mfix n decl : cofix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> - All (fun d => {s & Σ ;;; Γ |- d.(dtype) ▹□ s}) mfix -> - All (fun d => Σ ;;; Γ ,,, fix_context mfix |- d.(dbody) ◃ lift0 #|fix_context mfix| d.(dtype)) mfix -> + All (on_def_type (lift_sorting1 (checking Σ) (infering_sort Σ)) Γ) mfix -> + All (on_def_body (lift_sorting1 (checking Σ) (infering_sort Σ)) (fix_context mfix) Γ) mfix -> wf_cofixpoint Σ mfix -> Σ ;;; Γ |- tCoFix mfix n ▹ dtype decl @@ -128,7 +126,7 @@ Inductive infering `{checker_flags} (Σ : global_env_ext) (Γ : context) : term primitive_typing_hyps checking Σ Γ p -> Σ ;;; Γ |- tPrim p ▹ prim_type p prim_ty -with infering_sort `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> Universe.t -> Type := +with infering_sort `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> sort -> Type := | infer_sort_Sort t T u: Σ ;;; Γ |- t ▹ T -> red Σ Γ T (tSort u) -> @@ -157,8 +155,8 @@ and " Σ ;;; Γ |- t ▹□ u " := (@infering_sort _ Σ Γ t u) : type_scope and " Σ ;;; Γ |- t ▹Π ( na , A , B ) " := (@infering_prod _ Σ Γ t na A B) : type_scope and " Σ ;;; Γ |- t ▹{ ind } ( u , args ) " := (@infering_indu _ Σ Γ ind t u args) : type_scope and " Σ ;;; Γ |- t ◃ T " := (@checking _ Σ Γ t T) : type_scope -and "'wf_local_bd' Σ Γ" := (All_local_env (lift_sorting checking infering_sort Σ) Γ) -and "'wf_local_bd_rel' Σ Γ Γ'" := (All_local_rel (lift_sorting checking infering_sort Σ) Γ Γ'). +and "'wf_local_bd' Σ Γ" := (All_local_env (lift_sorting1 (checking Σ) (infering_sort Σ)) Γ) +and "'wf_local_bd_rel' Σ Γ Γ'" := (All_local_rel (lift_sorting1 (checking Σ) (infering_sort Σ)) Γ Γ'). Definition tybranches {cf} Σ Γ ci mdecl idecl p ptm n ctors brs := @@ -178,7 +176,7 @@ Definition branches_size {cf} {Σ Γ ci mdecl idecl p ptm brs} (all2i_size _ (fun i cdecl br p => (Nat.max - (All_local_rel_sorting_size checking_size infering_size _ _ _ p.2.1) + (All_local_rel_sorting_size (checking_size Σ) (infering_size _) _ _ p.2.1) (checking_size _ _ _ _ p.2.2))) a). @@ -195,11 +193,12 @@ Proof. | H : infering_prod _ _ _ _ _ _ |- _ => apply infering_prod_size in H | H : infering_indu _ _ _ _ _ _ |- _ => apply infering_indu_size in H | H : checking _ _ _ _ |- _ => apply checking_size in H - | H : wf_local_bd _ _ |- _ => apply (All_local_env_sorting_size _ _ (checking_size _) (infering_sort_size _) _ _) in H - | H : wf_local_bd_rel _ _ _ |- _ => apply (All_local_rel_sorting_size (checking_size _) (infering_sort_size _) _ _) in H + | H : wf_local_bd _ _ |- _ => apply (All_local_env_sorting_size _ _ (checking_size _ _) (infering_sort_size _ _) _) in H + | H : wf_local_bd_rel _ _ _ |- _ => apply (All_local_rel_sorting_size (checking_size _ _) (infering_sort_size _ _) _) in H | H : primitive_typing_hyps _ _ _ _ |- _ => apply (primitive_typing_hyps_size _ (checking_size _)) in H end ; match goal with + | H : lift_sorting _ _ _, H' : size |- _ => exact (S (Nat.max H' (lift_sorting_size (checking_size _ _ _) (infering_sort_size _ _ _) _ H))) | H : All2i _ _ _ _ |- _ => idtac | H : All _ _ |- _ => idtac | H1 : size, H2 : size, H3 : size |- _ => exact (S (Nat.max H1 (Nat.max H2 H3))) @@ -208,10 +207,12 @@ Proof. | |- _ => exact 1 end. - exact (S (Nat.max a0 (Nat.max i (Nat.max i0 (Nat.max (ctx_inst_size _ (checking_size _) c1) (branches_size (checking_size _) (infering_sort_size _) a2)))))). - - exact (S (Nat.max (all_size _ (fun d p => infering_sort_size _ _ _ _ _ p.π2) a) - (all_size _ (fun x p => checking_size _ _ _ _ _ p) a0))). - - exact (S (Nat.max (all_size _ (fun d p => infering_sort_size _ _ _ _ _ p.π2) a) - (all_size _ (fun x => checking_size _ _ _ _ _) a0))). + - exact (S (Nat.max + (all_size _ (fun x p => on_def_type_sorting_size (infering_sort_size _ Σ) _ _ p) a) + (all_size (on_def_body _ _ _) (fun x p => on_def_body_sorting_size (checking_size _ _) (infering_sort_size _ Σ) _ _ _ p) a0))). + - exact (S (Nat.max + (all_size _ (fun x p => on_def_type_sorting_size (infering_sort_size _ Σ) _ _ p) a) + (all_size (on_def_body _ _ _) (fun x p => on_def_body_sorting_size (checking_size _ _) (infering_sort_size _ Σ) _ _ _ p) a0))). Defined. Fixpoint infering_size_pos `{checker_flags} {Σ Γ t T} (d : Σ ;;; Γ |- t ▹ T) @@ -233,22 +234,24 @@ Arguments lexprod [A B]. Section BidirectionalInduction. - #[local] Notation wfl_size := (All_local_env_sorting_size (@checking_size _) (@infering_sort_size _) _ _). - #[local] Notation wfl_size_rel := (All_local_rel_sorting_size (@checking_size _) (@infering_sort_size _) _ _ _). + #[local] Notation wfl_size := (All_local_env_sorting_size (@checking_size _ _) (@infering_sort_size _ _) _). + #[local] Notation wfl_size_rel := (All_local_rel_sorting_size (@checking_size _ _) (@infering_sort_size _ _) _ _). Context `{cf : checker_flags}. Context (Σ : global_env_ext). Context (wfΣ : wf Σ). Context (Pcheck : context -> term -> term -> Type). Context (Pinfer : context -> term -> term -> Type). - Context (Psort : context -> term -> Universe.t -> Type). + Context (Psort : context -> term -> sort -> Type). Context (Pprod : context -> term -> aname -> term -> term -> Type). Context (Pind : context -> inductive -> term -> Instance.t -> list term -> Type). + Context (Pj : context -> judgment -> Type). Context (PΓ : context -> Type). Context (PΓ_rel : context -> context -> Type). (** This is what we wish to prove mutually given the previous predicates *) Definition env_prop_bd := + (forall Γ j, lift_sorting1 (checking Σ) (infering_sort Σ) Γ j -> Pj Γ j) × (forall Γ , wf_local_bd Σ Γ -> PΓ Γ) × (forall Γ Γ', wf_local_bd_rel Σ Γ Γ' -> PΓ_rel Γ Γ') × (forall Γ t T, Σ ;;; Γ |- t ◃ T -> Pcheck Γ t T) × @@ -262,6 +265,7 @@ Section BidirectionalInduction. *) Inductive typing_sum : Type := + | judgment_cons : forall (Γ : context) j, lift_sorting1 (checking Σ) (infering_sort Σ) Γ j -> typing_sum | context_cons : forall (Γ : context) (wfΓ : wf_local_bd Σ Γ), typing_sum | context_rel_cons : forall (Γ Γ' : context) (wfΓ' : wf_local_bd_rel Σ Γ Γ'), typing_sum | check_cons : forall (Γ : context) T t, Σ ;;; Γ |- t ◃ T -> typing_sum @@ -272,6 +276,7 @@ Section BidirectionalInduction. Definition typing_sum_size (d : typing_sum) := match d with + | judgment_cons _ _ d => (lift_sorting_size (@checking_size _ _ _) (@infering_sort_size _ _ _) _ d) | context_cons Γ wfΓ => wfl_size wfΓ | context_rel_cons _ _ wfΓ' => wfl_size_rel wfΓ' | check_cons _ _ _ d => (checking_size d) @@ -283,6 +288,7 @@ Section BidirectionalInduction. Definition Ptyping_sum (d : typing_sum) := match d with + | judgment_cons Γ j _ => Pj Γ j | context_cons Γ _ => PΓ Γ | context_rel_cons Γ Γ' _ => PΓ_rel Γ Γ' | check_cons Γ T t _ => Pcheck Γ t T @@ -293,6 +299,8 @@ Section BidirectionalInduction. end. Ltac applyIH := match goal with + | IH : forall d', _ -> Ptyping_sum d' |- Pj ?Γ ?j => + unshelve eapply (IH (judgment_cons Γ j _)) | IH : forall d', _ -> Ptyping_sum d' |- PΓ ?Γ => unshelve eapply (IH (context_cons Γ _)) | IH : forall d', _ -> Ptyping_sum d' |- PΓ_rel ?Γ ?Γ' => @@ -313,52 +321,58 @@ Section BidirectionalInduction. | |- dlexprod _ _ _ _ => constructor 2 ; cbn ; try lia | _ => assumption - | _ => cbn ; lia + | wfΓ' : wf_local_bd_rel _ _ _ |- _ => fold (wfl_size_rel wfΓ'); cbn -[Nat.max] ; lia + | _ => cbn -[Nat.max] ; lia | _ => idtac end. Lemma bidir_ind_env : - let Pdecl_check := fun Σ Γ wfΓ t T tyT => Pcheck Γ t T in - let Pdecl_sort := fun Σ Γ wfΓ t u tyT => Psort Γ t u in - let Pdecl_check_rel Γ := fun _ Δ _ t T _ => Pcheck (Γ,,,Δ) t T in - let Pdecl_sort_rel Γ := fun _ Δ _ t u _ => Psort (Γ,,,Δ) t u in + let Pdecl_check := fun Γ wfΓ t T tyT => Pcheck Γ t T in + let Pdecl_sort := fun Γ wfΓ t u tyT => Psort Γ t u in + let Pdecl_check_rel Γ := fun Δ _ t T _ => Pcheck (Γ,,,Δ) t T in + let Pdecl_sort_rel Γ := fun Δ _ t u _ => Psort (Γ,,,Δ) t u in + + (forall (Γ : context) j, + lift_sorting1 (fun Γ t T => checking Σ Γ t T × Pcheck Γ t T) (fun Γ t u => infering_sort Σ Γ t u × Psort Γ t u) Γ j -> Pj Γ j) -> (forall (Γ : context) (wfΓ : wf_local_bd Σ Γ), - All_local_env_over_sorting checking infering_sort Pdecl_check Pdecl_sort Σ Γ wfΓ -> PΓ Γ) -> + All_local_env_over_sorting (checking Σ) (infering_sort Σ) Pdecl_check Pdecl_sort Γ wfΓ -> + All_local_env (lift_sorting1 Pcheck Psort) Γ -> + PΓ Γ) -> (forall (Γ Γ' : context) (wfΓ' : wf_local_bd_rel Σ Γ Γ'), - All_local_env_over_sorting (fun Σ Δ => checking Σ (Γ,,,Δ)) (fun Σ Δ => infering_sort Σ (Γ,,,Δ)) (Pdecl_check_rel Γ) (Pdecl_sort_rel Γ) Σ Γ' wfΓ' -> PΓ_rel Γ Γ') -> + All_local_env_over_sorting (fun Δ => checking Σ (Γ,,,Δ)) (fun Δ => infering_sort Σ (Γ,,,Δ)) (Pdecl_check_rel Γ) (Pdecl_sort_rel Γ) Γ' wfΓ' -> + All_local_rel (lift_sorting1 Pcheck Psort) Γ Γ' -> + PΓ_rel Γ Γ') -> (forall (Γ : context) (n : nat) decl, nth_error Γ n = Some decl -> Pinfer Γ (tRel n) (lift0 (S n) (decl_type decl))) -> - (forall (Γ : context) (s : Universe.t), - wf_universe Σ s-> - Pinfer Γ (tSort s) (tSort (Universe.super s))) -> - - (forall (Γ : context) (n : aname) (t b : term) (s1 s2 : Universe.t), - Σ ;;; Γ |- t ▹□ s1 -> - Psort Γ t s1 -> - Σ ;;; Γ,, vass n t |- b ▹□ s2 -> - Psort (Γ,, vass n t) b s2 -> - Pinfer Γ (tProd n t b) (tSort (Universe.sort_of_product s1 s2))) -> - - (forall (Γ : context) (n : aname) (t b : term) (s : Universe.t) (bty : term), - Σ ;;; Γ |- t ▹□ s -> - Psort Γ t s -> - Σ ;;; Γ,, vass n t |- b ▹ bty -> - Pinfer (Γ,, vass n t) b bty -> - Pinfer Γ (tLambda n t b) (tProd n t bty)) -> - - (forall (Γ : context) (n : aname) (b B t : term) (s : Universe.t) (A : term), - Σ ;;; Γ |- B ▹□ s -> - Psort Γ B s -> - Σ ;;; Γ |- b ◃ B -> - Pcheck Γ b B -> - Σ ;;; Γ,, vdef n b B |- t ▹ A -> - Pinfer (Γ,, vdef n b B) t A -> - Pinfer Γ (tLetIn n b B t) (tLetIn n b B A)) -> + (forall (Γ : context) (s : sort), + wf_sort Σ s-> + Pinfer Γ (tSort s) (tSort (Sort.super s))) -> + + (forall (Γ : context) (na : aname) (t b : term) (s1 s2 : sort), + lift_sorting (checking Σ Γ) (infering_sort Σ Γ) (j_vass_s na t s1) -> + Pj Γ (j_vass_s na t s1) -> + Σ ;;; Γ,, vass na t |- b ▹□ s2 -> + Psort (Γ,, vass na t) b s2 -> + Pinfer Γ (tProd na t b) (tSort (Sort.sort_of_product s1 s2))) -> + + (forall (Γ : context) (na : aname) (t b bty : term), + lift_sorting (checking Σ Γ) (infering_sort Σ Γ) (j_vass na t) -> + Pj Γ (j_vass na t) -> + Σ ;;; Γ,, vass na t |- b ▹ bty -> + Pinfer (Γ,, vass na t) b bty -> + Pinfer Γ (tLambda na t b) (tProd na t bty)) -> + + (forall (Γ : context) (na : aname) (b B t A : term), + lift_sorting (checking Σ Γ) (infering_sort Σ Γ) (j_vdef na b B) -> + Pj Γ (j_vdef na b B) -> + Σ ;;; Γ,, vdef na b B |- t ▹ A -> + Pinfer (Γ,, vdef na b B) t A -> + Pinfer Γ (tLetIn na b B t) (tLetIn na b B A)) -> (forall (Γ : context) (t : term) na A B u, Σ ;;; Γ |- t ▹Π (na, A, B) -> @@ -401,13 +415,12 @@ Section BidirectionalInduction. wf_local_bd_rel Σ Γ predctx -> PΓ_rel Γ predctx -> is_allowed_elimination Σ (ind_kelim idecl) ps -> - ctx_inst checking Σ Γ (pparams p) + ctx_inst (checking Σ) Γ (pparams p) (List.rev mdecl.(ind_params)@[p.(puinst)]) -> - ctx_inst (fun _ => Pcheck) Σ Γ p.(pparams) + ctx_inst Pcheck Γ p.(pparams) (List.rev (subst_instance p.(puinst) mdecl.(ind_params))) -> isCoFinite mdecl.(ind_finite) = false -> - R_global_instance Σ (eq_universe Σ) (leq_universe Σ) - (IndRef ci) #|args| u (puinst p) -> + cmp_ind_universes Σ ci #|args| u (puinst p) -> All2 (convAlgo Σ Γ) (firstn (ci_npar ci) args) (pparams p) -> (* case_branch_typing *) wf_branches idecl brs -> @@ -432,19 +445,20 @@ Section BidirectionalInduction. (forall (Γ : context) (mfix : mfixpoint term) (n : nat) decl, fix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> - All (fun d => {s & (Σ ;;; Γ |- d.(dtype) ▹□ s) × Psort Γ d.(dtype) s}) mfix -> - All (fun d => (Σ ;;; Γ ,,, fix_context mfix |- d.(dbody) ◃ lift0 #|fix_context mfix| d.(dtype)) × - Pcheck (Γ ,,, fix_context mfix) d.(dbody) (lift0 #|fix_context mfix| d.(dtype))) mfix -> + All (on_def_type (lift_sorting1 (checking Σ) (infering_sort Σ)) Γ) mfix -> + All (on_def_type Pj Γ) mfix -> + All (on_def_body (lift_sorting1 (checking Σ) (infering_sort Σ)) (fix_context mfix) Γ) mfix -> + All (on_def_body Pj (fix_context mfix) Γ) mfix -> wf_fixpoint Σ mfix -> Pinfer Γ (tFix mfix n) (dtype decl)) -> (forall (Γ : context) (mfix : mfixpoint term) (n : nat) decl, cofix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> - All (fun d => {s & (Σ ;;; Γ |- d.(dtype) ▹□ s) × Psort Γ d.(dtype) s}) mfix -> - All (fun d => (Σ ;;; Γ ,,, fix_context mfix |- d.(dbody) ◃ lift0 #|fix_context mfix| d.(dtype)) × - Pcheck (Γ ,,, fix_context mfix) d.(dbody) (lift0 #|fix_context mfix| d.(dtype))) mfix -> - wf_cofixpoint Σ mfix -> + All (on_def_type (lift_sorting1 (checking Σ) (infering_sort Σ)) Γ) mfix -> + All (on_def_type Pj Γ) mfix -> + All (on_def_body (lift_sorting1 (checking Σ) (infering_sort Σ)) (fix_context mfix) Γ) mfix -> + All (on_def_body Pj (fix_context mfix) Γ) mfix -> wf_cofixpoint Σ mfix -> Pinfer Γ (tCoFix mfix n) (dtype decl)) -> (forall (Γ : context) p prim_ty cdecl, @@ -455,7 +469,7 @@ Section BidirectionalInduction. primitive_typing_hyps (fun _ => Pcheck) Σ Γ p -> Pinfer Γ (tPrim p) (prim_type p prim_ty)) -> - (forall (Γ : context) (t T : term) (u : Universe.t), + (forall (Γ : context) (t T : term) (u : sort), Σ ;;; Γ |- t ▹ T -> Pinfer Γ t T -> red Σ Γ T (tSort u) -> @@ -482,13 +496,14 @@ Section BidirectionalInduction. env_prop_bd. Proof using Type. - intros Pdecl_check Pdecl_sort Pdecl_check_rel Pdecl_sort_rel HΓ HΓRel HRel HSort HProd HLambda HLetIn HApp HConst HInd HConstruct HCase + intros Pdecl_check Pdecl_sort Pdecl_check_rel Pdecl_sort_rel Hj HΓ HΓRel HRel HSort HProd HLambda HLetIn HApp HConst HInd HConstruct HCase HProj HFix HCoFix HPrim HiSort HiProd HiInd HCheck ; unfold env_prop_bd. pose (@Fix_F typing_sum (precompose lt typing_sum_size) Ptyping_sum) as p. forward p. 2:{ enough (HP : forall x : typing_sum, Ptyping_sum x). - repeat split ; intros. + + exact (HP (judgment_cons Γ j X)). + exact (HP (context_cons Γ X)). + exact (HP (context_rel_cons Γ Γ' X)). + exact (HP (check_cons Γ T t X)). @@ -502,54 +517,67 @@ Section BidirectionalInduction. clear p. intros d IH. destruct d ; simpl. - 4: destruct i. + 5: destruct i. + + - eapply Hj; tea. + eapply lift_sorting_size_impl with (csize := @checking_size _ _ _) (ssize := @infering_sort_size _ _ _) (tu := l). + all: intros ?? Hty Hlt; split; tas. + all: applyIH. - eapply HΓ. + 2: { clear -IH wfΣ. induction wfΓ; cbn in *; simpl in *; constructor. + 1,3: apply IHwfΓ; intros; apply IH; lia. + all: assert (0 < wfl_size wfΓ) as Hpos by apply All_local_env_size_pos. + all: eapply lift_sorting_size_impl with (csize := @checking_size _ _ _) (ssize := @infering_sort_size _ _ _) (tu := t0). + all: intros ?? Hty Hlt; unfold lift_sorting_size in Hlt; simpl in Hlt. + all: applyIH. } dependent induction wfΓ. + constructor. - + destruct t0 as (u & d). - constructor. + + constructor. * apply IHwfΓ. - intros ; apply IH. - cbn in H |- *. lia. - * simpl. red. - applyIH. - cbn. + intros ; apply IH. + cbn in H |- *. lia. + * destruct t0 as (? & s & Hty & ?); cbn in *. assert (0 < wfl_size wfΓ) by apply All_local_env_size_pos. - lia. - + destruct t0 as [u h]. + red. + applyIH. + + destruct t0 as (Htm & s & Hty & _e); cbn in *. + assert (0 < wfl_size wfΓ) by apply All_local_env_size_pos. constructor. * apply IHwfΓ. intros ; apply IH. - cbn in H |- *. pose proof (infering_sort_size_pos h). lia. - * red. applyIH. pose proof (infering_sort_size_pos h). cbn in H |- *. lia. - * red. applyIH. pose proof (checking_size_pos t1). cbn in H |- *. lia. + cbn in H |- *. lia. + * red. cbn. applyIH. + * red. cbn. applyIH. - eapply HΓRel. - dependent induction wfΓ'. + 2: { clear -IH wfΣ. induction wfΓ'; cbn in *; simpl in *; constructor; fold (wf_local_bd_rel Σ Γ Γ0) in wfΓ'. + 1,3: apply IHwfΓ'; intros; apply IH; fold (wfl_size_rel wfΓ'); lia. + all: assert (0 < wfl_size_rel wfΓ') as Hpos by apply All_local_env_size_pos. + all: eapply lift_sorting_size_impl with (csize := @checking_size _ _ _) (ssize := @infering_sort_size _ _ _) (tu := t0). + all: intros ?? Hty Hlt; unfold lift_sorting_size in Hlt; simpl in Hlt. + all: applyIH. } + dependent induction wfΓ'; try fold (wf_local_bd_rel Σ Γ Γ0) in wfΓ'. + + constructor. + constructor. - + destruct t0 as (u & d). - constructor. * apply IHwfΓ'. intros ; apply IH. cbn in H |- *. - unfold wfl_size_rel in H. + fold (wfl_size_rel wfΓ'). lia. - * cbn. red. - applyIH. - cbn. + * destruct t0 as (? & s & Hty & ?); cbn in *. red. assert (0 < wfl_size_rel wfΓ') by apply All_local_env_size_pos. - unfold wfl_size_rel in H. - lia. - + destruct t0 as [u h]. + applyIH. + + destruct t0 as (Htm & s & Hty & _e); cbn in *. + assert (0 < wfl_size_rel wfΓ') by apply All_local_env_size_pos. constructor. * apply IHwfΓ'. intros ; apply IH. cbn in H |- *. fold (wfl_size_rel wfΓ'). - pose proof (infering_sort_size_pos h). lia. - * red. applyIH. pose proof (infering_sort_size_pos h). cbn in H |- *. lia. - * red. applyIH. pose proof (checking_size_pos t1). cbn in H |- *. lia. + lia. + * red. cbn. applyIH. + * red. cbn. applyIH. - destruct c. @@ -634,55 +662,48 @@ Section BidirectionalInduction. dependent induction a. 1: by constructor. constructor ; cbn. - * destruct p ; eexists ; split. - 1: eassumption. - applyIH. + * unfold on_def_type. applyIH. * apply (IHa s). intros. apply IH. - cbn. lia. + cbn [all_size]. lia. - + remember (all_size _ _ a) as s. + + remember (all_size _ _ a) as size. clear -IH. induction a0. 1: by constructor. constructor. - * intuition. - applyIH. + * unfold on_def_body. applyIH. * apply IHa0. intros. apply IH. - cbn. lia. + cbn in H |- *. lia. - unshelve eapply HCoFix ; eauto. - all: simpl in IH. + all: cbn [typing_sum_size infering_size] in IH. all: remember (fix_context mfix) as mfixcontext. - + remember (all_size _ _ a0) as s. clear -IH. dependent induction a. 1: by constructor. constructor ; cbn. - * destruct p ; eexists ; split. - 1: eassumption. - applyIH. + * unfold on_def_type. applyIH. * apply (IHa s). intros. apply IH. - cbn. lia. + cbn [all_size]. lia. - + remember (all_size _ _ a) as s. + + remember (all_size _ _ a) as size. clear -IH. - induction a0 as [| ? ? ?]. + induction a0. 1: by constructor. constructor. - * intuition. - applyIH. + * unfold on_def_body. applyIH. * apply IHa0. intros. apply IH. - cbn. lia. + cbn in H |- *. lia. - unshelve eapply HPrim; eauto. simpl in IH. diff --git a/pcuic/theories/Bidirectional/BDUnique.v b/pcuic/theories/Bidirectional/BDUnique.v index 1c489c92f..13bb91a57 100644 --- a/pcuic/theories/Bidirectional/BDUnique.v +++ b/pcuic/theories/Bidirectional/BDUnique.v @@ -46,16 +46,19 @@ Let Pind Γ ind t u args := Let Pcheck (Γ : context) (t T : term) := True. +Let Pj (Γ : context) (j : judgment) := lift_sorting (Pcheck Γ) (Psort Γ) j. + Let PΓ (Γ : context) := True. Let PΓ_rel (Γ Γ' : context) := True. -Theorem bidirectional_unique : env_prop_bd Σ Pcheck Pinfer Psort Pprod Pind PΓ PΓ_rel. +Theorem bidirectional_unique : env_prop_bd Σ Pcheck Pinfer Psort Pprod Pind Pj PΓ PΓ_rel. Proof using wfΣ. apply bidir_ind_env. all: intros ; red ; auto. + { apply lift_sorting_impl with (1 := X) => ?? [] //. } 1-9,11-13: intros ? T' ty_T' ; inversion_clear ty_T'. 14-17: intros. @@ -69,29 +72,28 @@ Proof using wfΣ. - eexists ; split. all: eapply closed_red_refl ; fvs. - - apply H in X2 => //. - apply H0 in X3. - 2:{ constructor ; auto. now eapply infering_sort_isType. } + - apply unlift_TypUniv in X0, X3. + apply X0 in X3 => //. + apply H in X4. + 2:{ constructor ; auto. now eapply lift_sorting_forget_univ, lift_sorting_lift_typing. } subst. eexists ; split. all: eapply closed_red_refl ; fvs. - - apply X1 in X4 as [bty' []]. - 2:{ constructor ; auto. now eapply infering_sort_isType. } - exists (tProd n t bty') ; split. + - apply X2 in X5 as [bty' []]. + 2:{ constructor ; auto. now apply lift_sorting_lift_typing. } + exists (tProd na t bty') ; split. all: now eapply closed_red_prod_codom. - - apply X2 in X6 as [A' []]. - 2:{ constructor ; auto. 2: eapply checking_typing ; tea. all: now eapply infering_sort_isType. } - exists (tLetIn n b B A'). - assert (Σ ;;; Γ |- b : B) - by (eapply checking_typing ; tea ; now eapply infering_sort_isType). + - apply lift_sorting_lift_typing in X, X4; tas. + apply lift_typing_is_open_term in X4 as [X4_1 X4_2]; tas; cbn in *. + apply X2 in X5 as [A' []]. + 2: now constructor. + apply wf_local_closed_context in X3. + exists (tLetIn na b B A'). split. all: eapply closed_red_letin ; tea. - all: apply closed_red_refl. - all: try now apply wf_local_closed_context. - 1,3: now eapply subject_is_open_term. - all: now eapply type_is_open_term. + all: now apply closed_red_refl. - unshelve epose proof (X0 _ _ _ _ X3) as (A''&B''&[]) ; tea. subst. @@ -109,8 +111,8 @@ Proof using wfΣ. eapply checking_typing ; tea. now eapply isType_tProd, validity, infering_prod_typing. - - unshelve epose proof (declared_constant_to_gen isdecl); eauto. - unshelve epose proof (declared_constant_to_gen H); eauto. + - pose proof (declared_constant_to_gen (wfΣ := wfΣ) isdecl). + pose proof (declared_constant_to_gen (wfΣ := wfΣ) H). replace decl0 with decl by (eapply declared_constant_inj ; eassumption). eexists ; split. all: eapply closed_red_refl. @@ -118,8 +120,8 @@ Proof using wfΣ. all: rewrite on_free_vars_subst_instance. all: now eapply closed_on_free_vars, declared_constant_closed_type. - - unshelve epose proof (declared_inductive_to_gen isdecl); eauto. - unshelve epose proof (declared_inductive_to_gen H); eauto. + - pose proof (declared_inductive_to_gen (wfΣ := wfΣ) isdecl). + pose proof (declared_inductive_to_gen (wfΣ := wfΣ) H). replace idecl0 with idecl by (eapply declared_inductive_inj ; eassumption). eexists ; split. all: eapply closed_red_refl. @@ -127,8 +129,8 @@ Proof using wfΣ. all: rewrite on_free_vars_subst_instance. all: now eapply closed_on_free_vars, declared_inductive_closed_type. - - unshelve epose proof (declared_constructor_to_gen isdecl); eauto. - unshelve epose proof (declared_constructor_to_gen H); eauto. + - pose proof (declared_constructor_to_gen (wfΣ := wfΣ) isdecl). + pose proof (declared_constructor_to_gen (wfΣ := wfΣ) H). replace cdecl0 with cdecl by (eapply declared_constructor_inj ; eassumption). replace mdecl0 with mdecl by (eapply declared_constructor_inj ; eassumption). eexists ; split. @@ -136,8 +138,8 @@ Proof using wfΣ. 1,3: fvs. all: now eapply closed_on_free_vars, declared_constructor_closed_type. - - unshelve epose proof (declared_projection_to_gen H1); eauto. - unshelve eapply declared_projection_to_gen in H; eauto. + - pose proof (declared_projection_to_gen (wfΣ := wfΣ) H1). + apply (declared_projection_to_gen (wfΣ := wfΣ)) in H. eapply declared_projection_inj in H as (?&?&?&?); tea. subst. move: (X2) => tyc'. @@ -202,8 +204,8 @@ Proof using wfΣ. - intros ? T' ty_T'. inversion ty_T' ; subst. - unshelve eapply declared_inductive_to_gen in H13; eauto. - unshelve eapply declared_inductive_to_gen in H; eauto. + unshelve eapply declared_inductive_to_gen in H13; cycle -2; eauto. + unshelve eapply declared_inductive_to_gen in H; cycle -2; eauto. move: (H) => /declared_inductive_inj /(_ H13) [? ?]. subst. assert (op' : is_open_term Γ (mkApps ptm0 (skipn (ci_npar ci) args0 ++ [c]))). @@ -237,11 +239,13 @@ Proof using wfΣ. * now eapply type_is_open_term, infering_typing. - depelim X; depelim X0. - * depelim X2. rewrite e in H; noconf H. eexists. split; eapply closed_red_refl; fvs. - * depelim X2. rewrite e in H; noconf H. eexists. split; eapply closed_red_refl; fvs. - * depelim X2. rewrite e in H; noconf H. eexists. split; eapply closed_red_refl; fvs. + * inversion X2 ; subst. rewrite H3 in H; noconf H. eexists. split; eapply closed_red_refl; fvs. + * inversion X2 ; subst. rewrite H3 in H; noconf H. eexists. split; eapply closed_red_refl; fvs. + * inversion X2 ; subst. rewrite H3 in H; noconf H. eexists. split; eapply closed_red_refl; fvs. all:simp prim_type; cbn. cbn in hty. - all:eapply type_is_open_term, checking_typing; tea; eexists; eapply checking_typing; tea; eexists; econstructor; tea. + all:eapply type_is_open_term, checking_typing; tea. + all:eapply has_sort_isType; eapply checking_typing; tea. + all:eapply has_sort_isType; econstructor; tea. - inversion X3 ; subst. eapply X0 in X4 as [T'' []]; subst ; tea. diff --git a/pcuic/theories/Conversion/PCUICClosedConv.v b/pcuic/theories/Conversion/PCUICClosedConv.v index ef0ee3569..f0afdf3ce 100644 --- a/pcuic/theories/Conversion/PCUICClosedConv.v +++ b/pcuic/theories/Conversion/PCUICClosedConv.v @@ -20,7 +20,7 @@ Proof. induction Δ; simpl; auto. destruct a as [na [b|] ty]; intros wfΓ wfctx; constructor; intuition auto. - exists s; auto. + now eapply lift_sorting_forget_univ. Qed. Lemma sorts_local_ctx_All_local_env {cf} P Σ Γ Δ s : @@ -32,64 +32,59 @@ Proof. destruct a as [na [b|] ty]; intros wfΓ wfctx; constructor; intuition eauto. destruct s => //. destruct wfctx; eauto. - destruct s => //. destruct wfctx. exists t; auto. + destruct s => //. destruct wfctx. now eapply lift_sorting_forget_univ. Qed. Lemma type_local_ctx_Pclosed Σ Γ Δ s : - type_local_ctx (lift_typing Pclosed) Σ Γ Δ s -> + type_local_ctx (fun _ => lift_wfb_term1 (fun Γ t => closedn #|Γ| t)) Σ Γ Δ s -> Alli (fun i d => closed_decl (#|Γ| + i) d) 0 (List.rev Δ). Proof. induction Δ; simpl; auto; try constructor. destruct a as [? [] ?]; intuition auto. - apply Alli_app_inv; auto. constructor. simpl. rewrite List.rev_length. 2:constructor. - unfold closed_decl. unfold Pclosed in b0. simpl. - rewrite app_context_length in b0. now rewrite Nat.add_comm. + rewrite app_context_length Nat.add_comm in b. + apply b. - apply Alli_app_inv; auto. constructor. simpl. rewrite List.rev_length. 2:constructor. - unfold closed_decl. unfold Pclosed in b. simpl. - rewrite app_context_length in b. rewrite Nat.add_comm. - now rewrite andb_true_r in b. + rewrite app_context_length Nat.add_comm in b. + apply b. Qed. Lemma sorts_local_ctx_Pclosed Σ Γ Δ s : - sorts_local_ctx (lift_typing Pclosed) Σ Γ Δ s -> + sorts_local_ctx (fun _ => lift_wfb_term1 (fun Γ t => closedn #|Γ| t)) Σ Γ Δ s -> Alli (fun i d => closed_decl (#|Γ| + i) d) 0 (List.rev Δ). Proof. induction Δ in s |- *; simpl; auto; try constructor. destruct a as [? [] ?]; intuition auto. - apply Alli_app_inv; eauto. constructor. simpl. rewrite List.rev_length. 2:constructor. - unfold closed_decl. unfold Pclosed in b0. simpl. - rewrite app_context_length in b0. now rewrite Nat.add_comm. + rewrite app_context_length Nat.add_comm in b. + apply b. - destruct s as [|u us]; auto. destruct X as [X b]. apply Alli_app_inv; eauto. constructor. simpl. rewrite List.rev_length. 2:constructor. - unfold closed_decl. unfold Pclosed in b. simpl. - rewrite app_context_length in b. rewrite Nat.add_comm. - now rewrite andb_true_r in b. + rewrite app_context_length Nat.add_comm in b. + apply b. Qed. -Lemma All_local_env_Pclosed Σ Γ : - All_local_env ( lift_typing Pclosed Σ) Γ -> +Lemma All_local_env_Pclosed Γ : + All_local_env (lift_wfb_term1 (fun Γ t => closedn #|Γ| t)) Γ -> Alli (fun i d => closed_decl i d) 0 (List.rev Γ). Proof. induction Γ; simpl; auto; try constructor. intros all; depelim all; intuition auto. - apply Alli_app_inv; auto. constructor. simpl. rewrite List.rev_length. 2:constructor. - unfold closed_decl. unfold Pclosed in l. simpl. red in l. - destruct l as [s H]. - now rewrite andb_true_r in H. + assumption. - apply Alli_app_inv; auto. constructor. simpl. rewrite List.rev_length. 2:constructor. - now simpl. + assumption. Qed. Lemma weaken_env_prop_closed {cf} : - weaken_env_prop cumulSpec0 (lift_typing typing) (lift_typing (fun (_ : global_env_ext) (Γ : context) (t T : term) => - closedn #|Γ| t && closedn #|Γ| T)). -Proof. repeat red. intros. destruct t; red in X0; eauto. Qed. + weaken_env_prop cumulSpec0 (lift_typing typing) (fun _ => lift_wfb_term1 (fun Γ t => closedn #|Γ| t)). +Proof. intros ?. auto. Qed. Lemma closedn_ctx_alpha {k ctx ctx'} : @@ -103,12 +98,11 @@ Proof. Qed. Lemma closedn_All_local_env (ctx : list context_decl) : - All_local_env - (fun (Γ : context) (b : term) (t : typ_or_sort) => - closedn #|Γ| b && typ_or_sort_default (closedn #|Γ|) t true) ctx -> + All_local_env (lift_wfb_term1 (fun Γ t => closedn #|Γ| t)) ctx -> closedn_ctx 0 ctx. Proof. - induction 1; auto; rewrite closedn_ctx_cons IHX /=; now move/andP: t0 => []. + induction 1 using All_local_env_ind1; auto. + rewrite closedn_ctx_cons IHX //=. Qed. Lemma declared_minductive_closed_inds {cf} {Σ ind mdecl u} {wfΣ : wf Σ} : @@ -116,10 +110,10 @@ Lemma declared_minductive_closed_inds {cf} {Σ ind mdecl u} {wfΣ : wf Σ} : forallb (closedn 0) (inds (inductive_mind ind) u (ind_bodies mdecl)). Proof. intros h. - eapply declared_minductive_to_gen in h. + unshelve eapply declared_minductive_to_gen in h; tea. red in h. eapply lookup_on_global_env in h. 2: eauto. - destruct h as [Σ' [ext wfΣ' decl']]. + destruct h as (Σ' & ext & wfΣ' & decl'). red in decl'. destruct decl' as [h ? ? ?]. rewrite inds_spec. rewrite forallb_rev. unfold mapi. @@ -127,7 +121,6 @@ Proof. induction h in n, m |- *. - reflexivity. - simpl. eauto. - Unshelve. all:eauto. Qed. Lemma closed_cstr_branch_context_gen {cf : checker_flags} {Σ} {wfΣ : wf Σ} {c mdecl cdecl} : @@ -145,25 +138,13 @@ Proof. Qed. Lemma closedn_All_local_closed: - forall (cf : checker_flags) (Σ : global_env_ext) (Γ : context) (ctx : list context_decl) - (wfΓ' : wf_local Σ (Γ ,,, ctx)), - All_local_env_over typing - (fun (Σ0 : global_env_ext) (Γ0 : context) (_ : wf_local Σ0 Γ0) (t T : term) (_ : Σ0;;; Γ0 |- t : T) => - closedn #|Γ0| t && closedn #|Γ0| T) Σ (Γ ,,, ctx) wfΓ' -> - closedn_ctx 0 Γ && closedn_ctx #|Γ| ctx. + forall (cf : checker_flags) (Σ : global_env_ext) (Γ : context) (wfΓ : wf_local Σ Γ), + All_local_env_over (typing Σ) (fun Γ _ t T _ => closedn #|Γ| t && closedn #|Γ| T) Γ wfΓ -> + closed_ctx Γ. Proof. - intros cf Σ Γ ctx wfΓ' al. - remember (Γ ,,, ctx) as Γ'. revert Γ' wfΓ' ctx HeqΓ' al. - induction Γ. simpl. intros. subst. unfold app_context in *. rewrite app_nil_r in wfΓ' al. + intros cf Σ Γ wfΓ al. induction al; try constructor; - rewrite closedn_ctx_cons /=; cbn. - move/andP: Hs => [] /= -> _. now rewrite IHal. - now rewrite IHal /= /test_decl /=. - intros. - unfold app_context in *. subst Γ'. - specialize (IHΓ (ctx ++ a :: Γ) wfΓ' (ctx ++ [a])). - rewrite -app_assoc in IHΓ. specialize (IHΓ eq_refl al). - rewrite closedn_ctx_app /= Nat.add_1_r andb_assoc in IHΓ. - now rewrite closedn_ctx_cons /=. + rewrite closedn_ctx_cons IHal /= /test_decl //; cbn. + now move/andP: Hs => [] /= -> _. Qed. diff --git a/pcuic/theories/Conversion/PCUICInstConv.v b/pcuic/theories/Conversion/PCUICInstConv.v index 32c8f3e5c..7552dc8b5 100644 --- a/pcuic/theories/Conversion/PCUICInstConv.v +++ b/pcuic/theories/Conversion/PCUICInstConv.v @@ -296,12 +296,13 @@ Proof. - unfold inst_context, snoc. rewrite fold_context_k_snoc0. unfold snoc. f_equal. all: auto. unfold map_decl. simpl. unfold vass. f_equal. - destruct t0 as [s ht]. eapply typed_inst. all: eauto. + destruct t0 as (_ & s & ht & _). eapply typed_inst. all: eauto. - unfold inst_context, snoc. rewrite fold_context_k_snoc0. unfold snoc. f_equal. all: auto. + destruct t0 as (hb & s & ht & _). cbn in hb. unfold map_decl. simpl. unfold vdef. f_equal. + f_equal. eapply typed_inst. all: eauto. - + eapply typed_inst in t1 as [? _]. all: eauto. + + eapply typed_inst. all: eauto. Qed. Lemma inst_destArity : @@ -1603,29 +1604,24 @@ Proof using Type. Qed. Lemma wf_local_app_inst (Σ : global_env_ext) {wfΣ : wf Σ} Γ Δ : - All_local_env (lift_typing (fun (Σ : global_env_ext) (Γ' : context) (t T : term) => + All_local_rel (fun Γ j => forall Δ σ, wf_local Σ Δ -> - Σ ;;; Δ ⊢ σ : (Γ ,,, Γ') -> - Σ ;;; Δ |- t.[σ] : T.[σ]) Σ) Δ -> + Σ ;;; Δ ⊢ σ : Γ -> + lift_typing0 (fun t T => Σ ;;; Δ |- t.[σ] : T.[σ]) j) Γ Δ -> forall Δ' σ, Σ ;;; Δ' ⊢ σ : Γ -> wf_local Σ Δ' -> wf_local Σ (Δ' ,,, inst_context σ Δ). Proof using Type. intros. - induction X. + induction X using All_local_rel_ind1. - now simpl. - - rewrite inst_context_snoc /=. constructor; auto. - apply infer_typing_sort_impl with id t0; intros Hs. - eapply (Hs (Δ' ,,, inst_context σ Γ0) (⇑^#|Γ0| σ)) => //. + - rewrite inst_context_snoc /=. + apply All_local_env_snoc; auto. + apply lift_typing_map with (j := j_decl _) => //. + eapply X; tas. eapply well_subst_app; auto. - - rewrite inst_context_snoc /=. constructor; auto. - * apply infer_typing_sort_impl with id t0; intros Hs. - eapply (Hs (Δ' ,,, inst_context σ Γ0) (⇑^#|Γ0| σ)) => //. - eapply well_subst_app; auto. - * simpl. apply t1 => //. - eapply well_subst_app; eauto. Qed. Lemma usubst_up_vass {Γ Δ σ na A} : @@ -1873,14 +1869,16 @@ Proof. eapply red_primArray_type; cbn; eauto. Defined. -Lemma eq_term_upto_univ_inst Σ : - forall Re Rle napp u v σ, - Reflexive Re -> Reflexive Rle -> - eq_term_upto_univ_napp Σ Re Rle napp u v -> - eq_term_upto_univ_napp Σ Re Rle napp u.[σ] v.[σ]. +Lemma eq_term_upto_univ_inst Σ cmp_universe cmp_sort pb napp u v σ : + RelationClasses.Reflexive (cmp_universe Conv) -> + RelationClasses.Reflexive (cmp_universe pb) -> + RelationClasses.Reflexive (cmp_sort Conv) -> + RelationClasses.Reflexive (cmp_sort pb) -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp u v -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp u.[σ] v.[σ]. Proof using Type. - intros Re Rle napp u v σ hRe hRle h. - induction u in v, napp, Re, Rle, hRe, hRle, σ, h |- * using term_forall_list_ind. + intros refl_univ_conv refl_univ_pb refl_sort_conv refl_sort_pb h. + induction u in v, pb, napp, σ, refl_univ_pb, refl_sort_pb, h |- * using term_forall_list_ind. all: dependent destruction h. all: try solve [ simpl ; constructor ; eauto @@ -1894,24 +1892,18 @@ Proof using Type. - simpl. constructor. all: eauto. * rewrite /inst_predicate. destruct X; destruct e as [? [? [ectx ?]]]. - rewrite (All2_fold_length ectx). red. + rewrite (All2_length ectx). red. intuition auto; simpl; solve_all. - * induction X0 in a, brs' |- *. - + inversion a. constructor. - + inversion a. subst. simpl. - destruct X1 as [a0 e0], p0. - constructor; eauto. - split; eauto. - simpl. - rewrite (All2_fold_length a0). - now eapply e1. + * unfold eq_branches, eq_branch in *; solve_all. + rewrite (All2_length a1). + now eapply b0. - simpl. constructor. - apply All2_length in a as e. rewrite <- e. - generalize #|m|. intro k. + apply All2_length in e as eq. rewrite <- eq. + generalize #|m|. intro k. unfold eq_mfixpoint in *. eapply All2_map. simpl. solve_all. - simpl. constructor. - apply All2_length in a as e. rewrite <- e. - generalize #|m|. intro k. + apply All2_length in e as eq. rewrite <- eq. + generalize #|m|. intro k. unfold eq_mfixpoint in *. eapply All2_map. simpl. solve_all. - simpl. constructor. eapply onPrims_map_prop; tea. cbn; intuition eauto. diff --git a/pcuic/theories/Conversion/PCUICNamelessConv.v b/pcuic/theories/Conversion/PCUICNamelessConv.v index a69528123..60a56bc2e 100644 --- a/pcuic/theories/Conversion/PCUICNamelessConv.v +++ b/pcuic/theories/Conversion/PCUICNamelessConv.v @@ -39,7 +39,7 @@ Local Ltac anonify := Local Ltac ih := lazymatch goal with - | ih : forall (v : term) (napp : nat), _ -> _ -> eq_term_upto_univ_napp _ _ _ _ _ _ -> ?u = _ + | ih : forall (v : term) (napp : nat), _ -> _ -> eq_term_upto_univ_napp _ _ _ _ _ _ _ -> ?u = _ |- ?u = ?v => eapply ih ; eassumption end. @@ -58,25 +58,23 @@ Qed. Lemma nameless_eqctx_IH ctx ctx' : forallb (nameless_decl nameless) ctx -> forallb (nameless_decl nameless) ctx' -> - eq_context_gen eq eq ctx ctx' -> + eq_context_upto_names ctx ctx' -> ctx = ctx'. Proof. + intros H H0 X. + apply All2_eq. solve_all. - induction X; auto. - all:destruct p; depelim H; depelim H0; auto; f_equal; subst; auto. - - unfold nameless_decl in i, i0; rtoProp. - f_equal. - eapply banon_eq_binder_annot; eauto. - - unfold nameless_decl in i, i0; rtoProp. - f_equal. - eapply banon_eq_binder_annot; eauto. + destruct b0; unfold nameless_decl in *; cbn in *; rtoProp. + all: f_equal. + all: eapply banon_eq_binder_annot; eauto. Qed. -Lemma eq_context_gen_upto ctx ctx' : - eq_context_gen eq eq ctx ctx' -> - eq_context_upto empty_global_env eq eq ctx ctx'. +Lemma eq_context_upto_names_eq_context_alpha ctx ctx' : + eq_context_upto_names ctx ctx' -> + eq_context_upto empty_global_env (fun _ => eq) (fun _ => eq) Conv ctx ctx'. Proof. - intros a; eapply All2_fold_impl; tea. + move/All2_fold_All2. + intros a; eapply All2_fold_impl; tea; cbn. intros. destruct X; subst; constructor; auto; try reflexivity. Qed. @@ -84,7 +82,7 @@ Lemma nameless_eq_term_spec : forall u v napp, nameless u -> nameless v -> - eq_term_upto_univ_napp empty_global_env eq eq napp u v -> + eq_term_upto_univ_napp empty_global_env (fun _ => eq) (fun _ => eq) Conv napp u v -> u = v. Proof. intros u v napp hu hv e. @@ -106,81 +104,41 @@ Proof. * eapply H0 ; eauto. * eapply IHl ; assumption. - f_equal ; try solve [ ih ]. - eapply eq_univ_make. assumption. + apply cmp_universe_instance_eq. assumption. - f_equal ; try solve [ ih ]. - eapply eq_univ_make. assumption. + apply cmp_universe_instance_eq. assumption. - f_equal ; try solve [ ih ]. - eapply eq_univ_make. assumption. + apply cmp_universe_instance_eq. assumption. - f_equal ; try solve [ ih ]. * destruct e as [eqpar [eqinst [eqctx eqret]]]. destruct X as [? [? ?]]. destruct p, p'; simpl in *. f_equal. + apply All2_eq; solve_all. - + red in eqinst. - apply Forall2_eq. eapply Forall2_map_inv in eqinst. - eapply (Forall2_impl eqinst); solve_all. - now apply Universe.make_inj in H. + + apply cmp_universe_instance_eq. assumption. + simpl in *. eapply nameless_eqctx_IH; eauto. + ih. - * revert brs' H3 H0 a. - induction l ; intros brs' h1 h2 h. - + destruct brs' ; inversion h. reflexivity. - + destruct brs' ; inversion h. subst. - cbn in h1, h2. destruct_andb. - inversion X. subst. simpl in H5. - move/andb_and: H5 => [/andb_and [Hac Hab] Hl]. - apply forallb_All in Hac. - f_equal. - ++ destruct a, b. cbn in *. destruct X1. - depelim h. destruct p0. depelim X0. simpl in *. - destruct p0 as []. - destruct X4. - f_equal; try ih. - { eapply nameless_eqctx_IH; eauto; solve_all. } - ++ eapply IHl ; tas. now depelim X0. + * apply All2_eq. unfold eq_branches, eq_branch in *. solve_all. + destruct x, y; cbn in *; f_equal. + + eapply nameless_eqctx_IH; eauto. + all: solve_all. + + ih. - f_equal ; try solve [ ih ]. - revert mfix' H1 H2 H H0 a. - induction m ; intros m' h1 h2 h3 h4 h. - + destruct m' ; inversion h. reflexivity. - + destruct m' ; inversion h. subst. - inversion X. subst. - cbn in h1, h2, h3, h4. destruct_andb. - f_equal. - * 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. - f_equal. - -- now apply banon_eq_binder_annot. - -- eapply Hty; eassumption. - -- eapply Hbod ; eassumption. - -- eassumption. - * eapply IHm ; assumption. + unfold eq_mfixpoint in *. apply All2_eq. solve_all. + destruct x, y; unfold test_def in *; destruct_andb; cbn in *; f_equal; try solve [ ih ]. + 2: assumption. + now apply banon_eq_binder_annot. - f_equal ; try solve [ ih ]. - revert mfix' H1 H2 H H0 a. - induction m ; intros m' h1 h2 h3 h4 h. - + destruct m' ; inversion h. reflexivity. - + destruct m' ; inversion h. subst. - inversion X. subst. - cbn in h1, h2, h3, h4. destruct_andb. - f_equal. - * 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. - -- now apply banon_eq_binder_annot. - -- eapply Hty; eassumption. - -- eapply Hbod ; eassumption. - -- assumption. - * eapply IHm ; assumption. + unfold eq_mfixpoint in *. apply All2_eq. solve_all. + destruct x, y; unfold test_def in *; destruct_andb; cbn in *; f_equal; try solve [ ih ]. + 2: assumption. + now apply banon_eq_binder_annot. - f_equal. destruct o; auto. f_equal. f_equal. cbn in X, hu, hv. rtoProp. - destruct a, a'; cbn in *. eapply Universe.make_inj in e. f_equal; intuition eauto. - solve_all. induction a0 => //. f_equal; eauto. - eapply r; intuition eauto. + destruct X as (hty & hdef & harr). eapply Universe.make'_inj in e. + destruct a, a'; cbn in *. f_equal; intuition eauto. + apply All2_eq. solve_all. Qed. Lemma banon_list l : forallb (banon ∘ anonymize) l. @@ -200,7 +158,7 @@ Proof. * eapply All_forallb, All_map; assumption. * rewrite forallb_map. eapply All_forallb. unfold ondecl in *. solve_all. - rewrite /nameless_decl /= a0. + rewrite /nameless_decl /= b. destruct (decl_body x); simpl in *; auto. * induction l. + reflexivity. @@ -208,7 +166,7 @@ Proof. repeat (eapply andb_true_intro ; split) ; try assumption. ++ rewrite forallb_map. eapply All_forallb. unfold ondecl in *; solve_all. - rewrite /nameless_decl /= a2. + rewrite /nameless_decl /= b. destruct (decl_body x); simpl in *; auto. ++ eapply IHl. assumption. - simpl ; repeat (eapply andb_true_intro ; split) ; try assumption. @@ -291,85 +249,63 @@ Proof. destruct nth_error => /= //. Qed. -Lemma R_global_instance_nl Σ Re Rle gr napp : - CRelationClasses.subrelation (R_global_instance Σ Re Rle gr napp) - (R_global_instance (nl_global_env Σ) Re Rle gr napp). +Lemma cmp_global_instance_nl Σ cmp_universe pb gr napp : + CRelationClasses.subrelation (cmp_global_instance Σ cmp_universe pb gr napp) + (cmp_global_instance (nl_global_env Σ) cmp_universe pb gr napp). Proof. intros t t'. - unfold R_global_instance, R_global_instance_gen. + unfold cmp_global_instance, cmp_global_instance_gen. now rewrite global_variance_nl_sigma_mon. Qed. -Lemma eq_context_nl_IH Σ Re ctx ctx' : - (forall (napp : nat) (t t' : term) - (Rle : Universe.t -> Universe.t -> Prop), - eq_term_upto_univ_napp Σ Re Rle napp t t' -> - eq_term_upto_univ_napp (nl_global_env Σ) Re Rle napp - (nl t) (nl t')) -> - eq_context_gen eq eq ctx ctx' -> - eq_context_gen eq eq +Lemma eq_context_nl_IH ctx ctx' : + eq_context_upto_names ctx ctx' -> + eq_context_upto_names (map (map_decl_anon nl) ctx) (map (map_decl_anon nl) ctx'). Proof. - intros aux H. - induction H; simpl; constructor; simpl; destruct p; simpl; - intuition (constructor; auto); subst; reflexivity. + intros H. apply All2_map, All2_impl with (1 := H); clear H. + intros ?? []; constructor. + all: assumption. Defined. -Lemma nl_eq_term_upto_univ : - forall Σ Re Rle napp t t', - eq_term_upto_univ_napp Σ Re Rle napp t t' -> - eq_term_upto_univ_napp (nl_global_env Σ) Re Rle napp (nl t) (nl t'). +Lemma nl_eq_term_upto_univ Σ cmp_universe cmp_sort pb napp t t' : + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp t t' -> + eq_term_upto_univ_napp (nl_global_env Σ) cmp_universe cmp_sort pb napp (nl t) (nl t'). Proof. - intros Σ Re Rle napp t t' h. - revert napp t t' Rle h. fix aux 5. - intros napp t t' Rle h. - destruct h. + induction t in napp, pb, t' |- * using term_forall_list_ind; intro e. + all: dependent destruction e. all: simpl. all: try solve [ econstructor ; eauto ]. - econstructor. - induction a. - + constructor. - + simpl. econstructor. all: eauto. + solve_all. - econstructor. all: try solve [ eauto ]. - eapply R_global_instance_nl; eauto. + eapply cmp_global_instance_nl; eauto. - econstructor. all: try solve [ eauto ]. - eapply R_global_instance_nl; eauto. + eapply cmp_global_instance_nl; eauto. - econstructor; eauto. + red. destruct e; intuition auto; simpl. - * induction a0; constructor; auto. + * solve_all. * eapply eq_context_nl_IH; tea. - * apply aux. auto. - + induction a; constructor; auto. + * solve_all. + + unfold eq_branches in *. solve_all. unfold eq_branch in *. intuition auto; simpl. * eapply eq_context_nl_IH; tea. - * destruct x, y; simpl in *. apply aux; auto. - - econstructor; eauto. - induction a; constructor; auto. - intuition auto. - * destruct x, y; simpl in *. apply aux; auto. - * destruct x, y; simpl in *. apply aux; auto. - - econstructor; eauto. - induction a; constructor; auto. - intuition auto. - * destruct x, y; simpl in *. apply aux; auto. - * destruct x, y; simpl in *. apply aux; auto. + * solve_all. + - econstructor; eauto. unfold eq_mfixpoint in *. + solve_all. + - econstructor; eauto. unfold eq_mfixpoint in *. + solve_all. - constructor. - destruct i as [? []], i' as [? []]; cbn in o; depelim o; cbn in *; constructor; eauto. - + now eapply aux. - + now eapply aux. - + cbn. induction a2. - ++ constructor. - ++ cbn; constructor; [now eapply aux|]. eapply IHa2. -Qed. - -Lemma eq_context_nl Σ Re Rle ctx ctx' : - eq_context_gen (eq_term_upto_univ Σ Re Re) - (eq_term_upto_univ Σ Re Rle) ctx ctx' -> - eq_context_gen - (eq_term_upto_univ (nl_global_env Σ) Re Re) - (eq_term_upto_univ (nl_global_env Σ) Re Rle) - (nlctx ctx) (nlctx ctx'). + destruct p as [? []], i' as [? []]; depelim o; try now constructor. + destruct X as (hty & hdef & harr). + constructor; cbn; eauto. + solve_all. +Qed. + +Lemma eq_context_nl Σ cmp_universe cmp_sort pb ctx ctx' : + eq_context_gen (fun pb => eq_term_upto_univ Σ cmp_universe cmp_sort pb) pb ctx ctx' -> + eq_context_gen (fun pb => eq_term_upto_univ (nl_global_env Σ) cmp_universe cmp_sort pb) pb (nlctx ctx) (nlctx ctx'). Proof. intros H. induction H; constructor; simpl; destruct p; intuition @@ -392,19 +328,18 @@ Proof. intros. apply nl_eq_term_upto_univ. assumption. Qed. -Lemma nl_compare_term {cf:checker_flags} le Σ: - forall φ t t', - compare_term le Σ φ t t' -> - compare_term le (nl_global_env Σ) φ (nl t) (nl t'). +Lemma nl_compare_term {cf:checker_flags} Σ φ pb t t' : + compare_term Σ φ pb t t' -> + compare_term (nl_global_env Σ) φ pb (nl t) (nl t'). Proof. - destruct le; intros. + destruct pb; intros. - apply nl_eq_term. assumption. - apply nl_leq_term. assumption. Qed. Corollary eq_term_nl_eq : forall u v, - eq_term_upto_univ empty_global_env eq eq u v -> + eq_term_upto_univ empty_global_env (fun _ => eq) (fun _ => eq) Conv u v -> nl u = nl v. Proof. intros u v h. @@ -424,12 +359,12 @@ Local Ltac ih3 := (*Lemma eq_context_nl_inv_IH Σ Re ctx ctx' : onctx (fun u : term => - forall (Rle : Universe.t -> Universe.t -> Prop) + forall (Rle : sort -> sort -> Prop) (napp : nat) (v : term), eq_term_upto_univ_napp Σ Re Rle napp (nl u) (nl v) -> eq_term_upto_univ_napp Σ Re Rle napp u v) ctx -> - eq_context_gen eq eq (map (map_decl_anon nl) ctx) (map (map_decl_anon nl) ctx') -> - eq_context_gen eq eq ctx ctx'. + eq_context_upto_names (map (map_decl_anon nl) ctx) (map (map_decl_anon nl) ctx') -> + eq_context_upto_names ctx ctx'. Proof. intros Hctx. unfold ondecl in *. induction ctx as [|[na [b|] ty] Γ] in ctx', Hctx |- *; @@ -488,7 +423,7 @@ Lemma binder_anonymize n : eq_binder_annot n (anonymize n). Proof. destruct n; reflexivity. Qed. #[global] Hint Resolve binder_anonymize : core. -#[global] Hint Constructors compare_decls : core. +#[global] Hint Constructors eq_decl_upto_names : core. Local Hint Unfold map_decl_anon : core. (* Lemma eq_term_upto_univ_tm_nl : @@ -543,8 +478,8 @@ Corollary eq_term_tm_nl : Proof. intros flags Σ G u. eapply eq_term_upto_univ_tm_nl. - - intro. eapply eq_universe_refl. - - intro. eapply eq_universe_refl. + - intro. eapply eq_sort_refl. + - intro. eapply eq_sort_refl. Qed. *) Lemma nl_decompose_prod_assum : @@ -800,31 +735,31 @@ Proof. Qed. Lemma nl_eq_decl {cf:checker_flags} : - forall le Σ φ d d', - compare_decl le Σ φ d d' -> - compare_decl le (nl_global_env Σ) φ (map_decl nl d) (map_decl nl d'). + forall Σ φ pb d d', + compare_decl Σ φ pb d d' -> + compare_decl (nl_global_env Σ) φ pb (map_decl nl d) (map_decl nl d'). Proof. - intros le Σ φ d d' []; constructor; destruct le; + intros Σ φ pb d d' []; constructor; destruct pb; intuition auto using nl_eq_term, nl_leq_term. Qed. Lemma nl_eq_decl' {cf:checker_flags} : - forall le Σ φ d d', - compare_decl le Σ φ d d' -> - compare_decl le (nl_global_env Σ) φ (map_decl_anon nl d) (map_decl_anon nl d'). + forall Σ φ pb d d', + compare_decl Σ φ pb d d' -> + compare_decl (nl_global_env Σ) φ pb (map_decl_anon nl d) (map_decl_anon nl d'). Proof. - intros le Σ φ d d' []; constructor; destruct le; + intros Σ φ pb d d' []; constructor; destruct pb; intuition auto using nl_eq_term, nl_leq_term. Qed. Lemma nl_eq_context {cf:checker_flags} : - forall le Σ φ Γ Γ', - compare_context le Σ φ Γ Γ' -> - compare_context le (nl_global_env Σ) φ (nlctx Γ) (nlctx Γ'). + forall Σ φ pb Γ Γ', + compare_context Σ φ pb Γ Γ' -> + compare_context (nl_global_env Σ) φ pb (nlctx Γ) (nlctx Γ'). Proof. - intros le Σ φ Γ Γ' h. + intros Σ φ pb Γ Γ' h. unfold eq_context, nlctx. - destruct le; now eapply eq_context_nl. + now eapply eq_context_nl. Qed. Lemma nl_decompose_app : @@ -1398,7 +1333,7 @@ Lemma nl_conv {cf:checker_flags} : Proof. intros ? Σ ? Γ A B h. induction h. - - constructor. unfold leq_term_ext. rewrite global_ext_constraints_nlg. + - constructor. rewrite global_ext_constraints_nlg. unfold nlg. destruct Σ. apply nl_eq_term. assumption. - eapply cumul_red_l. 2: eassumption. @@ -1414,7 +1349,7 @@ Lemma nl_cumul {cf:checker_flags} : Proof. intros ? Σ ? Γ A B h. induction h. - - constructor. unfold leq_term_ext. rewrite global_ext_constraints_nlg. + - constructor. rewrite global_ext_constraints_nlg. unfold nlg. destruct Σ. apply nl_leq_term. assumption. - eapply cumul_red_l. 2: eassumption. diff --git a/pcuic/theories/Conversion/PCUICRenameConv.v b/pcuic/theories/Conversion/PCUICRenameConv.v index d2c947fdf..4d63a6f23 100644 --- a/pcuic/theories/Conversion/PCUICRenameConv.v +++ b/pcuic/theories/Conversion/PCUICRenameConv.v @@ -220,13 +220,12 @@ Section Renaming. Context `{cf : checker_flags}. -Lemma eq_term_upto_univ_rename Σ : - forall Re Rle napp u v f, - eq_term_upto_univ_napp Σ Re Rle napp u v -> - eq_term_upto_univ_napp Σ Re Rle napp (rename f u) (rename f v). +Lemma eq_term_upto_univ_rename Σ cmp_universe cmp_sort pb napp u v f : + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp u v -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp (rename f u) (rename f v). Proof using Type. - intros Re Rle napp u v f h. - induction u in v, napp, Rle, f, h |- * using term_forall_list_ind. + intros h. + induction u in v, napp, pb, f, h |- * using term_forall_list_ind. all: dependent destruction h. all: try solve [ simpl ; constructor ; eauto @@ -239,29 +238,19 @@ Proof using Type. - simpl. constructor. all: eauto. * rewrite /rename_predicate. destruct X; destruct e as [? [? [ectx ?]]]. - rewrite (All2_fold_length ectx). red. + rewrite (All2_length ectx). red. intuition auto; simpl; solve_all. - * red in X0. solve_all. - rewrite -(All2_fold_length a). + * red in X0. unfold eq_branches, eq_branch in *. solve_all. + rewrite -(All2_length a1). now eapply b0. - - simpl. constructor. - apply All2_length in a as e. rewrite <- e. + - simpl. constructor. unfold eq_mfixpoint in *. + apply All2_length in e as eq. rewrite <- eq. generalize #|m|. intro k. - induction X in mfix', a, f, k |- *. - + inversion a. constructor. - + inversion a. subst. - simpl. constructor. - * unfold map_def. intuition eauto. - * eauto. - - simpl. constructor. - apply All2_length in a as e. rewrite <- e. + solve_all. + - simpl. constructor. unfold eq_mfixpoint in *. + apply All2_length in e as eq. rewrite <- eq. generalize #|m|. intro k. - induction X in mfix', a, f, k |- *. - + inversion a. constructor. - + inversion a. subst. - simpl. constructor. - * unfold map_def. intuition eauto. - * eauto. + solve_all. - simpl. constructor. eapply onPrims_map_prop; tea. cbn; intuition eauto. Qed. diff --git a/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v b/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v index e16019e6b..be8df6415 100644 --- a/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v +++ b/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v @@ -6,7 +6,7 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICOnOne PCUICAstUtils PCUICInducti PCUICLiftSubst PCUICEquality PCUICUnivSubst PCUICCases PCUICCumulativity PCUICTyping PCUICReduction PCUICWeakeningEnv - PCUICClosed PCUICPosition PCUICGuardCondition. + PCUICClosed PCUICPosition. Require Import Equations.Prop.DepElim. From Equations Require Import Equations. @@ -36,13 +36,13 @@ Qed. Lemma eq_valuation 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 : Universe.t, Universe.to_cuniv v u = Universe.to_cuniv v' u. + : forall u : sort, Sort.to_csort v u = Sort.to_csort v' u. Proof. intros [| | u]; cbnr. f_equal. assert (He : forall e : LevelExpr.t, val v e = val v' e). { intros [[] b]; cbnr; rewrite ?H1 ?H2; reflexivity. } rewrite !val_fold_right. - induction ((List.rev (LevelAlgExpr.exprs u).2)); cbn; congruence. + induction ((List.rev (Universe.exprs u).2)); cbn; congruence. Qed. (* Lemma is_prop_subst_instance_level u l @@ -66,10 +66,10 @@ Proof. destruct nth_error; cbnr. Qed. -Lemma subst_instance_univ0_val u exprs v v' +Lemma subst_instance_universe_val u exprs v v' (H1 : forall s, valuation_mono v s = valuation_mono v' s) (H2 : forall n, val v (nth n u Level.lzero) = valuation_poly v' n) - : val v (subst_instance_univ0 u exprs) = val v' exprs. + : val v (subst_instance_universe u exprs) = val v' exprs. Proof. symmetry. apply val_caract. split. @@ -84,13 +84,13 @@ Proof. symmetry; now apply subst_instance_level_expr_val. Qed. -Lemma subst_instance_univ_val u s v v' +Lemma subst_instance_sort_val u s v v' (H1 : forall s, valuation_mono v s = valuation_mono v' s) (H2 : forall n, val v (nth n u Level.lzero) = valuation_poly v' n) - : Universe.to_cuniv v (subst_instance_univ u s) = Universe.to_cuniv v' s. + : Sort.to_csort v (subst_instance_sort u s) = Sort.to_csort v' s. Proof. destruct s as [ | | exprs]; cbnr. - f_equal; now apply subst_instance_univ0_val. + f_equal; now apply subst_instance_universe_val. Qed. Definition subst_instance_valuation (u : Instance.t) (v : valuation) := @@ -105,120 +105,102 @@ Proof. Qed. -Lemma subst_instance_univ0_val' u exprs v - : val v (subst_instance_univ0 u exprs) = val (subst_instance_valuation u v) exprs. +Lemma subst_instance_universe_val' u exprs v + : val v (subst_instance_universe u exprs) = val (subst_instance_valuation u v) exprs. Proof. - now apply subst_instance_univ0_val. + now apply subst_instance_universe_val. Qed. -Lemma subst_instance_univ_val' u l v - : Universe.to_cuniv v (subst_instance_univ u l) = Universe.to_cuniv (subst_instance_valuation u v) l. +Lemma subst_instance_sort_val' u l v + : Sort.to_csort v (subst_instance_sort u l) = Sort.to_csort (subst_instance_valuation u v) l. Proof. - now apply subst_instance_univ_val. + now apply subst_instance_sort_val. Qed. -Lemma subst_instance_univ_make' (l : LevelExpr.t) u : - subst_instance u (LevelAlgExpr.make l) = LevelAlgExpr.make (subst_instance_level_expr u l). +Lemma subst_instance_universe_make' (l : LevelExpr.t) u : + subst_instance u (Universe.make l) = Universe.make (subst_instance_level_expr u l). Proof. reflexivity. Qed. -Lemma subst_instance_univ_make l u : - subst_instance_univ u (Universe.make l) - = Universe.make (subst_instance_level u l). +Lemma subst_instance_universe_make l u : + subst_instance_universe u (Universe.make' l) + = Universe.make' (subst_instance 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). +Class SubstUnivPreserving eq_universe {A} `{UnivSubst A} (eqA : A -> A -> Prop) := Build_SubstUnivPreserving : + forall u u1 u2, cmp_universe_instance eq_universe u1 u2 -> + eqA (subst_instance u1 u) (subst_instance u2 u). -Lemma subst_equal_inst_inst Re : - SubstUnivPreserving Re -> - forall u u1 u2, R_universe_instance Re u1 u2 -> - R_universe_instance Re (subst_instance u1 u) +Lemma subst_equal_inst_inst Re Re' : + SubstUnivPreserving Re Re' -> + forall u u1 u2, cmp_universe_instance Re u1 u2 -> + cmp_universe_instance Re' (subst_instance u1 u) (subst_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. + intros u1 u2; unfold cmp_universe_instance; cbn; constructor. + - pose proof (hRe (Universe.make' a) u1 u2 H) as HH. + now rewrite /subst_instance !subst_instance_universe_make in HH. - exact (IHu u1 u2 H). Qed. -Lemma subst_equal_inst_global_inst Σ Re Rle gr napp : - RelationClasses.Reflexive Re -> - SubstUnivPreserving Re -> - RelationClasses.subrelation Re Rle -> - forall u u1 u2, R_universe_instance Re u1 u2 -> - R_global_instance Σ Re Rle gr napp (subst_instance u1 u) +Lemma subst_equal_inst_global_inst Σ cmp_universe pb gr napp : + SubstUnivPreserving (cmp_universe Conv) (cmp_universe Conv) -> + forall u u1 u2, cmp_universe_instance (cmp_universe Conv) u1 u2 -> + cmp_global_instance Σ cmp_universe pb gr napp (subst_instance u1 u) (subst_instance u2 u). Proof. - intros reflRe hRe subr u u1 u2 Ru1u2. - unfold R_global_instance, R_global_instance_gen, R_opt_variance. - destruct global_variance_gen as [v|]; auto using subst_equal_inst_inst. - induction u in v |- *; cbnr; try now constructor. - - destruct v; simpl; auto. - split; auto. - destruct t; simpl; auto. - * pose proof (hRe (Universe.make a) u1 u2 Ru1u2) as HH. - now rewrite !subst_instance_univ_make in HH. - * pose proof (hRe (Universe.make a) u1 u2 Ru1u2) as HH. - now rewrite !subst_instance_univ_make in HH. + intros subst_conv u u1 u2 Ru1u2. + unfold cmp_global_instance, cmp_global_instance_gen, cmp_opt_variance. + destruct global_variance_gen as [| |v]. + - now eapply subst_equal_inst_inst. + - len. + - left. now eapply subst_equal_inst_inst. Qed. -Lemma eq_term_upto_univ_subst_instance Σ Re Rle napp : - RelationClasses.Reflexive Re -> - SubstUnivPreserving Re -> - RelationClasses.subrelation Re Rle -> +Lemma eq_term_upto_univ_subst_instance Σ cmp_universe cmp_sort pb napp : + SubstUnivPreserving (cmp_universe Conv) (cmp_universe Conv) -> + SubstUnivPreserving (cmp_universe Conv) (cmp_sort Conv) -> + SubstUnivPreserving (cmp_universe Conv) (cmp_sort pb) -> forall t u1 u2, - R_universe_instance Re u1 u2 -> - eq_term_upto_univ_napp Σ Re Rle napp (subst_instance u1 t) + cmp_universe_instance (cmp_universe Conv) u1 u2 -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp (subst_instance u1 t) (subst_instance u2 t). Proof. - intros ref hRe subr t. - induction t in napp, Re, Rle, ref, hRe, subr |- * using term_forall_list_ind; intros u1 u2 hu. - all: cbn; try constructor; eauto using subst_equal_inst_inst. - all: try eapply All2_map, All_All2; tea; cbn; intros; rdest; eauto. - all: try (eapply X0 || eapply IHt || eapply IHt1 || eapply IHt2 || eapply e || eapply e0); try typeclasses eauto; auto. - all: eauto using subst_equal_inst_global_inst. - - rewrite /eq_predicate /=. intuition auto. - * solve_all. eapply All_All2; tea; cbn; intros; rdest; eauto. - eapply X; eauto. tc. - * eapply subst_equal_inst_inst => //. - * solve_all. reflexivity. - * eapply X => //. - - solve_all. reflexivity. - - destruct p as [? []]; cbn in X; constructor; cbn; intuition eauto; cbn. - * rewrite -!subst_instance_univ_make. now eapply hRe. - * now eapply a1. - * now eapply a0. - * solve_all. eapply All_All2; tea. intuition eauto. now apply X. + intros hsubst_conv hsubst_sort_conv hsubst_sort_pb t. + induction t in napp, pb, hsubst_sort_pb |- * using term_forall_list_ind; intros u1 u2 hu. + all: cbn; try constructor; eauto using subst_equal_inst_inst, subst_equal_inst_global_inst. + all: solve_all; unfold eq_predicate, eq_branches, eq_branch, eq_mfixpoint in *. + all: try eapply All2_map, All_All2; tea; cbn; intros; rdest; eauto using subst_equal_inst_inst. + - solve_all. + - reflexivity. + - reflexivity. + - destruct p as [? []]; try now constructor. + destruct X as (hty & hdef & harr). + constructor; cbn; eauto. + * rewrite /= -!subst_instance_universe_make. + now eapply hsubst_conv. + * solve_all. Qed. #[global] Instance eq_universe_SubstUnivPreserving {cf:checker_flags} φ : - SubstUnivPreserving (eq_universe φ). + SubstUnivPreserving (eq_universe φ) (eq_universe φ). Proof. - intros [| | exprs]; cbnr. - intros u1 u2 hu. + intros exprs u1 u2 hu. unfold_univ_rel. 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. + - intros l1 X. eapply Forall2_nth_error_Some_l in hu; tea. + destruct hu as [l2 [-> H2]]. specialize (H2 v Hv). - destruct l1, l2'; cbn in *; noconf H2; 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. + cbn in *. lia. + - intros X. eapply Forall2_nth_error_None_l in hu; tea. destruct (nth_error u2 n); [discriminate|reflexivity]. } simpl. apply val_caract; split. @@ -235,25 +217,19 @@ Qed. #[global] Instance leq_universe_SubstUnivPreserving {cf:checker_flags} φ : - SubstUnivPreserving (leq_universe φ). + SubstUnivPreserving (leq_universe φ) (leq_universe φ). Proof. - intros [| | exprs] u1 u2 hu; cbnr. + intros exprs u1 u2 hu. unfold_univ_rel. 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. + - intros l1 X. eapply Forall2_nth_error_Some_l in hu; tea. + destruct hu as [l2 [-> H2]]. specialize (H2 v Hv). - destruct l1, l2'; cbn in *; noconf H2; 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. + cbn in *. lia. + - intros X. eapply Forall2_nth_error_None_l in hu; tea. destruct (nth_error u2 n); [discriminate|reflexivity]. } simpl. rewrite Z.sub_0_r. @@ -276,9 +252,21 @@ Proof. Qed. #[global] -Instance compare_universe_substu {cf} le Σ : SubstUnivPreserving (compare_universe le Σ). +Instance compare_universe_substu {cf} φ pb : SubstUnivPreserving (eq_universe φ) (compare_universe φ pb). +Proof. + destruct pb; tc. + intros u ui ui' H. + apply leq_universe_SubstUnivPreserving. + eapply cmp_universe_instance_impl'; tea. tc. +Qed. + +#[global] +Instance compare_sort_substu {cf:checker_flags} φ pb : + SubstUnivPreserving (eq_universe φ) (compare_sort φ pb). Proof. - destruct le; tc. + intros s u1 u2 hu. + destruct s as [| | u]; cbnr. rewrite compare_sort_type. + now eapply Build_SubstUnivPreserving. Qed. Global Instance subst_instance_def {A} `(UnivSubst A) : UnivSubst (def A) @@ -312,10 +300,10 @@ Proof. Qed. Lemma subst_instance_univ0_two u1 u2 exprs : - subst_instance_univ0 u1 (subst_instance_univ0 u2 exprs) - = subst_instance_univ0 (subst_instance u1 u2) exprs. + subst_instance_universe u1 (subst_instance_universe u2 exprs) + = subst_instance_universe (subst_instance u1 u2) exprs. Proof. - unfold subst_instance_univ0. + unfold subst_instance_universe. eapply eq_univ'. intro l; split; intro Hl; apply map_spec in Hl as [l' [H1 H2]]; apply map_spec; subst. @@ -327,8 +315,8 @@ Qed. Lemma subst_instance_univ_two u1 u2 s : - subst_instance_univ u1 (subst_instance_univ u2 s) - = subst_instance_univ (subst_instance u1 u2) s. + subst_instance_sort u1 (subst_instance_sort u2 s) + = subst_instance_sort (subst_instance u1 u2) s. Proof. destruct s; cbnr. f_equal. apply subst_instance_univ0_two. @@ -440,27 +428,33 @@ Proof. + now apply In_subst_instance_cstrs'. Qed. +Lemma is_propositional_subst_instance_univ u l + : Sort.is_propositional (subst_instance_sort u l) = Sort.is_propositional l. +Proof. + destruct l; cbnr. +Qed. + Lemma is_prop_subst_instance_univ u l - : Universe.is_prop (subst_instance_univ u l) = Universe.is_prop l. + : Sort.is_prop (subst_instance_sort u l) = Sort.is_prop l. Proof. destruct l; cbnr. Qed. Lemma is_sprop_subst_instance_univ u l - : Universe.is_sprop (subst_instance_univ u l) = Universe.is_sprop l. + : Sort.is_sprop (subst_instance_sort u l) = Sort.is_sprop l. Proof. destruct l; cbnr. Qed. Lemma is_prop_subst_instance u x0 : - Universe.is_prop x0 -> Universe.is_prop (subst_instance_univ u x0). + Sort.is_prop x0 -> Sort.is_prop (subst_instance_sort u x0). Proof. now rewrite is_prop_subst_instance_univ. Qed. -Lemma sup_subst_instance_univ0 u s1 s2 : - subst_instance u (LevelAlgExpr.sup s1 s2) - = LevelAlgExpr.sup (subst_instance u s1) (subst_instance u s2). +Lemma sup_subst_instance_univ0 ui u1 u2 : + subst_instance ui (Universe.sup u1 u2) + = Universe.sup (subst_instance ui u1) (subst_instance ui u2). Proof. apply eq_univ'. cbn. intro x; split; intro Hx. @@ -475,8 +469,8 @@ Proof. 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). + subst_instance_sort u (Sort.sup s1 s2) + = Sort.sup (subst_instance_sort u s1) (subst_instance_sort u s2). Proof. destruct s1, s2; cbnr. f_equal. apply sup_subst_instance_univ0. @@ -785,36 +779,52 @@ Qed. Global Instance leq_universe_subst_instance {cf : checker_flags} : SubstUnivPreserved leq_universe. Proof. - intros φ φ' u HH [| | exprs] [| | exprs'] Hle; cbnr; trivial. + intros φ φ' u HH exprs exprs' Hle. unfold_univ_rel eqn:H. - rewrite !subst_instance_univ0_val'; tas. + rewrite !subst_instance_universe_val'; tas. apply Hle. - eapply satisfies_subst_instance; tea. + eapply satisfies_subst_instance_ctr; tea. Qed. Global Instance eq_universe_subst_instance {cf : checker_flags} : SubstUnivPreserved eq_universe. Proof. - intros φ φ' u HH [| | exprs] [| | exprs'] Hle; cbnr; trivial. + intros φ φ' u HH exprs exprs' Hle; cbnr; trivial. unfold_univ_rel eqn:H. - rewrite !subst_instance_univ0_val'; tas. + rewrite !subst_instance_universe_val'; tas. apply Hle. - eapply satisfies_subst_instance; tea. + eapply satisfies_subst_instance_ctr; tea. +Qed. + +Global Instance leq_sort_subst_instance {cf : checker_flags} : SubstUnivPreserved leq_sort. +Proof. + intros φ φ' u HH [| | exprs] [| | exprs'] Hle; cbnr; trivial. + eapply Build_SubstUnivPreserved; tea. +Qed. + +Global Instance eq_sort_subst_instance {cf : checker_flags} : SubstUnivPreserved eq_sort. +Proof. + intros φ φ' u HH [| | exprs] [| | exprs'] Hle; cbnr; trivial. + eapply Build_SubstUnivPreserved; tea. +Qed. + +Global Instance compare_universe_subst_instance {cf : checker_flags} pb : SubstUnivPreserved (fun φ => compare_universe φ pb). +Proof. + destruct pb; tc. +Qed. + +Global Instance compare_sort_subst_instance {cf : checker_flags} pb : SubstUnivPreserved (fun φ => compare_sort φ pb). +Proof. + destruct pb; tc. Qed. -Lemma precompose_subst_instance Rle u i i' : - precompose (R_universe_instance Rle) (subst_instance u) i i' - <~> R_universe_instance (precompose Rle (subst_instance_univ u)) i i'. +Lemma precompose_subst_instance cmp_universe u i i' : + precompose (cmp_universe_instance cmp_universe) (subst_instance u) i i' + <~> cmp_universe_instance (precompose cmp_universe (subst_instance_universe u)) i i'. Proof. - unfold R_universe_instance, subst_instance. - replace (List.map Universe.make (List.map (subst_instance_level u) i)) - with (List.map (subst_instance_univ u) (List.map Universe.make i)). - 1: replace (List.map Universe.make (List.map (subst_instance_level u) i')) - with (List.map (subst_instance_univ u) (List.map Universe.make i')). - 1: split. - 1: apply Forall2_map_inv. - 1: apply Forall2_map. - all: rewrite !map_map; apply map_ext. - all: intro; apply subst_instance_univ_make. + unfold cmp_universe_instance, subst_instance, on_rel. + split; intro H; [apply Forall2_map_inv in H | apply Forall2_map]; apply Forall2_impl with (1 := H); intros. + - rewrite !subst_instance_universe_make //. + - rewrite -!subst_instance_universe_make //. Qed. Definition precompose_subst_instance__1 Rle u i i' @@ -831,91 +841,112 @@ Proof. Qed. Lemma subst_instance_make'_make u l : - subst_instance u (LevelAlgExpr.make (LevelExpr.make l)) = - LevelAlgExpr.make (LevelExpr.make (subst_instance_level u l)). + subst_instance u (Universe.make (LevelExpr.make l)) = + Universe.make (LevelExpr.make (subst_instance_level u l)). Proof. - now rewrite subst_instance_univ_make' subst_instance_level_expr_make. + now rewrite subst_instance_universe_make' subst_instance_level_expr_make. Qed. -Lemma precompose_subst_instance_global Σ Re Rle gr napp u i i' : - precompose (R_global_instance Σ Re Rle gr napp) (subst_instance u) i i' - <~> R_global_instance Σ (precompose Re (subst_instance_univ u)) - (precompose Rle (subst_instance_univ u)) gr napp i i'. +Lemma precompose_subst_instance_global Σ cmp_universe pb gr napp u i i' : + precompose (cmp_global_instance Σ cmp_universe pb gr napp) (subst_instance u) i i' + <~> cmp_global_instance Σ (fun pb => precompose (cmp_universe pb) (subst_instance_universe u)) + pb gr napp i i'. Proof. - unfold R_global_instance, R_global_instance_gen, R_opt_variance, subst_instance. - destruct global_variance_gen as [v|]; eauto using precompose_subst_instance. - induction i in i', v |- *; destruct i', v; simpl; try split; auto. - - destruct (IHi i' []). intros; auto. - - destruct (IHi i' []). intros; auto. - - destruct (IHi i' v). intros []; split; auto. - destruct t0; simpl in *; auto. - * now rewrite !subst_instance_make'_make. - * now rewrite !subst_instance_make'_make. - - destruct (IHi i' v). intros []; split; auto. - destruct t0; simpl in *; auto. - * now rewrite !subst_instance_make'_make in H. - * now rewrite !subst_instance_make'_make in H. + unfold cmp_global_instance, cmp_global_instance_gen, cmp_opt_variance, subst_instance. + destruct global_variance_gen as [| |v]. + - apply precompose_subst_instance. + - len. + - split. all: intros [H|H]; [left|right]. + 1,3: now apply precompose_subst_instance. + all: [> rewrite -(map_id v) in H; apply Forall3_map_inv in H | rewrite -(map_id v); apply Forall3_map]; + apply Forall3_impl with (1 := H); clear v i i' H; intros v ??. + all: destruct v => //=; unfold on_rel in *. + all: rewrite !subst_instance_universe_make //. Qed. -Definition precompose_subst_instance_global__1 Σ Re Rle gr napp u i i' - := fst (precompose_subst_instance_global Σ Re Rle gr napp u i i'). +Definition precompose_subst_instance_global__1 Σ cmp_universe pb gr napp u i i' + := fst (precompose_subst_instance_global Σ cmp_universe pb gr napp u i i'). -Definition precompose_subst_instance_global__2 Σ Re Rle gr napp u i i' - := snd (precompose_subst_instance_global Σ Re Rle gr napp u i i'). +Definition precompose_subst_instance_global__2 Σ cmp_universe pb gr napp u i i' + := snd (precompose_subst_instance_global Σ cmp_universe pb gr napp u i i'). Global Instance eq_term_upto_univ_subst_preserved {cf : checker_flags} Σ - (Re Rle : ConstraintSet.t -> Universe.t -> Universe.t -> Prop) napp - {he: SubstUnivPreserved Re} {hle: SubstUnivPreserved Rle} - : SubstUnivPreserved (fun φ => eq_term_upto_univ_napp Σ (Re φ) (Rle φ) napp). + (cmp_universe : forall _ _ (_ _ : Universe.t), Prop) (cmp_sort : forall _ _ (_ _ : sort), Prop) pb napp + {huniverse : SubstUnivPreserved (fun φ => cmp_universe φ Conv)} {huniverse2 : SubstUnivPreserved (fun φ => cmp_universe φ pb)} + {hsort : SubstUnivPreserved (fun φ => cmp_sort φ Conv)} {hsort2 : SubstUnivPreserved (fun φ => cmp_sort φ pb)} + : SubstUnivPreserved (fun φ => eq_term_upto_univ_napp Σ (cmp_universe φ) (cmp_sort φ) pb napp). Proof. intros φ φ' u HH t t'. - specialize (he _ _ _ HH). - specialize (hle _ _ _ HH). - clear HH. cbn in he. - induction t in napp, t', Rle, hle |- * using term_forall_list_ind; + specialize (huniverse _ _ _ HH). + specialize (huniverse2 _ _ _ HH). + specialize (hsort _ _ _ HH). + specialize (hsort2 _ _ _ HH). + clear HH. cbn in huniverse, huniverse2, hsort, hsort2. + induction t in napp, pb, huniverse2, hsort2, t' |- * using term_forall_list_ind; inversion 1; subst; cbn; constructor; - eauto using precompose_subst_instance__2, R_universe_instance_impl'. + eauto using precompose_subst_instance__2, cmp_universe_instance_impl'. + all: unfold eq_predicate, eq_branches, eq_branch, eq_mfixpoint in *. all: try (apply All2_map; eapply All2_impl'; tea; eapply All_impl; eauto; cbn; intros; aa). - - inv X. - eapply precompose_subst_instance_global__2. - eapply R_global_instance_impl_same_napp; eauto. - - inv X. - eapply precompose_subst_instance_global__2. - eapply R_global_instance_impl_same_napp; eauto. + - eapply precompose_subst_instance_global__2. + eapply cmp_global_instance_impl_same_napp; eauto. + - eapply precompose_subst_instance_global__2. + eapply cmp_global_instance_impl_same_napp; eauto. - destruct X2 as [? [? [? ?]]]. repeat split; simpl; eauto; solve_all. * eapply precompose_subst_instance. - eapply R_universe_instance_impl; eauto. - - destruct p as [? []]; depelim X1; cbn in X; try constructor; intuition eauto. - * cbn. rewrite -!subst_instance_univ_make. now eapply he. - * now eapply a2. - * now eapply a0. - * cbn. solve_all. + eapply cmp_universe_instance_impl; eauto. + - destruct p as [? []]; depelim X1; try now constructor. + destruct X as (hty & hdef & harr). + constructor; cbn; eauto. + * rewrite /= -!subst_instance_universe_make. + now eapply huniverse. + * solve_all. Qed. -Lemma leq_term_subst_instance {cf : checker_flags} Σ : SubstUnivPreserved (leq_term Σ). +Lemma leq_term_subst_instance {cf : checker_flags} Σ : SubstUnivPreserved (fun φ => leq_term Σ φ). Proof. apply eq_term_upto_univ_subst_preserved; cbn; apply _. Qed. -Lemma eq_term_subst_instance {cf : checker_flags} Σ : SubstUnivPreserved (eq_term Σ). +Lemma eq_term_subst_instance {cf : checker_flags} Σ : SubstUnivPreserved (fun φ => eq_term Σ φ). Proof. apply eq_term_upto_univ_subst_preserved; cbn; exact _. Qed. -Lemma compare_term_subst_instance {cf : checker_flags} pb Σ : SubstUnivPreserved (compare_term pb Σ). -Proof. apply eq_term_upto_univ_subst_preserved; cbn; try destruct pb; exact _. Qed. +Lemma compare_term_subst_instance {cf : checker_flags} Σ pb : SubstUnivPreserved (fun φ => compare_term Σ φ pb). +Proof. destruct pb; [apply eq_term_subst_instance|apply leq_term_subst_instance]. Qed. + +Global Instance compare_decls_subst_preserved {cf : checker_flags} Σ + (cmp_universe : forall _ _ (_ _ : Universe.t), Prop) (cmp_sort : forall _ _ (_ _ : sort), Prop) pb + {huniverse : SubstUnivPreserved (fun φ => cmp_universe φ Conv)} {huniverse2 : SubstUnivPreserved (fun φ => cmp_universe φ pb)} + {hsort : SubstUnivPreserved (fun φ => cmp_sort φ Conv)} {hsort2 : SubstUnivPreserved (fun φ => cmp_sort φ pb)} + : SubstUnivPreserved (fun φ => compare_decls (fun pb => eq_term_upto_univ Σ (cmp_universe φ) (cmp_sort φ) pb) pb). +Proof. + intros φ φ' u HH d d' []; constructor; cbn; auto. + all: now eapply eq_term_upto_univ_subst_preserved. +Qed. + +Global Instance eq_context_upto_subst_preserved {cf : checker_flags} Σ + (cmp_universe : forall _ _ (_ _ : Universe.t), Prop) (cmp_sort : forall _ _ (_ _ : sort), Prop) pb + {huniverse : SubstUnivPreserved (fun φ => cmp_universe φ Conv)} {huniverse2 : SubstUnivPreserved (fun φ => cmp_universe φ pb)} + {hsort : SubstUnivPreserved (fun φ => cmp_sort φ Conv)} {hsort2 : SubstUnivPreserved (fun φ => cmp_sort φ pb)} + : SubstUnivPreserved (fun φ => eq_context_upto Σ (cmp_universe φ) (cmp_sort φ) pb). +Proof. + intros φ φ' u HH Γ Δ. + induction 1; cbn; constructor; auto. + now eapply compare_decls_subst_preserved. +Qed. (** Now routine lemmas ... *) -Lemma In_subst_instance x u (l : LevelAlgExpr.t) : +Lemma In_subst_instance x u (l : Universe.t) : LevelExprSet.In x (subst_instance u l) <-> (exists x', LevelExprSet.In x' l /\ x = subst_instance u x'). Proof. unfold subst_instance; cbn. - unfold subst_instance_univ0. + unfold subst_instance_universe. now rewrite map_spec. Qed. Lemma subst_instance_univ_super l u - : subst_instance_univ u (Universe.super l) = Universe.super (subst_instance u l). + : subst_instance_sort u (Sort.super l) = Sort.super (subst_instance u l). Proof. destruct l; cbnr. f_equal. apply eq_univ'. @@ -967,11 +998,9 @@ Qed. Lemma product_subst_instance u s1 s2 - : subst_instance_univ u (Universe.sort_of_product s1 s2) - = Universe.sort_of_product (subst_instance_univ u s1) (subst_instance_univ u s2). + : subst_instance_sort u (Sort.sort_of_product s1 s2) + = Sort.sort_of_product (subst_instance_sort u s1) (subst_instance_sort u s2). Proof. - unfold Universe.sort_of_product. - rewrite is_prop_subst_instance_univ; tas. destruct s2; cbn; try reflexivity. destruct s1; cbn; try reflexivity. f_equal. @@ -1488,23 +1517,23 @@ Lemma is_allowed_elimination_subst_instance {cf : checker_flags} (Σ : global_en valid_constraints (global_ext_constraints (Σ.1, univs)) (subst_instance_cstrs inst Σ) -> is_allowed_elimination Σ al u -> - is_allowed_elimination (global_ext_constraints (Σ.1, univs)) al (subst_instance_univ inst u). + is_allowed_elimination (global_ext_constraints (Σ.1, univs)) al (subst_instance_sort inst u). Proof. destruct al, u as [| | u]; trivial. - intros val [H|isal]; [cbn in H; discriminate | right]. + intros val [H|isal] => //; right. cbn in isal |- *. unfold_univ_rel eqn:H. - eapply satisfies_subst_instance in Hv; eauto. - specialize (isal _ Hv). - rewrite subst_instance_univ0_val'; auto. + eapply satisfies_subst_instance_ctr in val; eauto. + specialize (isal _ val). + rewrite subst_instance_universe_val'; auto. Qed. -Global Instance compare_decl_subst_instance {cf : checker_flags} pb Σ : SubstUnivPreserved (compare_decl pb Σ). +Global Instance compare_decl_subst_instance {cf : checker_flags} pb Σ : SubstUnivPreserved (fun φ => compare_decl Σ φ pb). Proof. intros φ1 φ2 u HH ? ? [] => /=; constructor; auto; eapply compare_term_subst_instance; tea. Qed. -Global Instance compare_context_subst_instance {cf : checker_flags} pb Σ : SubstUnivPreserved (compare_context pb Σ). +Global Instance compare_context_subst_instance {cf : checker_flags} pb Σ : SubstUnivPreserved (fun φ => compare_context Σ φ pb). Proof. intros φ φ' u HH Γ Γ' X. eapply All2_fold_map, All2_fold_impl; tea. intros. eapply compare_decl_subst_instance; eassumption. @@ -1513,7 +1542,7 @@ Qed. Lemma subst_instance_destArity Γ A u : destArity (subst_instance u Γ) (subst_instance u A) = match destArity Γ A with - | Some (ctx, s) => Some (subst_instance u ctx, subst_instance_univ u s) + | Some (ctx, s) => Some (subst_instance u ctx, subst_instance_sort u s) | None => None end. Proof. @@ -1614,20 +1643,20 @@ Proof. Qed. Lemma All_local_env_over_subst_instance {cf : checker_flags} Σ Γ (wfΓ : wf_local Σ Γ) : - All_local_env_over typing - (fun Σ0 Γ0 (_ : wf_local Σ0 Γ0) t T (_ : Σ0;;; Γ0 |- t : T) => - forall u univs, wf_ext_wk Σ0 -> - consistent_instance_ext (Σ0.1, univs) Σ0.2 u -> - (Σ0.1, univs) ;;; subst_instance u Γ0 + All_local_env_over (typing Σ) + (fun Γ0 (_ : wf_local Σ Γ0) t T (_ : Σ;;; Γ0 |- t : T) => + forall u univs, wf_ext_wk Σ -> + consistent_instance_ext (Σ.1, univs) Σ.2 u -> + (Σ.1, univs) ;;; subst_instance u Γ0 |- subst_instance u t : subst_instance u T) - Σ Γ wfΓ -> + Γ wfΓ -> forall u univs, wf_ext_wk Σ -> consistent_instance_ext (Σ.1, univs) Σ.2 u -> wf_local (Σ.1, univs) (subst_instance u Γ). Proof. induction 1; simpl; rewrite /subst_instance /=; constructor; cbn in *; auto. - all: eapply infer_typing_sort_impl with _ tu; cbn in *; eauto. + all: eapply lift_sorting_fu_it_impl with (tu := tu); cbn in *; eauto. Qed. #[global] Hint Resolve All_local_env_over_subst_instance : univ_subst. @@ -1655,13 +1684,12 @@ Proof. - apply hin. Qed. -Lemma wf_universe_subst_instance {cf : checker_flags} (Σ : global_env_ext) univs u s : +Lemma wf_universe_subst_instance {cf : checker_flags} (Σ : global_env_ext) univs ui u : wf Σ -> - wf_universe Σ s -> - consistent_instance_ext (Σ.1, univs) Σ.2 u -> - wf_universe (Σ.1, univs) (subst_instance u s). + wf_universe Σ u -> + consistent_instance_ext (Σ.1, univs) Σ.2 ui -> + wf_universe (Σ.1, univs) (subst_instance ui u). Proof. - destruct s as [| | t]; cbnr. intros wfΣ Hl Hu e [[l n] [inl ->]]%In_subst_instance. destruct l as [|s|n']; simpl; auto. - apply global_ext_levels_InSet. @@ -1677,7 +1705,7 @@ Proof. destruct u; try discriminate. lsets. * destruct Hu as [declu [us vc]]. unfold subst_instance. simpl. - destruct (nth_error u n') eqn:hnth. + destruct (nth_error ui n') eqn:hnth. 2: simpl; apply global_ext_levels_InSet. eapply forallb_Forall in declu. eapply nth_error_forall in declu; eauto. @@ -1685,6 +1713,16 @@ Proof. + now apply not_var_global_levels in Hl. Qed. +Lemma wf_sort_subst_instance {cf : checker_flags} (Σ : global_env_ext) univs ui s : + wf Σ -> + wf_sort Σ s -> + consistent_instance_ext (Σ.1, univs) Σ.2 ui -> + wf_sort (Σ.1, univs) (subst_instance ui s). +Proof. + destruct s as [| | u]; cbnr. + apply wf_universe_subst_instance. +Qed. + Definition global_context_set Σ : ContextSet.t := universes Σ. Lemma global_context_set_sub_ext Σ φ : @@ -1903,8 +1941,8 @@ Section SubstIdentity. Lemma consistent_instance_ext_subst_abs_univ Σ u : wf_ext_wk Σ -> - wf_universe Σ u -> - subst_instance_univ (abstract_instance Σ.2) u = u. + wf_sort Σ u -> + subst_instance_sort (abstract_instance Σ.2) u = u. Proof using Type. intros wf cu. destruct u; simpl; auto. f_equal. @@ -1935,7 +1973,7 @@ Section SubstIdentity. now rewrite [subst_instance_instance _ _](consistent_instance_ext_subst_abs _ _ _ wf cu). Qed. - Lemma wf_universe_type1 Σ : wf_universe Σ Universe.type1. + Lemma wf_sort_type1 Σ : wf_sort Σ Sort.type1. Proof using Type. simpl. intros l hin%LevelExprSet.singleton_spec. @@ -1943,10 +1981,10 @@ Section SubstIdentity. apply global_ext_levels_InSet. Qed. - Lemma wf_universe_super {Σ u} : wf_universe Σ u -> wf_universe Σ (Universe.super u). + Lemma wf_sort_super {Σ u} : wf_sort Σ u -> wf_sort Σ (Sort.super u). Proof using Type. destruct u; cbn. - 1-2:intros _ l hin%LevelExprSet.singleton_spec; subst l; apply wf_universe_type1; + 1-2:intros _ l hin%LevelExprSet.singleton_spec; subst l; apply wf_sort_type1; now apply LevelExprSet.singleton_spec. intros Hl. intros l hin. @@ -1969,31 +2007,42 @@ Section SubstIdentity. wf_ext_wk Σ -> let u := abstract_instance (snd Σ) in subst_instance u t = t × subst_instance u T = T) + (fun Σ _ j => wf_ext_wk Σ -> + let u := abstract_instance (snd Σ) in + lift_wfu_term (fun t => subst_instance u t = t) (fun t => subst_instance u t = t) j) (fun Σ Γ => wf_ext_wk Σ -> let u := abstract_instance (snd Σ) in subst_instance u Γ = Γ). Proof using Type. - eapply typing_ind_env; intros; simpl in *; auto; try ((subst u || subst u0); split; [f_equal|]; intuition eauto). - 1:{ induction X; simpl; auto; unfold snoc. - * f_equal; auto. - unfold map_decl. simpl. unfold vass. f_equal. intuition auto. - * unfold map_decl. simpl. unfold vdef. repeat f_equal; intuition auto. } + eapply typing_ind_env; intros; simpl in *; auto. + { destruct X as (X & s & (_ & (Hty & Hu)) & e); tas; repeat split; tas. + 1: destruct j_term => //; destruct X as (_ & (X & _)); tas. + destruct j_univ => //; rewrite e; cbn in Hu; now depelim Hu. } + + { apply All_local_env_cst in X0. clear -X0 X1. + induction X0 => //=. cbn. + f_equal; tas. destruct x as [na bo t]; cbv [map_decl]; simpl in *. + specialize (p X1) as (ptm & pty & _); cbn in *. f_equal; tas. + destruct bo => //. cbn in *. f_equal. apply ptm. } + + all: try ((subst u || subst u0); split; [f_equal|]; intuition eauto). 1:{ rewrite subst_instance_lift. f_equal. generalize H. rewrite -H1 /subst_instance /= nth_error_map H /= => [=]. intros Hdecl. now rewrite -{2}Hdecl. } + all: try match goal with H : lift_wfu_term _ _ _ |- _ => destruct H as (Htm & Hty & Hs); cbn in Htm, Hty, Hs end. all:try (solve [f_equal; eauto; try congruence]). all:try (rewrite ?subst_instance_two; f_equal; eapply consistent_instance_ext_subst_abs; eauto). - now rewrite consistent_instance_ext_subst_abs_univ. - rewrite consistent_instance_ext_subst_abs_univ //. - now apply wf_universe_super. + now apply wf_sort_super. - - rewrite product_subst_instance. f_equal; - intuition eauto; now noconf b0; noconf b1. + - rewrite product_subst_instance. do 2 f_equal; tas. + now noconf b0. - intuition auto. noconf a; noconf b; noconf b0. rewrite subst_instance_subst /= /subst1. @@ -2037,19 +2086,26 @@ Section SubstIdentity. rewrite /subst_instance_list. now rewrite map_rev Hpars. * rewrite [subst_instance_constr _ _]subst_instance_two. noconf Hi. now rewrite [subst_instance _ u]H. - - solve_all. destruct a as [s [? ?]]; solve_all. - - clear X0. eapply nth_error_all in X as [s [Hs [IHs _]]]; eauto. - - solve_all. destruct a as [s [? ?]]. solve_all. - - clear X0. eapply nth_error_all in X as [s [Hs [IHs _]]]; eauto. - - destruct p as [? []]; cbn => //. do 2 f_equal. + - solve_all. + + destruct a0 as (_ & X & _); tas. + + destruct b as (X & _); tas. + - eapply nth_error_all in X0 as (_ & X0 & _); tea. + - solve_all. + + destruct a0 as (_ & X & _); tas. + + destruct b as (X & _); tas. + - eapply nth_error_all in X0 as (_ & X0 & _); tea. + - destruct p as [? []]; cbnr. do 2 f_equal. depelim X0. specialize (hty X1); specialize (hdef X1). - unfold mapu_array_model; destruct a; cbn -[Universe.make] in * => //=; f_equal; intuition eauto. - * destruct array_level => //. - rewrite subst_instance_univ_make in b. now injection b. + unfold mapu_array_model; destruct a; cbn -[Universe.make'] in *. + f_equal; intuition eauto. + * rewrite /subst_instance subst_instance_universe_make in b. + now injection b as e. * solve_all. - depelim X0; cbn => //=. depelim X. simp prim_type. cbn. f_equal; intuition eauto. - do 2 f_equal. cbn -[Universe.make] in b. rewrite subst_instance_univ_make in b. - now injection b. + do 2 f_equal. + cbn -[Universe.make'] in b. + rewrite /subst_instance subst_instance_universe_make in b. + now injection b as e. Qed. Lemma typed_subst_abstract_instance Σ Γ t T : diff --git a/pcuic/theories/Conversion/PCUICWeakeningConfigConv.v b/pcuic/theories/Conversion/PCUICWeakeningConfigConv.v index a3a9c9e6b..8dadbaf65 100644 --- a/pcuic/theories/Conversion/PCUICWeakeningConfigConv.v +++ b/pcuic/theories/Conversion/PCUICWeakeningConfigConv.v @@ -12,24 +12,21 @@ Require Import ssreflect. Set Default Goal Selector "!". Implicit Types (cf : checker_flags). -Lemma compare_term_config_impl {cf1 cf2} pb Σ φ t t' +Lemma compare_term_config_impl {cf1 cf2} Σ φ pb t t' : config.impl cf1 cf2 - -> @compare_term cf1 pb Σ φ t t' -> @compare_term cf2 pb Σ φ t t'. + -> @compare_term cf1 Σ φ pb t t' -> @compare_term cf2 Σ φ pb t t'. Proof. intro H. apply eq_term_upto_univ_impl; auto. - all: repeat change (@eq_universe ?cf) with (@compare_universe cf Conv). - all: repeat change (@leq_universe ?cf) with (@compare_universe cf Cumul). - 3: destruct pb. - 4: transitivity (@compare_universe cf1 Cumul φ); tc. - all: intros ??; now eapply cmp_universe_config_impl. + 1,2: intros ??; now eapply cmp_universe_config_impl. + 1,2: intros ??; now eapply cmp_sort_config_impl. Qed. Lemma eq_term_config_impl {cf1 cf2} Σ φ t t' - : config.impl cf1 cf2 -> @compare_term cf1 Conv Σ φ t t' -> @compare_term cf2 Conv Σ φ t t'. + : config.impl cf1 cf2 -> @compare_term cf1 Σ φ Conv t t' -> @compare_term cf2 Σ φ Conv t t'. Proof. eapply compare_term_config_impl with (pb := Conv). Qed. Lemma leq_term_config_impl {cf1 cf2} Σ ctrs t u - : config.impl cf1 cf2 -> @compare_term cf1 Cumul Σ ctrs t u -> @compare_term cf2 Cumul Σ ctrs t u. + : config.impl cf1 cf2 -> @compare_term cf1 Σ ctrs Cumul t u -> @compare_term cf2 Σ ctrs Cumul t u. Proof. apply compare_term_config_impl with (pb := Cumul). Qed. Lemma compare_decl_config_impl {cf1 cf2} pb Σ φ d d' @@ -46,61 +43,6 @@ Proof. intros Hcf. induction 1; constructor; auto; eapply (@compare_decl_config_impl cf1 cf2); eassumption. Qed. -(* TODO: Factor with PCUICWeakeningEnvConv.R_global_instance_weaken_env *) -Lemma R_global_instance_weaken_subrel Σ Re Re' Rle Rle' gr napp : - RelationClasses.subrelation Re Re' -> - RelationClasses.subrelation Rle Rle' -> - RelationClasses.subrelation Re Rle' -> - subrelation (R_global_instance Σ Re Rle gr napp) (R_global_instance Σ Re' Rle' gr napp). -Proof. - intros he hle hele t t'. - rewrite /R_global_instance_gen /R_opt_variance. - destruct global_variance_gen as [v|] eqn:look. - - induction t in v, t' |- *; destruct v, t'; simpl; auto. - intros []; split; auto. - destruct t0; simpl; auto. - - eauto using R_universe_instance_impl'. -Qed. - -(* TODO: Factor with PCUICWeakeningEnvConv.eq_term_upto_univ_weaken_env *) -#[global] -Instance eq_term_upto_univ_weaken_subrel Σ Re Re' Rle Rle' napp : - RelationClasses.subrelation Re Re' -> - RelationClasses.subrelation Rle Rle' -> - RelationClasses.subrelation Re Rle' -> - CRelationClasses.subrelation (eq_term_upto_univ_napp Σ Re Rle napp) - (eq_term_upto_univ_napp Σ Re' Rle' napp). -Proof. - intros he hele hle t t'. - induction t in napp, t', Rle, Rle', hle, hele |- * using PCUICInduction.term_forall_list_ind; - try (inversion 1; subst; constructor; - eauto using R_universe_instance_impl'; fail). - - inversion 1; subst; constructor. - eapply All2_impl'; tea. - eapply All_impl; eauto. - - inversion 1; subst; constructor. - eapply R_global_instance_weaken_subrel; [ .. | eassumption ]. all:eauto. - - inversion 1; subst; constructor. - eapply R_global_instance_weaken_subrel; [ .. | eassumption ]. all:eauto. - - inversion 1; subst; destruct X as [? [? ?]]; constructor; eauto. - * destruct X2 as [? [? ?]]. - constructor; intuition auto; solve_all. - + eauto using R_universe_instance_impl'. - * eapply All2_impl'; tea. - eapply All_impl; eauto. - cbn. intros x [? ?] y [? ?]. split; eauto. - - inversion 1; subst; constructor. - eapply All2_impl'; tea. - eapply All_impl; eauto. - cbn. intros x [? ?] y [[[? ?] ?] ?]. repeat split; eauto. - - inversion 1; subst; constructor. - eapply All2_impl'; tea. - eapply All_impl; eauto. - cbn. intros x [? ?] y [[[? ?] ?] ?]. repeat split; eauto. - - intros h; depelim h; constructor; cbn in X; intuition eauto. - depelim o; cbn in X; constructor; intuition eauto. solve_all. -Qed. - Lemma weakening_config_cumul_gen {cf1 cf2} pb Σ Γ M N : config.impl cf1 cf2 -> @cumulAlgo_gen cf1 Σ Γ pb M N -> @@ -137,33 +79,25 @@ Proof. - eapply cumul_Evar. solve_all. - eapply cumul_Case. * cbv [cumul_predicate] in *; destruct_head'_prod. repeat split; tas. - + eapply R_universe_instance_impl'; - [ hnf; intros * ?; eapply (@cmp_universe_config_impl cf1 cf2) | ]; - eassumption. - * solve_all. - * solve_all. - - eapply cumul_Fix. solve_all. - - eapply cumul_CoFix; solve_all. + eapply cmp_universe_instance_impl'; + [ hnf; intros * ?; eapply (@cmp_universe_config_impl cf1 cf2) | ]; + eassumption. + * assumption. + * unfold cumul_branches, cumul_branch in *. solve_all. + - eapply cumul_Fix. unfold cumul_mfixpoint in *. solve_all. + - eapply cumul_CoFix. unfold cumul_mfixpoint in *. solve_all. - eapply cumul_Prim; solve_all. destruct X; constructor; eauto. * now eapply (@cmp_universe_config_impl cf1 cf2). * solve_all. - eapply cumul_Ind; eauto. 2:solve_all. - eapply @R_global_instance_weaken_subrel; [ .. | eassumption ]. - all: repeat change (@eq_universe ?cf) with (@compare_universe cf Conv). - all: repeat change (@leq_universe ?cf) with (@compare_universe cf Cumul). - 3: destruct pb. - all: try (hnf; intros *; eapply (@cmp_universe_config_impl cf1 cf2); eassumption). - all: now etransitivity; [ | hnf; intros *; eapply (@cmp_universe_config_impl cf1 cf2); eassumption ]; tc. + eapply @cmp_global_instance_impl; [ .. | eassumption ]. + 3: auto with arith. all: intros ??; now apply (@cmp_universe_config_impl cf1 cf2). - eapply cumul_Construct; eauto. 2:solve_all. - eapply @R_global_instance_weaken_subrel; [ .. | eassumption ]. - all: repeat change (@eq_universe ?cf) with (@compare_universe cf Conv). - all: repeat change (@leq_universe ?cf) with (@compare_universe cf Cumul). - 3: destruct pb. - all: try (hnf; intros *; eapply (@cmp_universe_config_impl cf1 cf2); eassumption). - all: now etransitivity; [ | hnf; intros *; eapply (@cmp_universe_config_impl cf1 cf2); eassumption ]; tc. - - eapply cumul_Sort. eapply (@cmp_universe_config_impl cf1 cf2); eassumption. - - eapply cumul_Const. eapply R_universe_instance_impl'; eauto; tc. + eapply @cmp_global_instance_impl; [ .. | eassumption ]. + 3: auto with arith. all: intros ??; now apply (@cmp_universe_config_impl cf1 cf2). + - eapply cumul_Sort. eapply (@cmp_sort_config_impl cf1 cf2); eassumption. + - eapply cumul_Const. eapply cmp_universe_instance_impl'; eauto; tc. hnf; intros *; eapply (@cmp_universe_config_impl cf1 cf2); eassumption. Defined. diff --git a/pcuic/theories/Conversion/PCUICWeakeningConv.v b/pcuic/theories/Conversion/PCUICWeakeningConv.v index a1db2d285..aed3867e3 100644 --- a/pcuic/theories/Conversion/PCUICWeakeningConv.v +++ b/pcuic/theories/Conversion/PCUICWeakeningConv.v @@ -340,7 +340,7 @@ Qed. Lemma isType_on_free_vars {cf} {Σ : global_env_ext} {wfΣ : wf Σ} {Γ T} : isType Σ Γ T -> on_free_vars xpredT T. Proof. - intros [s Hs]. + intros (_ & s & Hs & _). eapply subject_closed in Hs. rewrite closedP_on_free_vars in Hs. eapply on_free_vars_impl; tea => //. @@ -349,7 +349,7 @@ Qed. Lemma isType_on_ctx_free_vars {cf} {Σ : global_env_ext} {wfΣ : wf Σ} {Γ T} : isType Σ Γ T -> on_ctx_free_vars xpredT Γ. Proof. - intros [s Hs]. + intros (_ & s & Hs & _). eapply typing_wf_local in Hs. eapply closed_wf_local in Hs; tea. eapply (closed_ctx_on_free_vars xpredT) in Hs. diff --git a/pcuic/theories/Conversion/PCUICWeakeningEnvConv.v b/pcuic/theories/Conversion/PCUICWeakeningEnvConv.v index 1f2ea9c14..335ef405c 100644 --- a/pcuic/theories/Conversion/PCUICWeakeningEnvConv.v +++ b/pcuic/theories/Conversion/PCUICWeakeningEnvConv.v @@ -14,15 +14,11 @@ Implicit Types (cf : checker_flags). Lemma compare_term_subset {cf} pb Σ φ φ' t t' : ConstraintSet.Subset φ φ' - -> compare_term pb Σ φ t t' -> compare_term pb Σ φ' t t'. + -> compare_term Σ φ pb t t' -> compare_term Σ φ' pb t t'. Proof. intro H. apply eq_term_upto_univ_impl; auto. - all: change eq_universe with (compare_universe Conv). - all: change leq_universe with (compare_universe Cumul). - 3: destruct pb. - 4: transitivity (compare_universe Cumul φ). - 4: tc. - all: intros ??; now eapply cmp_universe_subset. + 1,2: intros ??; now eapply cmp_universe_subset. + 1,2: intros ??; now eapply cmp_sort_subset. Qed. Lemma eq_term_subset {cf} Σ φ φ' t t' @@ -35,14 +31,14 @@ Proof. apply compare_term_subset with (pb := Cumul). Qed. Lemma compare_decl_subset {cf} pb Σ φ φ' d d' : ConstraintSet.Subset φ φ' - -> compare_decl pb Σ φ d d' -> compare_decl pb Σ φ' d d'. + -> compare_decl Σ φ pb d d' -> compare_decl Σ φ' pb d d'. Proof. intros Hφ []; constructor; eauto using compare_term_subset. Qed. Lemma compare_context_subset {cf} pb Σ φ φ' Γ Γ' : ConstraintSet.Subset φ φ' - -> compare_context pb Σ φ Γ Γ' -> compare_context pb Σ φ' Γ Γ'. + -> compare_context Σ φ pb Γ Γ' -> compare_context Σ φ' pb Γ Γ'. Proof. intros Hφ. induction 1; constructor; auto; eapply compare_decl_subset; eassumption. Qed. @@ -50,81 +46,89 @@ Qed. Section ExtendsWf. Context {cf : checker_flags}. Context {Pcmp: global_env_ext -> context -> conv_pb -> term -> term -> Type}. - Context {P: global_env_ext -> context -> term -> typ_or_sort -> Type}. + Context {P: global_env_ext -> context -> judgment -> Type}. Let wf := on_global_env Pcmp P. -Lemma global_variance_sigma_mon {Σ Σ' gr napp v} : +Lemma global_variance_sigma_mon {Σ Σ'} gr napp : wf Σ' -> extends Σ Σ' -> - global_variance Σ gr napp = Some v -> - global_variance Σ' gr napp = Some v. + match global_variance Σ gr napp with + | Variance v => global_variance Σ' gr napp = Variance v + | AllEqual => True + | AllIrrelevant => global_variance Σ' gr napp = AllIrrelevant + end. Proof using P Pcmp cf. intros wfΣ' ext. rewrite /global_variance_gen /lookup_constructor /lookup_constructor_gen /lookup_inductive /lookup_inductive_gen /lookup_minductive /lookup_minductive_gen. - destruct gr as [|ind|[ind i]|] => /= //. + destruct gr as [|ind|[ind i]|] => //=. - destruct (lookup_env Σ ind) eqn:look => //. eapply extends_lookup in look; eauto. rewrite look //. + destruct g => //=; destruct nth_error => //=. + destruct destArity as [[ctx s]|] => //=. + destruct Nat.leb => //=. + destruct ind_variance => //=. - destruct (lookup_env Σ (inductive_mind i)) eqn:look => //. eapply extends_lookup in look; eauto. rewrite look //. + destruct g => //=; destruct nth_error => //=; destruct nth_error => //=. + destruct Nat.leb => //=. Qed. -(** The definition of [R_global_instance] is defined so that it is weakenable. *) -Lemma R_global_instance_weaken_env Σ Σ' Re Re' Rle Rle' gr napp : +(** The definition of [cmp_global_instance] is defined so that it is weakenable. *) +Lemma cmp_global_instance_weaken_env Σ Σ' cmp_universe cmp_universe' pb pb' gr napp : wf Σ' -> extends Σ Σ' -> - RelationClasses.subrelation Re Re' -> - RelationClasses.subrelation Rle Rle' -> - RelationClasses.subrelation Re Rle' -> - subrelation (R_global_instance Σ Re Rle gr napp) (R_global_instance Σ' Re' Rle' gr napp). + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe' Conv) -> + RelationClasses.subrelation (cmp_universe pb) (cmp_universe' pb') -> + subrelation (cmp_global_instance Σ cmp_universe pb gr napp) (cmp_global_instance Σ' cmp_universe' pb' gr napp). Proof using P Pcmp cf. - intros wfΣ ext he hle hele t t'. - rewrite /R_global_instance_gen /R_opt_variance. - destruct global_variance_gen as [v|] eqn:look. - - rewrite (global_variance_sigma_mon wfΣ ext look). - induction t in v, t' |- *; destruct v, t'; simpl; auto. - intros []; split; auto. - destruct t0; simpl; auto. - - destruct (global_variance Σ' gr napp) => //. - * induction t in l, t' |- *; destruct l, t'; simpl; intros H; inv H; auto. - split; auto. destruct t0; simpl; auto. - * eauto using R_universe_instance_impl'. + intros wfΣ ext sub_conv sub_pb t t'. + unfold cmp_global_instance_gen. + move: (global_variance_sigma_mon gr napp wfΣ ext). + destruct global_variance_gen as [| |v] => //. + all: [> intros _ | intros ->; cbn ..]; auto. + 1: intro H; apply: cmp_instance_opt_variance; move: H => /=. + - now apply cmp_universe_instance_impl. + - intros [H | H]; [left | right]. + 1: eapply cmp_universe_instance_impl; tea. + eapply cmp_universe_instance_variance_impl; eassumption. Qed. #[global] -Instance eq_term_upto_univ_weaken_env Σ Σ' Re Re' Rle Rle' napp : +Instance eq_term_upto_univ_weaken_env Σ Σ' cmp_universe cmp_universe' cmp_sort cmp_sort' pb pb' napp : wf Σ' -> extends Σ Σ' -> - RelationClasses.subrelation Re Re' -> - RelationClasses.subrelation Rle Rle' -> - RelationClasses.subrelation Re Rle' -> - CRelationClasses.subrelation (eq_term_upto_univ_napp Σ Re Rle napp) - (eq_term_upto_univ_napp Σ' Re' Rle' napp). + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe' Conv) -> + RelationClasses.subrelation (cmp_universe pb) (cmp_universe' pb') -> + RelationClasses.subrelation (cmp_sort Conv) (cmp_sort' Conv) -> + RelationClasses.subrelation (cmp_sort pb) (cmp_sort' pb') -> + CRelationClasses.subrelation (eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp) + (eq_term_upto_univ_napp Σ' cmp_universe' cmp_sort' pb' napp). Proof using P Pcmp cf. - intros wfΣ ext he hele hle t t'. - induction t in napp, t', Rle, Rle', hle, hele |- * using PCUICInduction.term_forall_list_ind; + intros wfΣ ext univ_sub_conv univ_sub_pb sort_sub_conv sort_sub_pb t t'. + induction t in napp, t', pb, pb', univ_sub_pb, sort_sub_pb, t' |- * using PCUICInduction.term_forall_list_ind; try (inversion 1; subst; constructor; - eauto using R_universe_instance_impl'; fail). + eauto using cmp_universe_instance_impl'; fail). - inversion 1; subst; constructor. eapply All2_impl'; tea. eapply All_impl; eauto. - inversion 1; subst; constructor. - eapply R_global_instance_weaken_env. 6:eauto. all:eauto. + eapply cmp_global_instance_weaken_env. 5:eauto. all:eauto. - inversion 1; subst; constructor. - eapply R_global_instance_weaken_env. 6:eauto. all:eauto. + eapply cmp_global_instance_weaken_env. 5:eauto. all:eauto. - inversion 1; subst; destruct X as [? [? ?]]; constructor; eauto. * destruct X2 as [? [? ?]]. constructor; intuition auto; solve_all. - + eauto using R_universe_instance_impl'. + + eauto using cmp_universe_instance_impl'. * eapply All2_impl'; tea. eapply All_impl; eauto. cbn. intros x [? ?] y [? ?]. split; eauto. - inversion 1; subst; constructor. eapply All2_impl'; tea. eapply All_impl; eauto. - cbn. intros x [? ?] y [[[? ?] ?] ?]. repeat split; eauto. + cbn. intros x [? ?] y (?&?&?&?). repeat split; eauto. - inversion 1; subst; constructor. eapply All2_impl'; tea. eapply All_impl; eauto. - cbn. intros x [? ?] y [[[? ?] ?] ?]. repeat split; eauto. + cbn. intros x [? ?] y (?&?&?&?). repeat split; eauto. - inversion 1; subst; constructor. depelim X1; constructor; cbn in X; intuition eauto. solve_all. Qed. @@ -148,7 +152,7 @@ Proof using P Pcmp. intros wfΣ ext. induction 1; simpl. - econstructor. eapply compare_term_subset. - + now eapply global_ext_constraints_app. + + eapply global_ext_constraints_app. apply ext. + simpl in *. eapply eq_term_upto_univ_weaken_env in c; simpl; eauto. all:typeclasses eauto. - econstructor 2; eauto. eapply weakening_env_red1; eauto. @@ -177,13 +181,7 @@ Lemma weakening_env_cumulSpec0 Σ Σ' φ Γ pb M N : cumulSpec0 (Σ', φ) Γ pb M N. Proof. intros HΣ' Hextends Ind. - pose proof (subrelations_leq_extends _ _ φ Hextends). revert H. - assert (RelationClasses.subrelation - (eq_universe (global_ext_constraints (Σ,φ))) - (leq_universe (global_ext_constraints (Σ',φ)))). - { typeclasses eauto. } revert H. - generalize (leq_universe (global_ext_constraints (Σ',φ))); intros Rle Hlee Hle . - revert pb Γ M N Ind Σ' Rle Hle Hlee HΣ' Hextends. + revert pb Γ M N Ind Σ' HΣ' Hextends. induction 1. all:intros; try solve [econstructor; eauto with extends; intuition auto]. all: lazymatch goal with @@ -192,24 +190,24 @@ Proof. end. - eapply cumul_Evar. solve_all. - eapply cumul_Case. - * cbv [cumul_predicate] in *; destruct_head'_prod. repeat split; tas. + * cbv [cumul_predicate] in *; destruct_head'_prod. clear c0. repeat split; tas. + solve_all. - + eapply R_universe_instance_impl'; try apply subrelations_extends; eassumption. + + eapply cmp_universe_instance_impl'; tea. tc. + eauto. - * solve_all. - * solve_all. - - eapply cumul_Fix; solve_all. - - eapply cumul_CoFix; solve_all. + * eauto. + * unfold cumul_branches, cumul_branch in *. solve_all. + - eapply cumul_Fix; unfold cumul_mfixpoint in *; solve_all. + - eapply cumul_CoFix; unfold cumul_mfixpoint in *; solve_all. - eapply cumul_Prim; solve_all. destruct X; constructor; eauto. - * unfold compare_universe. eapply subrel_extends_eq; tea. + * eapply subrel_extends_cmp_universe; tea. * solve_all. - eapply cumul_Ind; eauto. 2:solve_all. - eapply @R_global_instance_weaken_env. 1,2,6:eauto. all: tc. + eapply @cmp_global_instance_weaken_env. 1,2,5:eauto. all: tc. - eapply cumul_Construct; eauto. 2:solve_all. - eapply @R_global_instance_weaken_env. 1,2,6:eauto. all: tc. + eapply @cmp_global_instance_weaken_env. 1,2,5:eauto. all: tc. - eapply cumul_Sort. eapply subrelations_compare_extends; tea. - - eapply cumul_Const. eapply R_universe_instance_impl'; eauto; tc. + - eapply cumul_Const. eapply cmp_universe_instance_impl'; eauto; tc. Defined. diff --git a/pcuic/theories/PCUICAlpha.v b/pcuic/theories/PCUICAlpha.v index b82cc459b..6a8a6efd0 100644 --- a/pcuic/theories/PCUICAlpha.v +++ b/pcuic/theories/PCUICAlpha.v @@ -18,8 +18,8 @@ Implicit Types cf : checker_flags. Notation "`≡α`" := upto_names. Infix "≡α" := upto_names (at level 60). -Notation "`≡Γ`" := (eq_context_upto empty_global_env eq eq). -Infix "≡Γ" := (eq_context_upto empty_global_env eq eq) (at level 20, no associativity). +Notation "`≡Γ`" := (eq_context_upto empty_global_env (fun _ => eq) (fun _ => eq) Conv). +Infix "≡Γ" := (eq_context_upto empty_global_env (fun _ => eq) (fun _ => eq) Conv) (at level 20, no associativity). #[global] Instance upto_names_terms_refl : CRelationClasses.Reflexive (All2 `≡α`). @@ -27,11 +27,11 @@ Proof. intro; apply All2_refl; reflexivity. Qed. Lemma eq_context_upto_empty_impl {cf} {Σ : global_env_ext} ctx ctx' : ctx ≡Γ ctx' -> - eq_context_upto Σ (eq_universe Σ) (eq_universe Σ) ctx ctx'. + eq_context Σ Σ ctx ctx'. Proof. intros; eapply All2_fold_impl; tea. intros ???? []; constructor; subst; auto; - eapply eq_term_upto_univ_empty_impl; tea; tc. + eapply upto_names_impl; tea; tc. Qed. Section Alpha. @@ -43,26 +43,16 @@ Section Alpha. wf Σ.1 -> wf_local Σ Γ -> nth_error Γ i = Some (vass na ty) -> - lift_typing typing Σ Γ (lift0 (S i) ty) Sort. + lift_typing typing Σ Γ (Typ (lift0 (S i) ty)). Proof using Type. intros Σ Γ i na ty hΣ hΓ e. simpl. - induction i in Γ, hΓ, e |- *. - - destruct Γ. 1: discriminate. - simpl in e. apply some_inj in e. subst. - inversion hΓ. subst. - apply infer_typing_sort_impl with id X0; intros Hs. - change (tSort _) with (lift0 1 (tSort X0.π1)). - eapply weakening with (Γ' := [ vass na ty ]). - all: assumption. - - destruct Γ. 1: discriminate. - simpl in e. - specialize IHi with (2 := e). - forward IHi by inversion hΓ; tas. - apply infer_typing_sort_impl with id IHi; intros Hs. - change (tSort _) with (lift0 1 (lift0 (S i) (tSort IHi.π1))). - rewrite simpl_lift0. - eapply weakening with (Γ' := [ c ]). - all: assumption. + eapply All_local_env_nth_error in e as hj; tea. + apply nth_error_Some_length in e. + rewrite -(firstn_skipn (S i) Γ) in hΓ |- *. + apply lift_typing_f_impl with (1 := hj) => // ?? HT. + eapply weakening with (Γ' := firstn (S i) Γ) in HT. + all: tas. + rewrite firstn_length_le // in HT. Qed. (* TODO MOVE *) @@ -81,25 +71,13 @@ Section Alpha. eauto using nth_error_Some_length. Qed. - (* TODO MOVE *) - Lemma nth_error_weak_context : - forall Γ Δ i d, - nth_error Δ i = Some d -> - nth_error (Γ ,,, Δ) i = Some d. - Proof using Type. - intros Γ Δ i d h. - rewrite -> nth_error_app_context_lt. - - assumption. - - apply nth_error_Some_length in h. assumption. - Qed. - - Lemma decompose_app_upto {Σ Re Rle x y hd tl} : - eq_term_upto_univ Σ Re Rle x y -> + Lemma decompose_app_upto {Σ cmp_universe cmp_sort pb x y hd tl} : + eq_term_upto_univ Σ cmp_universe cmp_sort pb x y -> decompose_app x = (hd, tl) -> - ∑ hd' tl', (y = mkApps hd' tl') * - eq_term_upto_univ_napp Σ Re Rle #|tl| hd hd' * - negb (isApp hd') * - All2 (eq_term_upto_univ Σ Re Re) tl tl'. + ∑ hd' tl', y = mkApps hd' tl' × + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb #|tl| hd hd' × + negb (isApp hd') × + All2 (eq_term_upto_univ Σ cmp_universe cmp_sort Conv) tl tl'. Proof using Type. intros eq dapp. pose proof (decompose_app_notApp _ _ _ dapp). @@ -133,9 +111,9 @@ Section Alpha. destruct t; congruence. Qed. - Lemma upto_names_destInd Re Rle t u : - eq_term_upto_univ empty_global_env Re Rle t u -> - rel_option (fun '(ind, u) '(ind', u') => (ind = ind') * R_universe_instance Re u u')%type (destInd t) (destInd u). + Lemma upto_names_destInd cmp_universe cmp_sort pb napp t u : + eq_term_upto_univ_napp empty_global_env cmp_universe cmp_sort pb napp t u -> + rel_option (fun '(ind, u) '(ind', u') => (ind = ind') * cmp_universe_instance (cmp_universe Conv) u u')%type (destInd t) (destInd u). Proof using Type. induction 1; simpl; constructor; try congruence. split; auto. @@ -159,20 +137,17 @@ Section Alpha. destruct X0 as [eqctx eqt]. apply (eq_context_upto_smash_context empty_global_env [] []) in eqctx; try constructor. apply eq_context_upto_rev' in eqctx. - eapply (eq_context_upto_nth_error empty_global_env _ _ _ _ rarg) in eqctx. + eapply (eq_context_upto_nth_error empty_global_env _ _ _ _ _ rarg) in eqctx. subst rarg'. destruct (nth_error (List.rev (smash_context [] c)) rarg). - inv eqctx. destruct X0. + all: inv eqctx => //. destruct X0. destruct (decompose_app) eqn:eqdec. - destruct (decompose_app_upto e eqdec) as [hd' [tl' [[[eq eqhd] napp] eqtl]]]. + destruct (decompose_app_upto e eqdec) as (hd' & tl' & eq & eqhd & napp & eqtl). rewrite eq. rewrite decompose_app_mkApps; auto. - eapply (eq_term_upto_univ_empty_impl _ Logic.eq Logic.eq Logic.eq Logic.eq #|l0| 0) in eqhd. - all:try typeclasses eauto. apply upto_names_destInd in eqhd. inv eqhd; auto. destruct a as [ind u], b0 as [ind' u']; simpl in *; auto. destruct X0 as [-> _]; auto. - now inv eqctx. Qed. Lemma upto_names_check_cofix mfix mfix' : @@ -191,10 +166,8 @@ Section Alpha. do 2 destruct decompose_prod_assum. destruct X as [_ eqt]. destruct (decompose_app) eqn:eqdec. - destruct (decompose_app_upto eqt eqdec) as [hd' [tl' [[[eq eqhd] napp] eqtl]]]. + destruct (decompose_app_upto eqt eqdec) as (hd' & tl' & eq & eqhd & napp & eqtl). rewrite eq. rewrite decompose_app_mkApps; auto. - eapply (eq_term_upto_univ_empty_impl _ Logic.eq Logic.eq Logic.eq Logic.eq #|l0| 0) in eqhd. - all:try typeclasses eauto. apply upto_names_destInd in eqhd. inv eqhd; auto. destruct a as [ind u], b as [ind' u']; simpl in *; auto. @@ -253,7 +226,7 @@ Section Alpha. Lemma All_decls_alpha_pb_ws_decl {le P} {Γ : context} {d d'} : (forall le t t', is_open_term Γ t -> is_open_term Γ t' -> upto_names' t t' -> P le t t') -> - compare_decls upto_names' upto_names' d d' -> + compare_decls (fun pb => eq_term_upto_univ empty_global_env (fun _ => eq) (fun _ => eq) pb) Conv d d' -> ws_decl Γ d -> ws_decl Γ d' -> All_decls_alpha_pb le P d d'. @@ -297,82 +270,36 @@ Section Alpha. Qed. - Lemma eq_context_upto_map2_set_binder_name Σ pctx pctx' Γ Δ : - pctx ≡Γ pctx' -> - eq_context_upto Σ eq eq Γ Δ -> - eq_context_upto Σ eq eq - (map2 set_binder_name (forget_types pctx) Γ) - (map2 set_binder_name (forget_types pctx') Δ). - Proof using Type. - intros eqp. - induction 1 in pctx, pctx', eqp |- *. - - induction eqp; cbn; constructor. - - depelim eqp. simpl. constructor. - simpl. constructor; auto. - destruct c, p; constructor; auto. - Qed. - - Lemma eq_context_upto_lift_context Σ Re Rle : - RelationClasses.subrelation Re Rle -> - forall u v n k, - eq_context_upto Σ Re Rle u v -> - eq_context_upto Σ Re Rle (lift_context n k u) (lift_context n k v). - Proof using Type. - intros re u v n k. - induction 1. - - constructor. - - rewrite !lift_context_snoc; constructor; eauto. - depelim p; constructor; simpl; intuition auto; - rewrite -(length_of X); - apply eq_term_upto_univ_lift; auto. - Qed. - Lemma eq_context_upto_subst_instance Σ : - forall u v i, + forall pb u v i, valid_constraints (global_ext_constraints Σ) (subst_instance_cstrs i Σ) -> - eq_context_upto Σ eq eq u v -> - eq_context_upto Σ eq eq (subst_instance i u) (subst_instance i v). + eq_context_upto Σ (fun _ => eq) (fun _ => eq) pb u v -> + eq_context_upto Σ (fun _ => eq) (fun _ => eq) pb (subst_instance i u) (subst_instance i v). Proof using Type. - intros u v i vc. - induction 1. - - constructor. - - rewrite !PCUICUnivSubst.subst_instance_cons. constructor; eauto. - depelim p; constructor; simpl; intuition auto. - eapply (PCUICUnivSubstitutionConv.eq_term_upto_univ_subst_preserved Σ (fun _ => eq) (fun _ => eq)). - intros x y u v ? ? ->; reflexivity. - intros x y u v ? ? ->; reflexivity. exact vc. - assumption. - eapply (PCUICUnivSubstitutionConv.eq_term_upto_univ_subst_preserved Σ (fun _ => eq) (fun _ => eq)). - intros x y u v ? ? ->; reflexivity. - intros x y u v ? ? ->; reflexivity. exact vc. - assumption. - eapply (PCUICUnivSubstitutionConv.eq_term_upto_univ_subst_preserved Σ (fun _ => eq) (fun _ => eq)). - intros x y u v ? ? ->; reflexivity. - intros x y u v ? ? ->; reflexivity. exact vc. - assumption. + intros pb u v i vc. + eapply PCUICUnivSubstitutionConv.eq_context_upto_subst_preserved with (cmp_universe := fun _ _ => eq) (cmp_sort := fun _ _ => eq). + 5: eassumption. + all: intros ??????[]; reflexivity. Qed. - Lemma eq_context_gen_upto ctx ctx' : - eq_context_gen eq eq ctx ctx' -> - eq_context_upto empty_global_env eq eq ctx ctx'. + Lemma eq_context_upto_names_eq_context_alpha ctx ctx' : + eq_context_upto_names ctx ctx' -> + ctx ≡Γ ctx'. Proof using Type. - intros a; eapply All2_fold_impl; tea. - intros. destruct X; subst; constructor; auto; try reflexivity. + apply eq_context_upto_names_eq_context_upto; tc. Qed. Lemma case_predicate_context_equiv {ind mdecl idecl p p'} : eq_predicate upto_names' eq p p' -> - eq_context_upto empty_global_env eq eq - (case_predicate_context ind mdecl idecl p) - (case_predicate_context ind mdecl idecl p'). + case_predicate_context ind mdecl idecl p ≡Γ case_predicate_context ind mdecl idecl p'. Proof using Type. intros [eqpars [eqinst [eqctx eqret]]]. rewrite /case_predicate_context /case_predicate_context_gen. - apply eq_context_upto_map2_set_binder_name => //. - now eapply eq_context_gen_upto. + eapply eq_context_gen_map2_set_binder_name => //. + now eapply eq_context_upto_names_eq_context_alpha. rewrite /pre_case_predicate_context_gen. - eapply R_universe_instance_eq in eqinst. rewrite -eqinst. + eapply cmp_universe_instance_eq in eqinst. rewrite -eqinst. apply eq_context_upto_subst_context; tea; tc. reflexivity. now apply All2_rev. @@ -385,9 +312,9 @@ Section Alpha. (case_branch_context ind mdecl p' (forget_types bctx') cdecl). Proof using Type. intros [eqpars [eqinst [eqctx eqret]]] eqctx'. - eapply R_universe_instance_eq in eqinst. + eapply cmp_universe_instance_eq in eqinst. rewrite /case_branch_context /case_branch_context_gen -eqinst. - apply eq_context_upto_map2_set_binder_name => //. + eapply eq_context_gen_map2_set_binder_name => //; tea. rewrite /pre_case_branch_context_gen. apply eq_context_upto_subst_context; tea; tc. reflexivity. @@ -405,7 +332,7 @@ Section Alpha. (case_branch_type ind mdecl idecl p' br' ptm' c cdecl).2. Proof using Type. intros [eqpars [eqinst [eqctx eqret]]] eqctx'. - eapply R_universe_instance_eq in eqinst. + eapply cmp_universe_instance_eq in eqinst. intros ptm ptm'. rewrite /case_branch_type /case_branch_type_gen -eqinst. cbn. eapply eq_term_mkApps. @@ -417,67 +344,46 @@ Section Alpha. - eapply All2_app. * eapply All2_map, All2_refl. intros. - eapply eq_term_upto_univ_empty_impl; tea; tc. - eapply eq_term_upto_univ_substs. tc. + eapply eq_term_upto_univ_empty_impl; cycle -1. + eapply eq_term_upto_univ_substs; tc. reflexivity. now eapply All2_rev. + all: tc. * constructor. 2:constructor. - eapply eq_term_upto_univ_empty_impl; tea; tc. - eapply eq_term_upto_univ_mkApps. len. - reflexivity. + eapply eq_term_upto_univ_empty_impl with (pb := Conv); cycle -1. + eapply eq_term_upto_univ_mkApps; cycle -1. eapply All2_app. - + eapply All2_map. eapply (All2_impl eqpars). - intros. now eapply eq_term_upto_univ_lift. - + eapply All2_refl. reflexivity. - Qed. + eapply All2_map. eapply (All2_impl eqpars). + intros. now eapply eq_term_upto_univ_lift. - Lemma eq_binder_annots_eq_context_gen_ctx {Δ : context} {nas} : - All2 (fun x y => eq_binder_annot x y.(decl_name)) nas Δ -> - eq_context_gen eq eq (map2 set_binder_name nas Δ) Δ. - Proof using Type. - induction Δ in nas |- * using PCUICInduction.ctx_length_rev_ind; simpl; intros hlen. - - depelim hlen. simpl. reflexivity. - - destruct nas as [|nas na] using rev_case => //; - pose proof (All2_length hlen) as hlen';len in hlen'; simpl in hlen'; try lia. - eapply All2_app_inv_l in hlen as (l1'&l2'&heq&alnas&allna). - depelim allna. depelim allna. - rewrite map2_app => /= //; try lia. unfold aname. lia. - eapply app_inj_tail in heq as [<- <-]. - simpl. eapply All2_fold_app; auto. - constructor. constructor. - destruct d as [na' [d|] ty]; constructor; cbn in *; auto; - try reflexivity. + eapply All2_refl. reflexivity. + + len. reflexivity. + + all: tc. Qed. Import PCUICUnivSubstitutionConv PCUICUnivSubstitutionTyp. Lemma inst_case_predicate_context_eq {Σ : global_env_ext} {wfΣ : wf Σ} {ind mdecl idecl p} : wf_predicate mdecl idecl p -> - eq_context_gen eq eq (ind_predicate_context ind mdecl idecl) (pcontext p) -> - eq_context_gen eq eq (case_predicate_context ind mdecl idecl p) (inst_case_predicate_context p). + eq_context_upto_names (ind_predicate_context ind mdecl idecl) (pcontext p) -> + eq_context_upto_names (case_predicate_context ind mdecl idecl p) (inst_case_predicate_context p). Proof using Type. intros wfp eq. rewrite /case_predicate_context /case_predicate_context_gen. epose proof (wf_pre_case_predicate_context_gen wfp). etransitivity. - now apply eq_binder_annots_eq_context_gen_ctx. + now apply eq_binder_annots_eq_ctx. rewrite /pre_case_predicate_context_gen /inst_case_predicate_context. rewrite /inst_case_context. - eapply eq_context_gen_subst_context. - now eapply eq_context_gen_eq_univ_subst_preserved. - Qed. - - Lemma ctx_inst_impl {Σ P Q Γ s Δ}: - (forall Σ Γ u U, P Σ Γ u U -> Q Σ Γ u U) -> - PCUICTyping.ctx_inst P Σ Γ s Δ -> PCUICTyping.ctx_inst Q Σ Γ s Δ. - Proof using Type. - intros HP. - induction 1; constructor; auto. + eapply eq_context_upto_names_subst_context. + now eapply eq_context_upto_names_univ_subst_preserved. Qed. - Lemma ctx_inst_eq_context {Σ P Q Γ Γ' s Δ}: - (forall Σ Γ Γ' u U, (Γ ≡Γ Γ') -> P Σ Γ u U -> Q Σ Γ' u U) -> + Lemma ctx_inst_eq_context {P Q Γ Γ' s Δ}: + (forall Γ Γ' u U, (Γ ≡Γ Γ') -> P Γ u U -> Q Γ' u U) -> Γ ≡Γ Γ' -> - PCUICTyping.ctx_inst P Σ Γ s Δ -> PCUICTyping.ctx_inst Q Σ Γ' s Δ. + PCUICTyping.ctx_inst P Γ s Δ -> PCUICTyping.ctx_inst Q Γ' s Δ. Proof using Type. intros HP e. induction 1; constructor; eauto. @@ -485,20 +391,17 @@ Section Alpha. Lemma wf_local_eq_context_upto_names {Σ : global_env_ext} {wfΣ : wf Σ} {Γ Δ Δ'} : wf_local Σ (Γ,,, Δ) -> - eq_context_gen eq eq Δ' Δ -> + eq_context_upto_names Δ' Δ -> wf_local Σ (Γ ,,, Δ'). Proof using Type. intros a eq. - eapply All2_fold_All2 in eq. induction eq; depelim a; cbn; try solve [constructor; auto]; depelim r; subst; constructor; auto. - 1,2: apply infer_typing_sort_impl with id l; intros Hs. - 3: rename l0 into Hs. + all: apply lift_typing_impl with (1 := l) => ?? Hs. all: eapply (closed_context_cumulativity _ (pb:=Conv)); tea; [apply IHeq; pcuic|]. all: eapply ws_cumul_ctx_pb_rel_app. all: eapply eq_context_upto_conv_context_rel; fvs. - all: eapply eq_context_gen_upto. - all: now eapply All2_fold_All2 in eq. + all: now eapply eq_context_upto_names_eq_context_alpha. Qed. Lemma case_branch_type_eq_context_gen_1 {ind mdecl idecl cdecl p n br} {ctx ctx' ret} : @@ -523,7 +426,7 @@ Section Alpha. len. eapply eq_term_upto_univ_lift. eapply eq_term_upto_univ_impl; revgoals. eapply eq_term_upto_univ_it_mkLambda_or_LetIn; tea. - 2:reflexivity. 2:lia. all:tc. + reflexivity. lia. all:tc. Qed. Lemma eq_context_conversion {Σ : global_env_ext} {wfΣ : wf Σ} {Γ Δ} {t T} : @@ -544,7 +447,9 @@ Section Alpha. Lemma upto_names_conv_context (Σ : global_env_ext) Γ Δ : Γ ≡Γ Δ -> conv_context cumulAlgo_gen Σ Γ Δ. Proof using Type. - eapply eq_context_upto_empty_conv_context. + intro H. + apply compare_context_cumul_pb_context. + now eapply eq_context_upto_empty_impl. Qed. Lemma isType_eq_context_conversion {Σ : global_env_ext} {wfΣ : wf Σ} {Γ Δ} {T} : @@ -553,45 +458,66 @@ Section Alpha. wf_local Σ Δ -> isType Σ Δ T. Proof using Type. - intros Hty eq wfΔ; apply infer_typing_sort_impl with id Hty; intros Hs. + intros Hty eq wfΔ; apply lift_typing_impl with (1 := Hty); intros ?? Hs. now eapply eq_context_conversion. Qed. + Lemma lift_judgment_alpha {Σ : global_env_ext} {Γ tm tm' t t' u} : + lift_typing0 (fun t T : term => + forall u : term, upto_names' t u -> Σ;;; Γ |- u : T) + (Judge tm t u) -> + match tm', tm with None, _ => unit | Some tm', Some tm => upto_names' tm tm' | _, _ => False end -> + upto_names' t t' -> + lift_typing typing Σ Γ (Judge tm' t' u). + Proof. + intros tu Htm Hty. + apply lift_sorting_it_impl_gen with tu => //. + 2: intro HT; now apply HT. + destruct tm' as [tm'|], tm as [tm|] => // HT. + specialize (HT _ Htm). + eapply type_Cumul'; tea. + { eapply lift_sorting_forget_univ, lift_sorting_it_impl_gen with tu => // Ht. now apply Ht. } + eapply eq_term_upto_univ_cumulSpec. + eapply eq_term_leq_term. + now eapply upto_names_impl_eq_term. + Qed. + Lemma typing_alpha_prop : env_prop (fun Σ Γ u A => forall Δ v, - eq_term_upto_univ empty_global_env eq eq u v -> + u ≡α v -> Γ ≡Γ Δ -> Σ ;;; Δ |- v : A) - (fun Σ Γ => forall Δ, Γ ≡Γ Δ -> wf_local Σ Δ). + (fun Σ Γ j => + forall Δ, + Γ ≡Γ Δ -> + lift_typing0 (fun t T => + forall u, t ≡α u -> + Σ ;;; Δ |- u : T) j) + (fun Σ Γ => forall Δ, Γ ≡Γ Δ -> wf_local Σ Δ). Proof using Type*. eapply typing_ind_env. + 1:{ + intros Σ wfΣ Γ j Hj Δ HΔ. + apply lift_typing_impl with (1 := Hj); intros ?? [_ IH]. + intros; now apply IH. + } all: intros Σ wfΣ Γ wfΓ. - - induction 1. + - intros _; clear wfΓ. induction 1 using All_local_env_ind1. * intros Δ eq; destruct Δ; depelim eq; constructor. * intros Δ eq; depelim eq. depelim c. - constructor; auto. - apply infer_typing_sort_impl with id tu; intros _. eapply Hs; auto. - * intros Δ eq; depelim eq. depelim c. - constructor; auto. - apply infer_typing_sort_impl with id tu; intros _. eapply Hs; auto. - red. - specialize (Hc _ _ e0 eq). - specialize (Hs _ _ e1 eq). - eapply type_Cumul'; tea. now exists tu.π1. - eapply eq_term_upto_univ_cumulSpec. - eapply eq_term_leq_term. - now eapply upto_names_impl_eq_term. + all: constructor; auto. + all: eapply lift_judgment_alpha with (1 := X _ eq) => //. - intros n decl hnth ih Δ v e eqctx; invs e. assert (isType Σ Γ (lift0 (S n) (decl_type decl))). { eapply validity. econstructor; eauto. } specialize (ih _ eqctx). - epose proof (eq_context_upto_nth_error _ _ _ _ _ _ eqctx). + epose proof (eq_context_upto_nth_error _ _ _ _ _ _ _ eqctx). erewrite hnth in X0. depelim X0. eapply type_Cumul. eapply type_Rel ; tea. - eapply eq_context_conversion; tea. eapply X.π2. + eapply eq_context_conversion; tea. eapply X.2.π2.1. depelim e. destruct p. eapply eq_term_upto_univ_cumulSpec, eq_term_leq_term, eq_term_upto_univ_lift. eapply upto_names_impl_eq_term. now symmetry. @@ -601,58 +527,49 @@ Section Alpha. eapply type_Sort; assumption. - intros na A B s1 s2 ih hA ihA hB ihB Δ v e eqctx; invs e. econstructor. - + eapply ihA. assumption. eauto. + + eapply lift_judgment_alpha with (1 := ihA _ eqctx) => //. + eapply context_conversion. * eapply ihB. assumption. reflexivity. * constructor; eauto. - simpl. eexists. eapply ihA; tea. + eapply lift_sorting_forget_univ. + eapply lift_judgment_alpha with (1 := ihA _ eqctx) => //. * constructor. - -- now eapply eq_context_upto_empty_conv_context. + -- now eapply upto_names_conv_context. -- constructor. assumption. constructor. eapply upto_names_impl_eq_term. assumption. - - intros na A t s1 B ih hA ihA hB ihB Δ v e e'; invs e. + - intros na A t B ih hA ihA hB ihB Δ v e eqctx; invs e. eapply type_Cumul'. + econstructor. - * eapply ihA; tea. + * eapply lift_judgment_alpha with (1 := ihA _ eqctx) => //. * eapply eq_context_conversion. -- eapply ihB. assumption. reflexivity. -- constructor. 1: assumption. simpl. constructor; auto. - -- constructor; eauto. exists s1. - now eapply ihA. + -- constructor; eauto. + eapply lift_judgment_alpha with (1 := ihA _ eqctx) => //. + eapply validity in hB;tea. eapply isType_tProd; eauto. split; eauto with pcuic. + eapply lift_judgment_alpha with (1 := ihA _ eqctx) => //. reflexivity. eapply validity. eapply ihB; eauto. constructor; auto. constructor ; auto. reflexivity. + apply eq_term_upto_univ_cumulSpec, eq_term_leq_term. symmetry. constructor; auto. all: try (eapply upto_names_impl_eq_term ; assumption). all: reflexivity. - - intros na b B t s1 A ih hB ihB hb ihb hA ihA Δ v e e'; invs e. - specialize (ihB _ _ X0 e'). - specialize (ihb _ _ X e'). + - intros na b B t A ih hbB ihbB hA ihA Δ v e eqctx; invs e. assert (isType Σ Γ (tLetIn na b B A)). { eapply validity. econstructor; eauto. } eapply type_Cumul'. + econstructor. - * eapply ihB; trea. - * econstructor. - -- eapply ihb; trea. - -- eapply ihB; trea. - -- eapply eq_term_upto_univ_cumulSpec, eq_term_leq_term. - eapply upto_names_impl_eq_term. assumption. + * eapply lift_judgment_alpha with (1 := ihbB _ eqctx) => //. * eapply eq_context_conversion. -- eapply ihA; trea. -- constructor. ++ assumption. ++ constructor; auto. -- constructor; auto. - ++ exists s1; eapply ihB; eauto. - ++ eapply type_Cumul'; tea. - exists s1. eapply ihB; eauto. - eapply eq_term_upto_univ_cumulSpec, eq_term_leq_term. - now apply upto_names_impl_eq_term. - + apply infer_typing_sort_impl with id X2; intros Hs. + eapply lift_judgment_alpha with (1 := ihbB _ eqctx) => //. + + apply lift_typing_impl with (1 := X2) => ?? Hs. eapply eq_context_conversion; tea. eauto. + eapply eq_term_upto_univ_cumulSpec, eq_term_leq_term. symmetry; constructor. assumption. @@ -668,33 +585,33 @@ Section Alpha. all:typeclasses eauto. * eapply ihu; trea. * eapply ihty. reflexivity. auto. - + apply infer_typing_sort_impl with id X1; intros Hs. eapply eq_context_conversion; tea. eauto. + + apply lift_typing_impl with (1 := X1) => ?? Hs. eapply eq_context_conversion; tea. eauto. + eapply eq_term_upto_univ_cumulSpec, eq_term_leq_term. symmetry. eapply upto_names_impl_eq_term. eapply eq_term_upto_univ_subst ; now auto. - intros cst u decl ? ? hdecl hcons Δ v e e'; invs e. - eapply R_universe_instance_eq in H2. subst. + eapply cmp_universe_instance_eq in H2. subst. constructor; eauto. - intros ind u mdecl idecl isdecl ? ? hcons Δ v e e'; invs e. - eapply R_universe_instance_eq in H2. subst. + eapply cmp_universe_instance_eq in H2. subst. econstructor ; eauto. - intros ind i u mdecl idecl cdecl isdecl ? ? ? Δ v e e'; invs e. - eapply R_universe_instance_eq in H4. subst. + eapply cmp_universe_instance_eq in H4. subst. econstructor ; eauto. - intros ind p c brs args ps mdecl idecl isdecl X X0 H Hpctx cpc wfp cup wfpctx Hret IHret wfcpc kelim HIHctxi Hc IHc iscof ptm wfbrs Hbrs Δ v e e'; invs e. have eqp := X1. - eassert (ctx_inst _ _ _ _) as Hctxi by now eapply ctx_inst_impl with (2 := HIHctxi). - eassert (PCUICEnvTyping.ctx_inst _ _ _ _ _) as IHctxi. - { eapply ctx_inst_impl with (2 := HIHctxi). intros ? ? ? ? [? r]; exact r. } + eassert (ctx_inst _ _ _ _) as Hctxi by now eapply ctx_inst_impl with (1 := HIHctxi). + eassert (PCUICEnvTyping.ctx_inst _ _ _ _) as IHctxi. + { eapply ctx_inst_impl with (1 := HIHctxi). intros ? ? [? r]. pattern Γ, t, T in r. exact r. } destruct X1 as [eqpars [eqinst [eqctx eqret]]]. assert (wf_predicate mdecl idecl p'). { destruct wfp. split; auto. { now rewrite <-(All2_length eqpars). } eapply Forall2_All2 in H1. eapply All2_Forall2. - eapply All2_fold_All2 in eqctx. eapply All2_sym in eqctx. + eapply All2_sym in eqctx. eapply (All2_trans' (@eq_binder_annot name name)); tea. 2:{ eapply All2_map; tea. eapply All2_impl; tea. simpl; intros. destruct X1; simpl; now symmetry. } @@ -702,7 +619,7 @@ Section Alpha. have wfcpc' := wfcpc (Δ ,,, case_predicate_context ind mdecl idecl p'). forward wfcpc'. { eapply eq_context_upto_cat; auto. apply (case_predicate_context_equiv eqp). } - eapply R_universe_instance_eq in eqinst. + eapply cmp_universe_instance_eq in eqinst. assert (isType Σ Δ (mkApps ptm (args ++ [c]))). { eapply isType_eq_context_conversion. eapply validity. econstructor; eauto. constructor; eauto. constructor; eauto. @@ -711,13 +628,12 @@ Section Alpha. + have cu' : consistent_instance_ext Σ (ind_universes mdecl) (puinst p'). { now rewrite -eqinst. } have convctx' : eq_context_upto_names (pcontext p') (ind_predicate_context ind mdecl idecl). - { etransitivity; tea. symmetry. now eapply All2_fold_All2 in eqctx. } - assert (eqp' : eq_context_gen eq eq (inst_case_predicate_context p') + { etransitivity; tea. now symmetry. } + assert (eqp' : eq_context_upto_names (inst_case_predicate_context p') (case_predicate_context ind mdecl idecl p')). { rewrite /inst_case_predicate_context. rewrite /case_predicate_context /case_predicate_context_gen in wfcpc. - symmetry. apply inst_case_predicate_context_eq; tea. - eapply All2_fold_All2. now symmetry. } + symmetry. apply inst_case_predicate_context_eq; tea. now symmetry. } assert (wf_local Σ (Δ,,, inst_case_predicate_context p')). { eapply wf_local_eq_context_upto_names. exact wfcpc'. assumption. } @@ -732,7 +648,7 @@ Section Alpha. destruct eqp. eapply PCUICSpine.ctx_inst_eq_context; tea. rewrite List.rev_involutive. - * eapply weaken_wf_local; tea. now eapply wf_local_app_inv in X4 as []. + * eapply weaken_wf_local; tea. now eapply All_local_env_app_inv in X4 as []. eapply (on_minductive_wf_params_indices_inst isdecl _ cup). * eapply ctx_inst_eq_context; tea. cbn; eauto. * eapply ctx_inst_eq_context; tea. cbn; intros; eauto. @@ -747,17 +663,16 @@ Section Alpha. red. do 2 red in wfbr. eapply Forall2_All2 in wfbr. eapply All2_Forall2. eapply All2_map_left. - eapply All2_fold_All2 in eqbrctx. eapply All2_map_left_inv in wfbr. eapply All2_trans'; tea. 2:{ eapply All2_symP; tea. tc. } intros ??? [[] ?]; try constructor; simpl; auto; now transitivity na'. } - destruct (wf_local_app_inv X4) as [wfΔ _]. + destruct (All_local_env_app_inv X4) as [wfΔ _]. assert (clΔ := (wf_local_closed_context wfΔ)). econstructor; tea; eauto. 2,3: constructor; tea ; eauto. * eapply (type_ws_cumul_pb (pb:=Cumul)). eapply IHc; eauto. - eexists; eapply isType_mkApps_Ind; tea. + eapply has_sort_isType; eapply isType_mkApps_Ind; tea. unshelve eapply (ctx_inst_spine_subst _ ctxinst'). eapply weaken_wf_local; tea. now eapply (on_minductive_wf_params_indices_inst isdecl). @@ -783,7 +698,7 @@ Section Alpha. 2:now eapply Forall2_All2 in wfbrs. epose proof (wf_case_branches_types' (p:=p') ps args brs' isdecl) as a. forward a. - { eexists; eapply isType_mkApps_Ind; tea. + { eapply has_sort_isType; eapply isType_mkApps_Ind; tea. unshelve eapply (ctx_inst_spine_subst _ ctxinst'). eapply weaken_wf_local; tea. eapply (on_minductive_wf_params_indices_inst isdecl _ cu'). } @@ -791,15 +706,15 @@ Section Alpha. forward a. { eapply eq_context_conversion; tea. eapply eq_context_upto_cat; auto. reflexivity. - eapply eq_context_gen_upto. now symmetry. } + eapply eq_context_upto_names_eq_context_alpha. now symmetry. } do 2 forward a by auto. eapply (All2i_All2_All2i_All2i Hbrs X3 a). intros n cdecl br br' [wfbr [wfbrctx wfbrty]]. destruct wfbrty as (IHbrctx & (Hbbody & IHbbody) & (Hbty & IHbty)). intros [eqbctx eqbodies] [wfbr' wfcpars wfcparsn wfcbctx Hbr'ty]. split; intuition auto. - etransitivity. symmetry. eapply All2_fold_All2. exact eqbctx. assumption. - eapply eq_context_gen_upto in eqbctx. + etransitivity. symmetry. exact eqbctx. assumption. + eapply eq_context_upto_names_eq_context_alpha in eqbctx. assert (cbreq := case_branch_context_equiv (ind:=ind) (mdecl:=mdecl) (cdecl:=cdecl) eqp eqbctx). rewrite case_branch_type_fst. intuition auto. @@ -835,37 +750,33 @@ Section Alpha. eapply All2_same. intro. eapply eq_term_upto_univ_refl ; auto. - - intros mfix n decl types hguard hnth hwf ihmfix ihmfixb wffix Δ v e e'; invs e. + - intros mfix n decl types hguard hnth hwf hmfix ihmfix hmfixb ihmfixb wffix Δ v e e'; invs e. eapply All2_nth_error_Some in hnth as hnth' ; eauto. destruct hnth' as [decl' [hnth' hh]]. - destruct hh as [[[ety eqann] ebo] era]. + destruct hh as (ety & eqann & ebo & era). assert (hwf' : wf_local Σ (Γ ,,, fix_context mfix')). { apply PCUICWeakeningTyp.All_mfix_wf; auto. eapply (All2_All_mix_left ihmfix) in X. clear -X. induction X; constructor; simpl; auto. - destruct r as [Hty [[[eqty eqann] eqbod] eqrarg]]. - apply infer_typing_sort_impl with id Hty; intros [Hs IH]. - apply IH; eauto. reflexivity. } + destruct r as [Hty (eqty & eqann & eqbod & eqrarg)]. + eapply lift_judgment_alpha with (1 := Hty _ ltac:(reflexivity)) => //. } assert (convctx : conv_context cumulAlgo_gen Σ (Γ ,,, fix_context mfix) (Γ ,,, fix_context mfix')). { eapply eq_context_upto_univ_conv_context. - eapply (eq_context_impl _ eq). intros x y eqx. subst. reflexivity. - 1-2:typeclasses eauto. + eapply eq_context_upto_cat. 1: reflexivity. + eapply eq_context_upto_empty_impl. change (fix_context mfix) with (fix_context_gen 0 mfix). change (fix_context mfix') with (fix_context_gen 0 mfix'). + generalize 0 at 2 3. + unfold fix_context_gen. + eapply (All2_All_mix_left ihmfix) in X. + clear -X. + induction X; try constructor; simpl; intros n; auto. + destruct r as [Hty (eqty & eqbod & eqrarg & eqann)]. eapply eq_context_upto_cat. - * apply eq_context_upto_refl; typeclasses eauto. - * generalize 0 at 3 4. - unfold fix_context_gen. - eapply (All2_All_mix_left ihmfix) in X. - clear -X. - induction X; try constructor; simpl; intros n; auto. - destruct r as [Hty [[[eqty eqbod] eqrarg] eqann]]. - eapply eq_context_upto_cat. - + constructor; constructor; auto. - eapply eq_term_upto_univ_empty_impl; eauto. - 4:now eapply eq_term_upto_univ_lift. all:tc. - + apply IHX. } + + constructor; constructor; auto. + now eapply eq_term_upto_univ_lift. + + apply IHX. } assert(#|fix_context mfix| = #|fix_context mfix'|). { now rewrite !fix_context_length (All2_length X). } specialize (hwf (Δ ,,, types)) as hwfΔ. @@ -873,7 +784,7 @@ Section Alpha. { apply eq_context_upto_cat; auto. reflexivity. } specialize (hwf (Γ ,,, types)). forward hwf. { apply eq_context_upto_cat; auto; reflexivity. } - eapply wf_local_app_inv in hwfΔ as []. + eapply All_local_env_app_inv in hwfΔ as []. eapply eq_context_conversion; tea. eapply type_Cumul'. + econstructor. @@ -884,71 +795,55 @@ Section Alpha. * eapply (All2_All_mix_left ihmfix) in X. clear -X. induction X; constructor; simpl; auto. - destruct r as [Hty [[[eqty eqann] eqbod] eqrarg]]. - apply infer_typing_sort_impl with id Hty; intros [Hs IH]. - apply IH; eauto. reflexivity. - * solve_all. - destruct b0 as [Hb IHb]. - specialize (IHb (Γ ,,, types) _ b2). - forward IHb. { apply eq_context_upto_cat; reflexivity. } + destruct r as [Hty (eqty & eqann & eqbod & eqrarg)]. + eapply lift_judgment_alpha with (1 := Hty _ ltac:(reflexivity)) => //. + * unfold eq_mfixpoint in *. solve_all. + specialize (b1 (Γ ,,, types)) as IHb. + forward IHb by eapply eq_context_upto_cat; reflexivity. + eapply @lift_judgment_alpha with (tm' := Some _) in IHb; tea. + 1: apply lift_typing_impl with (1 := IHb) => t T HT. + 2: { rewrite -H. apply eq_term_upto_univ_lift; assumption. } eapply context_conversion; eauto. - eapply (type_Cumul' (lift0 #|fix_context mfix| (dtype x))); auto. - apply infer_typing_sort_impl with id a1; intros [Hs IH]. - specialize (IH _ _ a0 e'). - rewrite <-H. - eapply (weakening _ _ _ _ (tSort _)); eauto. - eapply eq_context_conversion; tea. now symmetry. - eapply eq_term_upto_univ_cumulSpec. - rewrite <- H. - eapply eq_term_upto_univ_lift. - eapply eq_term_upto_univ_empty_impl. - 4: intuition eauto. - all: intros ? ? []; reflexivity. * revert wffix. unfold wf_fixpoint, wf_fixpoint_gen. move/andP => [] hm ho. apply/andP; split. - { solve_all. move: b b2. + { unfold eq_mfixpoint in *. solve_all. move: b0 a4. generalize (dbody x) (dbody y). clear=> t t' isL eq. destruct t => //. now depelim eq. } move: ho; enough (map check_one_fix mfix = map check_one_fix mfix') as ->; auto. - apply upto_names_check_fix. solve_all. - + eapply All_nth_error in ihmfix; tea. - now apply infer_typing_sort_impl with id ihmfix; intros []. + apply upto_names_check_fix. unfold eq_mfixpoint in *. solve_all. + + eapply All_nth_error in hmfix; tea. + apply eq_term_upto_univ_cumulSpec, eq_term_leq_term, upto_names_impl_eq_term. now symmetry. - - intros mfix n decl types hguard hnth hwf ihmfix ihmfixb wffix Δ v e e'; invs e. + - intros mfix n decl types hguard hnth hwf hmfix ihmfix hmfixb ihmfixb wffix Δ v e e'; invs e. eapply All2_nth_error_Some in hnth as hnth' ; eauto. destruct hnth' as [decl' [hnth' hh]]. - destruct hh as [[[ety eqann] ebo] era]. + destruct hh as (ety & eqann & ebo & era). assert (hwf' : wf_local Σ (Γ ,,, fix_context mfix')). { apply PCUICWeakeningTyp.All_mfix_wf; auto. eapply (All2_All_mix_left ihmfix) in X. clear -X. induction X; constructor; simpl; auto. - destruct r as [Hty [[[eqty eqann] eqbod] eqrarg]]. - apply infer_typing_sort_impl with id Hty; intros [Hs IH]. - apply IH; eauto. reflexivity. } + destruct r as [Hty (eqty & eqann & eqbod & eqrarg)]. + eapply lift_judgment_alpha with (1 := Hty _ ltac:(reflexivity)) => //. } assert (convctx : conv_context cumulAlgo_gen Σ (Γ ,,, fix_context mfix) (Γ ,,, fix_context mfix')). { eapply eq_context_upto_univ_conv_context. - eapply (eq_context_impl _ eq). intros x y eqx. subst. reflexivity. - 1-2:typeclasses eauto. - change (fix_context mfix) with (fix_context_gen 0 mfix). - change (fix_context mfix') with (fix_context_gen 0 mfix'). - eapply eq_context_upto_cat. - * apply eq_context_upto_refl; typeclasses eauto. - * generalize 0 at 3 4. + eapply eq_context_upto_cat. 1: reflexivity. + eapply eq_context_upto_empty_impl. + change (fix_context mfix) with (fix_context_gen 0 mfix). + change (fix_context mfix') with (fix_context_gen 0 mfix'). + generalize 0 at 2 3. unfold fix_context_gen. eapply (All2_All_mix_left ihmfix) in X. clear -X. induction X; try constructor; simpl; intros n; auto. - destruct r as [Hty [[[eqty eqbod] eqrarg] eqann]]. + destruct r as [Hty (eqty & eqbod & eqrarg & eqann)]. eapply eq_context_upto_cat. + constructor; constructor; auto. - eapply eq_term_upto_univ_empty_impl; eauto. - 4:now eapply eq_term_upto_univ_lift. all:tc. + now eapply eq_term_upto_univ_lift. + apply IHX. } assert(#|fix_context mfix| = #|fix_context mfix'|). { now rewrite !fix_context_length (All2_length X). } @@ -957,7 +852,7 @@ Section Alpha. { apply eq_context_upto_cat; auto. reflexivity. } specialize (hwf (Γ ,,, types)). forward hwf. { apply eq_context_upto_cat; auto; reflexivity. } - eapply wf_local_app_inv in hwfΔ as []. + eapply All_local_env_app_inv in hwfΔ as []. eapply eq_context_conversion; tea. eapply type_Cumul'. + econstructor. @@ -968,36 +863,25 @@ Section Alpha. * eapply (All2_All_mix_left ihmfix) in X. clear -X. induction X; constructor; simpl; auto. - destruct r as [Hty [[[eqty eqann] eqbod] eqrarg]]. - apply infer_typing_sort_impl with id Hty; intros [Hs IH]. - apply IH; eauto. reflexivity. - * solve_all. - destruct b0 as [Hb IHb]. - specialize (IHb (Γ ,,, types) _ b2). - forward IHb. { apply eq_context_upto_cat; reflexivity. } + destruct r as [Hty (eqty & eqann & eqbod & eqrarg)]. + eapply lift_judgment_alpha with (1 := Hty _ ltac:(reflexivity)) => //. + * unfold eq_mfixpoint in *. solve_all. + specialize (b1 (Γ ,,, types)) as IHb. + forward IHb by eapply eq_context_upto_cat; reflexivity. + eapply @lift_judgment_alpha with (tm' := Some _) in IHb; tea. + 1: apply lift_typing_impl with (1 := IHb) => t T HT. + 2: { rewrite -H. apply eq_term_upto_univ_lift; assumption. } eapply context_conversion; eauto. - eapply (type_Cumul' (lift0 #|fix_context mfix| (dtype x))); auto. - apply infer_typing_sort_impl with id a1; intros [Hs IH]. - specialize (IH _ _ a0 e'). - rewrite <-H. - eapply (weakening _ _ _ _ (tSort _)); eauto. - eapply eq_context_conversion; tea. now symmetry. - apply eq_term_upto_univ_cumulSpec. rewrite <- H. - eapply eq_term_upto_univ_lift. - eapply eq_term_upto_univ_empty_impl. - 4: intuition eauto. - all: intros ? ? []; reflexivity. * revert wffix. unfold wf_cofixpoint, wf_cofixpoint_gen. enough (map check_one_cofix mfix = map check_one_cofix mfix') as ->; auto. - apply upto_names_check_cofix. solve_all. - + eapply All_nth_error in ihmfix; tea. - now apply infer_typing_sort_impl with id ihmfix; intros []. + apply upto_names_check_cofix. unfold eq_mfixpoint in *. solve_all. + + eapply All_nth_error in hmfix; tea. + apply eq_term_upto_univ_cumulSpec, eq_term_leq_term, upto_names_impl_eq_term. now symmetry. - intros p prim_ty cdecl IH prim decl pinv pty ptyIH Δ v e e'. depelim e. depelim o. 1-2:econstructor; eauto; constructor. - pose proof (validity (type_Prim Σ Γ _ _ _ wfΓ prim decl pinv pty)) as [s Hs]. + pose proof (validity (type_Prim Σ Γ _ _ _ wfΓ prim decl pinv pty)) as (_ & s & Hs & _). eapply type_Cumul. econstructor; eauto. * depelim ptyIH. constructor; eauto. now rewrite -e. rewrite -e; eauto. specialize (hty _ _ e1 e'); eauto. eapply type_Cumul; tea. eapply hdef; eauto. @@ -1010,7 +894,7 @@ Section Alpha. eapply eq_term_upto_univ_cumulSpec. eapply eq_term_leq_term. eapply e1. * eapply eq_context_conversion in Hs; eauto. - * simp prim_type. eapply Universe.make_inj in e. rewrite e. + * simp prim_type. eapply Universe.make'_inj in e. rewrite e. eapply eq_term_upto_univ_cumulSpec. eapply upto_names_impl_leq_term. constructor. constructor. reflexivity. now symmetry. @@ -1018,7 +902,7 @@ Section Alpha. - intros t A B X wf ht iht har ihar hcu Δ v e e'. eapply (type_ws_cumul_pb (pb:=Cumul)). + eapply iht; tea. - + exists X; eauto. + + eapply has_sort_isType. specialize (wf _ e'). now eapply eq_context_conversion. + eapply (ws_cumul_pb_ws_cumul_ctx (pb':=Conv)); tea. 2:eapply PCUICInversion.into_ws_cumul; tea. @@ -1037,42 +921,29 @@ Section Alpha. Local Ltac inv H := inversion H; subst; clear H. - Lemma eq_term_upto_univ_napp_0 n t t' : - eq_term_upto_univ_napp empty_global_env eq eq n t t' -> + Lemma eq_term_upto_univ_napp_0 pb n t t' : + eq_term_upto_univ_napp empty_global_env (fun _ => eq) (fun _ => eq) pb n t t' -> t ≡α t'. Proof using Type. apply eq_term_upto_univ_empty_impl; typeclasses eauto. Qed. - Lemma upto_names_eq_term_refl Σ Re n t t' : - RelationClasses.Reflexive Re -> - t ≡α t' -> - eq_term_upto_univ_napp Σ Re Re n t t'. - Proof using Type. - intros. - eapply eq_term_upto_univ_empty_impl; tea; tc. - all:intros x y ->; reflexivity. - Qed. - - Lemma upto_names_eq_term_upto_univ Σ Re Rle napp t u : - RelationClasses.Reflexive Re -> - RelationClasses.Reflexive Rle -> - RelationClasses.Symmetric Re -> - RelationClasses.Transitive Re -> - RelationClasses.Transitive Rle -> - RelationClasses.subrelation Re Rle -> - eq_term_upto_univ_napp Σ Re Rle napp t u -> + Lemma upto_names_eq_term_upto_univ Σ cmp_universe cmp_sort pb napp t u : + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe pb) -> + RelationClasses.PreOrder (cmp_universe Conv) -> + RelationClasses.PreOrder (cmp_universe pb) -> + RelationClasses.PreOrder (cmp_sort Conv) -> + RelationClasses.PreOrder (cmp_sort pb) -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp t u -> forall t' u', t ≡α t' -> u ≡α u' -> - eq_term_upto_univ_napp Σ Re Rle napp t' u'. + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp t' u'. Proof using Type. intros. - eapply (upto_names_eq_term_refl Σ Re) in X0; tea. - eapply (upto_names_eq_term_refl Σ Re) in X1; tea. - symmetry in X0. - eapply eq_term_upto_univ_trans; tea. - eapply eq_term_upto_univ_impl; tea. reflexivity. reflexivity. - eapply eq_term_upto_univ_trans; tea. - eapply eq_term_upto_univ_impl; tea. reflexivity. reflexivity. + eapply symmetry, upto_names_impl in X0; tea. + eapply upto_names_impl in X1; tea. + eapply eq_term_upto_univ_trans; cycle -1. + eapply eq_term_upto_univ_trans; cycle -1. + all: tea; tc. Qed. Lemma upto_names_leq_term Σ φ t u t' u' @@ -1125,7 +996,7 @@ Section Alpha. isType Σ Γ v. Proof using Type. intros Hty eq. - apply infer_typing_sort_impl with id Hty; intros Hs. + apply lift_sorting_it_impl_gen with Hty => // Hs. eapply typing_alpha; eauto. Qed. @@ -1136,7 +1007,7 @@ Section Alpha. isType Σ Δ v. Proof using Type. intros Hty eqctx eqtm. - apply infer_typing_sort_impl with id Hty; intros Hs. + apply lift_sorting_it_impl_gen with Hty => // Hs. eapply typing_alpha_prop; eauto. Qed. diff --git a/pcuic/theories/PCUICArities.v b/pcuic/theories/PCUICArities.v index ae49a145a..ff594b694 100644 --- a/pcuic/theories/PCUICArities.v +++ b/pcuic/theories/PCUICArities.v @@ -21,12 +21,13 @@ Implicit Types cf : checker_flags. Notation isWAT := (isWfArity typing). Lemma isType_Sort {cf:checker_flags} {Σ Γ s} : - wf_universe Σ s -> + wf_sort Σ s -> wf_local Σ Γ -> isType Σ Γ (tSort s). Proof. intros wfs wfΓ. - eexists; econstructor; eauto. + eapply has_sort_isType. + econstructor; eauto. Qed. #[global] Hint Resolve isType_Sort : pcuic. @@ -78,7 +79,7 @@ Proof. (Σ, ind_universes mdecl) ;;; [] |- tInd {| inductive_mind := inductive_mind ind; inductive_ind := i |} u : (ind_type x)) 0 (ind_bodies mdecl)). { apply forall_nth_error_Alli. intros. eapply Alli_nth_error in oind; eauto. simpl in oind. - destruct oind. destruct onArity as [s Hs]. + destruct oind. destruct onArity as (_ & s & Hs & _). unshelve eapply declared_inductive_to_gen in isdecl; eauto. eapply type_Cumul; eauto. econstructor; eauto. split; eauto with pcuic. @@ -142,7 +143,7 @@ Section WfEnv. ∑ T' ctx' s', [× Σ ;;; Γ ⊢ T ⇝ T', (destArity [] T' = Some (ctx', s')), Σ ⊢ Γ ,,, smash_context [] ctx = Γ ,,, ctx' & - leq_universe (global_ext_constraints Σ) s s'] + leq_sort (global_ext_constraints Σ) s s'] | None => unit end. Proof using wfΣ. @@ -213,14 +214,15 @@ Section WfEnv. isType Σ Γ (tProd na A B) <~> (isType Σ Γ A × isType Σ (Γ,, vass na A) B). Proof. split; intro HH. - - destruct HH as [s H]. + - destruct HH as (_ & s & H & _). apply inversion_Prod in H; tas. destruct H as [s1 [s2 [HA [HB Hs]]]]. - split. - * eexists; tea. - * eexists; tea. + apply lift_sorting_forget_univ in HA. + split; pcuic. - destruct HH as [HA HB]. - destruct HA as [sA HA], HB as [sB HB]. - eexists. econstructor; eassumption. + pose proof (lift_sorting_extract HA). + destruct HB as (_ & sB & HB & _). + eapply has_sort_isType. + econstructor; eassumption. Defined. Lemma isType_subst {Γ Δ A} s : @@ -229,9 +231,9 @@ Section WfEnv. isType Σ Γ (subst0 s A). Proof using wfΣ. intros sub HT. - apply infer_typing_sort_impl with id HT; intros Hs. + apply lift_typing_f_impl with (1 := HT) => // ?? Hs. have wf := typing_wf_local Hs. - now eapply (substitution (Δ := []) (T := tSort _)). + now eapply (substitution (Δ := [])). Qed. Lemma isType_subst_gen {Γ Δ Δ'} {A} s : @@ -240,8 +242,8 @@ Section WfEnv. isType Σ (Γ ,,, subst_context s 0 Δ') (subst s #|Δ'| A). Proof using wfΣ. intros sub HT. - apply infer_typing_sort_impl with id HT; intros Hs. - now eapply (substitution (T:=tSort _)). + apply lift_typing_f_impl with (1 := HT) => // ?? Hs. + now eapply substitution. Qed. Lemma type_ws_cumul_pb {pb Γ t} T {U} : @@ -251,7 +253,7 @@ Section WfEnv. Σ ;;; Γ |- t : U. Proof using Type. intros. - eapply type_Cumul; tea. apply X0.π2. + eapply type_Cumul; tea. apply X0.2.π2.1. destruct pb. - eapply ws_cumul_pb_eq_le in X1. now eapply cumulAlgo_cumulSpec in X1. @@ -262,11 +264,11 @@ Section WfEnv. : isType Σ Γ (tLetIn na t A B) -> isType Σ Γ (B {0:=t}). Proof using wfΣ. intro HH. - apply infer_typing_sort_impl with id HH; intros H. - assert (Hs := typing_wf_universe _ H). - apply inversion_LetIn in H; tas. destruct H as (s1 & A' & HA & Ht & HB & H). + apply lift_sorting_it_impl_gen with HH => // H. + assert (Hs := typing_wf_sort _ H). + apply inversion_LetIn in H; tas. destruct H as (A' & Ht & HB & H). eapply (type_ws_cumul_pb (pb:=Cumul)) with (A' {0 := t}). eapply substitution_let in HB; eauto. - * econstructor; eauto with pcuic. econstructor; eauto. + * pcuic. * eapply ws_cumul_pb_Sort_r_inv in H as [s' [H H']]. transitivity (tSort s'); eauto. eapply red_ws_cumul_pb. @@ -279,8 +281,8 @@ Section WfEnv. Lemma isType_tLetIn_dom {Γ} (HΓ : wf_local Σ Γ) {na t A B} : isType Σ Γ (tLetIn na t A B) -> Σ ;;; Γ |- t : A. Proof using wfΣ. - intros (s & H). - apply inversion_LetIn in H; tas. now destruct H as (s1 & A' & HA & Ht & HB & H). + intros (_ & s & H & _). + apply inversion_LetIn in H; tas. destruct H as (A' & Ht & HB & H). now eapply unlift_TermTyp. Qed. Lemma wf_local_ass {Γ na A} : @@ -298,6 +300,7 @@ Section WfEnv. wf_local Σ (Γ ,, vdef na d ty). Proof using Type. constructor; eauto with pcuic. + split; tas. apply X0. Qed. Hint Resolve wf_local_ass wf_local_def : pcuic. @@ -393,44 +396,45 @@ Section WfEnv. Σ ;;; Γ ,, d |- t : tSort s -> match decl_body d return Type with | Some b => Σ ;;; Γ |- mkProd_or_LetIn d t : tSort s - | None => Σ ;;; Γ |- mkProd_or_LetIn d t : tSort (Universe.sort_of_product u s) + | None => Σ ;;; Γ |- mkProd_or_LetIn d t : tSort (Sort.sort_of_product u s) end. Proof using wfΣ. destruct d as [na [b|] dty] => [Hd Ht|Hd Ht]; rewrite /mkProd_or_LetIn /=. - have wf := typing_wf_local Ht. - depelim wf. clear l. + depelim wf. eapply type_Cumul. econstructor; eauto. - econstructor; eauto. now eapply typing_wf_universe in Ht; pcuic. + econstructor; eauto. now eapply typing_wf_sort in Ht; pcuic. eapply convSpec_cumulSpec, red1_cumulSpec. constructor. - have wf := typing_wf_local Ht. depelim wf; clear l. eapply type_Prod; eauto. + pcuic. Qed. Lemma type_it_mkProd_or_LetIn {Γ Γ' u t s} : - wf_universe Σ u -> + wf_sort Σ u -> type_local_ctx (lift_typing typing) Σ Γ Γ' u -> Σ ;;; Γ ,,, Γ' |- t : tSort s -> - Σ ;;; Γ |- it_mkProd_or_LetIn Γ' t : tSort (Universe.sort_of_product u s). + Σ ;;; Γ |- it_mkProd_or_LetIn Γ' t : tSort (Sort.sort_of_product u s). Proof using wfΣ. revert Γ u s t. induction Γ'; simpl; auto; move=> Γ u s t wfu equ Ht. - eapply type_Cumul; eauto. econstructor; eauto using typing_wf_local with pcuic. - eapply typing_wf_universe in Ht; auto with pcuic. - eapply cumul_Sort. eapply leq_universe_product. - - specialize (IHΓ' Γ u (Universe.sort_of_product u s)); auto. + eapply typing_wf_sort in Ht; auto with pcuic. + eapply cumul_Sort. eapply leq_sort_product. + - specialize (IHΓ' Γ u (Sort.sort_of_product u s)); auto. unfold app_context in Ht. eapply type_Cumul. eapply IHΓ'; auto. destruct a as [na [b|] ty]; intuition auto. - destruct a as [na [b|] ty]; intuition auto. + destruct a as [na [b|] ty]; cbn; destruct equ as (? & (Hb & u' & Ht' & equ)); cbn in Hb, Ht', equ; try subst u'. { apply typing_wf_local in Ht as XX. inversion XX; subst. eapply (type_mkProd_or_LetIn {| decl_body := Some b |}); auto. - + simpl. exact X0.π2. + + simpl. exact X0.2.π2.1. + eapply type_Cumul; eauto. econstructor; eauto with pcuic. - eapply cumul_Sort. eapply leq_universe_product. } + eapply cumul_Sort. eapply leq_sort_product. } eapply (type_mkProd_or_LetIn {| decl_body := None |}) => /=; eauto. econstructor; eauto with pcuic. eapply typing_wf_local in Ht. @@ -441,17 +445,17 @@ Section WfEnv. Fixpoint sort_of_products us s := match us with | [] => s - | u :: us => sort_of_products us (Universe.sort_of_product u s) + | u :: us => sort_of_products us (Sort.sort_of_product u s) end. - Lemma leq_universe_sort_of_products_mon {u u' v v'} : - Forall2 (leq_universe Σ) u u' -> - leq_universe Σ v v' -> - leq_universe Σ (sort_of_products u v) (sort_of_products u' v'). + Lemma leq_sort_sort_of_products_mon {u u' v v'} : + Forall2 (leq_sort Σ) u u' -> + leq_sort Σ v v' -> + leq_sort Σ (sort_of_products u v) (sort_of_products u' v'). Proof using Type. intros hu; induction hu in v, v' |- *; simpl; auto with pcuic. intros lev. eapply IHhu. - eapply leq_universe_product_mon => //. + eapply leq_sort_product_mon => //. Qed. Lemma type_it_mkProd_or_LetIn_sorts {Γ Γ' us t s} : @@ -463,11 +467,11 @@ Section WfEnv. induction Γ'; simpl; auto; move=> Γ us s t equ Ht. - destruct us => //. - destruct a as [na [b|] ty]; intuition auto. - * destruct a0 as [s' Hs]. + * destruct b0 as (_ & s' & Hs & _). eapply IHΓ'; eauto. eapply (type_mkProd_or_LetIn {| decl_body := Some b |}); auto. simpl. exact Hs. - * destruct us => //. destruct equ. + * destruct us => //. destruct equ as (? & _ & s' & Hty & ->). simpl. eapply IHΓ'; eauto. apply (type_mkProd_or_LetIn {| decl_body := None |}) => /=; eauto. @@ -478,15 +482,10 @@ Section WfEnv. isType Σ (Γ ,,, Γ') t -> isType Σ Γ (it_mkProd_or_LetIn Γ' t). Proof using cf Σ wfΣ. - move=> equs [s ttyp]; exists (sort_of_products us s). + move=> equs [_ [s [ttyp _]]]; eapply @has_sort_isType with (s := sort_of_products us s). apply: type_it_mkProd_or_LetIn_sorts=> //. Qed. - Lemma app_context_push Γ Δ Δ' d : (Γ ,,, Δ ,,, Δ') ,, d = (Γ ,,, Δ ,,, (Δ' ,, d)). - Proof using Type. - reflexivity. - Qed. - Hint Extern 4 (_ ;;; _ |- _ <= _) => reflexivity : pcuic. Ltac pcuic := eauto 5 with pcuic. @@ -600,8 +599,8 @@ Section WfEnv. isType Σ Γ (subst0 s T). Proof using wfΣ. intros sub HT. - apply infer_typing_sort_impl with id HT; intros Hs. - destruct HT as (s' & t); cbn in Hs |- *; clear t. + apply lift_sorting_it_impl_gen with HT => // Hs. + set (s' := HT.2.π1) in *; clearbody s'. clear HT. revert Γ s sub Hs. generalize (le_n #|Δ|). generalize #|Δ| at 2. @@ -619,37 +618,37 @@ Section WfEnv. rewrite it_mkProd_or_LetIn_app /= /mkProd_or_LetIn /=. intros Hs. - assert (wfs' := typing_wf_universe wfΣ Hs). - eapply inversion_LetIn in Hs as (? & ? & ? & ? & ? & ?); auto. - eapply substitution_let in t1; auto. - eapply ws_cumul_pb_LetIn_l_inv in w; auto. + assert (wfs' := typing_wf_sort wfΣ Hs). + eapply inversion_LetIn in Hs as (T' & wfT' & HT' & hlt); auto. + eapply substitution_let in HT'; auto. + eapply ws_cumul_pb_LetIn_l_inv in hlt; auto. pose proof (subslet_app_inv sub) as [subl subr]. depelim subl. depelim subl. rewrite subst_empty in H0. rewrite H0 in subr. specialize (IHn (subst_context [b] 0 l) (subst [b] #|l| T) ltac:(rewrite subst_context_length; lia)). specialize (IHn _ _ subr). - rewrite /subst1 subst_it_mkProd_or_LetIn Nat.add_0_r in t1. - rewrite !subst_empty in t3. + rewrite /subst1 subst_it_mkProd_or_LetIn Nat.add_0_r in HT'. + rewrite !subst_empty in t0. forward IHn. - eapply type_Cumul. eapply t1. econstructor; intuition eauto using typing_wf_local with pcuic. - eapply (cumulAlgo_cumulSpec _ (pb:=Cumul)), w. rewrite {2}Hl in IHn. + eapply type_Cumul. eapply HT'. econstructor; intuition eauto using typing_wf_local with pcuic. + eapply (cumulAlgo_cumulSpec _ (pb:=Cumul)), hlt. rewrite {2}Hl in IHn. now rewrite -subst_app_simpl -H0 firstn_skipn in IHn. intros Hs. - assert (wfs' := typing_wf_universe wfΣ Hs). - eapply inversion_Prod in Hs as (? & ? & ? & ? & ?); auto. + assert (wfs' := typing_wf_sort wfΣ Hs). + eapply inversion_Prod in Hs as (s1 & s2 & wfty & Hty & hlt); auto. pose proof (subslet_app_inv sub) as [subl subr]. - depelim subl; depelim subl. rewrite subst_empty in t2. rewrite H0 in subr. - epose proof (substitution0 t0 t2). - specialize (IHn (subst_context [t1] 0 l) (subst [t1] #|l| T)). + depelim subl; depelim subl. rewrite subst_empty in t0. rewrite H0 in subr. + epose proof (substitution0 Hty t0). + specialize (IHn (subst_context [t] 0 l) (subst [t] #|l| T)). forward IHn. rewrite subst_context_length; lia. specialize (IHn _ _ subr). rewrite /subst1 subst_it_mkProd_or_LetIn Nat.add_0_r in X. forward IHn. eapply type_Cumul. simpl in X. eapply X. econstructor; eauto with pcuic. - eapply ws_cumul_pb_Sort_inv in w. eapply cumul_Sort. - transitivity (Universe.sort_of_product x x0). - eapply leq_universe_product. auto. + eapply ws_cumul_pb_Sort_inv in hlt. eapply cumul_Sort. + transitivity (Sort.sort_of_product s1 s2). + eapply leq_sort_product. auto. rewrite {2}Hl in IHn. now rewrite -subst_app_simpl -H0 firstn_skipn in IHn. Qed. @@ -674,18 +673,18 @@ Section WfEnv. + simpl. intros. now eapply typing_wf_local in X. + rewrite it_mkProd_or_LetIn_app. destruct x as [na [b|] ty]; cbn; move=> H. - * apply inversion_LetIn in H as (s1 & A & H0 & H1 & H2 & H3); auto. - eapply All_local_env_app; split; pcuic. - eapply All_local_env_app. split. repeat constructor. now exists s1. - auto. apply IHΔ in H2. - eapply All_local_env_app_inv in H2. intuition auto. + * apply inversion_LetIn in H as (A & wfty & HT & hlt); auto. + eapply All_local_env_app. 1: eapply unlift_TermTyp in wfty; pcuic. + eapply All_local_env_app. repeat (constructor; tea). + apply IHΔ in HT. + eapply All_local_env_app_inv in HT. intuition auto. eapply All_local_env_impl; eauto. simpl. intros. now rewrite app_context_assoc. - * apply inversion_Prod in H as (s1 & A & H0 & H1 & H2); auto. - eapply All_local_env_app; split; pcuic. - eapply All_local_env_app. split. repeat constructor. now exists s1. - apply IHΔ in H1. - eapply All_local_env_app_inv in H1. intuition auto. + * apply inversion_Prod in H as (s1 & s2 & wfty & HT & hlt); auto. + eapply All_local_env_app. 1: eapply unlift_TypUniv in wfty; cbn in wfty; pcuic. + eapply All_local_env_app. apply lift_sorting_forget_univ in wfty. repeat (constructor; tea). + apply IHΔ in HT. + eapply All_local_env_app_inv in HT. intuition auto. eapply All_local_env_impl; eauto. simpl. intros. now rewrite app_context_assoc. Qed. @@ -693,7 +692,7 @@ Section WfEnv. Lemma isType_it_mkProd_or_LetIn_wf_local {Γ Δ T} : isType Σ Γ (it_mkProd_or_LetIn Δ T) -> wf_local Σ (Γ ,,, Δ). Proof using wfΣ. - move=> [s Hs]. + intros (_ & s & Hs & _). now eapply it_mkProd_or_LetIn_wf_local in Hs. Qed. @@ -703,12 +702,8 @@ Section WfEnv. isType Σ Γ T. Proof using wfΣ. intros wfΓ HT. - apply infer_typing_sort_impl with id HT; intros hs. - unshelve epose proof (subject_closed hs); eauto. - eapply (weakening _ _ Γ) in hs => //. - rewrite lift_closed in hs => //. - now rewrite app_context_nil_l in hs. - now rewrite app_context_nil_l. + apply lift_typing_impl with (1 := HT) => ?? hs. + eapply (weaken_ctx Γ) in hs; eauto. Qed. Lemma subst_telescope_subst_instance u s k Γ : diff --git a/pcuic/theories/PCUICAst.v b/pcuic/theories/PCUICAst.v index 4260e3b4f..ecf543a07 100644 --- a/pcuic/theories/PCUICAst.v +++ b/pcuic/theories/PCUICAst.v @@ -196,7 +196,7 @@ 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.t) +| tSort (u : sort) | tProd (na : aname) (A B : term) | tLambda (na : aname) (A t : term) | tLetIn (na : aname) (b B t : term) (* let na := b : B in t *) @@ -405,7 +405,7 @@ Instance subst_instance_constr : UnivSubst term := match c with | tRel _ | tVar _ => c | tEvar ev args => tEvar ev (List.map (subst_instance_constr u) args) - | tSort s => tSort (subst_instance_univ u s) + | tSort s => tSort (subst_instance_sort 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') @@ -431,7 +431,7 @@ Instance subst_instance_constr : UnivSubst term := (** Tests that the term is closed over [k] universe variables *) Fixpoint closedu (k : nat) (t : term) : bool := match t with - | tSort univ => closedu_universe k univ + | tSort s => closedu_sort k s | tInd _ u => closedu_instance k u | tConstruct _ _ u => closedu_instance k u | tConst _ u => closedu_instance k u @@ -1315,6 +1315,14 @@ Qed. #[global] Hint Rewrite test_context_map : map. +Lemma test_context_app (p : term -> bool) Γ Δ : + test_context p (Γ ,,, Δ) = test_context p Γ && test_context p Δ. +Proof using Type. + induction Δ; simpl; auto. + - now rewrite andb_true_r. + - now rewrite IHΔ andb_assoc. +Qed. + Lemma onctx_test P (p q : term -> bool) ctx : onctx P ctx -> test_context p ctx -> diff --git a/pcuic/theories/PCUICCasesContexts.v b/pcuic/theories/PCUICCasesContexts.v index 1ca76ba3f..500bbee39 100644 --- a/pcuic/theories/PCUICCasesContexts.v +++ b/pcuic/theories/PCUICCasesContexts.v @@ -81,10 +81,11 @@ Lemma alpha_eq_smash_context Δ Δ' : Proof. induction 1. * constructor. - * destruct x; depelim r; simpl; auto. + * destruct r; simpl; auto. rewrite !(smash_context_acc _ [_]). eapply All2_app; auto; repeat constructor; subst; simpl; auto. rewrite (All2_length X) -(alpha_eq_extended_subst l l' 0) // (alpha_eq_context_assumptions l l') //. + now constructor. Qed. Lemma alpha_eq_lift_context n k Δ Δ' : @@ -95,7 +96,7 @@ Proof. * constructor. * rewrite !lift_context_snoc; destruct x; depelim r; simpl; subst; auto; constructor; auto; repeat constructor; subst; simpl; auto; - now rewrite (All2_length X). + rewrite (All2_length X); now constructor. Qed. Lemma alpha_eq_subst_context s k Δ Δ' : @@ -106,7 +107,7 @@ Proof. * constructor. * rewrite !subst_context_snoc; destruct x; depelim r; simpl; subst; auto; constructor; auto; repeat constructor; subst; simpl; auto; - now rewrite (All2_length X). + rewrite (All2_length X); now constructor. Qed. Lemma inst_case_predicate_context_eq {mdecl idecl ind p} : diff --git a/pcuic/theories/PCUICCasesHelper.v b/pcuic/theories/PCUICCasesHelper.v index fa4588074..2ff8c1ec9 100644 --- a/pcuic/theories/PCUICCasesHelper.v +++ b/pcuic/theories/PCUICCasesHelper.v @@ -67,13 +67,13 @@ Proof. induction Δ'; simpl; auto. destruct a as [na [b|] ty]; simpl; intuition auto. + simpl; rewrite PCUICLiftSubst.lift_context_snoc /= Nat.add_0_r; - repeat split; tas. - - apply infer_typing_sort_impl with id a0; intros Hs. - now eapply weakening_typing in Hs. - - now eapply weakening_typing in b1. + split; tas. + apply lift_typing_f_impl with (1 := b0) => // ?? Hs. + now eapply weakening_typing in Hs. + rewrite PCUICLiftSubst.lift_context_snoc /= Nat.add_0_r. - intuition auto. - eapply weakening_typing in b; eauto. + split; tas. + apply lift_typing_f_impl with (1 := b) => // ?? Hs. + now eapply weakening_typing in Hs. Qed. Lemma wf_local_expand_lets0 {cf : checker_flags} {Σ : global_env_ext} : @@ -115,9 +115,9 @@ Section WithCheckerFlags. intros wfΣ wfΓ wfty. rewrite <- (firstn_skipn n Γ) in wfΓ |- *. assert (n = #|firstn n Γ|). { rewrite firstn_length_le; auto with arith. } - apply infer_typing_sort_impl with id wfty; intros Hs. - rewrite {3 4}H. - eapply (weakening_typing (Γ := skipn n Γ) (Γ' := Δ) (Γ'' := firstn n Γ) (T := tSort _)); + apply lift_typing_f_impl with (1 := wfty) => // ?? Hs. + rewrite {3 4 5}H. + eapply (weakening_typing (Γ := skipn n Γ) (Γ' := Δ) (Γ'' := firstn n Γ)); eauto with wf. Qed. @@ -353,7 +353,7 @@ Section CaseBranchTypeBeta. rewrite PCUICLiftSubst.lift_it_mkLambda_or_LetIn Nat.add_0_r. set cstr_ctx := (pre_case_branch_context_gen _ _ _ _ _). set Δ := (lift_context _ _ _). - move=> wfΣ [ps retWty] wfcstr_ctx ?. + move=> wfΣ retWty wfcstr_ctx ?. destruct retWty as (_ & ps & retWty & _). apply: red_betas_typed; first by rewrite /Δ /instantiate_cstr_indices; len; lia. rewrite /Δ; have -> : #|cstr_args cb| = #|cstr_ctx| by rewrite /cstr_ctx /pre_case_branch_context_gen; len. apply: weakening_typing; eassumption. @@ -395,7 +395,7 @@ Section RandomLemmas. ∑ s, spine_subst Σ Γ (pparams ++ indices) s (ind_params mib,,, ind_indices oib)@[puinst]. Proof using cf. move=> wfΣ inddecl pparamslen discrtypok. - have Γwf : wf_local Σ Γ by apply: (typing_wf_local discrtypok.π2). + have Γwf : wf_local Σ Γ by apply: (typing_wf_local discrtypok.2.π2.1). move: (inddecl)=> /(PCUICWeakeningEnvTyp.on_declared_inductive (Σ:=Σ.1)) [on_ind on_body]. have wfΣ1 := (wfΣ : wf Σ.1). have wfΣwk := declared_inductive_wf_ext_wk _ _ _ wfΣ (proj1 inddecl). @@ -421,7 +421,7 @@ Section RandomLemmas. * rewrite -app_context_assoc-subst_instance_app_ctx; apply: weaken_wf_local=> //. set (Δ := _ ,,, _). destruct Σ as [Σ1 Σ2]. refine (PCUICUnivSubstitutionTyp.typing_subst_instance_wf_local Σ1 Σ2 Δ puinst (ind_universes mib) _ _ _)=> //. - move: (onArity on_body) => [lind]. + destruct (onArity on_body) as (_ & s & lind & _). move: lind => /=. rewrite (PCUICGlobalMaps.ind_arity_eq on_body). move=> /PCUICSpine.type_it_mkProd_or_LetIn_inv [uparams [? [slparams + _ _]]]. move=> /PCUICSpine.type_it_mkProd_or_LetIn_inv [uindxs [? [slindxs + _ _]]]. @@ -447,7 +447,7 @@ Lemma type_Case_helper (rettyp : term) (brs : list term) (indices : list term) - (ps : Universe.t) (Σ : global_env_ext) Γ : + (ps : sort) (Σ : global_env_ext) Γ : wf_ext Σ -> declared_inductive Σ ind mib oib -> @@ -497,7 +497,7 @@ Proof. { do 2 constructor=> //=. rewrite /expand_lets_ctx /expand_lets_k_ctx /subst_context /lift_context !forget_types_fold_context_k. - apply: PCUICEquality.eq_context_gen_binder_annot; reflexivity. + apply PCUICEquality.eq_context_upto_names_binder_annot; reflexivity. } pose proof (discrtypok := validity discrtyp). @@ -526,8 +526,7 @@ Proof. hnf=>/=. rewrite /cstr_branch_context /expand_lets_ctx /expand_lets_k_ctx. rewrite /subst_context /lift_context !forget_types_fold_context_k. - apply: PCUICEquality.eq_context_gen_binder_annot. - apply: All2_fold_refl; reflexivity. + apply PCUICEquality.eq_context_upto_names_binder_annot. reflexivity. } constructor=> //. @@ -566,7 +565,7 @@ Lemma type_Case_subst_helper (rettyp : term) (brs : list term) (indices : list term) - (ps : Universe.t) (Σ : global_env_ext) Γ : + (ps : sort) (Σ : global_env_ext) Γ : wf_ext Σ -> declared_inductive Σ ind mib oib -> @@ -614,7 +613,7 @@ Proof. { do 2 constructor=> //=. rewrite /expand_lets_ctx /expand_lets_k_ctx /subst_context /lift_context !forget_types_fold_context_k. - apply: PCUICEquality.eq_context_gen_binder_annot; reflexivity. + apply PCUICEquality.eq_context_upto_names_binder_annot; reflexivity. } @@ -681,7 +680,7 @@ Proof. econstructor. 1: exact: mk_caseWtyp. - apply: subject_reduction; last exact red_rettyp. - exact: (validity mk_caseWtyp).π2. + exact: (validity mk_caseWtyp).2.π2.1. - apply: convSpec_cumulSpec. apply: conv_betas=> //. rewrite forallb_app /= (PCUICClosedTyp.subject_closed discrtyp) !andb_true_r. @@ -699,7 +698,7 @@ Lemma type_Case_simple_subst_helper (rettyp : term) (brs : list term) (indices : list term) - (ps : Universe.t) (Σ : global_env_ext) Γ : + (ps : sort) (Σ : global_env_ext) Γ : wf_ext Σ -> declared_inductive Σ ind mib oib -> @@ -748,7 +747,7 @@ Proof. { do 2 constructor=> //=. rewrite /expand_lets_ctx /expand_lets_k_ctx /subst_context /lift_context !forget_types_fold_context_k. - apply: PCUICEquality.eq_context_gen_binder_annot; reflexivity. + apply PCUICEquality.eq_context_upto_names_binder_annot; reflexivity. } diff --git a/pcuic/theories/PCUICClassification.v b/pcuic/theories/PCUICClassification.v index a6a3da35d..c8587ef75 100644 --- a/pcuic/theories/PCUICClassification.v +++ b/pcuic/theories/PCUICClassification.v @@ -368,7 +368,7 @@ Qed. * elimtype False; depelim ass. * simpl. rewrite !it_mkProd_or_LetIn_app /= /mkProd_or_LetIn /= //. intros wf sp. - pose proof (wf_local_app_inv wf) as [_ wfty]. + pose proof (All_local_env_app_inv wf) as [_ wfty]. eapply All_local_env_app_inv in wfty as [wfty _]. depelim wfty. intros Hargs. eapply nth_error_None in Hargs. len in Hargs. len; simpl. @@ -473,7 +473,7 @@ Qed. - eapply typing_spine_nth_error in sp; eauto; cycle 1. * eapply smash_context_assumption_context; constructor. * eapply wf_local_smash_end; eauto. - destruct isty as [s Hs]. + destruct isty as (_ & s & Hs & _). eapply inversion_it_mkProd_or_LetIn in Hs; eauto. now eapply typing_wf_local. * destruct (decompose_app (decl_type c)) as [hd tl] eqn:da. @@ -486,7 +486,7 @@ Qed. - eapply typing_spine_nth_error_None_prod in sp; eauto. * eapply smash_context_assumption_context; constructor. * eapply wf_local_smash_end; eauto. - destruct isty as [s Hs]. + destruct isty as (_ & s & Hs & _). eapply inversion_it_mkProd_or_LetIn in Hs; eauto. now eapply typing_wf_local. Qed. @@ -534,7 +534,7 @@ Qed. * len in sp. * eapply smash_context_assumption_context; constructor. * eapply wf_local_smash_end; eauto. - destruct isty as [s Hs]. + destruct isty as (_ & s & Hs & _). eapply inversion_it_mkProd_or_LetIn in Hs; eauto. now eapply typing_wf_local. * len; simpl. eapply nth_error_None in hargs => //. @@ -669,7 +669,7 @@ Section classification. now eapply invert_cumul_prod_ind in cum. - move=> [hargs ccum']. rewrite expand_lets_mkApps subst_mkApps /= in ccum'. - eapply invert_cumul_ind_ind in ccum' as ((? & ?) & ?). + eapply invert_cumul_ind_ind in ccum' as ((? & r) & ?). len in r. eapply ReflectEq.eqb_eq in i0. now subst ind'. Qed. @@ -702,7 +702,7 @@ Section classification. rewrite oib.(ind_arity_eq) in t0. rewrite !subst_instance_it_mkProd_or_LetIn in t0. eapply typing_spine_arity_mkApps_Ind in t0; eauto. - eapply typing_spine_isType_dom in t0 as [s Hs]. + eapply typing_spine_isType_dom in t0 as (_ & s & Hs & _). eexists; split; [sq|]. now eapply wt_closed_red_refl. now do 2 eapply isArity_it_mkProd_or_LetIn. now eapply (declared_inductive_valid_type decl). @@ -803,7 +803,7 @@ Section classification. now eapply invert_cumul_sort_ind in e. - eapply inversion_Prod in typed as (? & ? & ? & ? & e); auto. now eapply invert_cumul_sort_ind in e. - - eapply inversion_Lambda in typed as (? & ? & ? & ? & e); auto. + - eapply inversion_Lambda in typed as (? & ? & ? & e); auto. now eapply invert_cumul_prod_ind in e. - now rewrite head_mkApps /= /head /=. - exfalso; eapply invert_ind_ind; eauto. @@ -933,36 +933,37 @@ Section classification. pose proof (subject_closed t1); auto. eapply eval_closed in H; eauto. eapply subject_reduction in IHHe1. 2-3:eauto. - eapply inversion_Lambda in IHHe1 as (? & ? & ? & ? & e); eauto. - eapply (substitution0 (Γ := [])) in t3; eauto. + eapply inversion_Lambda in IHHe1 as (? & ? & ? & e); eauto. + eapply (substitution0 (Γ := [])) in t2; eauto. eapply IHHe3. rewrite (closed_subst a' 0 b); auto. - rewrite /subst1 in t3. eapply t3. + rewrite /subst1 in t2. eapply t2. eapply (type_ws_cumul_pb (pb:=Conv)); eauto. - eapply subject_reduction in IHHe2; eauto. now eexists. + eapply subject_reduction in IHHe2; eauto. eapply ws_cumul_pb_Prod_Prod_inv in e as [eqna dom codom]; eauto. now symmetry. - - eapply inversion_LetIn in Ht as (? & ? & ? & ? & ? & ?); auto. + - eapply inversion_LetIn in Ht as (? & ? & ? & ?); auto. + apply unlift_TermTyp in l as t0'. redt (tLetIn na b0' t b1); eauto. eapply red_letin; eauto. redt (b1 {0 := b0'}); auto. do 2 econstructor. - specialize (IHHe1 _ t1). + specialize (IHHe1 _ t0'). rewrite /subst1. rewrite -(closed_subst b0' 0 b1). - eapply subject_reduction in t1; eauto. eapply subject_closed in t1; eauto. + eapply subject_reduction in t0'; eauto. eapply subject_closed in t0'; eauto. eapply IHHe2. rewrite closed_subst. - eapply subject_reduction in t1; eauto. eapply subject_closed in t1; eauto. - pose proof (subject_is_open_term t2). - eapply substitution_let in t2; eauto. - eapply subject_reduction in t2. + eapply subject_reduction in t0'; eauto. eapply subject_closed in t0'; eauto. + pose proof (subject_is_open_term t0). + eapply substitution_let in t0; eauto. + eapply subject_reduction in t0. 3:{ eapply (red_red (Δ := [vass na t]) (Γ' := [])); eauto. 3:repeat constructor. - cbn. rewrite addnP_shiftnP. eapply type_is_open_term in t1. now erewrite t1. + cbn. rewrite addnP_shiftnP. eapply type_is_open_term in t0'. now erewrite t0'. repeat constructor. cbn. rewrite addnP_shiftnP. - eapply subject_is_open_term in t1. cbn in t1. now setoid_rewrite shiftnP0 in t1. } - apply t2. auto. + eapply subject_is_open_term in t0'. cbn in t0'. now setoid_rewrite shiftnP0 in t0'. } + apply t0. auto. - redt (subst_instance u body); auto. eapply red1_red. econstructor; eauto. diff --git a/pcuic/theories/PCUICConfluence.v b/pcuic/theories/PCUICConfluence.v index 1060284ed..d143f756a 100644 --- a/pcuic/theories/PCUICConfluence.v +++ b/pcuic/theories/PCUICConfluence.v @@ -93,8 +93,8 @@ Proof. Qed. Lemma clos_rt_OnOne2_local_env_incl R : -inclusion (OnOne2_local_env (on_one_decl (fun Δ => clos_refl_trans (R Δ)))) - (clos_refl_trans (OnOne2_local_env (on_one_decl R))). +inclusion (OnOne2_local_env (fun Δ => on_one_decl (clos_refl_trans (R Δ)))) + (clos_refl_trans (OnOne2_local_env (fun Δ => on_one_decl1 R Δ))). Proof. intros x y H. induction H; firstorder; try subst na'. @@ -168,8 +168,8 @@ Qed. Lemma OnOne2_local_env_apply {B} {P : B -> context -> term -> term -> Type} {l l'} (f : context -> term -> term -> B) : - OnOne2_local_env (on_one_decl (fun Γ x y => forall a : B, P a Γ x y)) l l' -> - OnOne2_local_env (on_one_decl (fun Γ x y => P (f Γ x y) Γ x y)) l l'. + OnOne2_local_env (fun Γ => on_one_decl (fun x y => forall a : B, P a Γ x y)) l l' -> + OnOne2_local_env (fun Γ => on_one_decl (fun x y => P (f Γ x y) Γ x y)) l l'. Proof. intros; eapply OnOne2_local_env_impl; tea. intros Δ x y. eapply on_one_decl_impl; intros Γ ? ?; eauto. @@ -178,8 +178,8 @@ Qed. Lemma OnOne2_local_env_apply_dep {B : context -> term -> term -> Type} {P : context -> term -> term -> Type} {l l'} : (forall Γ' x y, B Γ' x y) -> - OnOne2_local_env (on_one_decl (fun Γ x y => B Γ x y -> P Γ x y)) l l' -> - OnOne2_local_env (on_one_decl (fun Γ x y => P Γ x y)) l l'. + OnOne2_local_env (fun Γ => on_one_decl (fun x y => B Γ x y -> P Γ x y)) l l' -> + OnOne2_local_env (fun Γ => on_one_decl (fun x y => P Γ x y)) l l'. Proof. intros; eapply OnOne2_local_env_impl; tea. intros Δ x y. eapply on_one_decl_impl; intros Γ ? ?; eauto. @@ -197,9 +197,9 @@ Proof. Qed. Lemma OnOne2_local_env_exist' (P Q R : context -> term -> term -> Type) (l l' : context) : - OnOne2_local_env (on_one_decl P) l l' -> + OnOne2_local_env (fun Γ => on_one_decl1 P Γ) l l' -> (forall Γ x y, P Γ x y -> ∑ z : term, Q Γ x z × R Γ y z) -> - ∑ r : context, OnOne2_local_env (on_one_decl Q) l r × OnOne2_local_env (on_one_decl R) l' r. + ∑ r : context, OnOne2_local_env (fun Γ => on_one_decl1 Q Γ) l r × OnOne2_local_env (fun Γ => on_one_decl1 R Γ) l' r. Proof. intros o Hp. induction o. - destruct p; subst. specialize (Hp _ _ _ p) as [? []]. @@ -213,9 +213,9 @@ Qed. Lemma OnOne2_local_env_All2_fold (P : context -> term -> term -> Type) (Q : context -> context -> context_decl -> context_decl -> Type) (l l' : context) : - OnOne2_local_env (on_one_decl P) l l' -> - (forall Γ x y, on_one_decl P Γ x y -> Q Γ Γ x y) -> - (forall Γ Γ' d, OnOne2_local_env (on_one_decl P) Γ Γ' -> Q Γ Γ' d d) -> + OnOne2_local_env (fun Γ => on_one_decl1 P Γ) l l' -> + (forall Γ x y, on_one_decl1 P Γ x y -> Q Γ Γ x y) -> + (forall Γ Γ' d, OnOne2_local_env (fun Γ => on_one_decl1 P Γ) Γ Γ' -> Q Γ Γ' d d) -> (forall Γ x, Q Γ Γ x x) -> All2_fold Q l l'. Proof. @@ -224,18 +224,6 @@ Proof. now eapply All2_fold_refl. Qed. -Lemma on_one_decl_compare_decl Σ Re Rle Γ x y : - RelationClasses.Reflexive Re -> - RelationClasses.Reflexive Rle -> - on_one_decl - (fun (_ : context) (y0 v' : term) => eq_term_upto_univ Σ Re Rle y0 v') Γ x y -> - compare_decls (eq_term_upto_univ Σ Re Rle) (eq_term_upto_univ Σ Re Rle) x y. -Proof. - intros heq hle. - destruct x as [na [b|] ty], y as [na' [b'|] ty']; cbn; intuition (subst; auto); - constructor; auto; reflexivity. -Qed. - Lemma OnOne2_disj {A} (P Q : A -> A -> Type) (l l' : list A) : OnOne2 (fun x y => P x y + Q x y)%type l l' <~> OnOne2 P l l' + OnOne2 Q l l'. @@ -247,27 +235,26 @@ Qed. Notation red1_ctx_rel Σ Δ := (OnOne2_local_env - (on_one_decl - (fun (Γ : context) (x0 y0 : term) => red1 Σ (Δ,,, Γ) x0 y0))). + (fun (Γ : context) => on_one_decl + (fun (t u : term) => red1 Σ (Δ ,,, Γ) t u))). -Notation eq_one_decl Σ Re Rle := +Notation eq_one_decl Σ cmp_universe cmp_sort pb := (OnOne2_local_env - (on_one_decl - (fun _ (x0 y0 : term) => - eq_term_upto_univ Σ Re Rle x0 y0))). - -Lemma red1_eq_context_upto_l {Σ Σ' Rle Re Γ Δ u v} : - RelationClasses.Reflexive Rle -> - SubstUnivPreserving Rle -> - RelationClasses.Reflexive Re -> - SubstUnivPreserving Re -> - RelationClasses.subrelation Re Rle -> + (fun _ => on_one_decl + (fun (t u : term) => + eq_term_upto_univ Σ cmp_universe cmp_sort pb t u))). + +Lemma red1_eq_context_upto_l {Σ Σ' cmp_universe cmp_sort pb Γ Δ u v} : + RelationClasses.Reflexive (cmp_universe Conv) -> + RelationClasses.Reflexive (cmp_universe pb) -> + RelationClasses.Reflexive (cmp_sort Conv) -> + RelationClasses.Reflexive (cmp_sort pb) -> red1 Σ Γ u v -> - eq_context_upto Σ' Re Rle Γ Δ -> + eq_context_upto Σ' cmp_universe cmp_sort pb Γ Δ -> ∑ v', red1 Σ Δ u v' * - eq_term_upto_univ Σ' Re Re v v'. + eq_term_upto_univ Σ' cmp_universe cmp_sort Conv v v'. Proof. - intros hle hle' he he' hlee h e. + intros refl_univ_conv refl_univ_pb refl_sort_conv refl_sort_pb h e. induction h in Δ, e |- * using red1_ind_all. all: try solve [ eexists ; split ; [ @@ -286,7 +273,7 @@ Proof. all: try solve [ match goal with | r : red1 _ (?Γ ,, ?d) _ _ |- _ => - assert (e' : eq_context_upto Σ' Re Rle (Γ,, d) (Δ,, d)) + assert (e' : eq_context_upto Σ' cmp_universe cmp_sort pb (Γ,, d) (Δ,, d)) ; [ constructor ; [ eauto | constructor; eauto ] ; eapply eq_term_upto_univ_refl ; eauto @@ -302,7 +289,7 @@ Proof. ]. - assert (h : ∑ b', (option_map decl_body (nth_error Δ i) = Some (Some b')) * - eq_term_upto_univ Σ' Re Re body b'). + eq_term_upto_univ Σ' cmp_universe cmp_sort Conv body b'). { induction i in Γ, Δ, H, e |- *. - destruct e. + cbn in *. discriminate. @@ -344,25 +331,25 @@ Proof. specialize (IH (Δ ,,, inst_case_branch_context p x)). forward IH by now apply eq_context_upto_cat. exact IH. } eapply (OnOne2_exist' _ (fun x y => on_Trel_eq (red1 Σ (Δ ,,, inst_case_branch_context p x)) bbody bcontext x y) - (fun x y => on_Trel_eq (eq_term_upto_univ Σ' Re Re) bbody bcontext x y)) in X as [brr [Hred Heq]]; tea. + (fun x y => on_Trel_eq (eq_term_upto_univ Σ' cmp_universe cmp_sort Conv) bbody bcontext x y)) in X as [brr [Hred Heq]]; tea. 2:{ intros x y [[v' [redv' eq]] eqctx]. exists {| bcontext := bcontext x; bbody := v' |}. intuition auto. } eexists; split. eapply case_red_brs; tea. econstructor; eauto; try reflexivity. - eapply OnOne2_All2; tea => /=; intuition eauto; try reflexivity. + eapply OnOne2_All2; tea => /=; unfold eq_branch; intuition eauto; try reflexivity. now rewrite b. - destruct (IHh _ e) as [x [redl redr]]. exists (tApp x M2). split. constructor; auto. - constructor. eapply eq_term_upto_univ_impl. 5:eauto. - all:auto. 1-3:typeclasses eauto. + constructor. eapply eq_term_upto_univ_impl. 6:eauto. + all:auto. 1-4:typeclasses eauto. reflexivity. - assert (h : ∑ ll, OnOne2 (red1 Σ Δ) l ll * - All2 (eq_term_upto_univ Σ' Re Re) l' ll + All2 (eq_term_upto_univ Σ' cmp_universe cmp_sort Conv) l' ll ). { induction X. - destruct p as [p1 p2]. @@ -391,11 +378,7 @@ Proof. (d'.(dname), d'.(dbody), d'.(rarg)) ) mfix0 mfix' * - All2 (fun x y => - eq_term_upto_univ Σ' Re Re (dtype x) (dtype y) * - eq_term_upto_univ Σ' Re Re (dbody x) (dbody y) * - (rarg x = rarg y) * - (eq_binder_annot (dname x) (dname y)))%type mfix1 mfix'). + eq_mfixpoint (eq_term_upto_univ Σ' cmp_universe cmp_sort Conv) mfix1 mfix'). { induction X. - destruct p as [[p1 p2] p3]. eapply p2 in e as hh. destruct hh as [? [? ?]]. @@ -426,51 +409,12 @@ Proof. (d.(dname), d.(dtype), d.(rarg)) = (d'.(dname), d'.(dtype), d'.(rarg)) ) mfix0 mfix' * - All2 (fun x y => - eq_term_upto_univ Σ' Re Re (dtype x) (dtype y) * - eq_term_upto_univ Σ' Re Re (dbody x) (dbody y) * - (rarg x = rarg y) * - eq_binder_annot (dname x) (dname y))%type mfix1 mfix'). - { (* Maybe we should use a lemma using firstn or skipn to keep - fix_context intact. Anything general? - *) - Fail induction X using OnOne2_ind_l. - (* This FAILs because it reduces the type of X before unifying - unfortunately... - *) - change ( - OnOne2 - ((fun L (x y : def term) => - (red1 Σ (Γ ,,, fix_context L) (dbody x) (dbody y) - × (forall Δ : context, - eq_context_upto Σ' Re Rle (Γ ,,, fix_context L) Δ -> - ∑ v' : term, - red1 Σ Δ (dbody x) v' × eq_term_upto_univ Σ' Re Re (dbody y) v')) - × (dname x, dtype x, rarg x) = (dname y, dtype y, rarg y)) mfix0) mfix0 mfix1 - ) in X. - Fail induction X using OnOne2_ind_l. - revert mfix0 mfix1 X. - refine (OnOne2_ind_l _ (fun (L : mfixpoint term) (x y : def term) => - ((red1 Σ (Γ ,,, fix_context L) (dbody x) (dbody y) - × (forall Δ0 : context, - eq_context_upto Σ' Re Rle (Γ ,,, fix_context L) Δ0 -> - ∑ v' : term, - red1 Σ Δ0 (dbody x) v' × eq_term_upto_univ Σ' Re Re (dbody y) v')) - × (dname x, dtype x, rarg x) = (dname y, dtype y, rarg y))) - (fun L mfix0 mfix1 o => ∑ mfix' : list (def term), - OnOne2 - (fun d d' : def term => - red1 Σ (Δ ,,, fix_context L) (dbody d) (dbody d') - × (dname d, dtype d, rarg d) = (dname d', dtype d', rarg d')) mfix0 mfix' - × All2 - (fun x y : def term => - ((eq_term_upto_univ Σ' Re Re (dtype x) (dtype y) - × eq_term_upto_univ Σ' Re Re (dbody x) (dbody y)) × - (rarg x = rarg y)) * - eq_binder_annot (dname x) (dname y)) mfix1 mfix') _ _). - - intros L x y l [[p1 p2] p3]. + eq_mfixpoint (eq_term_upto_univ Σ' cmp_universe cmp_sort Conv) mfix1 mfix'). + { set (Ξ := fix_context _) in *. clearbody Ξ. + induction X. + - destruct p as [[p1 p2] p3]. assert ( - e' : eq_context_upto Σ' Re Rle (Γ ,,, fix_context L) (Δ ,,, fix_context L) + e' : eq_context_upto Σ' cmp_universe cmp_sort pb (Γ ,,, Ξ) (Δ ,,, Ξ) ). { eapply eq_context_upto_cat ; eauto. reflexivity. } eapply p2 in e' as hh. destruct hh as [? [? ?]]. @@ -484,7 +428,7 @@ Proof. * eapply All2_same. intros. repeat split ; eauto. all: eapply eq_term_upto_univ_refl ; eauto. - - intros L x l l' h [? [? ?]]. + - destruct IHX as [? [? ?]]. eexists. constructor. + eapply OnOne2_tl. eassumption. + constructor ; eauto. @@ -501,11 +445,7 @@ Proof. (d.(dname), d.(dbody), d.(rarg)) = (d'.(dname), d'.(dbody), d'.(rarg)) ) mfix0 mfix' * - All2 (fun x y => - eq_term_upto_univ Σ' Re Re (dtype x) (dtype y) * - eq_term_upto_univ Σ' Re Re (dbody x) (dbody y) * - (rarg x = rarg y) * - eq_binder_annot (dname x) (dname y))%type mfix1 mfix' + eq_mfixpoint (eq_term_upto_univ Σ' cmp_universe cmp_sort Conv) mfix1 mfix' ). { induction X. - destruct p as [[p1 p2] p3]. @@ -537,50 +477,12 @@ Proof. (d.(dname), d.(dtype), d.(rarg)) = (d'.(dname), d'.(dtype), d'.(rarg)) ) mfix0 mfix' * - All2 (fun x y => - eq_term_upto_univ Σ' Re Re (dtype x) (dtype y) * - eq_term_upto_univ Σ' Re Re (dbody x) (dbody y) * - (rarg x = rarg y) * - eq_binder_annot (dname x) (dname y))%type mfix1 mfix'). - { (* Maybe we should use a lemma using firstn or skipn to keep - fix_context intact. Anything general? - *) - Fail induction X using OnOne2_ind_l. - (* This FAILs because it reduces the type of X before unifying - unfortunately... - *) - change ( - OnOne2 - ((fun L (x y : def term) => - (red1 Σ (Γ ,,, fix_context L) (dbody x) (dbody y) - × (forall Δ : context, - eq_context_upto Σ' Re Rle (Γ ,,, fix_context L) Δ -> - ∑ v' : term, - red1 Σ Δ (dbody x) v' × eq_term_upto_univ Σ' Re Re (dbody y) v' )) - × (dname x, dtype x, rarg x) = (dname y, dtype y, rarg y)) mfix0) mfix0 mfix1 - ) in X. - Fail induction X using OnOne2_ind_l. - revert mfix0 mfix1 X. - refine (OnOne2_ind_l _ (fun (L : mfixpoint term) (x y : def term) => - (red1 Σ (Γ ,,, fix_context L) (dbody x) (dbody y) - × (forall Δ0 : context, - eq_context_upto Σ' Re Rle (Γ ,,, fix_context L) Δ0 -> - ∑ v' : term, - red1 Σ Δ0 (dbody x) v' × eq_term_upto_univ Σ' Re Re (dbody y) v' )) - × (dname x, dtype x, rarg x) = (dname y, dtype y, rarg y)) (fun L mfix0 mfix1 o => ∑ mfix' : list (def term), - (OnOne2 - (fun d d' : def term => - red1 Σ (Δ ,,, fix_context L) (dbody d) (dbody d') - × (dname d, dtype d, rarg d) = (dname d', dtype d', rarg d')) mfix0 mfix' - × All2 - (fun x y : def term => - ((eq_term_upto_univ Σ' Re Re (dtype x) (dtype y) - × eq_term_upto_univ Σ' Re Re (dbody x) (dbody y)) × - rarg x = rarg y) * - eq_binder_annot (dname x) (dname y)) mfix1 mfix')) _ _). - - intros L x y l [[p1 p2] p3]. + eq_mfixpoint (eq_term_upto_univ Σ' cmp_universe cmp_sort Conv) mfix1 mfix'). + { set (Ξ := fix_context _) in *. clearbody Ξ. + induction X. + -destruct p as [[p1 p2] p3]. assert ( - e' : eq_context_upto Σ' Re Rle (Γ ,,, fix_context L) (Δ ,,, fix_context L) + e' : eq_context_upto Σ' cmp_universe cmp_sort pb (Γ ,,, Ξ) (Δ ,,, Ξ) ). { eapply eq_context_upto_cat ; eauto. reflexivity. } eapply p2 in e' as hh. destruct hh as [? [? ?]]. @@ -594,7 +496,7 @@ Proof. * eapply All2_same. intros. repeat split ; eauto. all: eapply eq_term_upto_univ_refl ; eauto. - - intros L x l l' h [? [? ?]]. + - destruct IHX as [? [? ?]]. eexists. constructor. + eapply OnOne2_tl. eassumption. + constructor ; eauto. @@ -607,7 +509,7 @@ Proof. + constructor; assumption. - assert (h : ∑ ll, OnOne2 (red1 Σ Δ) (array_value arr) ll * - All2 (eq_term_upto_univ Σ' Re Re) value ll + All2 (eq_term_upto_univ Σ' cmp_universe cmp_sort Conv) value ll ). { induction X. - destruct p as [p1 p2]. @@ -641,17 +543,35 @@ Proof. eapply All2_refl; reflexivity. Qed. -Lemma eq_context_gen_context_assumptions {eq leq Γ Δ} : - eq_context_gen eq leq Γ Δ -> +Lemma eq_context_upto_names_context_assumptions {Γ Δ} : + eq_context_upto_names Γ Δ -> + context_assumptions Γ = context_assumptions Δ. +Proof. + induction 1; simpl; auto. + destruct r => /= //; try lia. +Qed. + +Lemma eq_context_upto_names_extended_subst {Γ Δ n} : + eq_context_upto_names Γ Δ -> + extended_subst Γ n = extended_subst Δ n. +Proof. + induction 1 in n |- *; cbn; auto. + destruct r; subst; cbn. f_equal; auto. + rewrite IHX. + now rewrite (eq_context_upto_names_context_assumptions X). +Qed. + +Lemma eq_context_gen_context_assumptions {cmp pb Γ Δ} : + eq_context_gen cmp pb Γ Δ -> context_assumptions Γ = context_assumptions Δ. Proof. induction 1; simpl; auto; destruct p => /= //; try lia. Qed. -Lemma eq_context_extended_subst {Σ Re Rle Γ Δ k} : - eq_context_gen (eq_term_upto_univ Σ Re Re) (eq_term_upto_univ Σ Re Rle) Γ Δ -> - All2 (eq_term_upto_univ Σ Re Re) (extended_subst Γ k) (extended_subst Δ k). +Lemma eq_context_extended_subst {Σ cmp_universe cmp_sort pb Γ Δ k} : + eq_context_gen (fun pb => eq_term_upto_univ Σ cmp_universe cmp_sort pb) pb Γ Δ -> + All2 (eq_term_upto_univ Σ cmp_universe cmp_sort Conv) (extended_subst Γ k) (extended_subst Δ k). Proof. intros Heq. induction Heq in k |- *; simpl. @@ -661,58 +581,44 @@ Proof. * constructor. + rewrite (eq_context_gen_context_assumptions Heq). len. rewrite (All2_fold_length Heq). - eapply eq_term_upto_univ_substs; eauto. tc. - eapply eq_term_upto_univ_lift, e0. + eapply eq_term_upto_univ_substs; eauto; tc. + eapply eq_term_upto_univ_lift, eqb. + eapply IHHeq. Qed. -Lemma eq_context_gen_eq_context_upto Σ Re Rle Γ Γ' : - RelationClasses.Reflexive Re -> - RelationClasses.Reflexive Rle -> - eq_context_gen eq eq Γ Γ' -> - eq_context_upto Σ Re Rle Γ Γ'. +Lemma eq_context_upto_names_eq_context_upto Σ cmp_universe cmp_sort pb : + RelationClasses.Reflexive (cmp_universe Conv) -> + RelationClasses.Reflexive (cmp_universe pb) -> + RelationClasses.Reflexive (cmp_sort Conv) -> + RelationClasses.Reflexive (cmp_sort pb) -> + subrelation eq_context_upto_names (eq_context_upto Σ cmp_universe cmp_sort pb). Proof. - intros. + intros ** Γ Γ' X. + apply All2_fold_All2 in X. eapply All2_fold_impl; tea. cbn; intros ????. move => []; constructor; subst; auto; reflexivity. Qed. -Lemma red1_eq_context_upto_univ_l {Σ Σ' Re Rle Γ ctx ctx' ctx''} : - RelationClasses.Reflexive Re -> - RelationClasses.Reflexive Rle -> - RelationClasses.Transitive Re -> - RelationClasses.Transitive Rle -> - SubstUnivPreserving Re -> - SubstUnivPreserving Rle -> - RelationClasses.subrelation Re Rle -> - eq_context_gen (eq_term_upto_univ Σ' Re Re) - (eq_term_upto_univ Σ' Re Re) ctx ctx' -> - OnOne2_local_env (on_one_decl - (fun (Γ' : context) (u v : term) => - forall (Rle : Relation_Definitions.relation Universe.t) - (napp : nat) (u' : term), - RelationClasses.Reflexive Re -> - RelationClasses.Reflexive Rle -> - RelationClasses.Transitive Re -> - RelationClasses.Transitive Rle -> - SubstUnivPreserving Re -> - SubstUnivPreserving Rle -> - (forall x y : Universe.t, Re x y -> Rle x y) -> - eq_term_upto_univ_napp Σ' Re Rle napp u u' -> +Lemma red1_eq_context_upto_univ_l {Σ Σ' cmp_universe cmp_sort Γ ctx ctx' ctx''} : + RelationClasses.PreOrder (cmp_universe Conv) -> + RelationClasses.PreOrder (cmp_sort Conv) -> + eq_context_gen (fun pb => eq_term_upto_univ Σ' cmp_universe cmp_sort pb) Conv ctx ctx' -> + OnOne2_local_env (fun (Γ' : context) => on_one_decl + (fun (u v : term) => forall (pb : conv_pb) (napp : nat) (u' : term), + eq_term_upto_univ_napp Σ' cmp_universe cmp_sort pb napp u u' -> ∑ v' : term, - red1 Σ (Γ,,, Γ') u' v' - × eq_term_upto_univ_napp Σ' Re Rle napp v v')) ctx ctx'' -> + red1 Σ (Γ ,,, Γ') u' v' + × eq_term_upto_univ_napp Σ' cmp_universe cmp_sort pb napp v v')) ctx ctx'' -> ∑ pctx, red1_ctx_rel Σ Γ ctx' pctx * - eq_context_gen (eq_term_upto_univ Σ' Re Re) (eq_term_upto_univ Σ' Re Re) ctx'' pctx. + eq_context_gen (fun pb => eq_term_upto_univ Σ' cmp_universe cmp_sort pb) Conv ctx'' pctx. Proof. - intros. - rename X into e, X0 into X. + intros preorder_univ_conv preorder_sort_conv e X. induction X in e, ctx' |- *. - red in p. simpl in p. depelim e. depelim c. destruct p as [-> p]. - eapply p in e1 as hh ; eauto. + eapply p in eqt as hh ; eauto. destruct hh as [? [? ?]]. eapply red1_eq_context_upto_l in r; cycle -1. { eapply eq_context_upto_cat; tea. @@ -726,7 +632,7 @@ Proof. - depelim e. depelim c. destruct p as [-> [[p ->]|[p ->]]]. - { eapply p in e2 as hh ; eauto. + { eapply p in eqt as hh ; eauto. destruct hh as [? [? ?]]. eapply red1_eq_context_upto_l in r; cycle -1. { eapply eq_context_upto_cat; tea. @@ -738,7 +644,7 @@ Proof. left. split; tea. reflexivity. + constructor. all: eauto. constructor; auto. now transitivity x. } - { eapply p in e1 as hh ; eauto. + { eapply p in eqb as hh ; eauto. destruct hh as [? [? ?]]. eapply red1_eq_context_upto_l in r; cycle -1. { eapply eq_context_upto_cat; tea. @@ -766,9 +672,10 @@ Qed. *) Global Instance eq_context_upto_univ_subst_preserved {cf:checker_flags} Σ - (Re Rle : ConstraintSet.t -> Universe.t -> Universe.t -> Prop) - {he: SubstUnivPreserved Re} {hle: SubstUnivPreserved Rle} - : SubstUnivPreserved (fun φ => eq_context_upto Σ (Re φ) (Rle φ)). + (cmp_universe : forall _ _ (_ _ : Universe.t), Prop) (cmp_sort : forall _ _ (_ _ : sort), Prop) pb + {univ_conv: SubstUnivPreserved (fun φ => cmp_universe φ Conv)} {univ_pb: SubstUnivPreserved (fun φ => cmp_universe φ pb)} + {sort_conv: SubstUnivPreserved (fun φ => cmp_sort φ Conv)} {sort_pb: SubstUnivPreserved (fun φ => cmp_sort φ pb)} + : SubstUnivPreserved (fun φ => eq_context_upto Σ (cmp_universe φ) (cmp_sort φ) pb). Proof. intros φ φ' u vc Γ Δ eqc. eapply All2_fold_map. @@ -777,86 +684,113 @@ Proof. destruct X; constructor; cbn; auto; eapply eq_term_upto_univ_subst_preserved; tc; eauto. Qed. -Lemma eq_context_gen_eq_univ_subst_preserved u Γ Δ : - eq_context_gen eq eq Γ Δ -> - eq_context_gen eq eq (subst_instance u Γ) (subst_instance u Δ). +Lemma eq_context_upto_names_univ_subst_preserved u Γ Δ : + eq_context_upto_names Γ Δ -> + eq_context_upto_names (subst_instance u Γ) (subst_instance u Δ). Proof. intros onctx. - eapply All2_fold_map. - eapply All2_fold_impl; tea. + eapply All2_map. + eapply All2_impl; tea. cbn; intros. destruct X; constructor; cbn; auto; now subst. Qed. -Lemma eq_term_upto_univ_subst_instance' {cf:checker_flags} Σ Re Rle : - RelationClasses.Reflexive Re -> - SubstUnivPreserving Re -> - RelationClasses.Transitive Re -> - RelationClasses.Transitive Rle -> - RelationClasses.subrelation Re Rle -> - SubstUnivPreserved (fun _ => Re) -> - SubstUnivPreserved (fun _ => Rle) -> +Lemma eq_term_upto_univ_subst_instance' {cf:checker_flags} Σ cmp_universe cmp_sort pb : + RelationClasses.Transitive (cmp_universe Conv) -> + RelationClasses.Transitive (cmp_universe pb) -> + RelationClasses.Transitive (cmp_sort Conv) -> + RelationClasses.Transitive (cmp_sort pb) -> + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe pb) -> + SubstUnivPreserving (cmp_universe Conv) (cmp_universe Conv) -> + SubstUnivPreserving (cmp_universe Conv) (cmp_sort Conv) -> + SubstUnivPreserving (cmp_universe Conv) (cmp_sort pb) -> + SubstUnivPreserved (fun _ => cmp_universe Conv) -> + SubstUnivPreserved (fun _ => cmp_universe pb) -> + SubstUnivPreserved (fun _ => cmp_sort Conv) -> + SubstUnivPreserved (fun _ => cmp_sort pb) -> forall x y napp u1 u2, - R_universe_instance Re u1 u2 -> - eq_term_upto_univ_napp Σ Re Rle napp x y -> - eq_term_upto_univ_napp Σ Re Rle napp (subst_instance u1 x) (subst_instance u2 y). + cmp_universe_instance (cmp_universe Conv) u1 u2 -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp x y -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp (subst_instance u1 x) (subst_instance u2 y). Proof. intros. eapply eq_term_upto_univ_trans with (subst_instance u2 x); tc. now eapply eq_term_upto_univ_subst_instance. - eapply (eq_term_upto_univ_subst_preserved Σ (fun _ => Re) (fun _ => Rle) napp ConstraintSet.empty ConstraintSet.empty u2). + eapply (eq_term_upto_univ_subst_preserved Σ (fun _ => cmp_universe) (fun _ => cmp_sort) pb napp ConstraintSet.empty ConstraintSet.empty u2). red. destruct check_univs => //. assumption. Qed. -Lemma eq_context_upto_univ_subst_instance Σ Re Rle : - RelationClasses.Reflexive Re -> - SubstUnivPreserving Re -> - RelationClasses.subrelation Re Rle -> +Lemma eq_context_upto_univ_subst_instance Σ cmp_universe cmp_sort pb : + SubstUnivPreserving (cmp_universe Conv) (cmp_universe Conv) -> + SubstUnivPreserving (cmp_universe Conv) (cmp_sort Conv) -> + SubstUnivPreserving (cmp_universe Conv) (cmp_sort pb) -> forall x u1 u2, - R_universe_instance Re u1 u2 -> - eq_context_upto Σ Re Rle (subst_instance u1 x) (subst_instance u2 x). + cmp_universe_instance (cmp_universe Conv) u1 u2 -> + eq_context_upto Σ cmp_universe cmp_sort pb (subst_instance u1 x) (subst_instance u2 x). Proof. - intros ref hRe subr t. + intros substu_univ substu_sort_conv substu_sort_pb t. induction t. intros. - rewrite /subst_instance /=. constructor. - rewrite /subst_instance /=. constructor; auto. destruct a as [na [b|] ty]; cbn; constructor; cbn; eauto using eq_term_upto_univ_subst_instance. - apply eq_term_upto_univ_subst_instance; eauto. tc. Qed. -Lemma eq_context_upto_univ_subst_instance' Σ Re Rle : - RelationClasses.Reflexive Re -> - RelationClasses.Reflexive Rle -> - RelationClasses.Transitive Re -> - RelationClasses.Transitive Rle -> - SubstUnivPreserving Re -> - RelationClasses.subrelation Re Rle -> +Lemma eq_context_upto_univ_subst_instance' {cf:checker_flags} Σ cmp_universe cmp_sort pb : + RelationClasses.Transitive (cmp_universe Conv) -> + RelationClasses.Transitive (cmp_universe pb) -> + RelationClasses.Transitive (cmp_sort Conv) -> + RelationClasses.Transitive (cmp_sort pb) -> + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe pb) -> + SubstUnivPreserving (cmp_universe Conv) (cmp_universe Conv) -> + SubstUnivPreserving (cmp_universe Conv) (cmp_sort Conv) -> + SubstUnivPreserving (cmp_universe Conv) (cmp_sort pb) -> + SubstUnivPreserved (fun _ => cmp_universe Conv) -> + SubstUnivPreserved (fun _ => cmp_universe pb) -> + SubstUnivPreserved (fun _ => cmp_sort Conv) -> + SubstUnivPreserved (fun _ => cmp_sort pb) -> forall x y u1 u2, - R_universe_instance Re u1 u2 -> - eq_context_gen eq eq x y -> - eq_context_upto Σ Re Rle (subst_instance u1 x) (subst_instance u2 y). + cmp_universe_instance (cmp_universe Conv) u1 u2 -> + eq_context_upto Σ cmp_universe cmp_sort pb x y -> + eq_context_upto Σ cmp_universe cmp_sort pb (subst_instance u1 x) (subst_instance u2 y). Proof. - intros ref refl' tr trle hRe subr x y u1 u2 ru eqxy. + intros ???????????? x y u1 u2 ru eqxy. eapply All2_fold_trans. intros ?????????. eapply compare_decl_trans. eapply eq_term_upto_univ_trans; tc. eapply eq_term_upto_univ_trans; tc. - eapply eq_context_gen_eq_context_upto; tea. - eapply eq_context_gen_eq_univ_subst_preserved; tea. eapply eq_context_upto_univ_subst_instance; tc. tea. + eapply eq_context_upto_univ_subst_preserved with (cmp_universe := fun _ => cmp_universe) (cmp_sort := fun _ => cmp_sort); tea; tc. + unfold_univ_rel eqn:He. + instantiate (1 := CS.empty). instantiate (1 := CS.empty) in Hv. + apply Hv. +Qed. + +Lemma eq_context_upto_names_subst_instance Σ cmp_universe cmp_sort pb : + SubstUnivPreserving (cmp_universe Conv) (cmp_universe Conv) -> + SubstUnivPreserving (cmp_universe Conv) (cmp_sort Conv) -> + SubstUnivPreserving (cmp_universe Conv) (cmp_sort pb) -> + forall x y u1 u2, + cmp_universe_instance (cmp_universe Conv) u1 u2 -> + eq_context_upto_names x y -> + eq_context_upto Σ cmp_universe cmp_sort pb (subst_instance u1 x) (subst_instance u2 y). +Proof. + intros substu_univ substu_sort_conv substu_sort_pb x y u1 u2 ru eqxy. + induction eqxy; cbn; constructor; eauto. + destruct r; constructor; cbn; eauto. + all: now apply eq_term_upto_univ_subst_instance. Qed. (*Lemma eq_context_upto_eq_univ_subst_instance {cf:checker_flags} Σ φ Re Rle : RelationClasses.Reflexive (Re φ) -> RelationClasses.Transitive (Re φ) -> RelationClasses.Transitive (Rle φ) -> - SubstUnivPreserving (Re φ) -> + SubstUnivPreserving (Re φ) (Re φ) -> RelationClasses.subrelation (Re φ) (Rle φ) -> SubstUnivPreserved Re -> SubstUnivPreserved Rle -> forall x y u1 u2, - R_universe_instance (Re φ) u1 u2 -> + cmp_universe_instance (Re φ) u1 u2 -> valid_constraints φ (subst_instance_cstrs u1 φ) -> eq_context_upto Σ eq eq x y -> eq_context_upto Σ (Re φ) (Rle φ) (subst_instance u1 x) (subst_instance u2 y). @@ -872,21 +806,24 @@ Proof. eapply eq_context_upto_univ_subst_instance; tc. tea. Qed.*) -Lemma red1_eq_term_upto_univ_l {Σ Σ' : global_env} Re Rle napp Γ u v u' : - RelationClasses.Reflexive Re -> - RelationClasses.Reflexive Rle -> - RelationClasses.Transitive Re -> - RelationClasses.Transitive Rle -> - SubstUnivPreserving Re -> - SubstUnivPreserving Rle -> - RelationClasses.subrelation Re Rle -> - eq_term_upto_univ_napp Σ' Re Rle napp u u' -> +Lemma red1_eq_term_upto_univ_l {Σ Σ' : global_env} cmp_universe cmp_sort pb napp Γ u v u' : + RelationClasses.PreOrder (cmp_universe Conv) -> + RelationClasses.PreOrder (cmp_universe pb) -> + RelationClasses.PreOrder (cmp_sort Conv) -> + RelationClasses.PreOrder (cmp_sort pb) -> + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe pb) -> + RelationClasses.subrelation (cmp_sort Conv) (cmp_sort pb) -> + SubstUnivPreserving (cmp_universe Conv) (cmp_universe Conv) -> + SubstUnivPreserving (cmp_universe Conv) (cmp_sort Conv) -> + SubstUnivPreserving (cmp_universe Conv) (cmp_sort pb) -> + eq_term_upto_univ_napp Σ' cmp_universe cmp_sort pb napp u u' -> red1 Σ Γ u v -> ∑ v', red1 Σ Γ u' v' * - eq_term_upto_univ_napp Σ' Re Rle napp v v'. + eq_term_upto_univ_napp Σ' cmp_universe cmp_sort pb napp v v'. Proof. - intros he he' tRe tRle hle hle' hR e h. - induction h in napp, u', e, he, he', tRe, tRle, Rle, hle, hle', hR |- * using red1_ind_all. + intros preorder_univ_conv preorder_univ_pb preorder_sort_conv preorder_sort_pb + sub_univ sub_sort substu_univ substu_sort_conv substu_sort_pb e h. + induction h in pb, napp, u', e, preorder_univ_pb, preorder_sort_pb, sub_univ, sub_sort, substu_sort_pb |- * using red1_ind_all. all: try solve [ dependent destruction e ; edestruct IHh as [? [? ?]] ; [ .. | eassumption | ] ; eauto ; tc ; @@ -899,30 +836,24 @@ Proof. idtac. 10,15:solve [ dependent destruction e ; - edestruct (IHh Rle) as [? [? ?]] ; [ .. | tea | ] ; eauto; + edestruct (IHh pb) as [? [? ?]] ; [ .. | tea | ] ; eauto; clear h; lazymatch goal with | r : red1 _ (?Γ,, vass ?na ?A) ?u ?v, - e : eq_term_upto_univ_napp _ _ _ _ ?A ?B + e : eq_term_upto_univ_napp _ _ _ _ _ ?A ?B |- _ => let hh := fresh "hh" in eapply red1_eq_context_upto_l in r as hh ; revgoals; - [ - constructor (* with (nb := na) *) ; [ - eapply (eq_context_upto_refl _ Re Re); eauto + [ constructor ; [ + eapply (eq_context_upto_refl _ _ _ Conv); eauto ; tc | constructor; tea ] - | reflexivity - | assumption - | assumption - | assumption - | assumption - | destruct hh as [? [? ?]] - ] + | tc ..]; + destruct hh as [? [? ?]] end; eexists ; split; [ solve [ econstructor ; eauto ] | constructor ; eauto ; - eapply eq_term_upto_univ_trans ; eauto ; + eapply eq_term_upto_univ_trans ; eauto ; tc ; eapply eq_term_upto_univ_leq ; eauto ] ]. @@ -940,13 +871,13 @@ Proof. - dependent destruction e. eexists. split. + constructor. eassumption. - + eapply eq_term_upto_univ_refl ; assumption. + + eapply eq_term_upto_univ_refl ; tc. - dependent destruction e. apply eq_term_upto_univ_mkApps_l_inv in e0 as [? [? [[h1 h2] h3]]]. subst. dependent destruction h1. - eapply All2_nth_error_Some in a as [t' [hnth [eqctx eqbod]]]; tea. - have lenctxass := eq_context_gen_context_assumptions eqctx. - have lenctx := All2_fold_length eqctx. + eapply All2_nth_error_Some in e1 as [t' [hnth [eqctx eqbod]]]; tea. + have lenctxass := eq_context_upto_names_context_assumptions eqctx. + have lenctx := All2_length eqctx. eexists. split. + constructor; tea. epose proof (All2_length h2). congruence. @@ -963,18 +894,18 @@ Proof. eapply eq_context_extended_subst; tea. rewrite /inst_case_branch_context. eapply eq_context_upto_subst_context; tc. - eapply eq_context_upto_univ_subst_instance'. - 7,8:tea. all:tc. apply e. + eapply eq_context_upto_names_subst_instance. + 4,5:tea. all:tc. apply e. now eapply All2_rev, e. } now eapply All2_rev, All2_skipn. - apply eq_term_upto_univ_napp_mkApps_l_inv in e as [? [? [[h1 h2] h3]]]. subst. dependent destruction h1. unfold unfold_fix in H. case_eq (nth_error mfix idx) ; - try (intros e ; rewrite e in H ; discriminate H). - intros d e. rewrite e in H. inversion H. subst. clear H. - eapply All2_nth_error_Some in a as hh ; try eassumption. - destruct hh as [d' [e' [[[? ?] erarg] eann]]]. + try (intros hnth ; rewrite hnth in H ; discriminate H). + intros d hnth. rewrite hnth in H. inversion H. subst. clear H. + eapply All2_nth_error_Some in e as hh ; try eassumption. + destruct hh as [d' [e' (? & ? & erarg & eann)]]. unfold is_constructor in H0. case_eq (nth_error args (rarg d)) ; try (intros bot ; rewrite bot in H0 ; discriminate H0). @@ -989,9 +920,9 @@ Proof. eapply isConstruct_app_eq_term_l ; eassumption. + eapply eq_term_upto_univ_napp_mkApps. * eapply eq_term_upto_univ_substs ; eauto. - -- eapply (eq_term_upto_univ_leq _ _ _ 0) ; eauto with arith. + -- eapply (eq_term_upto_univ_leq _ _ _ _ 0) ; eauto with arith. -- unfold fix_subst. - apply All2_length in a as el. rewrite <- el. + apply All2_length in e as el. rewrite <- el. generalize #|mfix|. intro n. induction n. ++ constructor. @@ -1003,18 +934,18 @@ Proof. dependent destruction h1. unfold unfold_cofix in H. destruct (nth_error mfix idx) eqn:hnth; noconf H. - eapply All2_nth_error_Some in a0 as hh ; tea. - destruct hh as [d' [e' [[[? ?] erarg] eann]]]. + eapply All2_nth_error_Some in e0 as hh ; tea. + destruct hh as [d' [e' (? & ? & erarg & eann)]]. eexists. split. + eapply red_cofix_case. unfold unfold_cofix. rewrite e'. reflexivity. + constructor. all: eauto. eapply eq_term_upto_univ_mkApps. all: eauto. eapply eq_term_upto_univ_substs ; eauto; try exact _. - eapply (eq_term_upto_univ_leq _ _ _ 0); auto with arith. - typeclasses eauto. + eapply (eq_term_upto_univ_leq _ _ _ _ 0); auto with arith. + 1,2: typeclasses eauto. unfold cofix_subst. - apply All2_length in a0 as el. rewrite <- el. + apply All2_length in e0 as el. rewrite <- el. generalize #|mfix|. intro n. induction n. * constructor. @@ -1025,20 +956,20 @@ Proof. dependent destruction h1. unfold unfold_cofix in H. case_eq (nth_error mfix idx) ; - try (intros e ; rewrite e in H ; discriminate H). - intros d e. rewrite e in H. inversion H. subst. clear H. + try (intros hnth ; rewrite hnth in H ; discriminate H). + intros d hnth. rewrite hnth in H. inversion H. subst. clear H. eapply All2_nth_error_Some in e as hh ; try eassumption. - destruct hh as [d' [e' [[[? ?] erarg] eann]]]. + destruct hh as [d' [e' (? & ? & erarg & eann)]]. eexists. split. + eapply red_cofix_proj. unfold unfold_cofix. rewrite e'. reflexivity. + constructor. eapply eq_term_upto_univ_mkApps. all: eauto. eapply eq_term_upto_univ_substs ; eauto; try exact _. - eapply (eq_term_upto_univ_leq _ _ _ 0); auto with arith. - typeclasses eauto. + eapply (eq_term_upto_univ_leq _ _ _ _ 0); auto with arith. + 1,2: typeclasses eauto. unfold cofix_subst. - apply All2_length in a as el. rewrite <- el. + apply All2_length in e as el. rewrite <- el. generalize #|mfix|. intro n. induction n. * constructor. @@ -1047,7 +978,7 @@ Proof. - dependent destruction e. eexists. split. + econstructor. all: eauto. - + eapply (eq_term_upto_univ_leq _ _ _ 0); tas. auto. auto with arith. + + eapply (eq_term_upto_univ_leq _ _ _ _ 0); tas. auto. auto with arith. now apply eq_term_upto_univ_subst_instance. - dependent destruction e. apply eq_term_upto_univ_mkApps_l_inv in e as [? [? [[h1 h2] h3]]]. subst. @@ -1057,52 +988,46 @@ Proof. eexists. split. + constructor. eassumption. + eapply eq_term_upto_univ_leq ; eauto. - eapply eq_term_eq_term_napp; auto. typeclasses eauto. + eapply eq_term_eq_term_napp; auto. - dependent destruction e. edestruct IHh as [? [? ?]] ; [ .. | eassumption | ] ; eauto. clear h. lazymatch goal with | r : red1 _ (?Γ,, vdef ?na ?a ?A) ?u ?v, - e1 : eq_term_upto_univ _ _ _ ?A ?B, - e2 : eq_term_upto_univ _ _ _ ?a ?b + e1 : eq_term_upto_univ _ _ _ _ ?A ?B, + e2 : eq_term_upto_univ _ _ _ _ ?a ?b |- _ => let hh := fresh "hh" in eapply red1_eq_context_upto_l in r as hh ; revgoals ; [ constructor (* with (nb := na) *) ; [ - eapply (eq_context_upto_refl _ Re Re) ; eauto + eapply (eq_context_upto_refl _ _ _ Conv) ; eauto ; tc | econstructor; tea ] - | reflexivity - | assumption - | assumption - | assumption - | assumption - | destruct hh as [? [? ?]] - ] + | tc ..]; + destruct hh as [? [? ?]] end. eexists. split. + eapply letin_red_body ; eauto. + constructor ; eauto. - eapply eq_term_upto_univ_trans ; eauto. + eapply eq_term_upto_univ_trans ; eauto ; tc. eapply eq_term_upto_univ_leq ; eauto. - dependent destruction e. destruct e as [? [? [? ?]]]. eapply OnOne2_prod_inv in X as [_ X]. assert (h : ∑ args, OnOne2 (red1 Σ Γ) (pparams p') args * - All2 (eq_term_upto_univ Σ' Re Re) params' args + All2 (eq_term_upto_univ Σ' cmp_universe cmp_sort Conv) params' args ). { destruct p, p' as []; cbn in *. - induction X in a0, pparams, pparams0, X |- *. - - depelim a0. - eapply p in e as hh ; eauto. + induction X in a, pparams, pparams0, X |- *. + - depelim a. + eapply p in e as hh ; eauto; tc. destruct hh as [? [? ?]]. eexists. split. + constructor; tea. + constructor. all: eauto. - + tc. - - depelim a0. - destruct (IHX _ a0) as [? [? ?]]. + - depelim a. + destruct (IHX _ a) as [? [? ?]]. eexists. split. + eapply OnOne2_tl. eassumption. + constructor. all: eauto. @@ -1117,12 +1042,12 @@ Proof. eapply IHh in e => //. destruct e as [v' [red eq]]. eapply red1_eq_context_upto_l in red. - 7:{ eapply eq_context_upto_cat. + 6:{ eapply eq_context_upto_cat. 2:{ instantiate (1:=PCUICCases.inst_case_predicate_context p'). rewrite /inst_case_predicate_context /inst_case_context. eapply eq_context_upto_subst_context; tc. - eapply eq_context_upto_univ_subst_instance'. - 7,8:tea. all:tc. + eapply eq_context_upto_names_subst_instance. + 4,5:tea. all:tc. now eapply All2_rev. } eapply eq_context_upto_refl; tc. } all:tc. @@ -1137,22 +1062,19 @@ Proof. eapply OnOne2_prod_assoc in X as [_ X]. assert (h : ∑ brs0, OnOne2 (fun br br' => - on_Trel_eq (red1 Σ (Γ ,,, inst_case_branch_context p' br)) bbody bcontext br br') brs' brs0 * - All2 (fun x y => - eq_context_gen eq eq (bcontext x) (bcontext y) * - (eq_term_upto_univ Σ' Re Re (bbody x) (bbody y)) - )%type brs'0 brs0 + on_Trel_eq (red1 Σ (Γ ,,, inst_case_branch_context p' br)) bbody bcontext br br') brs' brs0 × + eq_branches (fun t u => eq_term_upto_univ Σ' cmp_universe cmp_sort Conv t u) brs'0 brs0 ). - { induction X in a, brs' |- *. + { induction X in e1, brs' |- *. - destruct p0 as [p2 p3]. - dependent destruction a. destruct p0 as [h1 h2]. + dependent destruction e1. destruct e1 as [h1 h2]. eapply p2 in h2 as hh ; eauto. destruct hh as [? [? ?]]. - eapply (red1_eq_context_upto_l (Re:=Re) (Rle:=Rle) (Δ := Γ ,,, inst_case_branch_context p' y)) in r; cycle -1. + eapply (red1_eq_context_upto_l (cmp_universe := cmp_universe) (cmp_sort := cmp_sort) (pb := pb) (Δ := Γ ,,, inst_case_branch_context p' y)) in r; cycle -1. { eapply eq_context_upto_cat; tea. reflexivity. rewrite /inst_case_branch_context /inst_case_context. eapply eq_context_upto_subst_context; tc. - eapply eq_context_upto_univ_subst_instance'. 7,8:tea. all:tc. + eapply eq_context_upto_names_subst_instance. 4,5:tea. all:tc. apply e. apply All2_rev, e. } all:tc. destruct r as [v' [redv' eqv']]. @@ -1162,8 +1084,8 @@ Proof. + constructor. all: eauto. split ; eauto. cbn. transitivity (bcontext hd) ; eauto. now rewrite p3. simpl. now transitivity x. - - dependent destruction a. - destruct (IHX _ a) as [? [? ?]]. + - dependent destruction e1. + destruct (IHX _ e2) as [? [? ?]]. eexists. split. + eapply OnOne2_tl. eassumption. + constructor. all: eauto. @@ -1176,17 +1098,16 @@ Proof. - dependent destruction e. assert (h : ∑ args, OnOne2 (red1 Σ Γ) args' args * - All2 (eq_term_upto_univ Σ' Re Re) l' args + All2 (eq_term_upto_univ Σ' cmp_universe cmp_sort Conv) l' args ). { induction X in a, args' |- *. - destruct p as [p1 p2]. dependent destruction a. - eapply p2 in e as hh ; eauto. + eapply p2 in e as hh ; eauto; tc. destruct hh as [? [? ?]]. eexists. split. + constructor. eassumption. + constructor. all: eauto. - + tc. - dependent destruction a. destruct (IHX _ a) as [? [? ?]]. eexists. split. @@ -1203,18 +1124,14 @@ Proof. red1 Σ Γ d0.(dtype) d1.(dtype) × (d0.(dname), d0.(dbody), d0.(rarg)) = (d1.(dname), d1.(dbody), d1.(rarg)) - ) mfix' mfix * - All2 (fun x y => - eq_term_upto_univ Σ' Re Re x.(dtype) y.(dtype) * - eq_term_upto_univ Σ' Re Re x.(dbody) y.(dbody) * - (x.(rarg) = y.(rarg)) * - eq_binder_annot (dname x) (dname y))%type mfix1 mfix + ) mfix' mfix × + eq_mfixpoint (eq_term_upto_univ Σ' cmp_universe cmp_sort Conv) mfix1 mfix ). - { induction X in a, mfix' |- *. + { induction X in e, mfix' |- *. - destruct p as [[p1 p2] p3]. - dependent destruction a. - destruct p as [[[h1 h2] h3] h4]. - eapply p2 in h1 as hh ; eauto. + dependent destruction e. + destruct p as (h1 & h2 & h3 & h4). + eapply p2 in h1 as hh ; eauto; tc. destruct hh as [? [? ?]]. eexists. split. + constructor. @@ -1223,9 +1140,8 @@ Proof. + constructor. all: eauto. simpl. inversion p3. repeat split ; eauto. - + tc. - - dependent destruction a. destruct p as [[[h1 h2] h3] h4]. - destruct (IHX _ a) as [? [? ?]]. + - dependent destruction e. destruct p as (h1 & h2 & h3 & h4). + destruct (IHX _ e) as [? [? ?]]. eexists. split. + eapply OnOne2_tl. eassumption. + constructor. all: eauto. @@ -1240,83 +1156,46 @@ Proof. red1 Σ (Γ ,,, fix_context mfix0) x.(dbody) y.(dbody) × (dname x, dtype x, rarg x) = (dname y, dtype y, rarg y) ) mfix' mfix * - All2 (fun x y => - eq_term_upto_univ Σ' Re Re x.(dtype) y.(dtype) * - eq_term_upto_univ Σ' Re Re x.(dbody) y.(dbody) * - (x.(rarg) = y.(rarg)) * - eq_binder_annot (dname x) (dname y)) mfix1 mfix - ). - { revert mfix' a. - refine (OnOne2_ind_l _ (fun L x y => (red1 Σ (Γ ,,, fix_context L) (dbody x) (dbody y) - × (forall Rle napp (u' : term), - RelationClasses.Reflexive Re -> - RelationClasses.Reflexive Rle -> - RelationClasses.Transitive Re -> - RelationClasses.Transitive Rle -> - SubstUnivPreserving Re -> - SubstUnivPreserving Rle -> - (forall u u'0 : Universe.t, Re u u'0 -> Rle u u'0) -> - eq_term_upto_univ_napp Σ' Re Rle napp (dbody x) u' -> - ∑ v' : term, - red1 Σ (Γ ,,, fix_context L) u' v' - × eq_term_upto_univ_napp Σ' Re Rle napp (dbody y) v' )) - × (dname x, dtype x, rarg x) = (dname y, dtype y, rarg y)) (fun L mfix0 mfix1 o => forall mfix', All2 - (fun x y : def term => - ((eq_term_upto_univ Σ' Re Re (dtype x) (dtype y) - × eq_term_upto_univ Σ' Re Re (dbody x) (dbody y)) × - rarg x = rarg y) * eq_binder_annot (dname x) (dname y)) mfix0 mfix' -> ∑ mfix : list (def term), - ( OnOne2 - (fun x y : def term => - red1 Σ (Γ ,,, fix_context L) (dbody x) (dbody y) - × (dname x, dtype x, rarg x) = (dname y, dtype y, rarg y)) mfix' mfix ) * - ( All2 - (fun x y : def term => - ((eq_term_upto_univ Σ' Re Re (dtype x) (dtype y) - × eq_term_upto_univ Σ' Re Re (dbody x) (dbody y)) × - (rarg x = rarg y)) * eq_binder_annot (dname x) (dname y)) mfix1 mfix )) _ _ _ _ X). - - clear X. intros L x y l [[p1 p2] p3] mfix' h. - dependent destruction h. destruct p as [[[h1 h2] h3] h4]. - eapply p2 in h2 as hh ; eauto. + eq_mfixpoint (fun t u => eq_term_upto_univ Σ' cmp_universe cmp_sort Conv t u) mfix1 mfix + ). + { set (Ξ := fix_context _) in *. clearbody Ξ. + induction X in e, mfix' |- *. + - destruct p as [[p1 p2] p3]. + dependent destruction e. + destruct p as (h1 & h2 & h3 & h4). + eapply p2 in h2 as hh ; eauto; tc. destruct hh as [? [? ?]]. eexists. split. - + constructor. constructor. + + constructor. instantiate (1 := mkdef _ _ _ _ _). - simpl. eauto. reflexivity. - + constructor. constructor; simpl. all: eauto. - inversion p3. - simpl. repeat split ; eauto. destruct y0. simpl in *. - congruence. - - clear X. intros L x l l' h ih mfix' ha. - dependent destruction ha. destruct p as [[h1 h2] h3]. - destruct (ih _ ha) as [? [? ?]]. + simpl. eauto. + + constructor. all: eauto. + simpl. inversion p3. + repeat split ; eauto. + - dependent destruction e. destruct p as (h1 & h2 & h3 & h4). + destruct (IHX _ e) as [? [? ?]]. eexists. split. - + eapply OnOne2_tl. eauto. - + constructor. constructor. all: eauto. + + eapply OnOne2_tl. eassumption. + + constructor. all: eauto. } - destruct h as [mfix [? ?]]. + destruct h as [mfix [? ?]]. clear X. assert (h : ∑ mfix, OnOne2 (fun x y => red1 Σ (Γ ,,, fix_context mfix') x.(dbody) y.(dbody) × (dname x, dtype x, rarg x) = (dname y, dtype y, rarg y) ) mfix' mfix × - All2 (fun x y => - eq_term_upto_univ Σ' Re Re x.(dtype) y.(dtype) * - eq_term_upto_univ Σ' Re Re x.(dbody) y.(dbody) * - (x.(rarg) = y.(rarg)) * - eq_binder_annot (dname x) (dname y) - ) mfix1 mfix %type + eq_mfixpoint (fun t u => eq_term_upto_univ Σ' cmp_universe cmp_sort Conv t u) mfix1 mfix ). - { clear X. - assert (hc : eq_context_upto Σ' - Re Rle + { assert (hc : eq_context_upto Σ' + cmp_universe cmp_sort pb (Γ ,,, fix_context mfix0) (Γ ,,, fix_context mfix') ). { eapply eq_context_upto_cat. - - eapply eq_context_upto_refl; assumption. - - clear -he hle tRe tRle hR a. induction a. + - eapply eq_context_upto_refl; tc. + - clear o. induction e. + constructor. - + destruct r as [[[? ?] ?] ?]. + + destruct r as (? & ? & ? & ?). eapply All2_eq_context_upto. eapply All2_rev. eapply All2_mapi. @@ -1328,26 +1207,26 @@ Proof. eapply eq_term_upto_univ_impl; eauto. all:typeclasses eauto. * eapply All2_impl ; eauto. - intros ? ? [[[? ?] ?] ?] i. split; [split|]. + intros ? ? (? & ? & ? & ?) i. split; [split|]. -- assumption. -- cbn. constructor. -- cbn. eapply eq_term_upto_univ_lift. eapply eq_term_upto_univ_impl; eauto. - typeclasses eauto. + all: typeclasses eauto. } - clear a. + clear e. eapply OnOne2_impl_exist_and_All ; try eassumption. - clear o a0. - intros x x' y [r e] [[[? ?] ?] ?]. - inversion e. clear e. + clear o e0. + intros x x' y [r e] (? & ? & ? & ?). + noconf e. eapply red1_eq_context_upto_l in r as [? [? ?]]. - 7: eassumption. all: tea. + 6: eassumption. all: tc. eexists. constructor. - instantiate (1 := mkdef _ _ _ _ _). simpl. + instantiate (1 := mkdef _ _ _ _ _). all: simpl. intuition eauto. intuition eauto. - - rewrite H1. eauto. - - eapply eq_term_upto_univ_trans; tea. + - rewrite H0. eauto. + - eapply eq_term_upto_univ_trans; tea; tc. - etransitivity ; eauto. - now simpl. } @@ -1362,17 +1241,13 @@ Proof. (d0.(dname), d0.(dbody), d0.(rarg)) = (d1.(dname), d1.(dbody), d1.(rarg)) ) mfix' mfix * - All2 (fun x y => - eq_term_upto_univ Σ' Re Re x.(dtype) y.(dtype) * - eq_term_upto_univ Σ' Re Re x.(dbody) y.(dbody) * - (x.(rarg) = y.(rarg)) * - eq_binder_annot (dname x) (dname y))%type mfix1 mfix + eq_mfixpoint (fun t u => eq_term_upto_univ Σ' cmp_universe cmp_sort Conv t u) mfix1 mfix ). - { induction X in a, mfix' |- *. + { induction X in e, mfix' |- *. - destruct p as [[p1 p2] p3]. - dependent destruction a. - destruct p as [[[h1 h2] h3] h4]. - eapply p2 in h1 as hh ; eauto. + dependent destruction e. + destruct p as (h1 & h2 & h3 & h4). + eapply p2 in h1 as hh ; eauto; tc. destruct hh as [? [? ?]]. eexists. split. + constructor. @@ -1381,9 +1256,8 @@ Proof. + constructor. all: eauto. simpl. inversion p3. repeat split ; eauto. - + tc. - - dependent destruction a. destruct p as [[h1 h2] h3]. - destruct (IHX _ a) as [? [? ?]]. + - dependent destruction e. destruct p as (h1 & h2 & h3 & h4). + destruct (IHX _ e) as [? [? ?]]. eexists. split. + eapply OnOne2_tl. eassumption. + constructor. all: eauto. @@ -1398,83 +1272,46 @@ Proof. red1 Σ (Γ ,,, fix_context mfix0) x.(dbody) y.(dbody) × (dname x, dtype x, rarg x) = (dname y, dtype y, rarg y) ) mfix' mfix * - All2 (fun x y => - eq_term_upto_univ Σ' Re Re x.(dtype) y.(dtype) * - eq_term_upto_univ Σ' Re Re x.(dbody) y.(dbody) * - (x.(rarg) = y.(rarg)) * - eq_binder_annot (dname x) (dname y) - ) mfix1 mfix + eq_mfixpoint (fun t u => eq_term_upto_univ Σ' cmp_universe cmp_sort Conv t u) mfix1 mfix ). - { revert mfix' a. - refine (OnOne2_ind_l _ (fun L x y => (red1 Σ (Γ ,,, fix_context L) (dbody x) (dbody y) - × (forall Rle napp (u' : term), - RelationClasses.Reflexive Re -> - RelationClasses.Reflexive Rle -> - RelationClasses.Transitive Re -> - RelationClasses.Transitive Rle -> - SubstUnivPreserving Re -> - SubstUnivPreserving Rle -> - (forall u u'0 : Universe.t, Re u u'0 -> Rle u u'0) -> - eq_term_upto_univ_napp Σ' Re Rle napp (dbody x) u' -> - ∑ v' : term, - red1 Σ (Γ ,,, fix_context L) u' v' - × eq_term_upto_univ_napp Σ' Re Rle napp (dbody y) v')) - × (dname x, dtype x, rarg x) = (dname y, dtype y, rarg y)) (fun L mfix0 mfix1 o => forall mfix', All2 - (fun x y : def term => - ((eq_term_upto_univ Σ' Re Re (dtype x) (dtype y) - × eq_term_upto_univ Σ' Re Re (dbody x) (dbody y)) × - rarg x = rarg y) * eq_binder_annot (dname x) (dname y)) mfix0 mfix' -> ∑ mfix : list (def term), - ( OnOne2 - (fun x y : def term => - red1 Σ (Γ ,,, fix_context L) (dbody x) (dbody y) - × (dname x, dtype x, rarg x) = (dname y, dtype y, rarg y)) mfix' mfix ) * - ( All2 - (fun x y : def term => - ((eq_term_upto_univ Σ' Re Re (dtype x) (dtype y) - × eq_term_upto_univ Σ' Re Re (dbody x) (dbody y)) × - rarg x = rarg y) * eq_binder_annot (dname x) (dname y)) mfix1 mfix )) _ _ _ _ X). - - clear X. intros L x y l [[p1 p2] p3] mfix' h. - dependent destruction h. destruct p as [[[h1 h2] h3] h4]. - eapply p2 in h2 as hh ; eauto. + { set (Ξ := fix_context _) in *. clearbody Ξ. + induction X in e, mfix' |- *. + - destruct p as [[p1 p2] p3]. + dependent destruction e. + destruct p as (h1 & h2 & h3 & h4). + eapply p2 in h2 as hh ; eauto; tc. destruct hh as [? [? ?]]. - noconf p3. hnf in H. noconf H. - eexists. split; simpl. + eexists. split. + constructor. instantiate (1 := mkdef _ _ _ _ _). simpl. eauto. + constructor. all: eauto. - simpl. repeat split ; eauto; congruence. - - clear X. intros L x l l' h ih mfix' ha. - dependent destruction ha. destruct p as [[h1 h2] h3]. - destruct (ih _ ha) as [? [? ?]]. + simpl. inversion p3. + repeat split ; eauto. + - dependent destruction e. destruct p as (h1 & h2 & h3 & h4). + destruct (IHX _ e) as [? [? ?]]. eexists. split. - + eapply OnOne2_tl. eauto. + + eapply OnOne2_tl. eassumption. + constructor. all: eauto. } - destruct h as [mfix [? ?]]. + destruct h as [mfix [? ?]]. clear X. assert (h : ∑ mfix, OnOne2 (fun x y => red1 Σ (Γ ,,, fix_context mfix') x.(dbody) y.(dbody) × (dname x, dtype x, rarg x) = (dname y, dtype y, rarg y) ) mfix' mfix × - All2 (fun x y => - eq_term_upto_univ Σ' Re Re x.(dtype) y.(dtype) * - eq_term_upto_univ Σ' Re Re x.(dbody) y.(dbody) * - (x.(rarg) = y.(rarg)) * - eq_binder_annot (dname x) (dname y) - ) mfix1 mfix + eq_mfixpoint (fun t u => eq_term_upto_univ Σ' cmp_universe cmp_sort Conv t u) mfix1 mfix ). - { clear X. - assert (hc : eq_context_upto Σ' - Re Rle + { assert (hc : eq_context_upto Σ' + cmp_universe cmp_sort pb (Γ ,,, fix_context mfix0) (Γ ,,, fix_context mfix') ). { eapply eq_context_upto_cat. - - eapply eq_context_upto_refl; assumption. - - clear -he he' hle hle' hR a. induction a. + - eapply eq_context_upto_refl; tc. + - clear o. induction e. + constructor. - + destruct r as [[[? ?] ?] ?]. + + destruct r as (? & ? & ? & ?). eapply All2_eq_context_upto. eapply All2_rev. eapply All2_mapi. @@ -1486,25 +1323,25 @@ Proof. eapply eq_term_upto_univ_impl; eauto. all:typeclasses eauto. * eapply All2_impl ; eauto. - intros ? ? [[[? ?] ?] ?] i. split; [split|]. + intros ? ? (? & ? & ? & ?) i. split; [split|]. -- assumption. -- cbn. constructor. -- cbn. eapply eq_term_upto_univ_lift. eapply eq_term_upto_univ_impl; eauto. all:typeclasses eauto. } - clear a. + clear e. eapply OnOne2_impl_exist_and_All ; try eassumption. - clear o a0. - intros x x' y [r e] [[? ?] ?]. - inversion e. clear e. + clear o e0. + intros x x' y [r e] (? & ? & ? & ?). + noconf e. eapply red1_eq_context_upto_l in r as [? [? ?]]. - 7: eassumption. all: tea. + 6: eassumption. all: tc. eexists. instantiate (1 := mkdef _ _ _ _ _). simpl. intuition eauto. - - rewrite H1. eauto. - - eapply eq_term_upto_univ_trans ; tea. + - rewrite H0. eauto. + - eapply eq_term_upto_univ_trans ; tea; tc. - etransitivity ; eauto. - congruence. } @@ -1515,19 +1352,18 @@ Proof. - depelim e. depelim o. assert (h : ∑ args, OnOne2 (red1 Σ Γ) (array_value a') args * - All2 (eq_term_upto_univ Σ' Re Re) value args + All2 (eq_term_upto_univ Σ' cmp_universe cmp_sort Conv) value args ). - { revert X a0. clear r e e0. + { revert X a0. clear c e e0. generalize (array_value arr), (array_value a'). intros l l' X a. induction X in l', a |- *. - destruct p as [p1 p2]. dependent destruction a. - eapply p2 in e as hh ; eauto. + eapply p2 in e as hh ; eauto; tc. destruct hh as [? [? ?]]. eexists. split. + constructor. eassumption. + constructor. all:eauto. - + tc. - dependent destruction a. destruct (IHX _ a) as [? [? ?]]. eexists. split. @@ -1548,123 +1384,95 @@ Proof. + do 2 constructor; eauto. Qed. -Lemma Forall2_flip {A} (R : A -> A -> Prop) (x y : list A) : - Forall2 (flip R) y x -> Forall2 R x y. -Proof. - induction 1; constructor; auto. -Qed. - -Lemma R_universe_instance_flip R u u' : - R_universe_instance (flip R) u' u -> - R_universe_instance R u u'. -Proof. - unfold R_universe_instance. - apply Forall2_flip. -Qed. - -Lemma eq_context_upto_flip {Σ Re Rle Γ Δ} - `{!RelationClasses.Reflexive Re} - `{!RelationClasses.Symmetric Re} - `{!RelationClasses.Transitive Re} - `{!RelationClasses.Reflexive Rle} - `{!RelationClasses.Transitive Rle} - `{!RelationClasses.subrelation Re Rle} : - eq_context_upto Σ Re Rle Γ Δ -> - eq_context_upto Σ Re (flip Rle) Δ Γ. -Proof. - induction 1; constructor; auto; depelim p; constructor; auto. - - now symmetry. - - now eapply eq_term_upto_univ_napp_flip; try typeclasses eauto. - - now symmetry. - - now eapply eq_term_upto_univ_napp_flip; try typeclasses eauto. - - now eapply eq_term_upto_univ_napp_flip; try typeclasses eauto. -Qed. - -Lemma red1_eq_context_upto_r {Σ Σ' Re Rle Γ Δ u v} : - RelationClasses.Equivalence Re -> - RelationClasses.PreOrder Rle -> - SubstUnivPreserving Re -> - SubstUnivPreserving Rle -> - RelationClasses.subrelation Re Rle -> +Lemma red1_eq_context_upto_r {Σ Σ' cmp_universe cmp_sort pb Γ Δ u v} : + RelationClasses.PreOrder (cmp_universe Conv) -> + RelationClasses.Reflexive (cmp_universe pb) -> + RelationClasses.PreOrder (cmp_sort Conv) -> + RelationClasses.Reflexive (cmp_sort pb) -> red1 Σ Γ u v -> - eq_context_upto Σ' Re Rle Δ Γ -> + eq_context_upto Σ' cmp_universe cmp_sort pb Δ Γ -> ∑ v', red1 Σ Δ u v' * - eq_term_upto_univ Σ' Re Re v' v. + eq_term_upto_univ Σ' cmp_universe cmp_sort Conv v' v. Proof. intros. - destruct (@red1_eq_context_upto_l Σ Σ' (flip Rle) Re Γ Δ u v); auto; try typeclasses eauto. - - intros x; red; reflexivity. - - intros s u1 u2 Ru. red. apply R_universe_instance_flip in Ru. now apply H2. - - intros x y rxy; red. now symmetry in rxy. - - now apply eq_context_upto_flip. - - exists x. intuition auto. - now eapply eq_term_upto_univ_sym. + destruct (@red1_eq_context_upto_l Σ Σ' (fun pb => flip (cmp_universe pb)) (fun pb => flip (cmp_sort pb)) pb Γ Δ u v) as (v' & r & e); tas. + - intro x. red. reflexivity. + - intro x. red. reflexivity. + - eapply eq_context_upto_flip; [..|eassumption]; tc. + - exists v'; split; tas. + eapply eq_term_upto_univ_napp_flip; [..|eassumption]; tc. Qed. -Lemma red1_eq_term_upto_univ_r {Σ Σ' Re Rle napp Γ u v u'} : - RelationClasses.Reflexive Re -> - RelationClasses.Reflexive Rle -> - RelationClasses.Symmetric Re -> - RelationClasses.Transitive Re -> - RelationClasses.Transitive Rle -> - SubstUnivPreserving Re -> - SubstUnivPreserving Rle -> - RelationClasses.subrelation Re Rle -> - eq_term_upto_univ_napp Σ' Re Rle napp u' u -> +Lemma red1_eq_term_upto_univ_r {Σ Σ' cmp_universe cmp_sort pb napp Γ u v u'} : + RelationClasses.PreOrder (cmp_universe Conv) -> + RelationClasses.PreOrder (cmp_universe pb) -> + RelationClasses.PreOrder (cmp_sort Conv) -> + RelationClasses.PreOrder (cmp_sort pb) -> + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe pb) -> + RelationClasses.subrelation (cmp_sort Conv) (cmp_sort pb) -> + SubstUnivPreserving (cmp_universe Conv) (cmp_universe Conv) -> + SubstUnivPreserving (cmp_universe Conv) (cmp_sort Conv) -> + SubstUnivPreserving (cmp_universe Conv) (cmp_sort pb) -> + eq_term_upto_univ_napp Σ' cmp_universe cmp_sort pb napp u' u -> red1 Σ Γ u v -> ∑ v', red1 Σ Γ u' v' × - eq_term_upto_univ_napp Σ' Re Rle napp v' v. + eq_term_upto_univ_napp Σ' cmp_universe cmp_sort pb napp v' v. Proof. - intros he he' hse hte htle sre srle hR h uv. - destruct (@red1_eq_term_upto_univ_l Σ Σ' Re (flip Rle) napp Γ u v u'); auto. - - now eapply flip_Transitive. + intros preorder_univ_conv preorder_sort_pb preorder_sort_conv preoder_sort_pb sub_univ sub_sort hsubst_univ hsubst_sort_conv hsubst_sort_pb h uv. + destruct (@red1_eq_term_upto_univ_l Σ Σ' (fun pb => flip (cmp_universe pb)) (fun pb => flip (cmp_sort pb)) pb napp Γ u v u') as (v' & r & e). + all: eauto using RelationClasses.flip_PreOrder. + 1,2: intros ??; unfold flip; cbn; eauto. + - red. intros s u1 u2 ru. + eapply cmp_universe_instance_flip in ru; cbnr. + now apply hsubst_univ. - red. intros s u1 u2 ru. - apply R_universe_instance_flip in ru. - now apply srle. - - intros x y X. symmetry in X. apply hR in X. apply X. - - eapply eq_term_upto_univ_napp_flip; eauto. - - exists x. intuition auto. - eapply (@eq_term_upto_univ_napp_flip Σ' Re (flip Rle) Rle); eauto. - + now eapply flip_Transitive. - + unfold flip. intros ? ? H. symmetry in H. eauto. + eapply cmp_universe_instance_flip in ru; cbnr. + now apply hsubst_sort_conv. + - red. intros s u1 u2 ru. + eapply cmp_universe_instance_flip in ru; cbnr. + now apply hsubst_sort_pb. + - eapply eq_term_upto_univ_napp_flip; [..|eassumption]. all: reflexivity. + - exists v'. split; tas. + eapply eq_term_upto_univ_napp_flip; [..|eassumption]; tc. Qed. Section RedEq. Context (Σ : global_env_ext). - Context {Re Rle : Universe.t -> Universe.t -> Prop} - {refl : RelationClasses.Reflexive Re} - {refl': RelationClasses.Reflexive Rle} - {pres : SubstUnivPreserving Re} - {pres' : SubstUnivPreserving Rle} - {sym : RelationClasses.Symmetric Re} - {trre : RelationClasses.Transitive Re} - {trle : RelationClasses.Transitive Rle}. - Context (inclre : RelationClasses.subrelation Re Rle). - - Lemma red_eq_term_upto_univ_r {Γ T V U} : - eq_term_upto_univ Σ Re Rle T U -> red Σ Γ U V -> - ∑ T', red Σ Γ T T' * eq_term_upto_univ Σ Re Rle T' V. - Proof using inclre pres pres' refl refl' sym trle trre. + Context {cmp_universe : conv_pb -> Universe.t -> Universe.t -> Prop} {cmp_sort : conv_pb -> sort -> sort -> Prop} {pb : conv_pb} + {preorder_univ_conv : RelationClasses.PreOrder (cmp_universe Conv)} + {preorder_univ_pb : RelationClasses.PreOrder (cmp_universe pb)} + {preorder_sort_conv : RelationClasses.PreOrder (cmp_sort Conv)} + {preorder_sort_pb : RelationClasses.PreOrder (cmp_sort pb)} + {sub_univ : RelationClasses.subrelation (cmp_universe Conv) (cmp_universe pb)} + {sub_sort : RelationClasses.subrelation (cmp_sort Conv) (cmp_sort pb)} + {hsubst_univ : SubstUnivPreserving (cmp_universe Conv) (cmp_universe Conv)} + {hsubst_sort_conv : SubstUnivPreserving (cmp_universe Conv) (cmp_sort Conv)} + {hsubst_sort_pb : SubstUnivPreserving (cmp_universe Conv) (cmp_sort pb)}. + + Lemma red_eq_term_upto_univ_l {Σ' Γ u v u'} : + eq_term_upto_univ Σ' cmp_universe cmp_sort pb u u' -> + red Σ Γ u v -> + ∑ v', + red Σ Γ u' v' * eq_term_upto_univ Σ' cmp_universe cmp_sort pb v v'. + Proof using preorder_univ_conv preorder_univ_pb preorder_sort_conv preorder_sort_conv preorder_sort_pb sub_univ sub_sort hsubst_univ hsubst_sort_conv hsubst_sort_pb. intros eq r. - induction r in T, eq |- *. - - eapply red1_eq_term_upto_univ_r in eq as [v' [r' eq']]; eauto. - - exists T; split; eauto. + induction r in u', eq |- *. + - eapply red1_eq_term_upto_univ_l in eq as [v' [r' eq']]; eauto. + - exists u'. split; auto. - case: (IHr1 _ eq) => T' [r' eq']. case: (IHr2 _ eq') => T'' [r'' eq'']. exists T''. split=>//. now transitivity T'. Qed. - Lemma red_eq_term_upto_univ_l {Γ u v u'} : - eq_term_upto_univ Σ Re Rle u u' -> + Lemma red_eq_term_upto_univ_r {Σ' Γ u v u'} : + eq_term_upto_univ Σ' cmp_universe cmp_sort pb u' u -> red Σ Γ u v -> - ∑ v', - red Σ Γ u' v' * - eq_term_upto_univ Σ Re Rle v v'. - Proof using inclre pres pres' refl refl' trle trre. + ∑ v', red Σ Γ u' v' * eq_term_upto_univ Σ' cmp_universe cmp_sort pb v' v. + Proof using preorder_univ_conv preorder_univ_pb preorder_sort_conv preorder_sort_conv preorder_sort_pb sub_univ sub_sort hsubst_univ hsubst_sort_conv hsubst_sort_pb. intros eq r. induction r in u', eq |- *. - - eapply red1_eq_term_upto_univ_l in eq as [v' [r' eq']]; eauto. + - eapply red1_eq_term_upto_univ_r in eq as [v' [r' eq']]; eauto. - exists u'. split; auto. - case: (IHr1 _ eq) => T' [r' eq']. case: (IHr2 _ eq') => T'' [r'' eq'']. @@ -1673,6 +1481,57 @@ Section RedEq. Qed. End RedEq. +Section RedEqContext. + Context (Σ : global_env_ext). + Context {cmp_universe : conv_pb -> Universe.t -> Universe.t -> Prop} {cmp_sort : conv_pb -> sort -> sort -> Prop} {pb : conv_pb} + {preorder_univ_conv : RelationClasses.PreOrder (cmp_universe Conv)} + {preorder_univ_pb : RelationClasses.PreOrder (cmp_universe pb)} + {preorder_sort_conv : RelationClasses.PreOrder (cmp_sort Conv)} + {preorder_sort_pb : RelationClasses.PreOrder (cmp_sort pb)} + {sub_univ : RelationClasses.subrelation (cmp_universe Conv) (cmp_universe pb)} + {sub_sort : RelationClasses.subrelation (cmp_sort Conv) (cmp_sort pb)} + {hsubst_univ : SubstUnivPreserving (cmp_universe Conv) (cmp_universe Conv)} + {hsubst_sort_conv : SubstUnivPreserving (cmp_universe Conv) (cmp_sort Conv)} + {hsubst_sort_pb : SubstUnivPreserving (cmp_universe Conv) (cmp_sort pb)}. + + Lemma red_eq_context_upto_l {Γ Δ} {u} {v} : + eq_context_upto Σ cmp_universe cmp_sort pb Γ Δ -> + red Σ Γ u v -> + ∑ v', + red Σ Δ u v' * + eq_term_upto_univ Σ cmp_universe cmp_sort Conv v v'. + Proof using preorder_univ_conv preorder_univ_pb preorder_sort_conv preorder_sort_conv preorder_sort_pb sub_univ sub_sort hsubst_univ hsubst_sort_conv hsubst_sort_pb. + intros eq r. + induction r. + - eapply red1_eq_context_upto_l in eq as [v' [r' eq']]; eauto; tc. + - exists x. split; auto. reflexivity. + - case: IHr1 => T' [r' eq']. + case: IHr2 => T'' [r'' eq'']. + eapply (red_eq_term_upto_univ_l _ (pb := Conv) (u:=y) (v:=T'') (u':=T')) in r'' as (T''' & r'' & eq'''); tc; tea. + exists T'''; split. + + now transitivity T'. + + now transitivity T''. + Qed. + + Lemma red_eq_context_upto_r {Γ Δ} {u} {v} : + eq_context_upto Σ cmp_universe cmp_sort pb Δ Γ -> + red Σ Γ u v -> + ∑ v', + red Σ Δ u v' * + eq_term_upto_univ Σ cmp_universe cmp_sort Conv v' v. + Proof using preorder_univ_conv preorder_univ_pb preorder_sort_conv preorder_sort_conv preorder_sort_pb sub_univ sub_sort hsubst_univ hsubst_sort_conv hsubst_sort_pb. + intros eq r. + induction r. + - eapply red1_eq_context_upto_r in eq as [v' [r' eq']]; eauto; tc. + - exists x. split; auto. reflexivity. + - case: IHr1 => T' [r' eq']. + case: IHr2 => T'' [r'' eq'']. + eapply (red_eq_term_upto_univ_r _ (pb := Conv) (u:=y) (v:=T'') (u':=T')) in r'' as (T''' & r'' & eq'''); tc; tea. + exists T'''; split. + + now transitivity T'. + + now transitivity T''. + Qed. +End RedEqContext. Polymorphic Derive Signature for Relation.clos_refl_trans. @@ -1991,7 +1850,7 @@ Section RedPred. Hint Extern 4 (All2_fold _ _ _) => constructor : pcuic. Lemma OnOne2_local_env_pred1_ctx_over Γ Δ Δ' : - OnOne2_local_env (on_one_decl (fun Δ M N => pred1 Σ (Γ ,,, Δ) (Γ ,,, Δ) M N)) Δ Δ' -> + OnOne2_local_env (fun Δ => on_one_decl (fun M N => pred1 Σ (Γ ,,, Δ) (Γ ,,, Δ) M N)) Δ Δ' -> pred1_ctx_over Σ Γ Γ Δ Δ'. Proof using Type. induction 1. @@ -2313,7 +2172,7 @@ Section PredRed. - depelim X1; try solve [repeat constructor]; eauto. depelim X2; cbn in H0; rtoProp. eapply red_primArray_congr; eauto. - + now eapply Universe.make_inj in e. + + now eapply Universe.make'_inj in e. + solve_all. Qed. @@ -2888,7 +2747,7 @@ Section RedConfluence. Definition lift_ws (R : context -> context -> Type) : ws_context xpred0 -> ws_context xpred0 -> Type := fun Γ Γ' => R Γ Γ'. - Definition ws_red1_ctx := (lift_ws (OnOne2_local_env (on_one_decl (red1 Σ)))). + Definition ws_red1_ctx := (lift_ws (OnOne2_local_env (fun Γ => on_one_decl (red1 Σ Γ)))). Definition ws_red_ctx := lift_ws (red_ctx Σ). Definition ws_pred1_ctx := (lift_ws (on_contexts (pred1 Σ))). @@ -2947,8 +2806,8 @@ Section RedConfluence. Set Firstorder Solver eauto with pcuic core typeclass_instances. Lemma clos_rt_OnOne2_local_env_ctx_incl R : - inclusion (clos_refl_trans (OnOne2_local_env (on_one_decl R))) - (clos_refl_trans_ctx (OnOne2_local_env (on_one_decl R))). + inclusion (clos_refl_trans (OnOne2_local_env (fun Γ => on_one_decl1 R Γ))) + (clos_refl_trans_ctx (OnOne2_local_env (fun Γ => on_one_decl1 R Γ))). Proof using wfΣ. intros x y H. induction H; firstorder; try solve[econstructor; eauto]. @@ -3083,7 +2942,7 @@ Section RedConfluence. Proof using Type. move/on_free_vars_ctx_All_fold => a eqctx. apply on_free_vars_ctx_All_fold. - eapply eq_context_upto_names_gen in eqctx. + eapply All2_fold_All2 in eqctx. eapply All2_fold_All_fold_mix_left in eqctx; tea. cbn in eqctx. induction eqctx. - constructor; auto. @@ -3293,8 +3152,7 @@ Section RedConfluence. exists body'; split => //. rewrite -lift0_inst. econstructor; eauto. destruct (nth_error Δ x) eqn:hnth' => //. - eapply eq_context_upto_names_gen in eqctx'. - eapply All2_fold_nth in eqctx' as [d' [hnth'' [eqctx'' eqd]]]; tea. + eapply All2_nth_error_Some in eqctx' as [d' [hnth'' eqd]]; tea. depelim eqd. subst. noconf eq. subst. noconf eq. rewrite hnth'' //. Qed. @@ -3604,9 +3462,10 @@ Section RedConfluence. eq_context_upto_names Γ Γ' -> option_map decl_body (nth_error Γ' n) = Some d. Proof using Type. - move: Γ' n d; induction Γ; destruct n; simpl; intros; try congruence. - noconf H. depelim X. depelim c; subst; simpl => //. - depelim X. apply IHΓ; auto. + destruct nth_error as [decl|] eqn:hnth => //. + intros [= <-] eqctx. + eapply All2_nth_error_Some in eqctx as (decl' & -> & eqd); tea. + destruct eqd; reflexivity. Qed. Lemma decl_type_eq_context_upto_names Γ Γ' n d : @@ -3614,9 +3473,10 @@ Section RedConfluence. eq_context_upto_names Γ Γ' -> option_map decl_type (nth_error Γ' n) = Some d. Proof using Type. - move: Γ' n d; induction Γ; destruct n; simpl; intros; try congruence. - noconf H. depelim X. depelim c; subst; simpl => //. - depelim X. simpl. apply IHΓ; auto. + destruct nth_error as [decl|] eqn:hnth => //. + intros [= <-] eqctx. + eapply All2_nth_error_Some in eqctx as (decl' & -> & eqd); tea. + destruct eqd; reflexivity. Qed. Lemma eq_context_upto_names_app Γ Γ' Δ : @@ -3634,9 +3494,9 @@ Section RedConfluence. Proof using Type. move=> Hctx. eapply context_pres_let_bodies_red1. - eapply eq_context_upto_names_gen in Hctx. + eapply All2_fold_All2 in Hctx. eapply All2_fold_impl; tea => /= _ _ ? ? [] /=; - rewrite /pres_let_bodies /= //; intros; congruence. + rewrite /pres_let_bodies /= //. Qed. Lemma clos_rt_red1_eq_context_upto_names Γ Γ' t u : @@ -3670,7 +3530,7 @@ Section RedConfluence. induction X in X0, Δ, Δ', X1 |- *. depelim X1. depelim X0. constructor. depelim X1. depelim X0. constructor. eapply IHX; tea. - depelim r; depelim c; subst; depelim a; try constructor; eauto. + destruct a; depelim e; depelim r; subst; try constructor; eauto. 1,3:now etransitivity. all:eapply red_eq_context_upto_names; eauto. Qed. @@ -3685,7 +3545,7 @@ Section RedConfluence. induction X in X0, Δ, Δ', X1 |- *. depelim X1. depelim X0. constructor. depelim X1. depelim X0. constructor. eapply IHX; tea. - depelim r; depelim c; subst; depelim a; try constructor; eauto. + destruct a; depelim e; depelim r; subst; try constructor; eauto. 1,3:now etransitivity. all:eapply red_eq_context_upto_names; eauto. Qed. @@ -3854,32 +3714,31 @@ Section RedConfluence. End RedConfluence. (** Currently provable, but not if we add eta / sprop *) -Lemma eq_term_upto_univ_napp_on_free_vars {cf:checker_flags} {Σ : global_env} {P eq leq napp} {t u} : - eq_term_upto_univ_napp Σ eq leq napp t u -> +Lemma eq_term_upto_univ_napp_on_free_vars {cf:checker_flags} {Σ : global_env} {P cmp_universe cmp_sort pb napp} {t u} : + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp t u -> on_free_vars P t -> on_free_vars P u. Proof. - intros eqt ont. revert P t ont u eq leq napp eqt. + intros eqt ont. revert P t ont u pb napp eqt. apply: term_on_free_vars_ind; intros; depelim eqt. all:simpl; auto. all:try solve [solve_all]. - destruct e as [? [? [? ?]]]. - rewrite -(All2_fold_length a1). rewrite -(All2_length a0). + rewrite -(All2_length a). + unfold eq_branches, eq_branch in *. solve_all. rewrite test_context_k_closed_on_free_vars_ctx. eapply eq_context_upto_names_on_free_vars; tea. - now eapply eq_context_upto_names_gen in a1. rewrite test_context_k_closed_on_free_vars_ctx. - destruct a. + destruct a1. eapply eq_context_upto_names_on_free_vars; tea. - now eapply eq_context_upto_names_gen in a2. - destruct a as [hctx ihctx hb ihb]. - rewrite -(All2_fold_length a2). now eapply ihb. - - rewrite -(All2_length a). solve_all. + destruct a1 as [hctx ihctx hb ihb]. + rewrite -(All2_length a2). now eapply ihb. + - rewrite -(All2_length e). unfold eq_mfixpoint in *. solve_all. apply/andP; split; eauto. len in b2. eapply b2. eauto. - - rewrite -(All2_length a). solve_all. + - rewrite -(All2_length e). unfold eq_mfixpoint in *. solve_all. apply/andP; split; eauto. len in b2. eapply b2. eauto. - depelim o; cbn in *; solve_all. diff --git a/pcuic/theories/PCUICConsistency.v b/pcuic/theories/PCUICConsistency.v index c649d80af..ae4c67e07 100644 --- a/pcuic/theories/PCUICConsistency.v +++ b/pcuic/theories/PCUICConsistency.v @@ -19,13 +19,11 @@ From MetaCoq.PCUIC Require Import PCUICTyping PCUICEquality PCUICAst PCUICAstUti From Equations Require Import Equations. -Definition Prop_univ := Universe.of_levels (inl PropLevel.lProp). - Definition False_oib : one_inductive_body := {| ind_name := "False"; ind_indices := []; - ind_sort := Prop_univ; - ind_type := tSort Prop_univ; + ind_sort := sProp; + ind_type := tSort sProp; ind_kelim := IntoAny; ind_ctors := []; ind_projs := []; @@ -77,4 +75,4 @@ Proof. unshelve eapply declared_minductive_to_gen in Hdecl; eauto. red in Hdecl. rewrite Hdecl in typ_false'. cbn in typ_false'. inversion typ_false'. -Qed. \ No newline at end of file +Qed. diff --git a/pcuic/theories/PCUICContextConversion.v b/pcuic/theories/PCUICContextConversion.v index 940f7b143..d5efe4fbd 100644 --- a/pcuic/theories/PCUICContextConversion.v +++ b/pcuic/theories/PCUICContextConversion.v @@ -391,8 +391,8 @@ Proof. now transitivity y. Qed. -Definition compare_context {cf} pb Σ := - eq_context_upto Σ (eq_universe Σ) (compare_universe pb Σ). +Definition compare_context {cf} Σ pb := + PCUICEquality.compare_context Σ Σ pb. Section ContextConversion. Context {cf : checker_flags}. @@ -404,9 +404,9 @@ Section ContextConversion. Hint Resolve conv_ctx_refl' cumul_ctx_refl' : pcuic. - Lemma fill_le {Γ : closed_context} {t u : open_term Γ} {t' u'} : - leq_term Σ.1 Σ t u -> red Σ Γ t t' -> red Σ Γ u u' -> - ∑ t'' u'', red Σ Γ t' t'' * red Σ Γ u' u'' * leq_term Σ Σ t'' u''. + Lemma fill_pb {Γ : closed_context} {pb} {t u : open_term Γ} {t' u'} : + compare_term Σ.1 Σ pb t u -> red Σ Γ t t' -> red Σ Γ u u' -> + ∑ t'' u'', red Σ Γ t' t'' * red Σ Γ u' u'' * compare_term Σ Σ pb t'' u''. Proof using wfΣ. intros tu tt' uu'. eapply red_eq_term_upto_univ_l in tu; try exact tt'. all:try tc. @@ -417,19 +417,19 @@ Section ContextConversion. exists t'', unf. intuition auto. Qed. + Lemma fill_le {Γ : closed_context} {t u : open_term Γ} {t' u'} : + leq_term Σ.1 Σ t u -> red Σ Γ t t' -> red Σ Γ u u' -> + ∑ t'' u'', red Σ Γ t' t'' * red Σ Γ u' u'' * leq_term Σ Σ t'' u''. + Proof using wfΣ. + apply fill_pb. + Qed. + Lemma fill_eq {Γ : closed_context} {t u : open_term Γ} {t' u'} : eq_term Σ.1 Σ t u -> red Σ Γ t t' -> red Σ Γ u u' -> ∑ t'' u'', red Σ Γ t' t'' * red Σ Γ u' u'' * eq_term Σ.1 Σ t'' u''. Proof using wfΣ. - intros tu tt' uu'. - pose proof tu as tu2. - eapply red_eq_term_upto_univ_l in tu; try exact tt'; try tc. - destruct tu as [u'' [uu'' t'u'']]. - destruct (red_confluence uu' uu'') as [unf [ul ur]]. - eapply red_eq_term_upto_univ_r in t'u''; try exact ur; try tc. - destruct t'u'' as [t'' [t't'' t''unf]]. - exists t'', unf. intuition auto. - Qed. + apply fill_pb. + Qed. Lemma red_ctx_ws_cumul_ctx_pb {l Γ Γ'} : Σ ⊢ Γ ⇝ Γ' -> Σ ⊢ Γ ≤[l] Γ'. Proof using wfΣ. @@ -456,18 +456,18 @@ Section ContextConversion. Hint Resolve red_ctx_closed_left red_ctx_closed_right : fvs. Lemma red_compare_term_l {pb Γ} {u v u' : term} : - compare_term pb Σ Σ u u' -> + compare_term Σ Σ pb u u' -> red Σ Γ u v -> - ∑ v' : term, red Σ Γ u' v' × compare_term pb Σ Σ v v'. + ∑ v' : term, red Σ Γ u' v' × compare_term Σ Σ pb v v'. Proof using Type. destruct pb; cbn; apply red_eq_term_upto_univ_l; tc. Qed. Lemma red_compare_term_r {pb Γ} {u v u' : term} : - compare_term pb Σ Σ u u' -> + compare_term Σ Σ pb u u' -> red Σ Γ u' v -> - ∑ v' : term, red Σ Γ u v' × compare_term pb Σ Σ v' v. + ∑ v' : term, red Σ Γ u v' × compare_term Σ Σ pb v' v. Proof using Type. destruct pb; cbn; apply red_eq_term_upto_univ_r; tc. @@ -475,9 +475,9 @@ Section ContextConversion. Lemma closed_red_compare_term_l {pb Γ} {u v u' : term} : is_open_term Γ u' -> - compare_term pb Σ Σ u u' -> + compare_term Σ Σ pb u u' -> Σ ;;; Γ ⊢ u ⇝ v -> - ∑ v' : term, Σ ;;; Γ ⊢ u' ⇝ v' × compare_term pb Σ Σ v v'. + ∑ v' : term, Σ ;;; Γ ⊢ u' ⇝ v' × compare_term Σ Σ pb v v'. Proof using Type. intros isop comp [clΓ clu red]. destruct (red_compare_term_l comp red) as [nf [r eq]]. @@ -486,9 +486,9 @@ Section ContextConversion. Lemma closed_red_compare_term_r {pb Γ} {u v u' : term} : is_open_term Γ u -> - compare_term pb Σ Σ u u' -> + compare_term Σ Σ pb u u' -> Σ ;;; Γ ⊢ u' ⇝ v -> - ∑ v' : term, Σ ;;; Γ ⊢ u ⇝ v' × compare_term pb Σ Σ v' v. + ∑ v' : term, Σ ;;; Γ ⊢ u ⇝ v' × compare_term Σ Σ pb v' v. Proof using Type. intros isop comp [clΓ clu red]. destruct (red_compare_term_r comp red) as [nf [r eq]]. @@ -513,7 +513,7 @@ Section ContextConversion. Lemma ws_cumul_pb_red {pb} {Γ t u} : Σ ;;; Γ ⊢ t ≤[pb] u <~> ∑ v v', [× Σ ;;; Γ ⊢ t ⇝ v, Σ ;;; Γ ⊢ u ⇝ v' & - compare_term pb Σ (global_ext_constraints Σ) v v']. + compare_term Σ (global_ext_constraints Σ) pb v v']. Proof using wfΣ. split. - move/ws_cumul_pb_alt; intros (v & v' & [clΓ clt clu red red' leq]). @@ -588,63 +588,14 @@ Section ContextConversion. now exists v, v'. Qed. - Lemma red_eq_context_upto_l {R Re} {Γ Δ} {u} {v} - `{RelationClasses.Reflexive _ R} `{RelationClasses.Transitive _ R} `{SubstUnivPreserving R} - `{RelationClasses.Reflexive _ Re} `{RelationClasses.Transitive _ Re} `{SubstUnivPreserving Re} - `{RelationClasses.subrelation _ Re R} : - red Σ Γ u v -> - eq_context_upto Σ Re R Γ Δ -> - ∑ v', - red Σ Δ u v' * - eq_term_upto_univ Σ Re Re v v'. - Proof using Type. - intros r HΓ. - induction r. - - eapply (red1_eq_context_upto_l _ (Rle:=R)) in r; eauto. - destruct r as [v [? ?]]. exists v. intuition pcuic. - - exists x. split; auto. reflexivity. - - destruct IHr1 as [v' [? ?]]; eauto with fvs. - destruct IHr2 as [v'' [? ?]]; eauto with fvs. - eapply (red_eq_term_upto_univ_l _ _ (u:=y) (v:=v'') (u':=v')) in e; try tc. all:pcuic. - destruct e as [? [? ?]]. - exists x0; split; eauto. - now transitivity v'. - eapply eq_term_upto_univ_trans with v''; auto. - Qed. - - Lemma red_eq_context_upto_r {R Re Γ Δ} {u} {v} - `{RelationClasses.Equivalence _ Re} `{SubstUnivPreserving Re} - `{RelationClasses.PreOrder _ R} `{SubstUnivPreserving R} - `{RelationClasses.subrelation _ Re R} : - red Σ Δ u v -> - eq_context_upto Σ Re R Γ Δ -> - ∑ v', - red Σ Γ u v' * - eq_term_upto_univ Σ Re Re v v'. - Proof using Type. - intros r HΓ. - induction r. - - eapply (@red1_eq_context_upto_r Σ Σ Re R) in r; eauto. - destruct r as [v [? ?]]. exists v. intuition pcuic. - now symmetry. - - exists x. split; auto. reflexivity. - - destruct IHr1 as [v' [? ?]]. - destruct IHr2 as [v'' [? ?]]. - unshelve eapply (red_eq_term_upto_univ_l Σ _ (Γ := Γ) (u:=y) (v:=v'') (u':=v')) in e. all:pcuic. - destruct e as [? [? ?]]. - exists x0; split; eauto. - transitivity v'; auto. - eapply eq_term_upto_univ_trans with v''; auto; tc. - Qed. - Lemma closed_red_eq_context_upto_l {pb Γ Δ} {u} {v} : is_closed_context Δ -> Σ ;;; Γ ⊢ u ⇝ v -> - compare_context pb Σ Γ Δ -> + compare_context Σ pb Γ Δ -> ∑ v', Σ ;;; Δ ⊢ u ⇝ v' × eq_term Σ Σ v v'. Proof using Type. intros clΔ [onΓ onu r] c. - destruct (red_eq_context_upto_l r c) as [nf [red eq]]. + destruct (red_eq_context_upto_l _ c r) as [nf [red eq]]. exists nf. split; auto. split; eauto with fvs. now rewrite -(All2_fold_length c). Qed. @@ -652,26 +603,27 @@ Section ContextConversion. Lemma closed_red_eq_context_upto_r {pb Γ Δ} {u} {v} : is_closed_context Γ -> Σ ;;; Δ ⊢ u ⇝ v -> - compare_context pb Σ Γ Δ -> + compare_context Σ pb Γ Δ -> ∑ v', Σ ;;; Γ ⊢ u ⇝ v' × eq_term Σ Σ v v'. Proof using Type. intros clΔ [onΓ onu r] c. - destruct (red_eq_context_upto_r r c) as [nf [red eq]]. + destruct (red_eq_context_upto_r _ c r) as [nf [red eq]]. + symmetry in eq. exists nf. split; auto. split; eauto with fvs. now rewrite (All2_fold_length c). Qed. Lemma cumul_trans_red_leqterm {Γ : closed_context} {t u v : open_term Γ} : Σ ;;; Γ |- t <= u -> Σ ;;; Γ |- u <= v -> - ∑ l o r, red Σ Γ t l * - red Σ Γ u o * - red Σ Γ v r * - leq_term Σ.1 Σ l o * leq_term Σ.1 Σ o r. + ∑ l o r, red Σ Γ t l × + red Σ Γ u o × + red Σ Γ v r × + leq_term Σ.1 Σ l o × leq_term Σ.1 Σ o r. Proof using wfΣ. intros X X0. intros. - eapply cumul_alt in X as [t0 [u0 [[redl redr] eq]]]. - eapply cumul_alt in X0 as [u1 [v0 [[redl' redr'] eq']]]. + eapply cumul_alt in X as [t0 [u0 (redl & redr & eq)]]. + eapply cumul_alt in X0 as [u1 [v0 (redl' & redr' & eq')]]. destruct (red_confluence redr redl') as [unf [nfl nfr]]. eapply red_eq_term_upto_univ_r in eq; try tc. 2:tea. destruct eq as [t1 [red'0 eq2]]. @@ -684,36 +636,64 @@ Section ContextConversion. transitivity v0; auto. eapply eq2. eapply eq1. Qed. - Lemma conv_eq_context_upto {Γ} {Δ} {T U} : - eq_context_upto Σ (eq_universe Σ) (eq_universe Σ) Γ Δ -> - Σ ;;; Γ |- T = U -> - Σ ;;; Δ |- T = U. + Lemma cumul_pb_compare_context {pb pb' Γ Δ T U} : + compare_context Σ pb Γ Δ -> + Σ ;;; Δ |- T <=[pb'] U -> + Σ ;;; Γ |- T <=[pb'] U. Proof using Type. intros eqctx cum. - eapply conv_alt_red in cum as [nf [nf' [[redl redr] ?]]]. - eapply (red_eq_context_upto_l (R:=eq_universe _) (Re:=eq_universe _)) in redl; tea; tc. - eapply (red_eq_context_upto_l (R:=eq_universe _) (Re:=eq_universe _)) in redr; tea; tc. + eapply cumul_alt in cum as (nf & nf' & redl & redr & ?). + eapply red_eq_context_upto_r in redl; tea; tc. + eapply red_eq_context_upto_r in redr; tea; tc. destruct redl as [v' [redv' eqv']]. destruct redr as [v'' [redv'' eqv'']]. - eapply conv_alt_red. exists v', v''; intuition auto. + eapply cumul_alt. exists v', v''; repeat split; auto. transitivity nf. - now symmetry. now transitivity nf'. + eapply eq_term_upto_univ_leq; tc; auto. + transitivity nf' => //. + eapply eq_term_upto_univ_leq; tc; auto. + now symmetry. Qed. - Lemma conv_leq_context_upto {Γ Δ T U} : - eq_context_upto Σ (eq_universe Σ) (leq_universe Σ) Γ Δ -> - Σ ;;; Δ |- T = U -> - Σ ;;; Γ |- T = U. + (* Conversion is untyped so this currently holds as compare_context + just allows cumulativity on types, which do not participate in reduction. + However the useful lemma is the one above that shows we can lift a + conversion from a large context to a smaller one (contravariance). + *) + Local Remark cumul_pb_compare_context_inv {pb pb' Γ Δ T U} : + compare_context Σ pb Γ Δ -> + Σ ;;; Γ |- T <=[pb'] U -> + Σ ;;; Δ |- T <=[pb'] U. Proof using Type. intros eqctx cum. - eapply conv_alt_red in cum as [nf [nf' [[redl redr] ?]]]. - eapply (red_eq_context_upto_r (Re:=eq_universe _) (R:=leq_universe _)) in redl; tea. - eapply (red_eq_context_upto_r (Re:=eq_universe _) (R:=leq_universe _)) in redr; tea. + eapply cumul_alt in cum as (nf & nf' & redl & redr & ?). + eapply red_eq_context_upto_l in redl; tea; tc. + eapply red_eq_context_upto_l in redr; tea; tc. destruct redl as [v' [redv' eqv']]. destruct redr as [v'' [redv'' eqv'']]. - eapply conv_alt_red. exists v', v''; intuition auto. + eapply cumul_alt. exists v', v''; repeat split; auto. transitivity nf. - now symmetry. now transitivity nf'. + eapply eq_term_upto_univ_leq; tc; auto. + now symmetry. + transitivity nf' => //. + eapply eq_term_upto_univ_leq; tc; auto. + Qed. + + Lemma conv_eq_context_upto {Γ} {Δ} {T U} : + eq_context Σ Σ Γ Δ -> + Σ ;;; Γ |- T = U -> + Σ ;;; Δ |- T = U. + Proof using Type. + intro eqctx. symmetry in eqctx. + now eapply cumul_pb_compare_context. + Qed. + + Lemma conv_leq_context_upto {Γ Δ T U} : + leq_context Σ Σ Γ Δ -> + Σ ;;; Δ |- T = U -> + Σ ;;; Γ |- T = U. + Proof using Type. + apply cumul_pb_compare_context. Qed. (* Conversion is untyped so this currently holds as context ws_cumul_pb @@ -721,42 +701,24 @@ Section ContextConversion. However the useful lemma is the one above that shows we can lift a conversion from a large context to a smaller one (contravariance). *) - Local Remark conv_eq_context_upto_leq_inv {Γ Δ T U} : - eq_context_upto Σ (eq_universe Σ) (leq_universe Σ) Γ Δ -> + Local Remark conv_leq_context_upto_inv {Γ Δ T U} : + leq_context Σ Σ Γ Δ -> Σ ;;; Γ |- T = U -> Σ ;;; Δ |- T = U. Proof using Type. - intros eqctx cum. - eapply conv_alt_red in cum as [nf [nf' [[redl redr] ?]]]. - eapply (red_eq_context_upto_l (Re:=eq_universe _) (R:=leq_universe _)) in redl; tea. - eapply (red_eq_context_upto_l (Re:=eq_universe _) (R:=leq_universe _)) in redr; tea. - destruct redl as [v' [redv' eqv']]. - destruct redr as [v'' [redv'' eqv'']]. - eapply conv_alt_red. exists v', v''; intuition auto. - transitivity nf. - now symmetry. now transitivity nf'. + apply cumul_pb_compare_context_inv. Qed. Lemma cumul_leq_context_upto {Γ Δ T U} : - eq_context_upto Σ (eq_universe Σ) (leq_universe Σ) Δ Γ -> + leq_context Σ Σ Δ Γ -> Σ ;;; Γ |- T <= U -> Σ ;;; Δ |- T <= U. Proof using Type. - intros eqctx cum. - eapply cumul_alt in cum as [nf [nf' [[redl redr] ?]]]. - eapply (red_eq_context_upto_r (Re:=eq_universe Σ) (R:=leq_universe _)) in redl; tea. - eapply (red_eq_context_upto_r (Re:=eq_universe Σ) (R:=leq_universe _)) in redr; tea. - destruct redl as [v' [redv' eqv']]. - destruct redr as [v'' [redv'' eqv'']]. - eapply cumul_alt. exists v', v''; intuition auto. - unfold leq_term_ext. transitivity nf. - apply eq_term_leq_term. now symmetry. - transitivity nf'; auto. - now apply eq_term_leq_term. + apply cumul_pb_compare_context. Qed. Lemma ws_cumul_pb_compare_context {pb pb' Γ Δ T U} : - compare_context pb Σ Δ Γ -> + compare_context Σ pb Δ Γ -> is_closed_context Δ -> Σ ;;; Γ ⊢ T ≤[pb'] U -> Σ ;;; Δ ⊢ T ≤[pb'] U. @@ -777,7 +739,7 @@ Section ContextConversion. Qed. Local Remark ws_cumul_pb_compare_context_inv {pb pb' Γ Δ T U} : - compare_context pb Σ Γ Δ -> + compare_context Σ pb Γ Δ -> is_closed_context Δ -> Σ ;;; Γ ⊢ T ≤[pb'] U -> Σ ;;; Δ ⊢ T ≤[pb'] U. @@ -797,109 +759,82 @@ Section ContextConversion. now apply eq_term_leq_term. Qed. - (* Local Remark cumul_leq_context_upto_inv {Γ Δ T U} : - eq_context_upto Σ (eq_universe Σ) (leq_universe Σ) Γ Δ -> - Σ ;;; Γ |- T <= U -> - Σ ;;; Δ |- T <= U. - Proof. - intros eqctx cum. - eapply cumul_alt in cum as [nf [nf' [[redl redr] ?]]]. - eapply (red_eq_context_upto_l (Re:=eq_universe Σ) (R:=leq_universe Σ) (Δ:=Δ)) in redl; tas. - eapply (red_eq_context_upto_l (Re:=eq_universe Σ) (R:=leq_universe Σ) (Δ:=Δ)) in redr; tas. - destruct redl as [v' [redv' eqv']]. - destruct redr as [v'' [redv'' eqv'']]. - eapply cumul_alt. exists v', v''; intuition auto. - eapply leq_term_trans with nf. - apply eq_term_leq_term. now apply eq_term_sym. - eapply leq_term_trans with nf'; auto. - now apply eq_term_leq_term. - Qed. *) - - Lemma eq_context_upto_impl {Re Rle} {Re' Rle'} {Γ Δ} - `{RelationClasses.subrelation _ Re Re'} - `{RelationClasses.subrelation _ Rle Rle'} - `{RelationClasses.subrelation _ Re' Rle'} : - eq_context_upto Σ Re Rle Γ Δ -> - eq_context_upto Σ Re' Rle' Γ Δ. - Proof using Type. - induction 1; constructor; auto. - eapply compare_decls_impl; eauto. - intros x y h. - eapply eq_term_upto_univ_impl. 5:eauto. all:try tc || auto. - intros x y h. - eapply eq_term_upto_univ_impl. 5:eauto. all:try tc || auto. - transitivity Re'; auto. - Qed. - Lemma eq_leq_context_upto Γ Δ : - eq_context_upto Σ (eq_universe Σ) (eq_universe Σ) Γ Δ -> - eq_context_upto Σ (eq_universe Σ) (leq_universe Σ) Γ Δ. - Proof using Type. apply eq_context_upto_impl. Qed. + eq_context Σ Σ Γ Δ -> + leq_context Σ Σ Γ Δ. + Proof using Type. apply eq_context_upto_impl. all: tc. Qed. Lemma cumul_eq_context_upto {Γ Δ T U} : - eq_context_upto Σ (eq_universe Σ) (eq_universe Σ) Γ Δ -> + eq_context Σ Σ Γ Δ -> Σ ;;; Γ |- T <= U -> Σ ;;; Δ |- T <= U. Proof using Type. - intros eqctx cum. symmetry in eqctx. - apply eq_leq_context_upto in eqctx. - eapply cumul_leq_context_upto; eauto. + intro H. symmetry in H. revert H. + apply cumul_pb_compare_context. Qed. Lemma ws_cumul_pb_eq_context_upto {pb Γ Δ T U} : - eq_context_upto Σ (eq_universe Σ) (eq_universe Σ) Γ Δ -> + eq_context Σ Σ Γ Δ -> is_closed_context Δ -> Σ ;;; Γ ⊢ T ≤[pb] U -> Σ ;;; Δ ⊢ T ≤[pb] U. Proof using wfΣ. - intros eqctx cl cum. symmetry in eqctx. - eapply (ws_cumul_pb_compare_context (pb:=Conv)) in cum; tea. + intros eqctx cl. symmetry in eqctx. + now eapply ws_cumul_pb_compare_context. Qed. - Lemma conv_alt_red_ctx {Γ : closed_context} {Γ'} {T U : open_term Γ} : - Σ ;;; Γ |- T = U -> - @red_ctx Σ Γ Γ' -> - Σ ;;; Γ' |- T = U. + Lemma cumul_pb_alt_red_ctx {Γ : closed_context} {Γ' pb} {T U : open_term Γ} : + Σ ;;; Γ |- T <=[pb] U -> + red_ctx Σ Γ Γ' -> + Σ ;;; Γ' |- T <=[pb] U. Proof using wfΣ. intros H Hctx. - eapply conv_alt_red in H. apply conv_alt_red. - destruct H as [T' [U' [[redv redv'] leqvv']]]. + eapply cumul_alt in H. apply cumul_alt. + destruct H as (T' & U' & redv & redv' & leqvv'). destruct (red_red_ctx _ _ redv Hctx) as [Tj [redTj redT'j]]. destruct (red_red_ctx _ _ redv' Hctx) as [Uj [redUUj redU'j]]. - destruct (fill_eq (Γ := inj_closed Γ' byfvs) (t := inj_open T' byfvs) (u := inj_open U' byfvs) leqvv' redT'j redU'j) as [Tnf [Unf [[redTnf redUnf] eqnf]]]. + destruct (fill_pb (Γ := inj_closed Γ' byfvs) (t := inj_open T' byfvs) (u := inj_open U' byfvs) leqvv' redT'j redU'j) as [Tnf [Unf [[redTnf redUnf] eqnf]]]. exists Tnf, Unf; intuition eauto. now transitivity Tj. now transitivity Uj. Qed. - Lemma conv_alt_red_ctx_inv {Γ Γ' : closed_context} {T U : open_term Γ} : - Σ ;;; Γ |- T = U -> + Lemma cumul_pb_alt_red_ctx_inv {Γ Γ' : closed_context} {pb} {T U : open_term Γ} : + Σ ;;; Γ |- T <=[pb] U -> red_ctx Σ Γ' Γ -> - Σ ;;; Γ' |- T = U. + Σ ;;; Γ' |- T <=[pb] U. Proof using wfΣ. intros H Hctx. - apply conv_alt_red in H as [v [v' [[redl redr] leq]]]. + apply cumul_alt in H as (v & v' & redl & redr & leq). pose proof (red_red_ctx_inv redl Hctx). pose proof (red_red_ctx_inv redr Hctx). - apply conv_alt_red. + apply cumul_alt. exists v, v'. split. pcuic. auto. Qed. + Lemma conv_alt_red_ctx {Γ : closed_context} {Γ'} {T U : open_term Γ} : + Σ ;;; Γ |- T = U -> + red_ctx Σ Γ Γ' -> + Σ ;;; Γ' |- T = U. + Proof using wfΣ. + apply cumul_pb_alt_red_ctx. + Qed. + + Lemma conv_alt_red_ctx_inv {Γ Γ' : closed_context} {T U : open_term Γ} : + Σ ;;; Γ |- T = U -> + red_ctx Σ Γ' Γ -> + Σ ;;; Γ' |- T = U. + Proof using wfΣ. + apply cumul_pb_alt_red_ctx_inv. + Qed. + Lemma cumul_alt_red_ctx {Γ : closed_context} {Γ'} {T U : open_term Γ} : Σ ;;; Γ |- T <= U -> @red_ctx Σ Γ Γ' -> Σ ;;; Γ' |- T <= U. Proof using wfΣ. - intros H Hctx. - eapply cumul_alt in H. apply cumul_alt. - destruct H as [T' [U' [[redv redv'] leqvv']]]. - destruct (red_red_ctx _ _ redv Hctx) as [Tj [redTj redT'j]]. - destruct (red_red_ctx _ _ redv' Hctx) as [Uj [redUUj redU'j]]. - destruct (fill_le (Γ := inj_closed Γ' byfvs) (t := inj_open T' byfvs) (u := inj_open U' byfvs) leqvv' redT'j redU'j) as [Tnf [Unf [[redTnf redUnf] eqnf]]]. - exists Tnf, Unf; intuition eauto. - now transitivity Tj. - now transitivity Uj. + apply cumul_pb_alt_red_ctx. Qed. Lemma cumul_alt_red_ctx_inv {Γ Γ' : closed_context} {T U : open_term Γ} : @@ -907,13 +842,7 @@ Section ContextConversion. red_ctx Σ Γ' Γ -> Σ ;;; Γ' |- T <= U. Proof using wfΣ. - intros H Hctx. - apply cumul_alt in H as [v [v' [[redl redr] leq]]]. - pose proof (red_red_ctx_inv redl Hctx). - pose proof (red_red_ctx_inv redr Hctx). - apply cumul_alt. - exists v, v'. - split. pcuic. auto. + apply cumul_pb_alt_red_ctx_inv. Qed. Lemma ws_cumul_pb_red_ctx_inv {pb Γ Γ'} {T U} : @@ -957,7 +886,7 @@ Section ContextConversion. Lemma ws_cumul_ctx_pb_red {pb} {Γ Γ' : context} : ws_cumul_ctx_pb pb Σ Γ Γ' -> ∑ Δ Δ', Σ ⊢ Γ ⇝ Δ × Σ ⊢ Γ' ⇝ Δ' × - eq_context_upto Σ (eq_universe Σ) (compare_universe pb Σ) Δ Δ'. + compare_context Σ pb Δ Δ'. Proof using wfΣ. intros Hctx. induction Hctx. @@ -1087,62 +1016,75 @@ End ContextConversion. #[global] Hint Resolve conv_ctx_refl' : pcuic. #[global] Hint Constructors All_decls_alpha_pb : pcuic. -Lemma eq_context_upto_conv_context {cf:checker_flags} (Σ : global_env_ext) Re : - RelationClasses.subrelation Re (eq_universe Σ) -> - subrelation (eq_context_upto Σ Re Re) (fun Γ Γ' => conv_context cumulAlgo_gen Σ Γ Γ'). +Lemma compare_context_cumul_pb_context {cf} Σ pb : + subrelation (compare_context Σ pb) (cumul_pb_context cumulAlgo_gen pb Σ). Proof. - intros HRe Γ Δ h. induction h. + intros Γ Δ h. induction h. - constructor. - constructor; tas. - depelim p; constructor; auto; constructor; tas; - eapply eq_term_upto_univ_impl; tea; auto. + destruct p; constructor; auto; constructor; tas. Qed. -Lemma eq_context_upto_cumul_context {cf:checker_flags} (Σ : global_env_ext) Re Rle : - RelationClasses.subrelation Re (eq_universe Σ) -> - RelationClasses.subrelation Rle (leq_universe Σ) -> - RelationClasses.subrelation Re Rle -> - subrelation (eq_context_upto Σ Re Rle) (fun Γ Γ' => cumul_context cumulAlgo_gen Σ Γ Γ'). +Lemma eq_context_upto_cumul_pb_context {cf:checker_flags} (Σ : global_env_ext) cmp_universe cmp_sort pb pb' : + RelationClasses.subrelation (cmp_universe Conv) (eq_universe Σ) -> + RelationClasses.subrelation (cmp_universe pb) (compare_universe Σ pb') -> + RelationClasses.subrelation (cmp_sort Conv) (eq_sort Σ) -> + RelationClasses.subrelation (cmp_sort pb) (compare_sort Σ pb') -> + subrelation (eq_context_upto Σ cmp_universe cmp_sort pb) (cumul_pb_context cumulAlgo_gen pb' Σ). Proof. - intros HRe HRle hR Γ Δ h. induction h. - - constructor. - - constructor; tas. - depelim p; constructor; auto; constructor; tas. - eapply eq_term_upto_univ_impl. 5:eauto. all:tea. - now transitivity Rle. auto. - eapply eq_term_upto_univ_impl; eauto. - eapply eq_term_upto_univ_impl. 5:eauto. all:tea. - now transitivity Rle. auto. + intros. + etransitivity. 2: apply compare_context_cumul_pb_context. + eapply eq_context_upto_impl; tc. +Qed. + +Lemma eq_context_upto_conv_context {cf:checker_flags} (Σ : global_env_ext) cmp_universe cmp_sort : + RelationClasses.subrelation (cmp_universe Conv) (eq_universe Σ) -> + RelationClasses.subrelation (cmp_sort Conv) (eq_sort Σ) -> + subrelation (eq_context_upto Σ cmp_universe cmp_sort Conv) (conv_context cumulAlgo_gen Σ). +Proof. + intros. + apply eq_context_upto_cumul_pb_context; tc. +Qed. + +Lemma eq_context_upto_cumul_context {cf:checker_flags} (Σ : global_env_ext) cmp_universe cmp_sort : + RelationClasses.subrelation (cmp_universe Conv) (eq_universe Σ) -> + RelationClasses.subrelation (cmp_universe Cumul) (leq_universe Σ) -> + RelationClasses.subrelation (cmp_sort Conv) (eq_sort Σ) -> + RelationClasses.subrelation (cmp_sort Cumul) (leq_sort Σ) -> + subrelation (eq_context_upto Σ cmp_universe cmp_sort Cumul) (fun Γ Γ' => cumul_context cumulAlgo_gen Σ Γ Γ'). +Proof. + intros. + apply eq_context_upto_cumul_pb_context; tc. Qed. #[global] Instance eq_subrel_eq_univ {cf:checker_flags} Σ : RelationClasses.subrelation eq (eq_universe Σ). Proof. intros x y []. reflexivity. Qed. -Lemma eq_context_upto_empty_conv_context {cf:checker_flags} (Σ : global_env_ext) : - subrelation (eq_context_upto empty_global_env eq eq) (fun Γ Γ' => conv_context cumulAlgo_gen Σ Γ Γ'). +#[global] +Instance eq_subrel_eq_sort {cf:checker_flags} Σ : RelationClasses.subrelation eq (eq_sort Σ). +Proof. intros x y []. reflexivity. Qed. + +Lemma eq_context_upto_names_conv_context {cf:checker_flags} (Σ : global_env_ext) : + subrelation eq_context_upto_names (fun Γ Γ' => conv_context cumulAlgo_gen Σ Γ Γ'). Proof. - intros Γ Δ h. induction h. - - constructor. - - constructor; tas. - depelim p; constructor; auto; constructor. - all:eapply eq_term_upto_univ_empty_impl; tea; try typeclasses eauto. + etransitivity. + 2: apply compare_context_cumul_pb_context. + eapply eq_context_upto_names_eq_context_upto; tc. Qed. Lemma eq_context_upto_univ_conv_context {cf:checker_flags} {Σ : global_env_ext} Γ Δ : - eq_context_upto Σ.1 (eq_universe Σ) (eq_universe Σ) Γ Δ -> + eq_context Σ.1 Σ Γ Δ -> conv_context cumulAlgo_gen Σ Γ Δ. Proof. - intros h. eapply eq_context_upto_conv_context; tea. - reflexivity. + apply compare_context_cumul_pb_context. Qed. Lemma eq_context_upto_univ_cumul_context {cf:checker_flags} {Σ : global_env_ext} Γ Δ : - eq_context_upto Σ.1 (eq_universe Σ) (leq_universe Σ) Γ Δ -> + leq_context Σ.1 Σ Γ Δ -> cumul_context cumulAlgo_gen Σ Γ Δ. Proof. - intros h. eapply eq_context_upto_cumul_context; tea. - reflexivity. tc. tc. + apply compare_context_cumul_pb_context. Qed. Lemma conv_context_app_same {cf:checker_flags} Σ Γ Γ' Δ : @@ -1165,54 +1107,26 @@ Proof. constructor; reflexivity. Qed. -#[global] Hint Extern 4 (eq_term_upto_univ _ _ _ _ _) => reflexivity : pcuic. - -(* Definition on_decl (P : context -> term -> term -> Type) - (Γ : context) (t : term) (t' : option term) := - match t' with - | Some (b, b') => (P Γ b b' * P Γ Γ' t t')%type - | None => P Γ Γ' t t' - end. *) -Definition on_local_decl (P : context -> term -> typ_or_sort -> Type) (Γ : context) (d : context_decl) := - match decl_body d with - | Some b => P Γ b (Typ (decl_type d)) * P Γ (decl_type d) Sort - | None => P Γ (decl_type d) Sort - end. - -Lemma nth_error_All_local_env {P Γ n} (isdecl : n < #|Γ|) : - All_local_env P Γ -> - on_some (on_local_decl P (skipn (S n) Γ)) (nth_error Γ n). -Proof. - induction 1 in n, isdecl |- *. red; simpl. - - destruct n; simpl; inv isdecl. - - destruct n. red; simpl. red. simpl. apply t0. - simpl. apply IHX. simpl in isdecl. lia. - - destruct n; simpl in *. - * rewrite skipn_S skipn_0. red; cbn. - split; auto. - * rewrite !skipn_S. apply IHX. lia. -Qed. +#[global] Hint Extern 4 (eq_term_upto_univ _ _ _ _ _ _) => reflexivity : pcuic. Lemma context_cumulativity_wf_app {cf:checker_flags} Σ Γ Γ' Δ : cumul_context cumulAlgo_gen Σ Γ' Γ -> wf_local Σ Γ' -> - All_local_env - (lift_typing - (fun (Σ : global_env_ext) (Γ : context) (t T : term) => - forall Γ' : context, - cumul_context cumulAlgo_gen Σ Γ' Γ -> wf_local Σ Γ' -> Σ;;; Γ' |- t : T) Σ) - (Γ,,, Δ) -> + All_local_env (fun Γ j => + forall Γ' : context, + cumul_context cumulAlgo_gen Σ Γ' Γ -> wf_local Σ Γ' -> + (lift_typing0 (fun t T => Σ;;; Γ' |- t : T)) j) + (Γ,,, Δ) -> wf_local Σ (Γ' ,,, Δ). Proof. intros. - eapply wf_local_app => //. + eapply All_local_env_app => //. eapply All_local_env_app_inv in X1 as []. eapply All_local_env_impl_ind; tea => /=. - intros Γ'' t' T H HT. - apply lift_typing_impl with (1 := HT); intros ? IH. - eapply IH. eapply All2_fold_app => //. + intros Γ'' j H HT. + eapply HT. eapply All2_fold_app => //. eapply All2_fold_refl. intros. eapply cumul_decls_refl. - eapply All_local_env_app; split; auto. + eapply All_local_env_app; auto. Qed. Lemma is_closed_context_cumul_app Γ Δ Γ' : diff --git a/pcuic/theories/PCUICContexts.v b/pcuic/theories/PCUICContexts.v index e4f354f3f..f90d157c1 100644 --- a/pcuic/theories/PCUICContexts.v +++ b/pcuic/theories/PCUICContexts.v @@ -127,7 +127,7 @@ Lemma type_local_ctx_wf_local {cf:checker_flags} Σ Γ Δ s : Proof. induction Δ; simpl; auto. destruct a as [na [b|] ty]; - intros wfΓ wfctx; constructor; intuition auto. exists s; auto. + intros wfΓ wfctx; constructor; intuition auto. now eapply lift_sorting_forget_univ. Qed. Lemma sorts_local_ctx_wf_local {cf:checker_flags} Σ Γ Δ s : @@ -139,7 +139,7 @@ Proof. destruct a as [na [b|] ty]; intros wfΓ wfctx; constructor; intuition auto. pcuic. destruct s => //. destruct wfctx; eauto. - destruct s => //. destruct wfctx. exists t; auto. + destruct s => //. destruct wfctx. now eapply lift_sorting_forget_univ. Qed. Lemma instantiate_minductive {cf:checker_flags} Σ ind mdecl u Γ t T : @@ -159,21 +159,18 @@ Lemma type_local_ctx_instantiate {cf:checker_flags} Σ ind mdecl Γ Δ u s : declared_minductive Σ.1 ind mdecl -> type_local_ctx (lift_typing typing) (Σ.1, ind_universes mdecl) Γ Δ s -> consistent_instance_ext Σ (ind_universes mdecl) u -> - type_local_ctx (lift_typing typing) Σ (subst_instance u Γ) (subst_instance u Δ) (subst_instance_univ u s). + type_local_ctx (lift_typing typing) Σ (subst_instance u Γ) (subst_instance u Δ) (subst_instance_sort u s). Proof. intros Hctx Hu. induction Δ; simpl in *; intuition auto. - { destruct Σ as [Σ univs]. eapply (wf_universe_subst_instance (Σ, ind_universes mdecl)); eauto. } - destruct a as [na [b|] ty]; simpl; [destruct X as (Hwfctx & Ht & Hb) | destruct X as (Hwfctx & Ht)]; repeat split. - - now apply IHΔ. - - eapply infer_typing_sort_impl with _ Ht; intros Hs. + { destruct Σ as [Σ univs]. eapply (wf_sort_subst_instance (Σ, ind_universes mdecl)); eauto. } + destruct a as [na [b|] ty]; simpl; destruct X as (Hwfctx & Hj); split; eauto. + - eapply lift_typing_fu_impl with (1 := Hj) => //= ?? Hs. + eapply instantiate_minductive in Hs; eauto. + now rewrite subst_instance_app in Hs. + - eapply lift_typing_fu_impl with (1 := Hj) => //= ?? Hs. eapply instantiate_minductive in Hs; eauto. now rewrite subst_instance_app in Hs. - - eapply instantiate_minductive in Hb; eauto. - now rewrite subst_instance_app in Hb. - - now apply IHΔ. - - eapply instantiate_minductive in Ht; eauto. - now rewrite subst_instance_app in Ht. Qed. Lemma sorts_local_ctx_instantiate {cf:checker_flags} Σ ind mdecl Γ Δ u s : @@ -182,20 +179,17 @@ Lemma sorts_local_ctx_instantiate {cf:checker_flags} Σ ind mdecl Γ Δ u s : sorts_local_ctx (lift_typing typing) (Σ.1, ind_universes mdecl) Γ Δ s -> consistent_instance_ext Σ (ind_universes mdecl) u -> sorts_local_ctx (lift_typing typing) Σ (subst_instance u Γ) (subst_instance u Δ) - (List.map (subst_instance_univ u) s). + (List.map (subst_instance_sort u) s). Proof. intros Hctx Hu. induction Δ in s |- *; simpl in *; intuition auto. - destruct s; simpl; intuition eauto. - destruct a as [na [b|] ty]; simpl. intuition eauto. - - eapply infer_typing_sort_impl with _ a0; intros Hs. - eapply instantiate_minductive in Hs; eauto. - now rewrite subst_instance_app in Hs. - - eapply instantiate_minductive in b1; eauto. - now rewrite subst_instance_app in b1. - - destruct s; simpl; intuition eauto. - eapply instantiate_minductive in b; eauto. - now rewrite subst_instance_app in b. + 1: destruct s; simpl; intuition eauto. + destruct a as [na [b|] ty]; simpl. + 2: destruct s => //=. + all: destruct X as (Hwfctx & Hj); split; eauto. + all: eapply lift_typing_fu_impl with (1 := Hj) => //= ?? Hs. + all: eapply instantiate_minductive in Hs; eauto. + all: now rewrite subst_instance_app in Hs. Qed. Lemma subst_type_local_ctx {cf:checker_flags} Σ Γ Δ Δ' s ctxs : @@ -205,16 +199,13 @@ Lemma subst_type_local_ctx {cf:checker_flags} Σ Γ Δ Δ' s ctxs : subslet Σ Γ s Δ -> type_local_ctx (lift_typing typing) Σ Γ (subst_context s 0 Δ') ctxs. Proof. + intros wfΣ wfΓ X Hs. revert X. induction Δ'; simpl; auto. - destruct a as [na [b|] ty]; simpl; intuition auto. - + simpl; rewrite subst_context_snoc /= /subst_decl /map_decl /= Nat.add_0_r; - repeat split; tas. - - apply infer_typing_sort_impl with id a0; intros Hs. - now eapply substitution in Hs. - - now eapply substitution in b1. - + rewrite subst_context_snoc /= /subst_decl /map_decl /= Nat.add_0_r. - intuition auto. - eapply substitution in b; eauto. + rewrite subst_context_snoc /= /subst_decl /map_decl /= Nat.add_0_r. + destruct a as [na [b|] ty]; simpl. + all: intros [X Hj]; split; eauto. + all: apply lift_typing_f_impl with (1 := Hj) => //= ?? HT. + all: now eapply substitution in HT. Qed. Lemma subst_sorts_local_ctx {cf:checker_flags} Σ Γ Δ Δ' s ctxs : @@ -224,18 +215,13 @@ Lemma subst_sorts_local_ctx {cf:checker_flags} Σ Γ Δ Δ' s ctxs : subslet Σ Γ s Δ -> sorts_local_ctx (lift_typing typing) Σ Γ (subst_context s 0 Δ') ctxs. Proof. + intros wfΣ wfΓ X Hs. revert X. induction Δ' in ctxs |- *; simpl; auto. - destruct a as [na [b|] ty]; simpl; intuition auto. - + simpl; rewrite subst_context_snoc /= /subst_decl /map_decl /= Nat.add_0_r. - repeat split. - - now apply IHΔ'. - - apply infer_typing_sort_impl with id a0; intros Hs. - now eapply substitution in Hs. - - now eapply substitution in b1. - + rewrite subst_context_snoc /= /subst_decl /map_decl /= Nat.add_0_r. - destruct ctxs; auto. - intuition auto. - eapply substitution in b; eauto. + rewrite subst_context_snoc /= /subst_decl /map_decl /= Nat.add_0_r. + destruct a as [na [b|] ty]; simpl. 2: destruct ctxs => //=. + all: intros [X Hj]; split; eauto. + all: apply lift_typing_f_impl with (1 := Hj) => //= ?? HT. + all: now eapply substitution in HT. Qed. Lemma weaken_type_local_ctx {cf:checker_flags} Σ Γ Γ' Δ ctxs : @@ -244,10 +230,12 @@ Lemma weaken_type_local_ctx {cf:checker_flags} Σ Γ Γ' Δ ctxs : type_local_ctx (lift_typing typing) Σ Γ' Δ ctxs -> type_local_ctx (lift_typing typing) Σ (Γ ,,, Γ') Δ ctxs. Proof. + intros wfΣ wfΓ. induction Δ; simpl; auto. - destruct a as [na [b|] ty]; simpl; intuition auto. - 1: apply infer_typing_sort_impl with id a0; intros Hs. - all: rewrite -app_context_assoc. + rewrite -app_context_assoc. + destruct a as [na [b|] ty]; simpl. + all: intros [X Hj]; split; eauto. + all: apply lift_typing_impl with (1 := Hj) => //= ?? HT. all: eapply (weaken_ctx Γ); auto. Qed. @@ -257,13 +245,13 @@ Lemma weaken_sorts_local_ctx {cf:checker_flags} Σ Γ Γ' Δ ctxs : sorts_local_ctx (lift_typing typing) Σ Γ' Δ ctxs -> sorts_local_ctx (lift_typing typing) Σ (Γ ,,, Γ') Δ ctxs. Proof. + intros wfΣ wfΓ. induction Δ in ctxs |- *; simpl; auto. - destruct a as [na [b|] ty]; simpl; intuition auto. - 1: apply infer_typing_sort_impl with id a0; intros Hs. - all: rewrite -app_context_assoc. - 3: destruct ctxs; auto. - 3: destruct X1; split; auto. - all: now eapply (weaken_ctx Γ). + rewrite -app_context_assoc. + destruct a as [na [b|] ty]; simpl. 2: destruct ctxs => //. + all: intros [X Hj]; split; eauto. + all: apply lift_typing_impl with (1 := Hj) => //= ?? HT. + all: eapply (weaken_ctx Γ); auto. Qed. Lemma reln_app acc Γ Δ k : reln acc k (Γ ++ Δ) = @@ -359,19 +347,18 @@ Section WfEnv. Proof using wfΣ. induction Γ in Δ |- *; simpl; auto. intros wfΓ wfΔ. depelim wfΓ; simpl. - - apply IHΓ; auto. eapply All_local_env_app. split; auto. - repeat constructor; auto. - eapply All_local_env_impl; eauto. simpl; intros. - now rewrite app_context_assoc. - - apply IHΓ. auto. eapply All_local_env_subst; eauto. simpl; intros. - destruct T; simpl in *; pcuicfo auto. - 2: apply infer_typing_sort_impl with id X; intros Hs. - 1: rename X into Hs. - all: rewrite Nat.add_0_r. - all: eapply (substitution (Γ':=[vdef na b t]) (s := [b])) in Hs; eauto. - all: rewrite -{1}(subst_empty 0 b). - all: repeat constructor. - all: now rewrite !subst_empty. + - apply IHΓ; auto. eapply All_local_env_app; auto. + + constructor; auto. + + eapply All_local_env_impl; eauto. simpl; intros. + now rewrite app_context_assoc. + - apply IHΓ; auto. eapply All_local_env_subst; eauto. simpl; intros. + apply lift_typing_f_impl with (1 := X) => // ?? HT. + rewrite Nat.add_0_r. + eapply (substitution (Γ':=[vdef na b t]) (s := [b])) in HT; eauto. + rewrite -{1}(subst_empty 0 b). + repeat constructor. + rewrite !subst_empty. + now eapply unlift_TermTyp. Qed. Lemma wf_local_rel_smash_context {Γ Δ} : @@ -385,7 +372,7 @@ Section WfEnv. wf_local Σ (Γ ,,, Δ) -> wf_local Σ (Γ ,,, smash_context [] Δ). Proof using wfΣ. intros wf. - apply All_local_env_app. split. + apply All_local_env_app. now apply All_local_env_app_inv in wf. eapply wf_local_rel_smash_context; auto. Qed. @@ -654,41 +641,39 @@ Lemma subslet_extended_subst {cf} {Σ} {wfΣ : wf Σ} Γ Δ : (lift_context (context_assumptions Δ) 0 Δ). Proof. move=> wfΔ. - eapply wf_local_app_inv in wfΔ as [wfΓ wfΔ]. + eapply All_local_env_app_inv in wfΔ as [wfΓ wfΔ]. induction Δ as [|[na [d|] ?] ?] in wfΔ |- *; simpl; try constructor. * depelim wfΔ. repeat red in l, l0. specialize (IHΔ wfΔ). rewrite lift_context_snoc /lift_decl /= /map_decl /=. len. constructor => //. - eapply (weakening_typing (Γ'' := smash_context [] Δ)) in l0. - len in l0. simpl in l0. simpl. + apply unlift_TermTyp in l. + eapply (weakening_typing (Γ'' := smash_context [] Δ)) in l. + len in l. simpl in l. simpl. 2:{ eapply wf_local_smash_end; pcuic. } - eapply (substitution (Δ := [])) in l0; tea. + eapply (substitution (Δ := [])) in l; tea. * rewrite smash_context_acc. simpl. rewrite /map_decl /= /map_decl /=. simpl. depelim wfΔ. specialize (IHΔ wfΔ). rewrite lift_context_snoc /lift_decl /= /map_decl /=. + assert (wf_local Σ (Γ ,,, smash_context [] Δ)). + { eapply wf_local_smash_end. now eapply All_local_env_app. } + assert (wf_local Σ (Γ ,,, smash_context [] Δ ,, mkdecl na None + (subst0 (extended_subst Δ 0) (lift (context_assumptions Δ) #|Δ| decl_type)))). + { simpl. constructor; tas. + eapply lift_typing_f_impl with (f := fun t => _ (_ t)) (1 := l) => // ?? Hs. + eapply (weakening_typing (Γ'' := smash_context [] Δ)) in Hs; tas. + len in Hs. + now eapply (substitution (Δ := [])) in Hs. + } constructor. - rewrite (lift_extended_subst _ 1). rewrite -(lift_context_lift_context 1 _). eapply (subslet_lift _ [_]); eauto. - constructor. - { pose proof l.π2. eapply wf_local_smash_end; pcuic. } - apply infer_typing_sort_impl with id l; intros Hs. - eapply (weakening_typing (Γ'' := smash_context [] Δ)) in Hs. - len in Hs. simpl in Hs. simpl. - 2:{ eapply wf_local_smash_end; pcuic. } - eapply (substitution (Δ := [])) in Hs; tea. - eapply meta_conv. - econstructor. constructor. apply wf_local_smash_end; auto. - eapply wf_local_app; eauto. - apply infer_typing_sort_impl with id l; intros Hs. - eapply (weakening_typing (Γ'' := smash_context [] Δ)) in Hs. - len in Hs. simpl in Hs. simpl. - 2:{ eapply wf_local_smash_end; pcuic. } - eapply (substitution (Δ := [])) in Hs; tea. + econstructor; tas. reflexivity. simpl. rewrite (lift_extended_subst _ 1). rewrite distr_lift_subst. f_equal. len. @@ -776,7 +761,7 @@ Proof. elim: Δ2 us2 Δ1 us1=> [|+ Δ2 ih]. 1: move=> [|u us2] //=. move=> [? [bdy|] ty] us2 Δ1 us1 /=. - - rewrite app_context_assoc=> ? [? [??]] ; split=> //. + - rewrite app_context_assoc=> ? [? ?] ; split=> //. by apply: ih. - case: us2=> [|u us2] // ? [??] /=. rewrite app_context_assoc; split=> //. diff --git a/pcuic/theories/PCUICConvCumInversion.v b/pcuic/theories/PCUICConvCumInversion.v index 53b5e80af..b0f9da422 100644 --- a/pcuic/theories/PCUICConvCumInversion.v +++ b/pcuic/theories/PCUICConvCumInversion.v @@ -20,11 +20,10 @@ Definition conv_cum {cf:checker_flags} pb Σ Γ u v := Lemma eq_term_eq_termp {cf:checker_flags} leq (Σ : global_env_ext) x y : eq_term Σ Σ x y -> - eq_termp leq Σ x y. + eq_termp Σ leq x y. Proof. destruct leq; [easy|]. - cbn. - apply eq_term_upto_univ_leq; cbn; auto. + apply eq_term_leq_term. Qed. Lemma alt_into_ws_cumul_pb_terms {cf Σ} {wfΣ : wf Σ} {Γ l l'} : @@ -43,7 +42,7 @@ Lemma red_ctx_rel_par_conv {cf Σ Γ Γ0 Γ0' Γ1 Γ1'} {wfΣ : wf Σ} : is_closed_context (Γ ,,, Γ1) -> red_ctx_rel Σ Γ Γ0 Γ0' -> red_ctx_rel Σ Γ Γ1 Γ1' -> - eq_context_upto Σ (eq_universe Σ) (eq_universe Σ) Γ0' Γ1' -> + eq_context Σ Σ Γ0' Γ1' -> ws_cumul_ctx_pb_rel Conv Σ Γ Γ0 Γ1. Proof. intros clΓ0 clΓ1 r0 r1 eq. @@ -57,10 +56,9 @@ Proof. eapply (red_ctx_ws_cumul_ctx_pb (l:=Conv)) in r0. eapply (red_ctx_ws_cumul_ctx_pb (l:=Conv)) in r1. apply ws_cumul_ctx_pb_rel_app. etransitivity; tea. - symmetry. etransitivity; tea. - eapply (eq_context_upto_cat _ _ _ Γ _ Γ) in eq. 2:reflexivity. - eapply (eq_context_upto_ws_cumul_ctx_pb (pb:=Conv)) in eq. 2-3:fvs. - now symmetry. + etransitivity. 2: now symmetry. + eapply eq_context_upto_ws_cumul_ctx_pb. 1-2:fvs. + eapply eq_context_upto_cat. 1: reflexivity. assumption. Qed. Lemma into_red_terms {Σ Γ ts ts'} : @@ -73,13 +71,6 @@ Proof. move=> /= clΓ /andP[clx cll]. constructor; eauto using into_closed_red. Qed. -Lemma alpha_eq_context_gen Γ Δ : - eq_context_upto_names Γ Δ -> - eq_context_gen eq eq Γ Δ. -Proof. - induction 1; constructor; auto. -Qed. - Lemma untyped_subslet_ass {Γ s Δ} : assumption_context Δ -> #|s| = context_assumptions Δ -> @@ -107,8 +98,8 @@ Lemma inst_case_ws_cumul_ctx_pb {cf Σ} {wfΣ : wf Σ} {ind mdecl idecl Γ pars on_free_vars_ctx (closedP #|pars'| xpredT) ctx' -> is_closed_context Γ -> ws_cumul_pb_terms Σ Γ pars pars' -> - R_universe_instance (eq_universe Σ) puinst puinst' -> - eq_context_gen eq eq ctx ctx' -> + cmp_universe_instance (eq_universe Σ) puinst puinst' -> + eq_context_upto_names ctx ctx' -> Σ ⊢ Γ,,, inst_case_context pars puinst ctx = Γ,,, inst_case_context pars' puinst' ctx'. Proof. intros decli wfp wfp' onp onp' clΓ eqpars eqinst eqctx. @@ -139,7 +130,7 @@ Proof. rewrite on_free_vars_ctx_subst_instance -lenpars. eapply on_free_vars_ctx_impl; tea. apply shiftnP_up. lia. } eapply eq_context_upto_cat; [reflexivity|]. - eapply eq_context_upto_univ_subst_instance'; tc. 1:reflexivity. + eapply eq_context_upto_names_subst_instance; tc. 1:reflexivity. assumption. - cbn. eapply subst_instance_ws_cumul_ctx_pb_rel => //. @@ -159,16 +150,15 @@ Qed. #[global] Hint Resolve sq : core. -Lemma conv_cum_alt {cf:checker_flags} {leq} {Σ : global_env_ext} {Γ t t'} (wfΣ : ∥ wf Σ ∥) : - conv_cum leq Σ Γ t t' <-> - ∥∑ v v', [× Σ ;;; Γ ⊢ t ⇝ v, Σ ;;; Γ ⊢ t' ⇝ v' & eq_termp leq Σ v v']∥. +Lemma conv_cum_alt {cf:checker_flags} {pb} {Σ : global_env_ext} {Γ t t'} (wfΣ : ∥ wf Σ ∥) : + conv_cum pb Σ Γ t t' <-> + ∥∑ v v', [× Σ ;;; Γ ⊢ t ⇝ v, Σ ;;; Γ ⊢ t' ⇝ v' & eq_termp Σ pb v v']∥. Proof. destruct wfΣ. assert (forall P Q, (P <~> Q) -> (∥P∥ <-> ∥Q∥)) by (intros P Q H; split; intros [p]; constructor; apply H in p; auto). - destruct leq; cbn; apply H. - * eapply (ws_cumul_pb_alt_closed (pb:=Conv)). - * eapply (ws_cumul_pb_alt_closed (pb:=Cumul)). + apply H. + apply ws_cumul_pb_alt_closed. Qed. Lemma conv_conv_cum_l {cf:checker_flags} : @@ -215,10 +205,10 @@ Section fixed. rewrite !decompose_app_mkApps; by easy. Qed. - Lemma eq_term_upto_univ_napp_nonind Re Rle napp t t' : - eq_term_upto_univ_napp Σ Re Rle napp t t' -> + Lemma eq_term_upto_univ_napp_nonind cmp_universe cmp_sort pb napp t t' : + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp t t' -> isIndConstructApp t = false -> - eq_term_upto_univ Σ Re Rle t t'. + eq_term_upto_univ Σ cmp_universe cmp_sort pb t t'. Proof using Type. intros eq not_ind. generalize 0. @@ -240,30 +230,30 @@ Section fixed. apply IHr. Qed. - Lemma eq_termp_mkApps_inv leq v args v' args' : + Lemma eq_termp_mkApps_inv pb v args v' args' : isApp v = false -> isApp v' = false -> - eq_termp leq Σ (mkApps v args) (mkApps v' args') -> - eq_termp_napp leq Σ #|args| v v' × All2 (fun x y => eq_term Σ Σ x y) args args'. + eq_termp Σ pb (mkApps v args) (mkApps v' args') -> + eq_termp_napp Σ pb #|args| v v' × All2 (fun x y => eq_term Σ Σ x y) args args'. Proof using Type. intros noapp1 noapp2 eq. apply eq_term_upto_univ_mkApps_inv in eq as (?&?) => //. Qed. - Definition conv_cum_napp leq Γ napp t t' := + Definition conv_cum_napp pb Γ napp t t' := match t with | tInd _ _ - | tConstruct _ _ _ => ∥eq_termp_napp leq Σ napp t t'∥ - | _ => conv_cum leq Σ Γ t t' + | tConstruct _ _ _ => ∥eq_termp_napp Σ pb napp t t'∥ + | _ => conv_cum pb Σ Γ t t' end. - Lemma conv_cum_mkApps_inv leq Γ hd args hd' args' : - conv_cum leq Σ Γ (mkApps hd args) (mkApps hd' args') -> + Lemma conv_cum_mkApps_inv pb Γ hd args hd' args' : + conv_cum pb Σ Γ (mkApps hd args) (mkApps hd' args') -> isApp hd = false -> isApp hd' = false -> whnf RedFlags.default Σ Γ (mkApps hd args) -> whnf RedFlags.default Σ Γ (mkApps hd' args') -> - ∥conv_cum_napp leq Γ #|args| hd hd' × ws_cumul_pb_terms Σ Γ args args'∥. + ∥conv_cum_napp pb Γ #|args| hd hd' × ws_cumul_pb_terms Σ Γ args args'∥. Proof using wfΣ. intros conv notapp notapp' wh wh'. eapply conv_cum_alt in conv as [(?&?&[r1 r2 e])]; auto. @@ -295,13 +285,13 @@ Section fixed. try (constructor; constructor; rewrite a; auto). sq. exists (tPrim i'), (tPrim i'0). split => //. all:eauto with pcuic. - eapply (eq_term_upto_univ_napp_nonind _ _ 0); eauto. now constructor. + now constructor. - eapply alt_into_ws_cumul_pb_terms => //. clear -a1 a a0. induction a1 in args, args', x2, a, x3, a0, a1 |- *; depelim a; depelim a0; [now constructor|]. constructor. - + apply conv_alt_red. + + apply cumul_alt. now exists x, y. + now apply IHa1. Qed. @@ -375,9 +365,9 @@ Section fixed. eapply on_free_vars_ctx_inst_case_context; tea => //. now rewrite test_context_k_closed_on_free_vars_ctx. + eapply red_on_free_vars in r1; tea. - { len. rewrite (All2_fold_length a5). + { len. rewrite (All2_length a4). now setoid_rewrite shiftnP_add in p1. } - len. rewrite -shiftnP_add (All2_fold_length a5). + len. rewrite -shiftnP_add (All2_length a4). eapply on_ctx_free_vars_inst_case_context; auto. 1:now rewrite test_context_k_closed_on_free_vars_ctx. now erewrite -> on_free_vars_ctx_on_ctx_free_vars. } @@ -390,8 +380,8 @@ Section fixed. all:eapply into_closed_red; eauto. - rename a0 into brsa1. rename a2 into brsa2. - rename a3 into brseq. - clear -wfΣ decli brsa1 brsa2 brseq clΓ wfp wfp' a a1 p0 p5 p4 p9 r3 eqpars. + rename e0 into brseq. + clear -wfΣ decli brsa1 brsa2 brseq clΓ wfp wfp' a a1 p0 p5 p4 p9 c eqpars. induction brseq in brs, brs', brsa1, brsa2, p4, p9 |- *; depelim brsa1; depelim brsa2; [constructor|]. destruct p0, p1, r. @@ -420,7 +410,7 @@ Section fixed. move/andP: fv' => []. len. now setoid_rewrite shiftnP_add. } move/andP: fv => [] fv fvx1. len. eapply red_on_free_vars in fvx1; tea. - { rewrite e (All2_fold_length a0) -e0. now setoid_rewrite shiftnP_add in fvx1. } + { rewrite e (All2_length a0) -e0. now setoid_rewrite shiftnP_add in fvx1. } rewrite shiftnP_add. relativize (#|bcontext x1| + _). 1:rewrite -> on_free_vars_ctx_on_ctx_free_vars. 2:now len. now eapply ws_cumul_ctx_pb_closed_right in eqctx. } @@ -465,20 +455,22 @@ Section fixed. { rewrite on_free_vars_ctx_app clΓ /=. apply on_free_vars_fix_context; solve_all. } have clmfixctx' : is_closed_context (Γ ,,, fix_context mfix'). { rewrite on_free_vars_ctx_app clΓ /=. apply on_free_vars_fix_context; solve_all. } + unfold eq_mfixpoint in *. solve_all. move: clmfixctx clmfixctx'. - clear -wfΣ a a0 a1 clΓ. + clear -wfΣ a a0 e clΓ. cut (#|mfix| = #|mfix'|); - [|now apply All2_length in a; apply All2_length in a0; apply All2_length in a1]. - revert a a0 a1. - generalize mfix at 1 2 4 5 6. - generalize mfix' at 1 2 4 5. - intros ctx_fix ctx_fix'. + [|now apply All2_length in a; apply All2_length in a0; apply All2_length in e]. + revert a a0 e. + set mfix as ctx_fix in |-. set mfix' as ctx_fix' in |-. + change (fix_context mfix) with (fix_context ctx_fix). change (fix_context mfix') with (fix_context ctx_fix'). + change #|mfix| with #|ctx_fix|. change #|mfix'| with #|ctx_fix'|. + clearbody ctx_fix ctx_fix'. intros all1 all2 all len_eq. induction all in mfix, mfix', all1, all2, all |- *; depelim all1; depelim all2; subst; [constructor|]. constructor; [|now auto]. - destruct r as ((?&(((? & ?) & ?)&?))&?), p as (?&?&?&?&?), p0 as (?&?&?&?&?). + destruct r as ((?&(? & ? & ? & ?))&?), p as (?&?&?&?&?), p0 as (?&?&?&?&?). split; auto; try congruence. - eapply ws_cumul_pb_alt_closed; exists (dtype x), (dtype y). split; eauto. all:eapply into_closed_red; eauto. @@ -525,20 +517,22 @@ Section fixed. { rewrite on_free_vars_ctx_app clΓ /=. apply on_free_vars_fix_context; solve_all. } have clmfixctx' : is_closed_context (Γ ,,, fix_context mfix'). { rewrite on_free_vars_ctx_app clΓ /=. apply on_free_vars_fix_context; solve_all. } + unfold eq_mfixpoint in *. solve_all. move: clmfixctx clmfixctx'. - clear -wfΣ a a0 a1 clΓ. + clear -wfΣ a a0 e clΓ. cut (#|mfix| = #|mfix'|); - [|now apply All2_length in a; apply All2_length in a0; apply All2_length in a1]. - revert a a0 a1. - generalize mfix at 1 2 4 5 6. - generalize mfix' at 1 2 4 5. - intros ctx_fix ctx_fix'. + [|now apply All2_length in a; apply All2_length in a0; apply All2_length in e]. + revert a a0 e. + set mfix as ctx_fix in |-. set mfix' as ctx_fix' in |-. + change (fix_context mfix) with (fix_context ctx_fix). change (fix_context mfix') with (fix_context ctx_fix'). + change #|mfix| with #|ctx_fix|. change #|mfix'| with #|ctx_fix'|. + clearbody ctx_fix ctx_fix'. intros all1 all2 all len_eq. induction all in mfix, mfix', all1, all2, all |- *; depelim all1; depelim all2; subst; [constructor|]. constructor; [|now auto]. - destruct r as ((?&(((? & ?) & ?)&?))&?), p as (?&?&?&?&?), p0 as (?&?&?&?&?). + destruct r as ((?&(? & ? & ? & ?))&?), p as (?&?&?&?&?), p0 as (?&?&?&?&?). split; auto; try congruence. - eapply ws_cumul_pb_alt_closed; exists (dtype x), (dtype y). split; eauto. all:eapply into_closed_red; eauto. diff --git a/pcuic/theories/PCUICConversion.v b/pcuic/theories/PCUICConversion.v index 23a5771aa..89c49a3b4 100644 --- a/pcuic/theories/PCUICConversion.v +++ b/pcuic/theories/PCUICConversion.v @@ -38,7 +38,7 @@ Ltac pcuic := intuition eauto 5 with pcuic || (try solve [repeat red; cbn in *; intuition auto; eauto 5 with pcuic || (try lia || congruence)]). #[global] -Hint Resolve eq_universe_leq_universe' : pcuic. +Hint Resolve eq_leq_sort' : pcuic. Derive Signature for assumption_context. Derive Signature for clos_refl_trans_1n. @@ -91,7 +91,7 @@ Section CumulSpecIsCumulAlgo. * change (fun Γ t u => Σ ;;; Γ ⊢ t ≤s[Conv] u) with (convSpec Σ). destruct p as [p x]. cbn in *. try repeat split; cbn; try reflexivity; eauto. * induction X; econstructor. - + destruct p0 as [ [ _ hbody ] hhd ]. rewrite hhd. split; eauto. reflexivity. + + unfold cumul_branch. destruct p0 as [ [ _ hbody ] hhd ]. rewrite hhd. split; eauto. reflexivity. + apply All2_reflexivity. eapply Prod_reflexivity; intro x; reflexivity. + split; reflexivity. + exact IHX. @@ -99,7 +99,7 @@ Section CumulSpecIsCumulAlgo. change (fun t u => Σ ;;; Γ ⊢ t ≤s[Conv] u) with (convSpec Σ Γ). induction X; econstructor; eauto; try reflexivity. * exact p.2. - - eapply cumul_Fix. set (mfixAbs := mfix0). unfold mfixAbs at 2. clearbody mfixAbs. + - eapply cumul_Fix. unfold cumul_mfixpoint. set (Ξ := fix_context mfix0). clearbody Ξ. induction X; econstructor; eauto; try reflexivity. * destruct p as [ [ _ hdtype ] e ]. pose proof (erarg := snd_eq e). pose proof (edbody := snd_eq (fst_eq e)). @@ -107,16 +107,15 @@ Section CumulSpecIsCumulAlgo. repeat split; eauto ; try reflexivity. * apply All2_reflexivity. repeat eapply Prod_reflexivity; intro x; reflexivity. * repeat split; reflexivity. - - eapply cumul_Fix. set (mfixAbs := mfix0). unfold mfixAbs at 2. - assert (Habs : mfixAbs = mfix0) by reflexivity. clearbody mfixAbs. - induction X; destruct Habs; econstructor; eauto; try reflexivity. + - eapply cumul_Fix. unfold cumul_mfixpoint. set (Ξ := fix_context mfix0) in *. clearbody Ξ. + induction X; econstructor; eauto; try reflexivity. * destruct p as [ [ _ hdtype ] e ]. pose proof (erarg := snd_eq e). pose proof (edbody := snd_eq (fst_eq e)). pose proof (edname := fst_eq (fst_eq e)). clear e. destruct erarg, edbody, edname. repeat split; eauto ; try reflexivity. * apply All2_reflexivity. repeat eapply Prod_reflexivity; intro x; reflexivity. * repeat split; reflexivity. - - eapply cumul_CoFix. set (mfixAbs := mfix0). unfold mfixAbs at 2. clearbody mfixAbs. + - eapply cumul_CoFix. unfold cumul_mfixpoint. set (Ξ := fix_context mfix0) in *. clearbody Ξ. induction X; econstructor; eauto; try reflexivity. * destruct p as [ [ _ hdtype ] e ]. pose proof (erarg := snd_eq e). pose proof (edbody := snd_eq (fst_eq e)). @@ -124,9 +123,8 @@ Section CumulSpecIsCumulAlgo. repeat split; eauto ; try reflexivity. * apply All2_reflexivity. repeat eapply Prod_reflexivity; intro x; reflexivity. * repeat split; reflexivity. - - eapply cumul_CoFix. set (mfixAbs := mfix0). unfold mfixAbs at 2. - assert (Habs : mfixAbs = mfix0) by reflexivity. clearbody mfixAbs. - induction X; destruct Habs; econstructor; eauto; try reflexivity. + - eapply cumul_CoFix. unfold cumul_mfixpoint. set (Ξ := fix_context mfix0) in *. clearbody Ξ. + induction X; econstructor; eauto; try reflexivity. * destruct p as [ [ _ hdtype ] e ]. pose proof (erarg := snd_eq e). pose proof (edbody := snd_eq (fst_eq e)). pose proof (edname := fst_eq (fst_eq e)). clear e. destruct erarg, edbody, edname. @@ -138,13 +136,11 @@ Section CumulSpecIsCumulAlgo. + apply X0. + reflexivity. - eapply cumul_Prim. constructor; eauto; trea. - eapply All2_reflexivity; intro; reflexivity. - eapply cumul_Prim; constructor; eauto; trea. - eapply All2_reflexivity; intro; reflexivity. Defined. - Proposition convSpec_cumulSpec (Γ : context) (M N : term) : - Σ ;;; Γ |- M =s N -> Σ ;;; Γ |- M <=s N. + Proposition convSpec_cumulSpec (Γ : context) (pb : conv_pb) (M N : term) : + Σ ;;; Γ |- M =s N -> Σ ;;; Γ ⊢ M ≤s[pb] N. Proof. intro Hconv. apply cumul_Sym. apply cumul_Sym. assumption. Defined. @@ -162,7 +158,7 @@ Section CumulSpecIsCumulAlgo. Ltac try_with_nil := match goal with H : _ |- _ => eapply (H _ _ _ [] []); eauto end. Proposition eq_term_upto_univ_napp_cumulSpec (Γ : context) {pb} M N args args' : - compare_term_napp pb Σ Σ #|args| M N -> + compare_term_napp Σ Σ pb #|args| M N -> All2 (cumulSpec0 Σ Γ Conv) args args' -> cumulSpec0 Σ Γ pb (mkApps M args) (mkApps N args'). Proof. @@ -184,7 +180,8 @@ Section CumulSpecIsCumulAlgo. + intuition. + try_with_nil. intuition. * try_with_nil. - * unfold tCaseBrsProp in X0. eapply All2_All_mix_left in X0. 2: tea. + * unfold tCaseBrsProp in X0. unfold cumul_branches, eq_branches, cumul_branch, eq_branch in *. + eapply All2_All_mix_left in X0. 2: tea. eapply All2_impl. 1: eassumption. cbn. intuition. try_with_nil. - apply cumul_mkApps; eauto. eapply cumul_Fix. unfold tFixProp in *. @@ -196,7 +193,7 @@ Section CumulSpecIsCumulAlgo. Defined. Proposition eq_term_upto_univ_cumulSpec (Γ : context) {pb} M N : - compare_term pb Σ Σ M N -> cumulSpec0 Σ Γ pb M N. + compare_term Σ Σ pb M N -> cumulSpec0 Σ Γ pb M N. Proof. intros. eapply (eq_term_upto_univ_napp_cumulSpec _ _ _ [] []); eauto. Defined. @@ -205,13 +202,9 @@ Section CumulSpecIsCumulAlgo. Σ ;;; Γ ⊢ M ≤[pb] N -> Σ ;;; Γ ⊢ M ≤s[pb] N. Proof. induction 1. - - destruct pb; eapply eq_term_upto_univ_cumulSpec; eauto. - - destruct pb; revgoals. - * eapply cumul_Trans; eauto. apply convSpec_cumulSpec. apply red1_cumulSpec ; assumption. - * eapply cumul_Trans; eauto. apply red1_cumulSpec ; assumption. - - destruct pb; revgoals. - * eapply cumul_Trans; eauto. apply cumul_Sym. apply red1_cumulSpec ; assumption. - * eapply cumul_Trans; eauto. apply cumul_Sym. apply red1_cumulSpec ; assumption. + - eapply eq_term_upto_univ_cumulSpec; eauto. + - eapply cumul_Trans; eauto. apply convSpec_cumulSpec. apply red1_cumulSpec ; assumption. + - eapply cumul_Trans; eauto. apply cumul_Sym. apply red1_cumulSpec ; assumption. Defined. End CumulSpecIsCumulAlgo. @@ -266,7 +259,7 @@ Qed. #[global] Hint Resolve closed_red_open_right : fvs. Ltac fvs := eauto 10 with fvs. -#[global] Hint Resolve eq_universe_leq_universe : core. +#[global] Hint Resolve eq_leq_sort : core. Section ConvCongruences. Context {cf:checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ}. @@ -319,10 +312,10 @@ Section ConvCongruences. Lemma ws_cumul_pb_Sort_inv {Γ s s'} pb : Σ ;;; Γ ⊢ tSort s ≤[pb] tSort s' -> - compare_universe pb Σ s s'. + compare_sort Σ pb s s'. Proof using Type. intros H; depind H. - - destruct pb; now inversion c. + - now inversion c. - depelim r. solve_discr. - depelim r. solve_discr. Qed. @@ -332,7 +325,7 @@ Section ConvCongruences. False. Proof using Type. intros H. depind H. - - destruct pb; now inversion c. + - now inversion c. - depelim r. solve_discr. - depelim r; solve_discr. + eapply IHws_cumul_pb. reflexivity. @@ -343,7 +336,7 @@ Section ConvCongruences. Σ ;;; Γ ⊢ tProd na dom codom ≤[pb] tSort s -> False. Proof using Type. intros H; depind H; auto. - - destruct pb; now inversion c. + - now inversion c. - depelim r. + solve_discr. + eapply IHws_cumul_pb; reflexivity. @@ -353,10 +346,10 @@ Section ConvCongruences. Lemma ws_cumul_pb_Sort_l_inv {Γ s T} pb : Σ ;;; Γ ⊢ tSort s ≤[pb] T -> - ∑ s', Σ ;;; Γ ⊢ T ⇝ tSort s' × compare_universe pb Σ s s'. + ∑ s', Σ ;;; Γ ⊢ T ⇝ tSort s' × compare_sort Σ pb s s'. Proof using Type. intros H. depind H. - - destruct pb; inversion c; eauto using into_closed_red. + - inversion c; eauto using into_closed_red. - depelim r. solve_discr. - destruct IHws_cumul_pb as [s' [redv leq]]. exists s'. split; auto. eapply into_closed_red; tea. @@ -365,10 +358,10 @@ Section ConvCongruences. Lemma ws_cumul_pb_Sort_r_inv {Γ s T} pb : Σ ;;; Γ ⊢ T ≤[pb] tSort s -> - ∑ s', Σ ;;; Γ ⊢ T ⇝ tSort s' × compare_universe pb Σ s' s. + ∑ s', Σ ;;; Γ ⊢ T ⇝ tSort s' × compare_sort Σ pb s' s. Proof using Type. intros H. depind H. - - destruct pb; inversion c; eauto using into_closed_red. + - inversion c; eauto using into_closed_red. - destruct IHws_cumul_pb as [s' [redv leq]]. exists s'. split; auto. eapply into_closed_red; tea. eapply red_step with v; eauto with fvs. @@ -1065,16 +1058,16 @@ Section Inversions. Qed. Lemma eq_term_upto_univ_conv_arity_l : - forall Re Rle Γ u v, + forall cmp_universe cmp_sort pb Γ u v, isArity u -> is_closed_context Γ -> is_open_term Γ u -> is_open_term Γ v -> - eq_term_upto_univ Σ Re Rle u v -> + eq_term_upto_univ Σ cmp_universe cmp_sort pb u v -> Is_conv_to_Arity Σ Γ v. Proof using Type. - intros Re Rle Γ u v a clΓ clu clv e. - induction u in Γ, clΓ, clv, clu, a, v, Rle, e |- *. all: try (cbn [isArity] in *; congruence). + intros cmp_universe cmp_sort pb Γ u v a clΓ clu clv e. + induction u in Γ, clΓ, clv, clu, a, v, pb, e |- *. all: try (cbn [isArity] in *; congruence). all: dependent destruction e. - eexists. split. + constructor. eapply closed_red_refl => //. @@ -1100,16 +1093,16 @@ Section Inversions. Qed. Lemma eq_term_upto_univ_conv_arity_r : - forall Re Rle Γ u v, + forall cmp_universe cmp_sort pb Γ u v, isArity u -> is_closed_context Γ -> is_open_term Γ u -> is_open_term Γ v -> - eq_term_upto_univ Σ Re Rle v u -> + eq_term_upto_univ Σ cmp_universe cmp_sort pb v u -> Is_conv_to_Arity Σ Γ v. Proof using Type. - intros Re Rle Γ u v a clΓ clu clv e. - induction u in Γ, clΓ, clv, clu, a, v, Rle, e |- *. all: try (cbn [isArity] in *; congruence). + intros cmp_universe cmp_sort pb Γ u v a clΓ clu clv e. + induction u in Γ, clΓ, clv, clu, a, v, pb, e |- *. all: try (cbn [isArity] in *; congruence). all: dependent destruction e. - eexists. split. + constructor. eapply closed_red_refl => //. @@ -1285,28 +1278,28 @@ Section Inversions. Notation invert_red_ind := red_mkApps_tInd. Lemma compare_term_mkApps_l_inv {pb} {u : term} {l : list term} {t : term} : - compare_term pb Σ Σ (mkApps u l) t -> + compare_term Σ Σ pb (mkApps u l) t -> ∑ (u' : term) (l' : list term), - [× eq_term_upto_univ_napp Σ (eq_universe Σ) (compare_universe pb Σ) #|l| u u', + [× compare_term_napp Σ Σ pb #|l| u u', All2 (eq_term Σ Σ) l l' & t = mkApps u' l']. Proof using wfΣ. - destruct pb => /= => /eq_term_upto_univ_mkApps_l_inv; firstorder. + move => /= => /eq_term_upto_univ_mkApps_l_inv; firstorder. Qed. Lemma compare_term_mkApps_r_inv {pb} {u : term} {l : list term} {t : term} : - compare_term pb Σ Σ t (mkApps u l) -> + compare_term Σ Σ pb t (mkApps u l) -> ∑ (u' : term) (l' : list term), - [× eq_term_upto_univ_napp Σ (eq_universe Σ) (compare_universe pb Σ) #|l| u' u, + [× compare_term_napp Σ Σ pb #|l| u' u, All2 (eq_term Σ Σ) l' l & t = mkApps u' l']. Proof using wfΣ. - destruct pb => /= => /eq_term_upto_univ_mkApps_r_inv; firstorder. + move => /= => /eq_term_upto_univ_mkApps_r_inv; firstorder. Qed. Lemma ws_cumul_pb_Ind_l_inv {pb Γ ind ui l T} : Σ ;;; Γ ⊢ mkApps (tInd ind ui) l ≤[pb] T -> ∑ ui' l', [× Σ ;;; Γ ⊢ T ⇝ (mkApps (tInd ind ui') l'), - R_global_instance Σ (eq_universe Σ) (compare_universe pb Σ) (IndRef ind) #|l| ui ui' & + cumul_Ind_univ Σ pb ind #|l| ui ui' & All2 (fun a a' => Σ ;;; Γ ⊢ a = a') l l']. Proof using wfΣ. move/ws_cumul_pb_red=> [v [v' [redv redv' leqvv']]]. @@ -1356,7 +1349,7 @@ Section Inversions. Σ ;;; Γ ⊢ T ≤[pb] mkApps (tInd ind ui) l -> ∑ ui' l', [× Σ ;;; Γ ⊢ T ⇝ (mkApps (tInd ind ui') l'), - R_global_instance Σ (eq_universe Σ) (compare_universe pb Σ) (IndRef ind) #|l| ui' ui & + cumul_Ind_univ Σ pb ind #|l| ui' ui & ws_cumul_pb_terms Σ Γ l' l]. Proof using wfΣ. move/ws_cumul_pb_red=> [v [v' [redv redv' leqvv']]]. @@ -1381,7 +1374,7 @@ Section Inversions. Lemma ws_cumul_pb_Ind_inv {pb Γ ind ind' ui ui' l l'} : Σ ;;; Γ ⊢ mkApps (tInd ind ui) l ≤[pb] mkApps (tInd ind' ui') l' -> [× ind = ind', - R_global_instance Σ (eq_universe Σ) (compare_universe pb Σ) (IndRef ind) #|l| ui ui', + cumul_Ind_univ Σ pb ind #|l| ui ui', is_closed_context Γ & ws_cumul_pb_terms Σ Γ l l']. Proof using wfΣ. move/ws_cumul_pb_red=> [v [v' [redv redv' leqvv']]]. @@ -1405,13 +1398,13 @@ Section Inversions. ∑ u' args', [× Σ ;;; Γ ⊢ T ⇝ mkApps (tConst cst u') args', All2 (fun args args' => Σ ;;; Γ ⊢ args ≤[Conv] args') args args' & - PCUICEquality.R_universe_instance (eq_universe Σ) u u']. + PCUICEquality.cmp_universe_instance (eq_universe Σ) u u']. Proof using wfΣ. intros hdecl hb H. eapply ws_cumul_pb_red in H as (v & v' & [tv tv' eqp]). epose proof (invert_red_axiom_app hdecl hb tv) as [args' [-> ?]]. eapply compare_term_mkApps_l_inv in eqp as [t' [l' []]]; subst v'. - depelim e. + depelim c. exists u', l'. split => //. eapply closed_red_open_right in tv'. rewrite on_free_vars_mkApps /= in tv'. solve_all. @@ -1427,13 +1420,13 @@ Section Inversions. ∑ u' args', [× Σ ;;; Γ ⊢ T ⇝ mkApps (tConst cst u') args', All2 (fun args args' => Σ ;;; Γ ⊢ args ≤[Conv] args') args' args & - PCUICEquality.R_universe_instance (eq_universe Σ) u' u]. + PCUICEquality.cmp_universe_instance (eq_universe Σ) u' u]. Proof using wfΣ. intros hdecl hb H. eapply ws_cumul_pb_red in H as (v & v' & [tv tv' eqp]). epose proof (invert_red_axiom_app hdecl hb tv') as [args' [-> ?]]. eapply compare_term_mkApps_r_inv in eqp as [t' [l' []]]; subst v. - depelim e. + depelim c. exists u0, l'. split => //. eapply closed_red_open_right in tv. rewrite on_free_vars_mkApps /= in tv. solve_all. @@ -1620,13 +1613,13 @@ Qed. Definition ws_cumul_pb_predicate {cf} Σ Γ p p' := [× ws_cumul_pb_terms Σ Γ p.(pparams) p'.(pparams), - R_universe_instance (eq_universe Σ) (puinst p) (puinst p'), - eq_context_gen eq eq (pcontext p) (pcontext p') & + cmp_universe_instance (eq_universe Σ) (puinst p) (puinst p'), + eq_context_upto_names (pcontext p) (pcontext p') & Σ ;;; Γ ,,, inst_case_predicate_context p ⊢ preturn p = preturn p']. Definition ws_cumul_pb_brs {cf} Σ Γ p := All2 (fun br br' => - eq_context_gen eq eq (bcontext br) (bcontext br') × + eq_context_upto_names (bcontext br) (bcontext br') × Σ ;;; Γ ,,, inst_case_branch_context p br ⊢ bbody br = bbody br'). Section Inversions. @@ -1712,7 +1705,7 @@ Section Inversions. Lemma invert_cumul_ind_ind {Γ ind ind' u u' args args'} : Σ ;;; Γ ⊢ mkApps (tInd ind u) args ≤ mkApps (tInd ind' u') args' -> - (eqb ind ind' * PCUICEquality.R_global_instance Σ (eq_universe Σ) (leq_universe Σ) (IndRef ind) #|args| u u' * + (eqb ind ind' * cumul_Ind_univ Σ Cumul ind #|args| u u' * ws_cumul_pb_terms Σ Γ args args'). Proof using wfΣ. intros ht; eapply ws_cumul_pb_Ind_l_inv in ht as (? & ? & [? ? ?]); auto. @@ -1929,7 +1922,7 @@ Section ConvRedConv. Qed. #[global] - Instance all_eq_term_refl : Reflexive (All2 (eq_term_upto_univ Σ.1 (eq_universe Σ) (eq_universe Σ))). + Instance all_eq_term_refl : Reflexive (All2 (eq_term_upto_univ Σ.1 (compare_universe Σ) (compare_sort Σ) Conv)). Proof using Type. intros x. apply All2_same. intros. reflexivity. Qed. @@ -1948,7 +1941,7 @@ Section ConvRedConv. ws_cumul_ctx_pb false Σ (Γ ,,, Δ) (Γ' ,,, Δ') -> #|Γ| = #|Γ'| -> ∑ Δ1 Δ1', red_ctx_rel Σ Γ Δ Δ1 × red_ctx_rel Σ Γ' Δ' Δ1' × - eq_context_upto Σ (eq_universe Σ) (eq_universe Σ) Δ1 Δ1'. + eq_context_upto Σ (compare_universe Σ) (compare_sort Σ) Conv Δ1 Δ1'. Proof. intros. pose proof (closed_con) @@ -2082,8 +2075,8 @@ Section ConvRedConv. Proof using Type. intros x; eapply All2_refl; reflexivity. Qed. Instance eqbrs_refl : Reflexive (All2 (fun x y : branch term => - eq_context_gen eq eq (bcontext x) (bcontext y) * - eq_term_upto_univ Σ.1 (eq_universe Σ) (eq_universe Σ) (bbody x) (bbody y))). + eq_context_upto_names (bcontext x) (bcontext y) * + eq_term_upto_univ Σ.1 (compare_universe Σ) (compare_sort Σ) Conv (bbody x) (bbody y))). Proof using Type. intros brs; eapply All2_refl; split; reflexivity. Qed. Lemma ws_cumul_pb_Case_p {Γ ci c brs p p'} : @@ -2192,27 +2185,27 @@ Section ConvRedConv. rewrite [is_open_term _ _]is_open_case_split onp onbrs /= andb_true_r //. Qed. - Lemma eq_context_gen_subst_context : + Lemma eq_context_upto_names_subst_context : forall u v n k, - eq_context_gen eq eq u v -> - eq_context_gen eq eq (subst_context n k u) (subst_context n k v). + eq_context_upto_names u v -> + eq_context_upto_names (subst_context n k u) (subst_context n k v). Proof using Type. - intros re u v n. + intros u v n k. induction 1. - constructor. - rewrite !subst_context_snoc; constructor; eauto. - depelim p; constructor; simpl; intuition auto; subst; - rewrite -(length_of X); auto. + rewrite -(All2_length X). + destruct r; constructor; assumption. Qed. - Lemma eq_context_gen_inst_case_context {Γ Δ Δ' pars puinst} : - eq_context_gen eq eq Δ Δ' -> - eq_context_gen eq eq (Γ ,,, inst_case_context pars puinst Δ) (Γ ,,, inst_case_context pars puinst Δ'). + Lemma eq_context_upto_names_inst_case_context {Γ Δ Δ' pars puinst} : + eq_context_upto_names Δ Δ' -> + eq_context_upto_names (Γ ,,, inst_case_context pars puinst Δ) (Γ ,,, inst_case_context pars puinst Δ'). Proof using Type. intros. - apply All2_fold_app; [reflexivity|]. + apply All2_app; [|reflexivity]. rewrite /inst_case_context. - now eapply eq_context_gen_subst_context, eq_context_gen_eq_univ_subst_preserved. + now eapply eq_context_upto_names_subst_context, eq_context_upto_names_univ_subst_preserved. Qed. Lemma ws_cumul_pb_Case_one_brs {Γ indn p c brs brs'} : @@ -2222,7 +2215,7 @@ Section ConvRedConv. is_open_brs Γ p brs -> is_open_brs Γ p brs' -> OnOne2 (fun u v => - eq_context_gen eq eq u.(bcontext) v.(bcontext) × + eq_context_upto_names u.(bcontext) v.(bcontext) × Σ ;;; (Γ ,,, inst_case_branch_context p u) ⊢ u.(bbody) = v.(bbody)) brs brs' -> Σ ;;; Γ ⊢ tCase indn p c brs = tCase indn p c brs'. Proof using wfΣ. @@ -2238,7 +2231,7 @@ Section ConvRedConv. + rewrite [is_open_term _ _]is_open_case_split onp onc /=. move: onbrs' i1. rewrite !forallb_app => /andP[] -> /= /andP[] => /andP[] => -> /= _ ->. - now rewrite andb_true_r shiftnP_add app_length inst_case_branch_context_length /= -(All2_fold_length a). + now rewrite andb_true_r shiftnP_add app_length inst_case_branch_context_length /= -(All2_length a). + constructor; try reflexivity. eapply All2_app; try reflexivity. constructor; try split; try reflexivity; cbn => //. @@ -2255,13 +2248,13 @@ Section ConvRedConv. 2:{ eapply IHh => //. move: onbrs' i2. rewrite !forallb_app => /andP[] -> /= /andP[] => /andP[] => -> /= _ ->. - now rewrite andb_true_r shiftnP_add app_length inst_case_branch_context_length /= (All2_fold_length a). } + now rewrite andb_true_r shiftnP_add app_length inst_case_branch_context_length /= (All2_length a). } eapply into_closed_red; eauto => //. { constructor. constructor. eapply OnOne2_app. constructor; auto. cbn. split; auto. eapply red1_eq_context_upto_names; tea. rewrite /inst_case_branch_context /=. - now eapply eq_context_upto_names_gen, eq_context_gen_inst_case_context. } + now eapply eq_context_upto_names_inst_case_context. } rewrite [is_open_term _ _]is_open_case_split onp onc /= //. Qed. @@ -2269,7 +2262,7 @@ Section ConvRedConv. Lemma is_open_brs_OnOne2 Γ p x y : is_open_brs Γ p x -> OnOne2 (fun u v : branch term => - eq_context_gen eq eq (bcontext u) (bcontext v) * + eq_context_upto_names (bcontext u) (bcontext v) * (Σ ;;; Γ,,, inst_case_branch_context p u ⊢ bbody u = bbody v)) y x -> is_open_brs Γ p y. Proof using wfΣ. @@ -2283,7 +2276,6 @@ Section ConvRedConv. now rewrite shiftnP_add. } rewrite !test_context_k_closed_on_free_vars_ctx in cl *. eapply eq_context_upto_names_on_free_vars; tea. - eapply eq_context_upto_names_gen. now symmetry. - now move: op => /= /andP[] => ->. Qed. @@ -2304,12 +2296,11 @@ Section ConvRedConv. split; [eapply ws_cumul_pb_is_open_term_left|eapply ws_cumul_pb_is_open_term_right]; tea. instantiate (1:=bbody x). instantiate (1 := Conv). eapply ws_cumul_pb_eq_context_upto; tea. - { eapply eq_context_gen_eq_context_upto; tc. - now eapply eq_context_gen_inst_case_context. } + { eapply eq_context_upto_names_eq_context_upto; tc. + now eapply eq_context_upto_names_inst_case_context. } eapply ws_cumul_pb_is_closed_context in cv. eapply eq_context_upto_names_on_free_vars; tea. - eapply eq_context_upto_names_gen. - now eapply eq_context_gen_inst_case_context. } + now eapply eq_context_upto_names_inst_case_context. } induction h. - apply ws_cumul_pb_compare; tas. 1-2:rewrite [is_open_term _ _]is_open_case_split onp onc /= //. @@ -2352,10 +2343,7 @@ Section ConvRedConv. (if b then tFix else tCoFix) mfix idx. Lemma eq_term_fix_or_cofix b mfix idx mfix' : - All2 (fun x y : def term => - ((eq_term_upto_univ Σ.1 (eq_universe Σ) (eq_universe Σ) (dtype x) (dtype y) - × eq_term_upto_univ Σ.1 (eq_universe Σ) (eq_universe Σ) (dbody x) (dbody y)) × rarg x = rarg y) * - eq_binder_annot (dname x) (dname y)) mfix mfix' -> + eq_mfixpoint (eq_term Σ Σ) mfix mfix' -> eq_term Σ Σ (fix_or_cofix b mfix idx) (fix_or_cofix b mfix' idx). Proof using Type. destruct b; constructor; auto. @@ -2453,8 +2441,8 @@ Section ConvRedConv. OnOne2 (fun u v => dtype u = dtype v × Σ ;;; Γ ,,, fix_context mfix ⊢ dbody u = dbody v × - (rarg u = rarg v) * - (eq_binder_annot (dname u) (dname v)) + rarg u = rarg v × + eq_binder_annot (dname u) (dname v) ) mfix mfix' -> Σ ;;; Γ ⊢ fix_or_cofix b mfix idx = fix_or_cofix b mfix' idx. Proof using wfΣ. @@ -2468,7 +2456,7 @@ Section ConvRedConv. cbn. eapply All2_app; [eapply All2_refl; reflexivity|]. constructor; cbn; [|constructor]. - constructor; auto. f_equal. auto. + rewrite -hty. constructor; auto. + cbn. eapply All2_app; eauto. constructor; auto. constructor; auto. } @@ -2546,21 +2534,20 @@ Section ConvRedConv. assert (thm : Σ ;;; Γ ⊢ fix_or_cofix b mfix idx = fix_or_cofix b mfix' idx × #|mfix| = #|mfix'| × - eq_context_upto Σ eq eq (Γ ,,, fix_context mfix) (Γ ,,, fix_context mfix') + eq_context_upto_names (Γ ,,, fix_context mfix) (Γ ,,, fix_context mfix') ). { eapply (All2_many_OnOne2_pres _ (fun x => True)) in h. 2:intuition. induction h. - - split; try reflexivity. - + constructor => //; rewrite ?is_open_fix_or_cofix //. - cbn; reflexivity. - + split; reflexivity. + - repeat split; try reflexivity. + constructor => //; rewrite ?is_open_fix_or_cofix //. + cbn; reflexivity. - destruct r as [hl [r _]]. assert (is_open_mfix Γ y). { eapply (is_open_fix_onone2) in r; intuition auto. now rewrite fix_context_length -(OnOne2_length r) -(rtrans_clos_length h). } destruct (IHh H) as [? []]. - split. + repeat split. 2: lia. + etransitivity. * eassumption. * apply ws_cumul_pb_fix_one_body; tea; eauto with fvs. @@ -2568,12 +2555,10 @@ Section ConvRedConv. intros [na ty bo ra] [na' ty' bo' ra'] [? [hh ?]]. simpl in *. intuition eauto. eapply ws_cumul_pb_eq_context_upto. 3: eassumption. - 1:eapply eq_context_impl. 4: eassumption. - all:tc. + 1: apply eq_context_upto_names_eq_context_upto; tas; tc. rewrite on_free_vars_ctx_app onΓ /=. apply on_free_vars_fix_context. solve_all. - + split; [lia|]. - etransitivity. + + etransitivity. * eassumption. * apply OnOne2_split in r as [[na ty bo ra] [[na' ty' bo' ra'] [l1 [l2 [[? [? [? ?]]] [? ?]]]]]]. @@ -2581,15 +2566,12 @@ Section ConvRedConv. rewrite 2!fix_context_fix_context_alt. rewrite 2!map_app. simpl. unfold def_sig at 2 5. simpl. - eapply eq_context_upto_cat. - -- eapply eq_context_upto_refl; auto. - -- eapply eq_context_upto_rev'. - rewrite 2!mapi_app. cbn. - eapply eq_context_upto_cat. - ++ constructor; tas; revgoals. - ** constructor; tas. eapply eq_term_upto_univ_refl. all: auto. - ** eapply eq_context_upto_refl; auto. - ++ eapply eq_context_upto_refl; auto. + eapply All2_app; try reflexivity. + eapply All2_rev. + rewrite 2!mapi_app. cbn. + eapply All2_app; try reflexivity. + constructor; try reflexivity. + constructor. assumption. } apply thm. Qed. @@ -3179,14 +3161,12 @@ Section ConvSubst. Lemma eq_context_upto_ws_cumul_ctx_pb {pb Γ Γ'} : is_closed_context Γ -> is_closed_context Γ' -> - eq_context_upto Σ (eq_universe Σ) (compare_universe pb Σ) Γ Γ' -> + compare_context Σ pb Γ Γ' -> ws_cumul_ctx_pb pb Σ Γ Γ'. Proof using wfΣ. intros cl cl' eq. apply into_ws_cumul_ctx_pb; auto. - destruct pb; revgoals. - { now eapply eq_context_upto_univ_cumul_context. } - { now eapply eq_context_upto_univ_conv_context. } + now apply compare_context_cumul_pb_context. Qed. Lemma substitution_ws_cumul_ctx_pb {Γ Δ Δ' Γ' s s'} : @@ -3394,7 +3374,7 @@ Qed. Lemma subst_instance_ws_cumul_ctx_pb {cf:checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} {Δ u u'} pb : wf_local Σ (subst_instance u Δ) -> - R_universe_instance (eq_universe (global_ext_constraints Σ)) u u' -> + cmp_universe_instance (eq_universe (global_ext_constraints Σ)) u u' -> ws_cumul_ctx_pb pb Σ (subst_instance u Δ) (subst_instance u' Δ). Proof. move=> wf equ. @@ -3402,10 +3382,10 @@ Proof. induction Δ as [|d Δ] in wf |- *. - constructor. - simpl. depelim wf. + all: destruct l as (Hb & s & Hs & _); cbn in Hb, Hs. * cbn; constructor. + apply IHΔ => //. + destruct d as [na [b|] ty]; constructor; cbn in *; auto; try congruence. - destruct l as [s Hs]. constructor. 1:eauto with fvs. { now eapply subject_closed in Hs; rewrite is_open_term_closed in Hs. } { erewrite on_free_vars_subst_instance. @@ -3415,15 +3395,13 @@ Proof. * cbn; constructor. + apply IHΔ => //. + destruct d as [na [b'|] ty]; constructor; cbn in *; auto; try congruence; noconf H. - { destruct l as [s Hs]. - constructor. 1:eauto with fvs. - { now eapply subject_closed in l0; rewrite is_open_term_closed in l0. } + { constructor. 1:eauto with fvs. + { now eapply subject_closed in Hb; rewrite is_open_term_closed in Hb. } { erewrite on_free_vars_subst_instance. - eapply subject_closed in l0; rewrite is_open_term_closed in l0. - now rewrite on_free_vars_subst_instance in l0. } + eapply subject_closed in Hb; rewrite is_open_term_closed in Hb. + now rewrite on_free_vars_subst_instance in Hb. } apply eq_term_upto_univ_subst_instance; try typeclasses eauto. auto. } - { destruct l as [s Hs]. - constructor. 1:eauto with fvs. + { constructor. 1:eauto with fvs. { now eapply subject_closed in Hs; rewrite is_open_term_closed in Hs. } { erewrite on_free_vars_subst_instance. eapply subject_closed in Hs; rewrite is_open_term_closed in Hs. @@ -3447,7 +3425,7 @@ Hint Rewrite is_closed_context_subst_instance is_open_term_subst_instance : fvs. Lemma eq_term_compare_term {cf} pb Σ t u : eq_term Σ Σ t u -> - compare_term pb Σ Σ t u. + compare_term Σ Σ pb t u. Proof. destruct pb; cbn; auto. now apply eq_term_leq_term. @@ -3455,7 +3433,7 @@ Qed. Lemma subst_instance_ws_cumul_ctx_pb_rel {cf:checker_flags} {Σ} {wfΣ : wf Σ} {Γ Δ u u' pb} : is_closed_context (Γ ,,, Δ) -> - R_universe_instance (eq_universe Σ) u u' -> + cmp_universe_instance (eq_universe Σ) u u' -> ws_cumul_ctx_pb_rel pb Σ Γ (subst_instance u Δ) (subst_instance u' Δ). Proof. move=> equ. @@ -3502,7 +3480,7 @@ Proof. Qed. Lemma eq_term_inds {cf:checker_flags} (Σ : global_env_ext) u u' ind mdecl : - R_universe_instance (eq_universe (global_ext_constraints Σ)) u u' -> + cmp_universe_instance (eq_universe (global_ext_constraints Σ)) u u' -> All2 (eq_term Σ Σ) (inds (inductive_mind ind) u (ind_bodies mdecl)) (inds (inductive_mind ind) u' (ind_bodies mdecl)). Proof. @@ -3511,13 +3489,11 @@ Proof. induction n; constructor; auto. clear IHn. repeat constructor. destruct ind; simpl in *. - eapply (R_global_instance_empty_impl _ _ _ _ _ _ 0). - 4:{ unfold R_global_instance, R_global_instance_gen. simpl. eauto. } - all:typeclasses eauto. + apply cmp_instance_opt_variance; simpl; assumption. Qed. Lemma conv_inds {cf:checker_flags} (Σ : global_env_ext) Γ u u' ind mdecl : - R_universe_instance (eq_universe (global_ext_constraints Σ)) u u' -> + cmp_universe_instance (eq_universe (global_ext_constraints Σ)) u u' -> is_closed_context Γ -> ws_cumul_pb_terms Σ Γ (inds (inductive_mind ind) u (ind_bodies mdecl)) (inds (inductive_mind ind) u' (ind_bodies mdecl)). @@ -3527,29 +3503,20 @@ Proof. induction n; constructor; auto. clear IHn. repeat constructor; auto. destruct ind; simpl in *. - eapply (R_global_instance_empty_impl _ _ _ _ _ _ 0). - 4:{ unfold R_global_instance, R_global_instance_gen. simpl. eauto. } - all:typeclasses eauto. + apply cmp_instance_opt_variance; simpl; assumption. Qed. -Lemma R_global_instance_length Σ Req Rle ref napp i i' : - R_global_instance Σ Req Rle ref napp i i' -> #|i| = #|i'|. +Lemma cmp_global_instance_length Σ Req Rle ref napp i i' : + cmp_global_instance Σ Req Rle ref napp i i' -> #|i| = #|i'|. Proof. - unfold R_global_instance, R_global_instance_gen. - destruct global_variance_gen. - { induction i in l, i' |- *; destruct l, i'; simpl; auto; try lia; try easy. - * specialize (IHi i' []). simpl in IHi. intuition auto with all. - * intros []. intuition eauto with all. - } - { unfold R_universe_instance. - intros H % Forall2_length. now rewrite !map_length in H. } + apply cmp_opt_variance_length. Qed. -Lemma R_universe_instance_variance_irrelevant Re Rle i i' : +Lemma cmp_universe_instance_variance_irrelevant cmp_universe pb i i' : #|i| = #|i'| -> - R_universe_instance_variance Re Rle [] i i'. + cmp_opt_variance cmp_universe pb AllIrrelevant i i'. Proof. - now induction i in i' |- *; destruct i'; simpl; auto. + auto. Qed. Lemma ws_cumul_pb_it_mkProd_or_LetIn {cf pb Σ} {wfΣ : wf Σ} (Δ Γ Γ' : context) (B B' : term) : @@ -3880,7 +3847,7 @@ Section MoreCongruenceLemmas. Defined. Lemma ws_cumul_pb_Ind {pb Γ ind ui ui' l l'} : - [× R_global_instance Σ (eq_universe Σ) (compare_universe pb Σ) (IndRef ind) #|l| ui ui', + [× cumul_Ind_univ Σ pb ind #|l| ui ui', is_closed_context Γ & ws_cumul_pb_terms Σ Γ l l'] -> Σ ;;; Γ ⊢ mkApps (tInd ind ui) l ≤[pb] mkApps (tInd ind ui') l'. intros [Rglob Rclosed Hll']. apply ws_cumul_pb_terms_alt in Hll'. destruct Hll' as [l0 [l0' [Hl0 Hl0' Hl0l0']]]. @@ -3906,7 +3873,7 @@ Section MoreCongruenceLemmas. Defined. Lemma ws_cumul_pb_Construct {pb Γ i k ui ui' l l'} : - [× R_global_instance Σ (eq_universe Σ) (compare_universe pb Σ) (ConstructRef i k) #|l| ui ui', + [× cumul_Construct_univ Σ pb i k #|l| ui ui', is_closed_context Γ & ws_cumul_pb_terms Σ Γ l l'] -> Σ ;;; Γ ⊢ mkApps (tConstruct i k ui) l ≤[pb] mkApps (tConstruct i k ui') l'. intros [Rglob Rclosed Hll']. apply ws_cumul_pb_terms_alt in Hll'. destruct Hll' as [l0 [l0' [Hl0 Hl0' Hl0l0']]]. @@ -4031,10 +3998,10 @@ Proof. unfold inst_case_predicate_context. apply PCUICOnFreeVarsConv.on_free_vars_ctx_inst_case_context ; eauto. ++ rewrite shiftnP_add in Hreturn. rewrite <- inst_case_predicate_context_length in Hreturn. rewrite <- app_length in Hreturn. eassumption. - ++ rewrite shiftnP_add in Hreturn'. rewrite <- (All2_fold_length Hpcon) in Hreturn'. + ++ rewrite shiftnP_add in Hreturn'. rewrite <- (All2_length Hpcon) in Hreturn'. rewrite <- inst_case_predicate_context_length in Hreturn'. rewrite <- app_length in Hreturn'. eassumption. - * unfold ws_cumul_pb_brs. + * unfold cumul_branches, cumul_branch, ws_cumul_pb_brs in *. repeat toAll. eapply All2_impl. 1: tea. cbn; intros; destruct_head'_prod. split; eauto. rewrite -> test_context_k_closed_on_free_vars_ctx in *. repeat match goal with H : _ |- _ => progress toProp H end; destruct_head'_and. @@ -4042,7 +4009,7 @@ Proof. + apply PCUICOnFreeVarsConv.on_free_vars_ctx_inst_case_context ; eauto; repeat toAll; eauto. + let H := multimatch goal with H : _ |- _ => H end in erewrite -> shiftnP_add, <- inst_case_branch_context_length, <- app_length in H; exact H. - + rewrite -> shiftnP_add in *. rewrite <- (All2_fold_length ltac:(eassumption)) in *. erewrite <- inst_case_branch_context_length in *. + + rewrite -> shiftnP_add in *. rewrite <- (All2_length ltac:(eassumption)) in *. erewrite <- inst_case_branch_context_length in *. rewrite <- app_length in *. tea. - intros; eapply ws_cumul_pb_Proj_c; eauto. - intros Γ mfix mfix' idx Hmfixmfix' Hmfixmfix'_dep HΓ H H'. cbn in *. diff --git a/pcuic/theories/PCUICCumulProp.v b/pcuic/theories/PCUICCumulProp.v index d012e7512..71bd00c95 100644 --- a/pcuic/theories/PCUICCumulProp.v +++ b/pcuic/theories/PCUICCumulProp.v @@ -24,14 +24,13 @@ Section no_prop_leq_type. Context `{cf : checker_flags}. Variable Hcf : prop_sub_type = false. -Variable Hcf' : check_univs. Lemma cumul_sort_confluence {Σ} {wfΣ : wf Σ} {Γ A u v} : Σ ;;; Γ ⊢ A ≤ tSort u -> Σ ;;; Γ ⊢ A ≤ tSort v -> ∑ v', (Σ ;;; Γ ⊢ A = tSort v') * - (leq_universe (global_ext_constraints Σ) v' u /\ - leq_universe (global_ext_constraints Σ) v' v). + (leq_sort (global_ext_constraints Σ) v' u /\ + leq_sort (global_ext_constraints Σ) v' v). Proof using Type. move=> H H'. eapply ws_cumul_pb_Sort_r_inv in H as [u'u ?]. @@ -50,8 +49,8 @@ Lemma cumul_ind_confluence {Σ : global_env_ext} {wfΣ : wf Σ} {Γ A ind u v l [× Σ ;;; Γ ⊢ A ⇝ (mkApps (tInd ind v') l''), ws_cumul_pb_terms Σ Γ l l'', ws_cumul_pb_terms Σ Γ l' l'', - R_global_instance Σ (eq_universe Σ) (leq_universe Σ) (IndRef ind) #|l| v' u & - R_global_instance Σ (eq_universe Σ) (leq_universe Σ) (IndRef ind) #|l'| v' v]. + cmp_ind_universes Σ ind #|l| v' u & + cmp_ind_universes Σ ind #|l'| v' v]. Proof using Type. move=> H H'. eapply ws_cumul_pb_Ind_r_inv in H as [u'u [l'u [redl ru ?]]]. @@ -84,33 +83,45 @@ Proof using Type. now eapply wf_local_closed_context. Qed. +Lemma is_propositional_bottom {Σ Γ T s s'} : + wf_ext Σ -> + Σ ;;; Γ ⊢ T ≤ tSort s -> + Σ ;;; Γ ⊢ T ≤ tSort s' -> + Sort.is_propositional s -> Sort.is_propositional s'. +Proof using Hcf. + intros wfΣ hs hs'. + destruct (cumul_sort_confluence hs hs') as [x' [conv [leq leq']]]. + move: leq leq'. + destruct x', s, s' => //=. + now rewrite Hcf. +Qed. + Lemma is_prop_bottom {Σ Γ T s s'} : wf_ext Σ -> Σ ;;; Γ ⊢ T ≤ tSort s -> Σ ;;; Γ ⊢ T ≤ tSort s' -> - Universe.is_prop s -> Universe.is_prop s'. -Proof using Hcf Hcf'. + Sort.is_prop s -> Sort.is_prop s'. +Proof using Hcf. intros wfΣ hs hs'. destruct (cumul_sort_confluence hs hs') as [x' [conv [leq leq']]]. - intros isp. - eapply leq_universe_prop_r in leq; eauto. - unshelve eapply (leq_universe_prop_no_prop_sub_type _ _ _ _ _ _ leq'); eauto. + move: leq leq'. + destruct x', s, s' => //=. + now rewrite Hcf. Qed. Lemma is_sprop_bottom {Σ Γ T s s'} : wf_ext Σ -> Σ ;;; Γ ⊢ T ≤ tSort s -> Σ ;;; Γ ⊢ T ≤ tSort s' -> - Universe.is_sprop s -> Universe.is_sprop s'. -Proof using Hcf'. + Sort.is_sprop s -> Sort.is_sprop s'. +Proof using Type. intros wfΣ hs hs'. destruct (cumul_sort_confluence hs hs') as [x' [conv [leq leq']]]. - intros isp. - eapply leq_universe_sprop_r in leq; eauto. - unshelve eapply (leq_universe_sprop_l _ _ _ _ _ leq'); eauto. + move: leq leq'. + destruct x', s, s' => //=. Qed. -Lemma prop_sort_eq {Σ Γ u u'} : Universe.is_prop u -> Universe.is_prop u' -> +Lemma prop_sort_eq {Σ Γ u u'} : Sort.is_prop u -> Sort.is_prop u' -> is_closed_context Γ -> Σ ;;; Γ ⊢ tSort u = tSort u'. Proof using Type. @@ -120,7 +131,7 @@ Proof using Type. red. red. constructor. Qed. -Lemma sprop_sort_eq {Σ Γ u u'} : Universe.is_sprop u -> Universe.is_sprop u' -> +Lemma sprop_sort_eq {Σ Γ u u'} : Sort.is_sprop u -> Sort.is_sprop u' -> is_closed_context Γ -> Σ ;;; Γ ⊢ tSort u = tSort u'. Proof using Type. @@ -132,7 +143,7 @@ Qed. Lemma conv_sort_inv {Σ : global_env_ext} {wfΣ : wf Σ} Γ s s' : Σ ;;; Γ ⊢ tSort s = tSort s' -> - eq_universe (global_ext_constraints Σ) s s'. + eq_sort (global_ext_constraints Σ) s s'. Proof using Type. intros H. eapply ws_cumul_pb_alt_closed in H as [v [v' [redv redv' eqvv']]]. @@ -141,36 +152,37 @@ Proof using Type. now depelim eqvv'. Qed. -Lemma is_prop_superE {Σ l} : wf_ext Σ -> Universe.is_prop (Universe.super l) -> False. -Proof using Hcf'. - intros wfΣ. - eapply is_prop_gt; eauto. - eapply leq_universe_refl. +Lemma is_prop_superE {l} : Sort.is_prop (Sort.super l) -> False. +Proof using Type. + destruct l => //. Qed. -Lemma is_sprop_superE {Σ l} : wf_ext Σ -> Universe.is_sprop (Universe.super l) -> False. +Lemma is_sprop_superE {l} : Sort.is_sprop (Sort.super l) -> False. Proof using Type. - intros wfΣ. destruct l => //. + destruct l => //. Qed. -Lemma is_prop_prod {s s'} : Universe.is_prop s' -> Universe.is_prop (Universe.sort_of_product s s'). +Lemma is_prop_prod {s s'} : Sort.is_prop s' -> Sort.is_prop (Sort.sort_of_product s s'). Proof using Type. - intros isp. - unfold Universe.sort_of_product. rewrite isp. auto. + destruct s' => //. Qed. -Lemma is_sprop_prod {s s'} : Universe.is_sprop s' -> Universe.is_sprop (Universe.sort_of_product s s'). +Lemma is_sprop_prod {s s'} : Sort.is_sprop s' -> Sort.is_sprop (Sort.sort_of_product s s'). Proof using Type. - intros isp. - unfold Universe.sort_of_product. rewrite isp orb_true_r. auto. + destruct s' => //. Qed. -Definition eq_univ_prop (u v : Universe.t) := - (Universe.is_prop u <-> Universe.is_prop v) /\ - (Universe.is_sprop u <-> Universe.is_sprop v). +Definition eq_univ_prop (u v : sort) := + match Sort.to_family u, Sort.to_family v with + | Sort.fSProp, Sort.fSProp => true + | Sort.fProp, Sort.fProp => true + | Sort.fType, Sort.fType => true + | Sort.fType, Sort.fProp | Sort.fProp, Sort.fType => prop_sub_type + | _, _ => false + end. Definition eq_term_prop (Σ : global_env) napp := - PCUICEquality.eq_term_upto_univ_napp Σ eq_univ_prop eq_univ_prop napp. + PCUICEquality.eq_term_upto_univ_napp Σ (fun _ _ _ => True) (fun _ => eq_univ_prop) Conv napp. Reserved Notation " Σ ;;; Γ |- t ~~ u " (at level 50, Γ, t, u at next level). @@ -195,103 +207,39 @@ Inductive cumul_prop `{checker_flags} (Σ : global_env_ext) (Γ : context) : ter where " Σ ;;; Γ |- t ~~ u " := (cumul_prop Σ Γ t u) : type_scope. -Lemma eq_term_prop_impl Σ Re Rle t u : - wf_ext Σ -> - forall n, - PCUICEquality.eq_term_upto_univ_napp Σ.1 Re Rle n t u -> - subrelation Re eq_univ_prop -> - subrelation Rle eq_univ_prop -> - eq_term_prop Σ n t u. +Lemma eq_term_prop_impl Σ cmp_universe cmp_sort pb napp t u : + RelationClasses.subrelation (cmp_sort Conv) eq_univ_prop -> + RelationClasses.subrelation (cmp_sort pb) eq_univ_prop -> + PCUICEquality.eq_term_upto_univ_napp Σ.1 cmp_universe cmp_sort pb napp t u -> + eq_term_prop Σ napp t u. Proof using Type. - intros wfΣ n eq. - intros. - eapply PCUICEquality.eq_term_upto_univ_impl in eq. eauto. + intros hsub_conv hsub_pb. + eapply PCUICEquality.eq_term_upto_univ_impl. all:auto. + all: now intros ??. Qed. -Lemma leq_universe_prop_spec Σ u1 u2 : - check_univs -> - wf_ext Σ -> - leq_universe Σ u1 u2 -> - match u1, u2 with - | Universe.lProp, Universe.lProp => True - | Universe.lSProp, Universe.lSProp => True - | Universe.lProp, Universe.lSProp => False - | Universe.lSProp, Universe.lProp => False - | Universe.lProp, Universe.lType _ => prop_sub_type - | Universe.lSProp, Universe.lType _ => False - | Universe.lType l, Universe.lType l' => True - | Universe.lType _, _ => False - end. +#[global] +Instance subrelation_compare_sort_eq_prop φ pb : + RelationClasses.subrelation (compare_sort φ pb) eq_univ_prop. Proof using Type. - intros cu wf leq. - apply wf_ext_consistent in wf. - apply (leq_universe_props _ _ _ cu wf leq). -Qed. - -Lemma subrelation_eq_universe_eq_prop Σ : - wf_ext Σ -> - subrelation (eq_universe Σ) eq_univ_prop. -Proof using Hcf Hcf'. - intros wfΣ x y eq'. red. - split; intros. - eapply eq_universe_leq_universe in eq'. - eapply leq_universe_prop_spec in eq'; auto. - destruct x, y; simpl in *; auto; cong. - eapply eq_universe_leq_universe in eq'. - eapply leq_universe_prop_spec in eq'; auto. - destruct x, y; simpl in *; auto; cong. + intros s s'. + destruct pb, s, s' => //=. Qed. -Lemma subrelation_leq_universe_eq_prop Σ : - wf_ext Σ -> - subrelation (leq_universe Σ) eq_univ_prop. -Proof using Hcf Hcf'. - intros wfΣ x y eq'. red. - split; intros; - eapply leq_universe_prop_spec in eq'; auto; - destruct x, y; simpl in *; auto; cong. -Qed. - -Lemma eq_term_eq_term_prop_impl Σ t u : - wf_ext Σ -> - forall n, - PCUICEquality.eq_term_upto_univ_napp Σ.1 (eq_universe Σ) (eq_universe Σ) n t u -> - eq_term_prop Σ n t u. -Proof using Hcf Hcf'. - intros wfΣ n eq. eapply eq_term_prop_impl; eauto. - now apply subrelation_eq_universe_eq_prop. - now apply subrelation_eq_universe_eq_prop. -Qed. - -Lemma leq_term_eq_term_prop_impl Σ t u : - wf_ext Σ -> - forall n, - PCUICEquality.eq_term_upto_univ_napp Σ.1 (eq_universe Σ) (leq_universe Σ) n t u -> - eq_term_prop Σ n t u. -Proof using Hcf Hcf'. - intros wfΣ n eq. eapply eq_term_prop_impl; eauto. - now apply subrelation_eq_universe_eq_prop. - now apply subrelation_leq_universe_eq_prop. -Qed. - -Lemma cumul_cumul_prop Σ Γ A B : - wf_ext Σ -> - Σ ;;; Γ ⊢ A ≤ B -> - Σ ;;; Γ |- A ~~ B. -Proof using Hcf Hcf'. - intros wfΣ. induction 1. - - constructor => //. now apply leq_term_eq_term_prop_impl in c. - - econstructor 2; eauto. - - econstructor 3; eauto. +Lemma eq_term_eq_term_prop_impl Σ φ pb napp t u : + compare_term_napp Σ φ pb napp t u -> + eq_term_prop Σ napp t u. +Proof using Type. + eapply eq_term_prop_impl. + all: apply subrelation_compare_sort_eq_prop. Qed. -Lemma conv_cumul_prop Σ Γ A B : - wf_ext Σ -> - Σ ;;; Γ ⊢ A = B -> +Lemma cumul_pb_cumul_prop Σ Γ pb A B : + Σ ;;; Γ ⊢ A ≤[pb] B -> Σ ;;; Γ |- A ~~ B. -Proof using Hcf Hcf'. - intros wfΣ. induction 1. +Proof using Type. + induction 1. - constructor => //. now apply eq_term_eq_term_prop_impl in c. - econstructor 2; eauto. - econstructor 3; eauto. @@ -327,46 +275,81 @@ Proof using Type. econstructor 2; eauto. Qed. -Lemma cumul_prop_props {Σ Γ u u'} {wfΣ : wf Σ}: - Universe.is_prop u -> - Σ ;;; Γ |- tSort u ~~ tSort u' -> - Universe.is_prop u'. -Proof using Type. - intros isp equiv. +Lemma cumul_prop_sort {Σ Γ s s'} {wfΣ : wf Σ} : + Σ ;;; Γ |- tSort s ~~ tSort s' -> + eq_univ_prop s s'. +Proof. + intros equiv. eapply cumul_prop_alt in equiv as [nf [nf' [redl redr eq]]]. eapply invert_red_sort in redl. apply invert_red_sort in redr. subst. - depelim eq. red in e. intuition auto. + now depelim eq. +Qed. + +Lemma cumul_prop_props {Σ Γ u u'} {wfΣ : wf Σ}: + Sort.is_prop u -> + Σ ;;; Γ |- tSort u ~~ tSort u' -> + Sort.is_prop u'. +Proof using Hcf. + move => isp /cumul_prop_sort. move: isp. + destruct u, u' => //=. cbn. rewrite Hcf //. Qed. Lemma cumul_sprop_props {Σ Γ u u'} {wfΣ : wf Σ} : - Universe.is_sprop u -> + Sort.is_sprop u -> Σ ;;; Γ |- tSort u ~~ tSort u' -> - Universe.is_sprop u'. + Sort.is_sprop u'. Proof using Type. - intros isp equiv. - eapply cumul_prop_alt in equiv as [nf [nf' [redl redr eq]]]. - eapply invert_red_sort in redl. apply invert_red_sort in redr. - subst. - depelim eq. red in e. intuition auto. + move => isp /cumul_prop_sort. move: isp. + destruct u, u' => //. Qed. Instance refl_eq_univ_prop : RelationClasses.Reflexive eq_univ_prop. Proof using Type. - intros x. red. intuition. + now intros []. Qed. Instance sym_eq_univ_prop : RelationClasses.Symmetric eq_univ_prop. Proof using Type. - intros x y; unfold eq_univ_prop; intuition. + now intros [] []. Qed. Instance trans_eq_univ_prop : RelationClasses.Transitive eq_univ_prop. Proof using Type. - intros x y; unfold eq_univ_prop; intuition. + now intros [] [] []. Qed. -Lemma LevelExprSet_For_all (P : LevelExpr.t -> Prop) (u : LevelAlgExpr.t) : +#[global] +Instance equiv_eq_univ_prop : RelationClasses.Equivalence eq_univ_prop. +Proof using Type. + split; exact _. +Qed. + +#[global] +Instance equiv_True {T} : @RelationClasses.Equivalence T (fun _ _ => True). +Proof using Type. + split; auto. +Qed. + +#[global] +Instance subrel_R_True {T} R : @RelationClasses.subrelation T R (fun _ _ => True). +Proof. + now intros ??. +Qed. + +#[global] +Instance substu_f_True {T} `{UnivSubst T} f : SubstUnivPreserving f (fun _ _ => True). +Proof using Type. + now intros ???. +Qed. + +#[global] +Instance substu_True_eq_univ_prop : SubstUnivPreserving (fun _ _ => True) eq_univ_prop. +Proof using Type. + now intros []???. +Qed. + +Lemma LevelExprSet_For_all (P : LevelExpr.t -> Prop) (u : Universe.t) : LevelExprSet.For_all P u <-> Forall P (LevelExprSet.elements u). Proof using Type. @@ -429,11 +412,6 @@ Proof using Type. apply univ_exprs_map_all. Qed. -Lemma univ_is_prop_make x : Universe.is_prop (Universe.make x) = false. -Proof using Type. - destruct x; simpl; auto. -Qed. - (* Lemma is_prop_subst_level_expr u1 u2 s : Forall2 (fun x y : Level.t => eq_univ_prop (Universe.make x) (Universe.make y)) u1 u2 -> LevelExpr.is_prop (subst_instance_level_expr u1 s) = LevelExpr.is_prop (subst_instance_level_expr u2 s). @@ -450,15 +428,6 @@ Proof. now rewrite hu. Qed. *) -Instance substuniv_eq_univ_prop : SubstUnivPreserving eq_univ_prop. -Proof using Type. - intros s u1 u2 hu. - red in hu. - eapply Forall2_map_inv in hu. - rewrite /subst_instance_univ. - destruct s; red; simpl; auto; try intuition reflexivity. -Qed. - Lemma cumul_prop_sym Σ Γ T U : wf Σ.1 -> Σ ;;; Γ |- T ~~ U -> @@ -495,45 +464,24 @@ Qed. Global Instance cumul_prop_transitive Σ Γ : wf Σ -> CRelationClasses.Transitive (cumul_prop Σ Γ). Proof using Type. intros. red. intros. now eapply cumul_prop_trans. Qed. -Lemma cumul_prop_cum_l {Σ Γ A T B} {wfΣ : wf_ext Σ} : +Lemma cumul_prop_cumul_pb_l {Σ Γ pb A T B} {wfΣ : wf_ext Σ} : Σ ;;; Γ |- A ~~ T -> - Σ ;;; Γ ⊢ A ≤ B -> + Σ ;;; Γ ⊢ A ≤[pb] B -> Σ ;;; Γ |- B ~~ T. -Proof using Hcf Hcf'. - intros HT cum. - eapply cumul_cumul_prop in cum; auto. - eapply CRelationClasses.transitivity ; eauto. - eapply cumul_prop_sym; eauto. -Qed. - -Lemma cumul_prop_cum_r {Σ Γ A T B} {wfΣ : wf_ext Σ} : - Σ ;;; Γ |- A ~~ T -> - Σ ;;; Γ ⊢ B ≤ A -> - Σ ;;; Γ |- B ~~ T. -Proof using Hcf Hcf'. - intros HT cum. - eapply cumul_cumul_prop in cum; auto. - eapply CRelationClasses.transitivity ; eauto. -Qed. - -Lemma cumul_prop_conv_l {Σ Γ A T B} {wfΣ : wf_ext Σ} : - Σ ;;; Γ |- A ~~ T -> - Σ ;;; Γ ⊢ A = B -> - Σ ;;; Γ |- B ~~ T. -Proof using Hcf Hcf'. +Proof using Type. intros HT cum. - eapply conv_cumul_prop in cum; auto. - eapply CRelationClasses.transitivity ; eauto. + eapply cumul_pb_cumul_prop in cum; auto. + etransitivity ; eauto. eapply cumul_prop_sym; eauto. Qed. -Lemma cumul_prop_conv_r {Σ Γ A T B} {wfΣ : wf_ext Σ} : +Lemma cumul_prop_cum_pb_r {Σ Γ pb A T B} {wfΣ : wf_ext Σ} : Σ ;;; Γ |- A ~~ T -> - Σ ;;; Γ ⊢ B = A -> + Σ ;;; Γ ⊢ B ≤[pb] A -> Σ ;;; Γ |- B ~~ T. -Proof using Hcf Hcf'. +Proof using Type. intros HT cum. - eapply conv_cumul_prop in cum; auto. + eapply cumul_pb_cumul_prop in cum; auto. eapply CRelationClasses.transitivity ; eauto. Qed. @@ -713,9 +661,9 @@ Qed. Lemma red_conv_prop {Σ Γ T U} {wfΣ : wf_ext Σ} : Σ ;;; Γ ⊢ T ⇝ U -> Σ ;;; Γ |- T ~~ U. -Proof using Hcf Hcf'. +Proof using Type. move/(red_ws_cumul_pb (pb:=Conv)). - now apply conv_cumul_prop. + now apply cumul_pb_cumul_prop. Qed. Lemma substitution_red_terms_conv_prop {Σ Γ Δ Γ' s s' M} {wfΣ : wf_ext Σ} : @@ -724,7 +672,7 @@ Lemma substitution_red_terms_conv_prop {Σ Γ Δ Γ' s s' M} {wfΣ : wf_ext Σ} untyped_subslet Γ s Δ -> red_terms Σ Γ s s' -> Σ ;;; (Γ ,,, subst_context s 0 Γ') |- (subst s #|Γ'| M) ~~ (subst s' #|Γ'| M). -Proof using Hcf Hcf'. +Proof using Type. intros. apply red_conv_prop. eapply closed_red_red_subst; tea. @@ -734,7 +682,7 @@ Lemma context_conversion_cumul_prop {Σ Γ Δ M N} {wfΣ : wf_ext Σ} : Σ ;;; Γ |- M ~~ N -> Σ ⊢ Γ = Δ -> Σ ;;; Δ |- M ~~ N. -Proof using Hcf Hcf'. +Proof using Type. induction 1; intros. - constructor => //. eauto with fvs. now rewrite -(All2_fold_length X). now rewrite -(All2_fold_length X). @@ -745,7 +693,7 @@ Proof using Hcf Hcf'. eapply conv_red_conv in X1. 2:exact X0. 3:{ eapply ws_cumul_pb_refl. fvs. now rewrite (All2_fold_length X0). } 2:{ eapply closed_red_refl. fvs. now rewrite (All2_fold_length X0). } - symmetry in X1. now eapply conv_cumul_prop. + symmetry in X1. now eapply cumul_pb_cumul_prop. - specialize (IHX X0). transitivity v => //. eapply red1_red in r. assert (Σ ;;; Γ ⊢ u ⇝ v) by (now apply into_closed_red). @@ -753,7 +701,7 @@ Proof using Hcf Hcf'. eapply conv_red_conv in X1. 2:exact X0. 3:{ eapply ws_cumul_pb_refl. fvs. now rewrite (All2_fold_length X0). } 2:{ eapply closed_red_refl. fvs. now rewrite (All2_fold_length X0). } - symmetry in X1. now eapply conv_cumul_prop. + symmetry in X1. now eapply cumul_pb_cumul_prop. Qed. (** Note: a more general version involving substitution in an extended context Γ ,,, Δ would be @@ -768,7 +716,7 @@ Lemma substitution_untyped_cumul_prop_cumul {Σ Γ Δ Δ' s s' M} {wfΣ : wf_ext untyped_subslet Γ s' Δ' -> All2 (cumul_prop Σ Γ) s s' -> Σ ;;; Γ |- subst0 s M ~~ subst0 s' M. -Proof using Hcf Hcf'. +Proof using Type. intros clctx clctx' clM subs subs' Heq. assert (lens' := All2_length Heq). destruct (cumul_prop_args Heq) as (nf & nf' & [redl redr eq]) => //. @@ -800,50 +748,15 @@ Proof using Type. repeat constructor. Qed. -Lemma is_prop_subst_instance_level u l - : Universe.is_prop (Universe.make (subst_instance_level u l)) = Universe.is_prop (Universe.make l). -Proof using Type. - destruct l; simpl; auto. -Qed. - -Lemma R_opt_variance_impl Re Rle v x y : - subrelation Re Rle -> - R_universe_instance Re x y -> - R_opt_variance Re Rle v x y. -Proof using Type. - intros sub. - destruct v; simpl; auto. - intros H. eapply Forall2_map_inv in H. - induction H in l |- *; simpl; auto. - destruct l. auto. - split. destruct t; simpl; auto. - eauto. -Qed. - -Lemma cumul_prop_subst_instance_level Σ univs u u' (l : Level.t) : +Lemma cmp_True_subst_instance Σ univs u u' (i : Instance.t) : wf Σ.1 -> consistent_instance_ext Σ univs u -> consistent_instance_ext Σ univs u' -> - eq_univ_prop (subst_instance u (Universe.make l)) - (subst_instance u' (Universe.make l)). + cmp_universe_instance (fun _ _ => True) (subst_instance u i) (subst_instance u' i). Proof using Type. intros wfΣ cu cu'. red. - unfold subst_instance; cbn. intuition. -Qed. - -Lemma cumul_prop_subst_instance_instance Σ univs u u' (i : Instance.t) : - wf Σ.1 -> - consistent_instance_ext Σ univs u -> - consistent_instance_ext Σ univs u' -> - R_universe_instance eq_univ_prop (subst_instance u i) - (subst_instance u' i). -Proof using Type. - intros wfΣ cu cu'. red. - eapply All2_Forall2, All2_map. - unfold subst_instance. - eapply All2_map. eapply All2_refl. - intros x. red. - rewrite !is_prop_subst_instance_level /=. split; reflexivity. + eapply All2_Forall2, All2_map, All2_refl => ui. + unfold on_rel. auto. Qed. Lemma cumul_prop_subst_instance {Σ Γ univs u u' T} {wfΣ : wf Σ} : @@ -866,41 +779,41 @@ Proof using Type. - constructor. eapply PCUICEquality.eq_term_upto_univ_impl in IHT1; eauto. all:try typeclasses eauto. apply IHT2. - - constructor. now eapply cumul_prop_subst_instance_instance. - - constructor. red. apply R_opt_variance_impl. intros x y; auto. - now eapply cumul_prop_subst_instance_instance. - - constructor. red. apply R_opt_variance_impl. intros x y; auto. - now eapply cumul_prop_subst_instance_instance. + - constructor. now eapply cmp_True_subst_instance. + - constructor. now eapply cmp_instance_opt_variance, cmp_True_subst_instance. + - constructor. now eapply cmp_instance_opt_variance, cmp_True_subst_instance. - cbn. constructor. splits; simpl; solve_all. - eapply cumul_prop_subst_instance_instance; tea. reflexivity. + eapply cmp_True_subst_instance; tea. reflexivity. apply IHT. eapply All2_map. - eapply All_All2; tea. cbn. + eapply All_All2; tea. cbn. unfold eq_branch. intuition auto. rewrite /id. reflexivity. + - constructor. unfold eq_mfixpoint. + eapply All2_map. eapply All_All2; tea. + cbn. intros d []. + intuition eauto. + - constructor. unfold eq_mfixpoint. + eapply All2_map. eapply All_All2; tea. + cbn. intros d []. + intuition eauto. - cbn. constructor; splits; simpl. destruct p as [? []]; constructor; cbn in *; intuition eauto. - unshelve eapply cumul_prop_subst_instance_level; tea. exact (array_level a). solve_all. Qed. -Lemma R_eq_univ_prop_consistent_instances Σ univs u u' : +Lemma cmp_True_instance Σ univs u u' : wf Σ.1 -> consistent_instance_ext Σ univs u -> consistent_instance_ext Σ univs u' -> - R_universe_instance eq_univ_prop u u'. + cmp_universe_instance (fun _ _ => True) u u'. Proof using Type. intros wfΣ cu cu'. destruct univs; simpl in *. - - destruct u, u' => /=; [|done..]. red. - simpl. constructor. - - intuition. - eapply Forall2_map. - eapply All2_Forall2. - solve_all. - eapply All2_impl. - eapply All_All_All2; eauto. lia. - simpl; intros. - intuition auto with *. + - destruct u, u' => //=. + - apply Forall2_triv. + destruct cu as (_ & ? & _). + destruct cu' as (_ & ? & _). + congruence. Qed. Lemma untyped_subslet_inds Γ ind u u' mdecl : @@ -927,7 +840,7 @@ Lemma cumul_prop_tProd {Σ : global_env_ext} {Γ na t ty na' t' ty'} {wfΣ : wf_ eq_term Σ.1 Σ t t' -> Σ ;;; Γ ,, vass na t |- ty ~~ ty' -> Σ ;;; Γ |- tProd na t ty ~~ tProd na' t' ty'. -Proof using Hcf Hcf'. +Proof using Type. intros eqann eq cum. eapply cumul_prop_alt in cum as (nf & nf' & [redl redr eq']). eapply cumul_prop_alt. eexists (tProd na t nf), (tProd na' t' nf'); split; eauto. @@ -942,7 +855,7 @@ Proof using Hcf Hcf'. now rewrite on_free_vars_ctx_snoc /= onΓ. repeat (constructor; auto). - repeat (constructor; auto). - eapply eq_term_eq_term_prop_impl; auto. + eapply eq_term_eq_term_prop_impl; eauto. Qed. Lemma cumul_prop_tLetIn (Σ : global_env_ext) {Γ na t d ty na' t' d' ty'} {wfΣ : wf_ext Σ} : @@ -951,11 +864,11 @@ Lemma cumul_prop_tLetIn (Σ : global_env_ext) {Γ na t d ty na' t' d' ty'} {wfΣ eq_term Σ.1 Σ d d' -> Σ ;;; Γ ,, vdef na d t |- ty ~~ ty' -> Σ ;;; Γ |- tLetIn na d t ty ~~ tLetIn na' d' t' ty'. -Proof using Hcf Hcf'. +Proof using Type. intros eqann eq eq' cum. eapply cumul_prop_alt in cum as (nf & nf' & [redl redr eq'']). eapply cumul_prop_alt. - assert(eq_context_upto Σ (eq_universe Σ) (eq_universe Σ) (Γ ,, vdef na d t) (Γ ,, vdef na' d' t')). + assert(eq_context Σ Σ (Γ ,, vdef na d t) (Γ ,, vdef na' d' t')). { repeat constructor; pcuic. eapply eq_context_upto_refl; typeclasses eauto. } eapply (closed_red_eq_context_upto_l (pb:=Conv)) in redr; eauto. 2:{ eapply clrel_ctx in redl. rewrite !on_free_vars_ctx_snoc in redl |- *. @@ -968,9 +881,9 @@ Proof using Hcf Hcf'. eexists (tLetIn na d t nf), (tLetIn na' d' t' v'); split. - now eapply closed_red_letin_body. - now eapply closed_red_letin_body. - - constructor; eauto using eq_term_eq_term_prop_impl. - apply eq_term_eq_term_prop_impl; auto. - apply eq_term_eq_term_prop_impl; auto. + - constructor; eauto. + eapply eq_term_eq_term_prop_impl; eauto. + eapply eq_term_eq_term_prop_impl; eauto. transitivity nf'. auto. now eapply eq_term_eq_term_prop_impl. Qed. @@ -981,7 +894,7 @@ Lemma cumul_prop_mkApps {Σ Γ f args f' args'} {wfΣ : wf_ext Σ} : eq_term Σ.1 Σ f f' -> All2 (cumul_prop Σ Γ) args args' -> Σ ;;; Γ |- mkApps f args ~~ mkApps f' args'. -Proof using Hcf Hcf'. +Proof using Type. intros clΓ clf clf' eq eq'. eapply cumul_prop_alt. eapply cumul_prop_args in eq' as (nf & nf' & [redl redr eq']). @@ -990,8 +903,9 @@ Proof using Hcf Hcf'. - eapply closed_red_mkApps; auto. - eapply eq_term_upto_univ_mkApps. eapply eq_term_upto_univ_impl. - 5:eapply eq. all:auto. 4:lia. - all:now eapply subrelation_eq_universe_eq_prop. + 6:eapply eq. 5: lia. + all: tc. + auto. Qed. Hint Resolve closed_red_open_right : fvs. @@ -1046,17 +960,10 @@ Proof using Type. Qed. Notation eq_term_napp Σ n x y := - (eq_term_upto_univ_napp Σ (eq_universe Σ) (eq_universe Σ) n x y). + (eq_term_upto_univ_napp Σ (eq_sort Σ) (eq_sort Σ) n x y). Notation leq_term_napp Σ n x y := - (eq_term_upto_univ_napp Σ (eq_universe Σ) (leq_universe Σ) n x y). - -Lemma eq_term_upto_univ_napp_leq {Σ : global_env_ext} {n x y} : - eq_term_napp Σ n x y -> - leq_term_napp Σ n x y. -Proof using Type. - eapply eq_term_upto_univ_impl; auto; typeclasses eauto. -Qed. + (eq_term_upto_univ_napp Σ (eq_sort Σ) (leq_sort Σ) n x y). Lemma cumul_prop_is_open {Σ Γ T U} : Σ ;;; Γ |- T ~~ U -> @@ -1074,53 +981,41 @@ Proof using Type. eapply on_free_vars_ctx_impl. discriminate. Qed. -Lemma eq_context_upto_map2_set_binder_name Σ eqterm leterm eq le pctx pctx' Γ Δ : - eq_context_gen eqterm leterm pctx pctx' -> - eq_context_upto Σ eq le Γ Δ -> - eq_context_upto Σ eq le - (map2 set_binder_name (forget_types pctx) Γ) - (map2 set_binder_name (forget_types pctx') Δ). +(** Well-typed terms in the leq_term relation live in the same sort hierarchy. *) +Lemma typing_leq_term_prop_gen : + env_prop + (fun Σ Γ t T => + forall t' T' : term, + on_udecl Σ.1 Σ.2 -> + Σ ;;; Γ |- t' : T' -> + forall pb n, compare_term_napp Σ Σ pb n t' t -> + Σ ;;; Γ |- T ~~ T')%type + (fun Σ Γ j => on_udecl Σ.1 Σ.2 -> + lift_typing0 (fun t T => + forall t' T' : term, + Σ ;;; Γ |- t' : T' -> + forall pb n, compare_term_napp Σ Σ pb n t' t -> + Σ ;;; Γ |- T ~~ T') j) + (fun Σ Γ => wf_local Σ Γ). Proof using Type. -intros eqp. -induction 1 in pctx, pctx', eqp |- *. -- induction eqp; cbn; constructor. -- depelim eqp. simpl. constructor. - simpl. constructor; auto. - destruct c, p; constructor; auto. -Qed. + eapply typing_ind_env. + { intros ???? H ?. apply lift_typing_impl with (1 := H) => ?? [] ?? ?? //. eauto. } + 1: now auto. -(** Well-typed terms in the leq_term relation live in the same sort hierarchy. *) -Lemma typing_leq_term_prop (Σ : global_env_ext) Γ t t' T T' : - wf Σ.1 -> - Σ ;;; Γ |- t : T -> - on_udecl Σ.1 Σ.2 -> - Σ ;;; Γ |- t' : T' -> - forall n, leq_term_napp Σ n t' t -> - Σ ;;; Γ |- T ~~ T'. -Proof using Hcf Hcf'. - intros wfΣ Ht. - revert Σ wfΣ Γ t T Ht t' T'. - eapply (typing_ind_env - (fun Σ Γ t T => - forall t' T' : term, - on_udecl Σ.1 Σ.2 -> - Σ;;; Γ |- t' : T' -> - forall n, leq_term_napp Σ n t' t -> - Σ ;;; Γ |- T ~~ T')%type - (fun Σ Γ => wf_local Σ Γ)); auto;intros Σ wfΣ Γ wfΓ; intros. + all: intros Σ wfΣ Γ wfΓ; intros. 1-13:match goal with - [ H : leq_term_napp _ _ _ _ |- _ ] => depelim H + [ H : compare_term_napp _ _ _ _ _ _ |- _ ] => depelim H end; assert (wf_ext Σ) by (split; assumption). - 15:{ assert (wf_ext Σ) by (split; assumption). specialize (X1 _ _ H X5 _ X6). - eapply cumul_prop_cum_l; tea. + 15:{ assert (wf_ext Σ) by (split; assumption). specialize (X1 _ _ H X5 _ _ X6). + eapply cumul_prop_cumul_pb_l; tea. eapply cumulSpec_cumulAlgo_curry in X4; tea; fvs. } 6:{ eapply inversion_App in X6 as (na' & A' & B' & hf & ha & cum); auto. - specialize (X3 _ _ H hf _ X7_1). - specialize (X5 _ _ H ha _ (eq_term_upto_univ_napp_leq X7_2)). - eapply cumul_cumul_prop in cum; auto. + specialize (X3 _ _ H hf _ _ X7_1). + specialize (X5 _ _ H ha _ _ X7_2). + eapply cumul_pb_cumul_prop in cum; auto. transitivity (B' {0 := u0}) => //. eapply cumul_prop_prod_inv in X3 => //. transitivity (B' {0 := u}). @@ -1132,63 +1027,56 @@ Proof using Hcf Hcf'. rewrite shiftnP_add. now eapply cumul_prop_is_open in X3 as []. eapply eq_term_eq_term_prop_impl => //. eapply PCUICEquality.eq_term_upto_univ_substs. - all:try typeclasses eauto. - eapply PCUICEquality.eq_term_upto_univ_refl. all:try typeclasses eauto. + all: try reflexivity. constructor. 2:constructor. now symmetry. } - eapply inversion_Rel in X0 as [decl' [wfΓ' [Hnth Hcum]]]; auto. - rewrite Hnth in H; noconf H. now eapply cumul_cumul_prop in Hcum. + rewrite Hnth in H; noconf H. now eapply cumul_pb_cumul_prop in Hcum. - eapply inversion_Sort in X0 as [wf [wfs Hs]]; auto. - apply subrelation_leq_universe_eq_prop in x => //. - apply cumul_cumul_prop in Hs => //. + apply subrelation_compare_sort_eq_prop in c => //. + apply cumul_pb_cumul_prop in Hs => //. eapply cumul_prop_trans; eauto. destruct (cumul_prop_is_open Hs) as []. - constructor => //. constructor. symmetry. - split; split; intros H'. 1,2:now eapply is_prop_superE in H'. - 1,2:now eapply is_sprop_superE in H'. + constructor => //. constructor. rewrite /eq_univ_prop !Sort.sType_super //. - eapply inversion_Prod in X4 as [s1' [s2' [Ha [Hb Hs]]]]; auto. - specialize (X1 _ _ H Ha). - specialize (X1 _ (eq_term_upto_univ_napp_leq X5_1)). + specialize (X1 H). apply unlift_TypUniv in X1, Ha. + specialize (X1 _ _ Ha _ _ X5_1). eapply context_conversion in Hb. 3:{ constructor. apply conv_ctx_refl. constructor. eassumption. constructor. eauto. } - 2:{ constructor; eauto. now exists s1. } - specialize (X3 _ _ H Hb _ X5_2). - eapply cumul_cumul_prop in Hs => //. + 2:{ pcuic. } + specialize (X3 _ _ H Hb _ _ X5_2). + eapply cumul_pb_cumul_prop in Hs => //. eapply cumul_prop_trans; eauto. constructor; fvs. constructor. - split. - * split; intros Hs'; apply is_prop_sort_prod in Hs'; eapply is_prop_prod; eapply cumul_prop_props; eauto. - now eapply cumul_prop_sym; eauto. - * split; intros Hs'; apply is_sprop_sort_prod in Hs'; eapply is_sprop_prod; eapply cumul_sprop_props; eauto. - now eapply cumul_prop_sym; eauto. - - - eapply inversion_Lambda in X4 as (s & B & dom & bod & cum). - specialize (X1 _ _ H dom _ (eq_term_upto_univ_napp_leq X5_1)). + apply cumul_prop_sort in X1, X3. move: X1 X3. clear. + unfold eq_univ_prop. + destruct s2, s2' => //=. + all: destruct s1, s1' => //=. + + - eapply inversion_Lambda in X4 as (B & dom & bod & cum); tas. specialize (X3 t0 B H). - assert(conv_context cumulAlgo_gen Σ (Γ ,, vass na ty) (Γ ,, vass n t)). + assert(conv_context cumulAlgo_gen Σ (Γ ,, vass na0 ty) (Γ ,, vass na t)). { repeat constructor; pcuic. } forward X3 by eapply context_conversion; eauto; pcuic. - specialize (X3 _ X5_2). eapply cumul_cumul_prop in cum; eauto. + specialize (X3 _ _ X5_2). eapply cumul_pb_cumul_prop in cum; eauto. eapply cumul_prop_trans; eauto. - eapply cumul_prop_tProd; eauto. now symmetry. now symmetry. auto. + eapply cumul_prop_tProd; eauto. now symmetry. now symmetry. - - eapply inversion_LetIn in X6 as (s1' & A & dom & bod & codom & cum); auto. - specialize (X1 _ _ H dom _ (eq_term_upto_univ_napp_leq X7_2)). - specialize (X3 _ _ H bod _ (eq_term_upto_univ_napp_leq X7_1)). - assert(conv_context cumulAlgo_gen Σ (Γ ,, vdef na t ty) (Γ ,, vdef n b b_ty)). + - eapply inversion_LetIn in X4 as (A & dombod & codom & cum); auto. + assert(conv_context cumulAlgo_gen Σ (Γ ,, vdef na0 t ty) (Γ ,, vdef na b b_ty)). { repeat constructor; pcuic. } - specialize (X5 u A H). - forward X5 by eapply context_conversion; eauto; pcuic. - specialize (X5 _ X7_3). - eapply cumul_cumul_prop in cum; eauto. + specialize (X3 u A H). + forward X3 by eapply context_conversion; eauto; pcuic. + specialize (X3 _ _ X5_3). + eapply cumul_pb_cumul_prop in cum; eauto. eapply cumul_prop_trans; eauto. eapply cumul_prop_tLetIn; auto; now symmetry. - eapply inversion_Const in X1 as [decl' [wf [declc [cu cum]]]]; auto. - eapply cumul_cumul_prop in cum; eauto. + eapply cumul_pb_cumul_prop in cum; eauto. eapply cumul_prop_trans; eauto. unshelve eapply declared_constant_to_gen in H, declc; eauto. pose proof (declared_constant_inj _ _ H declc); subst decl'. @@ -1199,7 +1087,7 @@ Proof using Hcf Hcf'. - eapply inversion_Ind in X1 as [decl' [idecl' [wf [declc [cu cum]]]]]; auto. unshelve eapply declared_inductive_to_gen in isdecl, declc; eauto. pose proof (declared_inductive_inj isdecl declc) as [-> ->]. - eapply cumul_cumul_prop in cum; eauto. + eapply cumul_pb_cumul_prop in cum; eauto. eapply cumul_prop_trans; eauto. eapply cumul_prop_subst_instance; tea. fvs. destruct (cumul_prop_is_open cum) as []. @@ -1209,7 +1097,7 @@ Proof using Hcf Hcf'. unshelve eapply declared_constructor_to_gen in declc; eauto. unshelve epose proof (isdecl' := declared_constructor_to_gen isdecl); eauto. pose proof (declared_constructor_inj isdecl' declc) as [-> [-> ->]]. - eapply cumul_cumul_prop in cum; eauto. + eapply cumul_pb_cumul_prop in cum; eauto. eapply cumul_prop_trans; eauto. unfold type_of_constructor. have clars : is_closed_context (arities_context (ind_bodies mdecl))@[u]. @@ -1230,10 +1118,9 @@ Proof using Hcf Hcf'. len. len in clty. * len. * len. - * generalize (ind_bodies mdecl). - induction l; simpl; constructor; auto. - constructor. simpl. eapply R_opt_variance_impl. now intros x. - eapply R_eq_univ_prop_consistent_instances; eauto. + * induction (ind_bodies mdecl) in |- *; simpl; constructor; auto. + constructor. simpl. eapply cmp_instance_opt_variance. + eapply cmp_True_instance; eauto. * simpl. eapply (@substitution_untyped_cumul_prop _ Γ (subst_instance u0 (arities_context mdecl.(ind_bodies))) []) => //. eapply on_free_vars_terms_inds. @@ -1243,13 +1130,13 @@ Proof using Hcf Hcf'. len in clty; len. - eapply inversion_Case in X9 as (mdecl' & idecl' & isdecl' & indices' & data & cum); auto. - eapply cumul_cumul_prop in cum; eauto. + eapply cumul_pb_cumul_prop in cum; eauto. eapply cumul_prop_trans; eauto. simpl. clear X8. unshelve eapply declared_inductive_to_gen in isdecl, isdecl'; eauto. destruct (declared_inductive_inj isdecl isdecl'). subst. destruct data. - specialize (X7 _ _ H5 scrut_ty _ (eq_term_upto_univ_napp_leq X10)). + specialize (X7 _ _ H5 scrut_ty _ _ X10). eapply cumul_prop_sym => //. destruct e as [eqpars [eqinst [eqpctx eqpret]]]. rewrite /ptm. @@ -1262,24 +1149,23 @@ Proof using Hcf Hcf'. rewrite /predctx. rewrite /case_predicate_context /case_predicate_context_gen. rewrite /pre_case_predicate_context_gen /inst_case_context. - eapply eq_context_upto_map2_set_binder_name; tea. - eapply eq_context_upto_subst_context; tea. tc. + eapply eq_context_upto_names_map2_set_binder_name; tea. + eapply eq_context_upto_subst_context; tea. 1,2: tc. eapply eq_context_upto_univ_subst_instance; try tc. tas. now eapply All2_rev. } - eapply All2_app. 2:(repeat constructor; auto using eq_term_eq_term_prop_impl). + eapply All2_app. 2:(repeat constructor; eauto using eq_term_eq_term_prop_impl). eapply cumul_prop_mkApps_Ind_inv in X7 => //. eapply All2_app_inv_l in X7 as (?&?&?&?&?). eapply All2_symP => //. typeclasses eauto. eapply app_inj in e as [eql ->] => //. move: (All2_length eqpars). - move: (All2_length a0). lia. fvs. now eapply subject_is_open_term in scrut_ty. + move: (All2_length a). lia. fvs. now eapply subject_is_open_term in scrut_ty. now apply subject_is_open_term in X6. - eapply inversion_Proj in X3 as (u' & mdecl' & idecl' & cdecl' & pdecl' & args' & inv); auto. intuition auto. - specialize (X2 _ _ H0 a0 _ (eq_term_upto_univ_napp_leq X4)). - eapply eq_term_upto_univ_napp_leq in X4. - eapply cumul_cumul_prop in b; eauto. + specialize (X2 _ _ H0 a0 _ _ X4). + eapply cumul_pb_cumul_prop in b; eauto. eapply cumul_prop_trans; eauto. eapply cumul_prop_mkApps_Ind_inv in X2 => //. unshelve epose proof (a' := declared_projection_to_gen a); eauto. @@ -1304,7 +1190,7 @@ Proof using Hcf Hcf'. * epose proof (projection_subslet Σ _ _ _ _ _ _ _ _ _ a wfΣ a0 (validity a0)). now eapply subslet_untyped_subslet. * constructor => //. symmetry; constructor => //; fvs. - { now eapply leq_term_eq_term_prop_impl. } + { now eapply eq_term_eq_term_prop_impl. } { now eapply All2_rev. } * eapply (@substitution_cumul_prop Σ Γ (projection_context p.(proj_ind) mdecl' idecl' u') []) => //. { apply (projection_subslet Σ _ _ _ _ _ _ _ _ _ a wfΣ a0 (validity a0)). } @@ -1315,37 +1201,61 @@ Proof using Hcf Hcf'. rewrite (declared_minductive_ind_npars a) in H1. rewrite closedn_on_free_vars //. eapply closed_upwards; tea. lia. - - eapply inversion_Fix in X2 as (decl' & fixguard' & Hnth & types' & bodies & wffix & cum); auto. - eapply cumul_cumul_prop in cum; eauto. + - eapply inversion_Fix in X4 as (decl' & fixguard' & Hnth & types' & bodies & wffix & cum); auto. + eapply cumul_pb_cumul_prop in cum; eauto. eapply cumul_prop_trans; eauto. - eapply All2_nth_error in a; eauto. - destruct a as [[[a _] _] _]. + eapply All2_nth_error in e; eauto. + destruct e as (a & _). constructor; [fvs|..]. - { eapply nth_error_all in X0 as [? [dty ?]]; tea. - now apply subject_is_open_term in dty. } + { eapply nth_error_all in X0; tea. + now apply isType_is_open_term in X0. } { now eapply cumul_prop_is_open in cum as []. } eapply eq_term_eq_term_prop_impl; eauto. now symmetry in a. - - eapply inversion_CoFix in X2 as (decl' & fixguard' & Hnth & types' & bodies & wfcofix & cum); auto. - eapply cumul_cumul_prop in cum; eauto. + - eapply inversion_CoFix in X4 as (decl' & fixguard' & Hnth & types' & bodies & wfcofix & cum); auto. + eapply cumul_pb_cumul_prop in cum; eauto. eapply cumul_prop_trans; eauto. - eapply All2_nth_error in a; eauto. - destruct a as [[[a _] _] _]. + eapply All2_nth_error in e; eauto. + destruct e as (a & _). constructor; [fvs|..]. - { eapply nth_error_all in X0 as [? [dty ?]]; tea. - now apply subject_is_open_term in dty. } + { eapply nth_error_all in X0; tea. + now apply isType_is_open_term in X0. } { now eapply cumul_prop_is_open in cum as []. } eapply eq_term_eq_term_prop_impl; eauto. now symmetry in a. - depelim X3. eapply inversion_Prim in X2 as [prim_ty' [cdecl' []]]; tea. depelim o. - 1-2:rewrite H in e; noconf e; eapply cumul_cumul_prop; eauto; pcuic. + 1-2:rewrite H in e; noconf e; eapply cumul_pb_cumul_prop; eauto; pcuic. depelim X1. - cbn in H, e2. rewrite H in e2. noconf e2. eapply cumul_cumul_prop; eauto; pcuic. + cbn in H, e2. rewrite H in e2. noconf e2. eapply cumul_pb_cumul_prop; eauto; pcuic. move: w; simp prim_type. intro. etransitivity; tea. constructor; fvs. cbn. depelim X0. fvs. eapply eq_term_leq_term. symmetry; repeat constructor; eauto. Qed. +Lemma typing_leq_term_prop (Σ : global_env_ext) Γ t t' T T' : + wf Σ.1 -> + Σ ;;; Γ |- t : T -> + on_udecl Σ.1 Σ.2 -> + Σ ;;; Γ |- t' : T' -> + forall pb n, compare_term_napp Σ Σ pb n t' t -> + Σ ;;; Γ |- T ~~ T'. +Proof using Type. + intros. + now eapply (env_prop_typing typing_leq_term_prop_gen). +Qed. + +Lemma typing_cumul_term_prop {Σ : global_env_ext} {wfΣ : wf_ext Σ} pb Γ t t' T T' : + Σ ;;; Γ |- t : T -> + Σ ;;; Γ |- t' : T' -> + Σ ;;; Γ |- t' <=[pb] t -> + Σ ;;; Γ |- T ~~ T'. +Proof using Type. + intros. + apply cumul_alt in X1 as (v & v' & r & r' & leq). + eapply typing_leq_term_prop. 5: eassumption. 1,3: apply wfΣ. + all: eapply subject_reduction; [ auto | | eassumption]; assumption. +Qed. + End no_prop_leq_type. diff --git a/pcuic/theories/PCUICCumulativity.v b/pcuic/theories/PCUICCumulativity.v index c6b5bcbb7..318be8f28 100644 --- a/pcuic/theories/PCUICCumulativity.v +++ b/pcuic/theories/PCUICCumulativity.v @@ -14,9 +14,7 @@ Set Default Goal Selector "!". Reserved Notation " Σ ;;; Γ |- t <=[ pb ] u" (at level 50, Γ, t, u at next level, format "Σ ;;; Γ |- t <=[ pb ] u"). -Definition leq_term_ext `{checker_flags} (Σ : global_env_ext) Rle t u := eq_term_upto_univ Σ (eq_universe Σ) Rle t u. - -Notation " Σ ⊢ t <===[ pb ] u" := (compare_term pb Σ Σ t u) (at level 50, t, u at next level). +Notation " Σ ⊢ t <===[ pb ] u" := (compare_term Σ Σ pb t u) (at level 50, t, u at next level). (** ** Cumulativity *) @@ -53,18 +51,17 @@ Instance conv_decls_refl {cf:checker_flags} Σ Γ Γ' : Reflexive (conv_decls cu #[global] Instance cumul_decls_refl {cf:checker_flags} Σ Γ Γ' : Reflexive (cumul_decls cumulAlgo_gen Σ Γ Γ') := _. -Lemma cumul_alt `{cf : checker_flags} Σ Γ t u : - Σ ;;; Γ |- t <= u <~> { v & { v' & (red Σ Γ t v * red Σ Γ u v' * - leq_term_ext Σ (leq_universe Σ) v v')%type } }. +Lemma cumul_alt `{cf : checker_flags} Σ Γ pb t u : + Σ ;;; Γ |- t <=[pb] u <~> ∑ v v', red Σ Γ t v × red Σ Γ u v' × Σ ⊢ v <===[ pb ] v'. Proof. split. - induction 1. + exists t, u. intuition auto. - + destruct IHX as (v' & v'' & (redv & redv') & leqv). + + destruct IHX as (v' & v'' & redv & redv' & leqv). exists v', v''. intuition auto. now eapply red_step. - + destruct IHX as (v' & v'' & (redv & redv') & leqv). + + destruct IHX as (v' & v'' & redv & redv' & leqv). exists v', v''. intuition auto. now eapply red_step. - - intros [v [v' [[redv redv'] Hleq]]]. + - intros (v & v' & redv & redv' & Hleq). apply clos_rt_rt1n in redv. apply clos_rt_rt1n in redv'. induction redv. @@ -83,7 +80,7 @@ Qed. #[global] Instance conv_refl' {cf:checker_flags} Σ Γ : Reflexive (convAlgo Σ Γ). Proof. - intro; constructor. unfold leq_term_ext. reflexivity. + intro; constructor. reflexivity. Qed. Lemma red_cumul `{cf : checker_flags} {Σ : global_env_ext} {Γ t u} : @@ -170,18 +167,18 @@ Qed. Lemma eq_term_eq_term_napp {cf:checker_flags} Σ ϕ napp t t' : eq_term Σ ϕ t t' -> - eq_term_upto_univ_napp Σ (eq_universe ϕ) (eq_universe ϕ) napp t t'. + eq_term_upto_univ_napp Σ (compare_universe ϕ) (compare_sort ϕ) Conv napp t t'. Proof. - intros. eapply eq_term_upto_univ_impl. 5:eauto. - 4:auto with arith. all:typeclasses eauto. + intros. eapply eq_term_upto_univ_impl. 6:eauto. + 5:auto with arith. all:typeclasses eauto. Qed. Lemma leq_term_leq_term_napp {cf:checker_flags} Σ ϕ napp t t' : leq_term Σ ϕ t t' -> - eq_term_upto_univ_napp Σ (eq_universe ϕ) (leq_universe ϕ) napp t t'. + eq_term_upto_univ_napp Σ (compare_universe ϕ) (compare_sort ϕ) Cumul napp t t'. Proof. - intros. eapply eq_term_upto_univ_impl. 5:eauto. - 4:auto with arith. all:typeclasses eauto. + intros. eapply eq_term_upto_univ_impl. 6:eauto. + 5:auto with arith. all:typeclasses eauto. Qed. Lemma eq_term_mkApps `{checker_flags} Σ φ f l f' l' : @@ -249,35 +246,19 @@ Proof. + eauto. Qed. -Lemma conv_alt_red {cf : checker_flags} {Σ : global_env_ext} {Γ : context} {t u : term} : - Σ;;; Γ |- t = u <~> (∑ v v' : term, (red Σ Γ t v × red Σ Γ u v') × - eq_term Σ (global_ext_constraints Σ) v v'). -Proof. - split. - - induction 1. - * exists t, u; intuition auto. - * destruct IHX as [? [? [? ?]]]. - exists x, x0; intuition auto. eapply red_step; eauto. - * destruct IHX as [? [? [? ?]]]. - exists x, x0; intuition auto. eapply red_step; eauto. - - destruct 1 as [? [? [[? ?] ?]]]. - eapply red_conv_conv; eauto. - eapply red_conv_conv_inv; eauto. now constructor. -Qed. - -Definition eq_termp_napp {cf:checker_flags} (pb: conv_pb) (Σ : global_env_ext) napp := - compare_term_napp pb Σ Σ napp. +Definition eq_termp_napp {cf:checker_flags} (Σ : global_env_ext) (pb: conv_pb) napp := + compare_term_napp Σ Σ pb napp. -Notation eq_termp pb Σ := (compare_term pb Σ Σ). +Notation eq_termp Σ pb := (compare_term Σ Σ pb). Lemma eq_term_eq_termp {cf:checker_flags} pb (Σ : global_env_ext) x y : eq_term Σ Σ x y -> - eq_termp pb Σ x y. + eq_termp Σ pb x y. Proof. destruct pb; [easy|]. cbn. apply eq_term_upto_univ_leq; auto. - typeclasses eauto. + all: typeclasses eauto. Qed. Lemma cumul_App_l {cf:checker_flags} : diff --git a/pcuic/theories/PCUICCumulativitySpec.v b/pcuic/theories/PCUICCumulativitySpec.v index 431eb9dc3..10164a8e6 100644 --- a/pcuic/theories/PCUICCumulativitySpec.v +++ b/pcuic/theories/PCUICCumulativitySpec.v @@ -10,38 +10,53 @@ Set Default Goal Selector "!". Implicit Types (cf : checker_flags). -Definition cumul_predicate (cumul : context -> term -> term -> Type) Γ Re p p' := - All2 (cumul Γ) p.(pparams) p'.(pparams) * - (R_universe_instance Re p.(puinst) p'.(puinst) * - ((eq_context_gen eq eq p.(pcontext) p'.(pcontext)) * - cumul (Γ ,,, inst_case_predicate_context p) p.(preturn) p'.(preturn))). - -Definition cumul_predicate_dep {cumul Γ Re p p'} - (H : cumul_predicate cumul Γ Re p p') +Definition cumul_predicate (cumul : context -> term -> term -> Type) cumul_universe Γ p p' := + All2 (cumul Γ) p.(pparams) p'.(pparams) × + cmp_universe_instance cumul_universe p.(puinst) p'.(puinst) × + eq_context_upto_names p.(pcontext) p'.(pcontext) × + cumul (Γ ,,, inst_case_predicate_context p) p.(preturn) p'.(preturn). + +Definition cumul_predicate_dep {cumul cumul_universe Γ p p'} + (H : cumul_predicate cumul cumul_universe Γ p p') (cumul' : forall Γ p p', cumul Γ p p' -> Type) - Re' + cumul_universe' := let '(Hparams, (Huinst, (Heq, Hcumul))) := H in - All2_dep (cumul' Γ) Hparams * - (R_universe_instance_dep Re Re' Huinst - * cumul' _ _ _ Hcumul). + All2_dep (cumul' Γ) Hparams × + cmp_universe_instance_dep cumul_universe cumul_universe' Huinst × + cumul' _ _ _ Hcumul. -Lemma cumul_predicate_undep {cumul Γ Re p p' H cumul' Re'} - : @cumul_predicate cumul' Γ Re' p p' <~> @cumul_predicate_dep cumul Γ Re p p' H (fun Γ p p' _ => cumul' Γ p p') (fun x y _ => Re' x y). +Lemma cumul_predicate_undep {cumul cumul_universe Γ p p' H cumul' cumul_universe'} : + @cumul_predicate cumul' cumul_universe' Γ p p' <~> + @cumul_predicate_dep cumul cumul_universe Γ p p' H (fun Γ p p' _ => cumul' Γ p p') (fun x y _ => on_rel cumul_universe' Universe.make' x y). Proof. - cbv [cumul_predicate cumul_predicate_dep R_universe_instance R_universe_instance_dep] in *. + cbv [cumul_predicate cumul_predicate_dep cmp_universe_instance cmp_universe_instance_dep] in *. split; intro; repeat destruct ?; subst; rdest; try assumption. all: repeat first [ assumption | toAll ]. Qed. +Definition cumul_branch (cumul_term : context -> term -> term -> Type) Γ p br br' := + eq_context_upto_names br.(bcontext) br'.(bcontext) × + cumul_term (Γ ,,, inst_case_branch_context p br) br.(bbody) br'.(bbody). + +Definition cumul_branches cumul_term Γ p brs brs' := All2 (cumul_branch cumul_term Γ p) brs brs'. + +Definition cumul_mfixpoint (cumul_term : context -> term -> term -> Type) Γ mfix mfix' := + All2 (fun d d' => + cumul_term Γ d.(dtype) d'.(dtype) × + cumul_term (Γ ,,, fix_context mfix) d.(dbody) d'.(dbody) × + d.(rarg) = d'.(rarg) × + eq_binder_annot d.(dname) d'.(dname) + ) mfix mfix'. + Reserved Notation " Σ ;;; Γ ⊢ t ≤s[ pb ] u" (at level 50, Γ, t, u at next level, format "Σ ;;; Γ ⊢ t ≤s[ pb ] u"). Definition cumul_Ind_univ {cf} (Σ : global_env_ext) pb i napp := - R_global_instance Σ (eq_universe Σ) (compare_universe pb Σ) (IndRef i) napp. + cmp_global_instance Σ (compare_universe Σ) pb (IndRef i) napp. Definition cumul_Construct_univ {cf} (Σ : global_env_ext) pb i k napp := - R_global_instance Σ (eq_universe Σ) (compare_universe pb Σ) (ConstructRef i k) napp. + cmp_global_instance Σ (compare_universe Σ) pb (ConstructRef i k) napp. (** * Definition of cumulativity and conversion relations *) Local Unset Elimination Schemes. @@ -79,11 +94,11 @@ Inductive cumulSpec0 {cf : checker_flags} (Σ : global_env_ext) Γ (pb : conv_pb Σ ;;; Γ ⊢ mkApps (tConstruct i k u) args ≤s[pb] mkApps (tConstruct i k u') args' | cumul_Sort : forall s s', - compare_universe pb Σ s s' -> + compare_sort Σ pb s s' -> Σ ;;; Γ ⊢ tSort s ≤s[pb] tSort s' | cumul_Const : forall c u u', - R_universe_instance (compare_universe Conv Σ) u u' -> + cmp_universe_instance (compare_universe Σ Conv) u u' -> Σ ;;; Γ ⊢ tConst c u ≤s[pb] tConst c u' (* congruence rules *) @@ -117,12 +132,9 @@ Inductive cumulSpec0 {cf : checker_flags} (Σ : global_env_ext) Γ (pb : conv_pb Σ ;;; Γ ⊢ tLetIn na t ty u ≤s[pb] tLetIn na' t' ty' u' | cumul_Case indn : forall p p' c c' brs brs', - cumul_predicate (fun Γ t u => Σ ;;; Γ ⊢ t ≤s[Conv] u) Γ (compare_universe Conv Σ) p p' -> + cumul_predicate (fun Γ t u => Σ ;;; Γ ⊢ t ≤s[Conv] u) (compare_universe Σ Conv) Γ p p' -> Σ ;;; Γ ⊢ c ≤s[Conv] c' -> - All2 (fun br br' => - eq_context_gen eq eq (bcontext br) (bcontext br') × - Σ ;;; Γ ,,, inst_case_branch_context p br ⊢ bbody br ≤s[Conv] bbody br' - ) brs brs' -> + cumul_branches (fun Γ t u => Σ ;;; Γ ⊢ t ≤s[Conv] u) Γ p brs brs' -> Σ ;;; Γ ⊢ tCase indn p c brs ≤s[pb] tCase indn p' c' brs' | cumul_Proj : forall p c c', @@ -130,25 +142,15 @@ Inductive cumulSpec0 {cf : checker_flags} (Σ : global_env_ext) Γ (pb : conv_pb Σ ;;; Γ ⊢ tProj p c ≤s[pb] tProj p c' | cumul_Fix : forall mfix mfix' idx, - All2 (fun x y => - Σ ;;; Γ ⊢ x.(dtype) ≤s[Conv] y.(dtype) × - Σ ;;; Γ ,,, fix_context mfix ⊢ x.(dbody) ≤s[Conv] y.(dbody) × - (x.(rarg) = y.(rarg)) × - eq_binder_annot x.(dname) y.(dname) - ) mfix mfix' -> + cumul_mfixpoint (fun Γ t u => Σ ;;; Γ ⊢ t ≤s[Conv] u) Γ mfix mfix' -> Σ ;;; Γ ⊢ tFix mfix idx ≤s[pb] tFix mfix' idx | cumul_CoFix : forall mfix mfix' idx, - All2 (fun x y => - Σ ;;; Γ ⊢ x.(dtype) ≤s[Conv] y.(dtype) × - Σ ;;; Γ ,,, fix_context mfix ⊢ x.(dbody) ≤s[Conv] y.(dbody) × - (x.(rarg) = y.(rarg)) × - eq_binder_annot x.(dname) y.(dname) - ) mfix mfix' -> + cumul_mfixpoint (fun Γ t u => Σ ;;; Γ ⊢ t ≤s[Conv] u) Γ mfix mfix' -> Σ ;;; Γ ⊢ tCoFix mfix idx ≤s[pb] tCoFix mfix' idx | cumul_Prim p p' : - onPrims (fun x y => Σ ;;; Γ ⊢ x ≤s[Conv] y) (compare_universe Conv Σ) p p' -> + onPrims (fun x y => Σ ;;; Γ ⊢ x ≤s[Conv] y) (compare_universe Σ Conv) p p' -> Σ ;;; Γ ⊢ tPrim p ≤s[pb] tPrim p' (** Reductions *) @@ -377,15 +379,10 @@ Lemma cumulSpec0_rect : (forall (Γ : context) (pb : conv_pb) (indn : case_info) (p p' : predicate term) (c c' : term) (brs brs' : list (branch term)) - (Hp : cumul_predicate (fun Γ => cumulSpec0 Σ Γ Conv) Γ (compare_universe Conv Σ) p p') - (_ : cumul_predicate_dep Hp (fun Γ => P cf Σ Γ Conv) (fun _ _ _ => True)) + (Hp : cumul_predicate (fun Γ => cumulSpec0 Σ Γ Conv) (compare_universe Σ Conv) Γ p p') + (_ : cumul_predicate_dep Hp (fun Γ => P cf Σ Γ Conv) (fun l l' _ => on_rel (fun _ _ => True) Universe.make' l l')) (Hc : cumulSpec0 Σ Γ Conv c c') (_ : P cf Σ Γ Conv c c' Hc) - (Hbody : All2 - (fun br br' : branch term => - eq_context_gen eq eq (bcontext br) (bcontext br') * - cumulSpec0 Σ (Γ,,, inst_case_branch_context p br) Conv - (bbody br) (bbody br')) - brs brs') + (Hbody : cumul_branches (fun Γ => cumulSpec0 Σ Γ Conv) Γ p brs brs') (_ : All2_dep (fun br br' Hc => P cf Σ (Γ,,, inst_case_branch_context p br) Conv (bbody br) (bbody br') (snd Hc)) Hbody), P cf Σ Γ pb (tCase indn p c brs) (tCase indn p' c' brs') @@ -437,7 +434,7 @@ Lemma cumulSpec0_rect : (forall (Γ : context) (pb : conv_pb) (i : inductive) (u u' : list Level.t) (args args' : list term) - (Hu : R_global_instance Σ (eq_universe Σ) (compare_universe pb Σ) (IndRef i) #|args| u u') + (Hu : cumul_Ind_univ Σ pb i #|args| u u') (Hargs : All2 (cumulSpec0 Σ Γ Conv) args args') (_ : All2_dep (P cf Σ Γ Conv) Hargs), P cf Σ Γ pb (mkApps (tInd i u) args) (mkApps (tInd i u') args') @@ -445,19 +442,19 @@ Lemma cumulSpec0_rect : (forall (Γ : context) (pb : conv_pb) (i : inductive) (k : nat) (u u' : list Level.t) (args args' : list term) - (Hu : R_global_instance Σ (eq_universe Σ) (compare_universe pb Σ) (ConstructRef i k) #|args| u u') + (Hu : cumul_Construct_univ Σ pb i k #|args| u u') (Hargs : All2 (cumulSpec0 Σ Γ Conv) args args') (_ : All2_dep (P cf Σ Γ Conv) Hargs), P cf Σ Γ pb (mkApps (tConstruct i k u) args) (mkApps (tConstruct i k u') args') (cumul_Construct _ _ _ _ _ _ _ _ _ Hu Hargs)) -> - (forall (Γ : context) (pb : conv_pb) (s s' : Universe.t) - (Hs : compare_universe pb Σ s s'), + (forall (Γ : context) (pb : conv_pb) (s s' : sort) + (Hs : compare_sort Σ pb s s'), P cf Σ Γ pb (tSort s) (tSort s') (cumul_Sort _ _ _ _ _ Hs)) -> (forall (Γ : context) (pb : conv_pb) (c : kername) (u u' : list Level.t) - (Hu : R_universe_instance (compare_universe Conv Σ) u u'), + (Hu : cmp_universe_instance (compare_universe Σ Conv) u u'), P cf Σ Γ pb (tConst c u) (tConst c u') (cumul_Const _ _ _ _ _ _ Hu)) -> @@ -492,23 +489,23 @@ Proof. * revert c0. generalize (pparams p), (pparams p'). fix aux' 3; destruct c0; constructor; auto. * apply Forall2_dep_from_nth_error; intros; exact I. - + revert brs brs' a. - fix aux' 3; destruct a; constructor; intuition auto. + + revert brs brs' c1. + fix aux' 3; destruct c1; constructor; intuition auto. - eapply X17 ; eauto. - eapply X18 ; eauto. - revert a. + revert c. unfold cumul_mfixpoint. set (mfixAbs_context := fix_context mfix). clearbody mfixAbs_context. revert mfix mfix'. - fix aux' 3; destruct a; constructor. + fix aux' 3; destruct c; constructor. + intuition auto. + auto. - eapply X19 ; eauto. - revert a. + revert c. unfold cumul_mfixpoint. set (mfixAbs_context := fix_context mfix). clearbody mfixAbs_context. revert mfix mfix'. - fix aux' 3; destruct a; constructor. + fix aux' 3; destruct c; constructor. + intuition auto. + auto. - eapply X20; eauto. @@ -656,15 +653,10 @@ Lemma convSpec0_ind_all : (forall (Γ : context) (indn : case_info) (p p' : predicate term) (c c' : term) (brs brs' : list (branch term)) - (Hp : cumul_predicate (fun Γ => cumulSpec0 Σ Γ Conv) Γ (compare_universe Conv Σ) p p') - (_ : cumul_predicate_dep Hp (fun Γ => P cf Σ Γ Conv) (fun _ _ _ => True)) + (Hp : cumul_predicate (fun Γ => cumulSpec0 Σ Γ Conv) (compare_universe Σ Conv) Γ p p') + (_ : cumul_predicate_dep Hp (fun Γ => P cf Σ Γ Conv) (fun l l' _ => on_rel (fun _ _ => True) Universe.make' l l')) (Hc : cumulSpec0 Σ Γ Conv c c') (_ : P cf Σ Γ Conv c c' Hc) - (Hbody : All2 - (fun br br' : branch term => - eq_context_gen eq eq (bcontext br) (bcontext br') * - cumulSpec0 Σ (Γ,,, inst_case_branch_context p br) Conv - (bbody br) (bbody br')) - brs brs') + (Hbody : cumul_branches (fun Γ => cumulSpec0 Σ Γ Conv) Γ p brs brs') (_ : All2_dep (fun br br' Hc => P cf Σ (Γ,,, inst_case_branch_context p br) Conv (bbody br) (bbody br') (snd Hc)) Hbody), P cf Σ Γ Conv (tCase indn p c brs) (tCase indn p' c' brs') @@ -678,13 +670,7 @@ Lemma convSpec0_ind_all : (forall (Γ : context) (mfix : mfixpoint term) (mfix' : list (def term)) (idx : nat) - (H : All2 - (fun x y : def term => - (cumulSpec0 Σ Γ Conv (dtype x) (dtype y)) - * ((cumulSpec0 Σ (Γ,,, fix_context mfix) Conv (dbody x) (dbody y)) - × (rarg x = rarg y) - × eq_binder_annot (dname x) (dname y))) - mfix mfix') + (H : cumul_mfixpoint (fun Γ => cumulSpec0 Σ Γ Conv) Γ mfix mfix') (_ : All2_dep (fun x y H => P cf Σ Γ Conv (dtype x) (dtype y) (fst H) × P cf Σ (Γ,,, fix_context mfix) Conv (dbody x) (dbody y) (fst (snd H))) @@ -694,13 +680,7 @@ Lemma convSpec0_ind_all : (forall (Γ : context) (mfix : mfixpoint term) (mfix' : list (def term)) (idx : nat) - (H : All2 - (fun x y : def term => - (cumulSpec0 Σ Γ Conv (dtype x) (dtype y)) - × ((cumulSpec0 Σ (Γ,,, fix_context mfix) Conv (dbody x) (dbody y)) - × (rarg x = rarg y) - × (eq_binder_annot (dname x) (dname y)))) - mfix mfix') + (H : cumul_mfixpoint (fun Γ => cumulSpec0 Σ Γ Conv) Γ mfix mfix') (_ : All2_dep (fun x y H => P cf Σ Γ Conv (dtype x) (dtype y) (fst H) × P cf Σ (Γ,,, fix_context mfix) Conv (dbody x) (dbody y) (fst (snd H))) @@ -716,7 +696,7 @@ Lemma convSpec0_ind_all : (forall (Γ : context) (i : inductive) (u u' : list Level.t) (args args' : list term) - (Hu : R_global_instance Σ (eq_universe Σ) (compare_universe Conv Σ) (IndRef i) #|args| u u') + (Hu : cumul_Ind_univ Σ Conv i #|args| u u') (Hargs : All2 (cumulSpec0 Σ Γ Conv) args args') (_ : All2_dep (P cf Σ Γ Conv) Hargs), P cf Σ Γ Conv (mkApps (tInd i u) args) (mkApps (tInd i u') args') @@ -724,19 +704,19 @@ Lemma convSpec0_ind_all : (forall (Γ : context) (i : inductive) (k : nat) (u u' : list Level.t) (args args' : list term) - (Hu : R_global_instance Σ (eq_universe Σ) (compare_universe Conv Σ) (ConstructRef i k) #|args| u u') + (Hu : cumul_Construct_univ Σ Conv i k #|args| u u') (Hargs : All2 (cumulSpec0 Σ Γ Conv) args args') (_ : All2_dep (P cf Σ Γ Conv) Hargs), P cf Σ Γ Conv (mkApps (tConstruct i k u) args) (mkApps (tConstruct i k u') args') (cumul_Construct _ _ _ _ _ _ _ _ _ Hu Hargs)) -> - (forall (Γ : context) (s s' : Universe.t) - (Hs : compare_universe Conv Σ s s'), + (forall (Γ : context) (s s' : sort) + (Hs : compare_sort Σ Conv s s'), P cf Σ Γ Conv (tSort s) (tSort s') (cumul_Sort _ _ _ _ _ Hs)) -> (forall (Γ : context) (c : kername) (u u' : list Level.t) - (Hu : R_universe_instance (compare_universe Conv Σ) u u'), + (Hu : cmp_universe_instance (compare_universe Σ Conv) u u'), P cf Σ Γ Conv (tConst c u) (tConst c u') (cumul_Const _ _ _ _ _ _ Hu)) -> diff --git a/pcuic/theories/PCUICElimination.v b/pcuic/theories/PCUICElimination.v index e0fb65836..b5341a278 100644 --- a/pcuic/theories/PCUICElimination.v +++ b/pcuic/theories/PCUICElimination.v @@ -19,8 +19,7 @@ Require Import ssreflect. Implicit Types (cf : checker_flags) (Σ : global_env_ext). -Definition Is_proof `{cf : checker_flags} Σ Γ t := ∑ T u, Σ ;;; Γ |- t : T × Σ ;;; Γ |- T : tSort u × - (Universe.is_prop u || Universe.is_sprop u). +Definition Is_proof `{cf : checker_flags} Σ Γ t := ∑ T u, Σ ;;; Γ |- t : T × Σ ;;; Γ |- T : tSort u × Sort.is_propositional u. (* TODO: Figure out whether [SingletonProp], [Computational], and [Subsingleton] should use [strictly_extends_decls] or [extends]. -Jason Gross *) Definition SingletonProp `{cf : checker_flags} (Σ : global_env_ext) (ind : inductive) := @@ -59,7 +58,7 @@ Lemma typing_spine_case_predicate {cf: checker_flags} {Σ : global_env_ext} {wf wf_local Σ Γ -> declared_inductive Σ ci mdecl idecl -> consistent_instance_ext Σ (ind_universes mdecl) u -> - wf_universe Σ ps -> + wf_sort Σ ps -> spine_subst Σ Γ params (List.rev params) (smash_context [] (subst_instance u (ind_params mdecl))) -> spine_subst Σ Γ indices (List.rev indices) @@ -85,7 +84,7 @@ Lemma pre_case_predicate_context_gen_eq {cf: checker_flags} {Σ : global_env_ext wf_local Σ Γ -> declared_inductive Σ ci mdecl idecl -> consistent_instance_ext Σ (ind_universes mdecl) (puinst p) -> - wf_universe Σ ps -> + wf_sort Σ ps -> spine_subst Σ Γ (pparams p) (List.rev (pparams p)) (smash_context [] (ind_params mdecl)@[puinst p]) -> Σ ⊢ Γ ,,, pre_case_predicate_context_gen ci mdecl idecl (pparams p) (puinst p) = @@ -97,7 +96,7 @@ Proof. rewrite /pre_case_predicate_context_gen /inst_case_context /ind_predicate_context. rewrite /case_predicate_context' /=. cbn. rewrite subst_context_snoc; len. - constructor. constructor; cbn. reflexivity. + constructor. unfold subst_decl, map_decl; cbn. relativize (subst _ _ _). 1: constructor. reflexivity. symmetry. rewrite subst_instance_mkApps subst_mkApps. f_equal. cbn. now rewrite [subst_instance_instance _ _](subst_instance_id_mdecl _ _ _ cu). rewrite [to_extended_list _]to_extended_list_k_app; len; rewrite !map_app. @@ -121,25 +120,23 @@ Qed. Lemma elim_restriction_works_kelim1 {cf : checker_flags} {Σ : global_env_ext} {Γ T ci p c brs mdecl idecl} : - check_univs -> wf_ext Σ -> declared_inductive Σ ci.(ci_ind) mdecl idecl -> Σ ;;; Γ |- tCase ci p c brs : T -> (Is_proof Σ Γ (tCase ci p c brs) -> False) -> ind_kelim idecl = IntoAny \/ ind_kelim idecl = IntoSetPropSProp. Proof. - intros cu wfΣ. pose wfΣ' := wfΣ.1. intros. + intros wfΣ. pose wfΣ' := wfΣ.1. intros. assert (HT := X). eapply inversion_Case in X as [mdecl' [idecl' [isdecl' [indices [data cum]]]]]; eauto. destruct data. unshelve epose proof (H_ := declared_inductive_to_gen H); eauto. unshelve epose proof (isdecl'_ := declared_inductive_to_gen isdecl'); eauto. eapply declared_inductive_inj in isdecl'_ as []. 2:exact H_. subst. - enough (~ (Universe.is_prop ps \/ Universe.is_sprop ps)). - { clear -cu wfΣ allowed_elim H1. + enough (~ (Sort.is_propositional ps)). + { clear -wfΣ allowed_elim H1. apply wf_ext_consistent in wfΣ as (val&sat). - unfold is_allowed_elimination, is_lSet, eq_universe, eq_levelalg in *. - rewrite cu in allowed_elim. + unfold is_allowed_elimination, is_lSet in *. destruct (ind_kelim idecl); auto; destruct ps; cbn in *; try discriminate; intuition congruence. } intros Huf. apply H0. @@ -154,7 +151,7 @@ Proof. pose proof (PCUICInductiveInversion.isType_mkApps_Ind_smash H (validity scrut_ty)). forward X1. apply (wf_predicate_length_pars wf_pred). simpl in X1. destruct X1 as [sppars [spargs cu']]. - assert (eqctx' : All2 (PCUICEquality.compare_decls eq eq) + assert (eqctx' : eq_context_upto_names (Γ,,, case_predicate_context' ci mdecl idecl p) (Γ,,, predctx)). { eapply All2_app. 2:eapply All2_refl; reflexivity. @@ -165,10 +162,10 @@ Proof. { transitivity (Γ ,,, case_predicate_context' ci mdecl idecl p); revgoals. * symmetry. eapply alpha_eq_context_ws_cumul_ctx_pb => //; fvs. now symmetry. * eapply pre_case_predicate_context_gen_eq; tea. pcuic. - now eapply PCUICWfUniverses.typing_wf_universe in pret_ty. } + now eapply PCUICWfUniverses.typing_wf_sort in pret_ty. } unshelve epose proof (typing_spine_case_predicate (ps:=ps) _ H cons _ sppars). 1-2:shelve. * pcuic. - * now eapply PCUICWfUniverses.typing_wf_universe in pret_ty. + * now eapply PCUICWfUniverses.typing_wf_sort in pret_ty. * rewrite -smash_context_subst_context_let_expand in X2. specialize (X2 spargs scrut_ty). eapply typing_spine_strengthen; tea; revgoals. @@ -178,11 +175,9 @@ Proof. apply ws_cumul_pb_refl; fvs. + eapply validity in pret_ty. eapply isType_it_mkProd_or_LetIn; tea. - * destruct Huf as [Huf|Huf]; rewrite Huf // orb_true_r //. Qed. -Lemma elim_sort_intype {cf:checker_flags} Σ mdecl ind idecl ind_indices ind_sort cdecls : - Universe.is_prop ind_sort -> +Lemma elim_sort_intype {cf:checker_flags} Σ mdecl ind idecl ind_indices cdecls : elim_sort_prop_ind cdecls = IntoAny -> on_constructors cumulSpec0 (lift_typing typing) (Σ, ind_universes mdecl) mdecl @@ -192,11 +187,11 @@ Lemma elim_sort_intype {cf:checker_flags} Σ mdecl ind idecl ind_indices ind_sor (∑ cdecl cdecl_sorts, (ind_ctors idecl = [cdecl]) * (cdecls = [cdecl_sorts]) * - (Forall is_propositional cdecl_sorts) * + (Forall Sort.is_propositional cdecl_sorts) * (on_constructor cumulSpec0 (lift_typing typing) (Σ, ind_universes mdecl) mdecl (inductive_ind ind) idecl ind_indices cdecl cdecl_sorts))%type. Proof. - intros uf lein onc. + intros lein onc. induction onc; simpl in *. left; auto. destruct l' as [|c cs]. @@ -211,7 +206,7 @@ Qed. Lemma typing_spine_it_mkProd_or_LetIn_full_inv {cf:checker_flags} Σ Γ Δ s args s' : wf Σ.1 -> typing_spine Σ Γ (it_mkProd_or_LetIn Δ (tSort s)) args (tSort s') -> - leq_universe (global_ext_constraints Σ) s s'. + leq_sort (global_ext_constraints Σ) s s'. Proof. intros wfΣ. induction Δ using PCUICInduction.ctx_length_rev_ind in args |- *. @@ -281,23 +276,16 @@ Proof. move=> wfe; apply wfe. Qed. Hint Resolve wf_ext_wf : core. Lemma is_propositional_subst_instance u s : - is_propositional (subst_instance_univ u s) = is_propositional s. + Sort.is_propositional (subst_instance_sort u s) = Sort.is_propositional s. Proof. destruct s => //. Qed. -Lemma leq_universe_propositional_l {cf:checker_flags} ϕ u1 u2 : - check_univs -> +Lemma leq_sort_propositional_l {cf:checker_flags} ϕ u1 u2 : prop_sub_type = false -> - consistent ϕ -> - leq_universe ϕ u1 u2 -> is_propositional u1 -> u1 = u2. + leq_sort ϕ u1 u2 -> Sort.is_propositional u1 -> u1 = u2. Proof. - intros Hcf ps cu le. - unfold is_propositional. - destruct (Universe.is_prop u1) eqn:eq. - apply leq_universe_prop_no_prop_sub_type in le; auto. - simpl. now destruct u1, u2. - simpl. intros sp. - apply leq_universe_sprop_l in le; auto. - now destruct u1, u2. + destruct u1 => //=. + 1: intros ->. + all: destruct u2 => //=. Qed. Lemma isType_ws_cumul_ctx_pb {cf Σ Γ Δ T} {wfΣ : wf Σ}: @@ -307,27 +295,26 @@ Lemma isType_ws_cumul_ctx_pb {cf Σ Γ Δ T} {wfΣ : wf Σ}: isType Σ Δ T. Proof. intros HT wf eq. - apply infer_sort_impl with id HT; intros Hs. + apply lift_typing_impl with (1 := HT); intros ?? Hs. eapply closed_context_conversion; tea. Qed. Lemma typing_spine_proofs {cf:checker_flags} Σ Γ Δ ind u args' args T' s : - check_univs -> wf_ext Σ -> Σ ;;; Γ |- T' : tSort s -> typing_spine Σ Γ (it_mkProd_or_LetIn Δ (mkApps (tInd ind u) args')) args T' -> ((All_local_assum (fun Γ' t => - (∑ s, (Σ ;;; Γ ,,, Γ' |- t : tSort s) * is_propositional s)%type) Δ -> + (∑ s, (Σ ;;; Γ ,,, Γ' |- t : tSort s) * Sort.is_propositional s)%type) Δ -> ∥ All (Is_proof Σ Γ) args ∥) * (forall mdecl idecl (Hdecl : declared_inductive Σ.1 ind mdecl idecl), consistent_instance_ext Σ (ind_universes mdecl) u -> - ((is_propositional s -> s = subst_instance_univ u idecl.(ind_sort)) /\ + ((Sort.is_propositional s -> s = subst_instance_sort u idecl.(ind_sort)) /\ (prop_sub_type = false -> - is_propositional (subst_instance_univ u idecl.(ind_sort)) -> - s = subst_instance_univ u idecl.(ind_sort)))))%type. + Sort.is_propositional (subst_instance_sort u idecl.(ind_sort)) -> + s = subst_instance_sort u idecl.(ind_sort)))))%type. Proof. - intros checku wfΣ Ht. pose wfΣ' := wfΣ.1. + intros wfΣ Ht. pose wfΣ' := wfΣ.1. induction Δ using PCUICInduction.ctx_length_rev_ind in Γ, args', args, T', Ht |- *; simpl; intros sp. - dependent elimination sp as [spnil _ _ e|spcons isty isty' e _ sp]. split; [repeat constructor|]. @@ -347,19 +334,11 @@ Proof. rewrite -it_mkProd_or_LetIn_app in sp. eapply typing_spine_it_mkProd_or_LetIn_full_inv in sp; auto. split. - intros Hs. - destruct s => //. - eapply leq_universe_prop_r in sp; auto. - rewrite (is_prop_subst_instance_univ ui') in sp => //. - now destruct (ind_sort idecl). - apply wfΣ. - eapply leq_universe_sprop_r in sp; auto. - rewrite (is_sprop_subst_instance_univ ui') in sp => //. - now destruct (ind_sort idecl). - apply wfΣ. + revert sp. + destruct s => //=; destruct ind_sort => //. intros propsub props. rewrite is_propositional_subst_instance in props. - apply leq_universe_propositional_l in sp; eauto. subst s. + apply leq_sort_propositional_l in sp; eauto. subst s. now destruct (ind_sort idecl). now destruct (ind_sort idecl). now eapply declared_inductive_valid_type. @@ -387,7 +366,7 @@ Proof. eapply subslet_def_tip. eapply typing_wf_local in Ht2. rewrite app_context_assoc in Ht2. eapply All_local_env_app_inv in Ht2 as [Ht2 _]. - depelim Ht2. apply l0. + depelim Ht2. now apply unlift_TermTyp in l. now rewrite app_context_assoc in Ht2. * intros mdecl idec decli. now apply H. @@ -398,13 +377,13 @@ Proof. eapply ws_cumul_pb_Prod_l_inv in e as [na' [dom' [codom' [red eqann conv cum]]]]; auto. eapply subject_reduction_closed in Ht; eauto. intros. - pose proof (PCUICWfUniverses.typing_wf_universe wfΣ Ht). + pose proof (PCUICWfUniverses.typing_wf_sort wfΣ Ht). eapply inversion_Prod in Ht as [s1 [s2 [dom [codom cum']]]]; auto. specialize (H Γ0 ltac:(reflexivity) (Γ ,, vass na' dom') args' []). eapply (type_Cumul _ _ _ _ (tSort s)) in codom; cycle 1; eauto. { econstructor; pcuic. } { eapply ws_cumul_pb_Sort_inv in cum'. - eapply cumul_Sort. etransitivity; eauto. eapply leq_universe_product. } + eapply cumul_Sort. etransitivity; eauto. eapply leq_sort_product. } specialize (H _ codom). have eqctx : Σ ⊢ Γ ,, vass na ty = Γ ,, vass na' dom'. { constructor. apply ws_cumul_ctx_pb_refl. fvs. constructor; auto. } @@ -433,7 +412,7 @@ Proof. intros prs;eapply All_local_assum_app in prs as [prd prs]. depelim prd. eapply (type_ws_cumul_pb (pb:=Conv) _ (U:=ty)) in tyhd. - 2:{ destruct s0 as [s' [Hs' _]]. exists s'; auto. } + 2:{ destruct s0 as [s' [Hs' _]]. eapply has_sort_isType; eauto. } 2:now symmetry. destruct H as [H _]. forward H. { @@ -455,15 +434,15 @@ Lemma check_ind_sorts_is_propositional {cf:checker_flags} (Σ : global_env_ext) (onib : on_ind_body cumulSpec0 (lift_typing typing) (Σ.1, ind_universes mdecl) (inductive_mind ind) mdecl (inductive_ind ind) idecl) : (ind_kelim idecl <> IntoPropSProp /\ ind_kelim idecl <> IntoSProp) -> - is_propositional (ind_sort idecl) -> + Sort.is_propositional (ind_sort idecl) -> check_ind_sorts (lift_typing typing) (Σ.1, ind_universes mdecl) (PCUICEnvironment.ind_params mdecl) (PCUICEnvironment.ind_kelim idecl) (ind_indices idecl) (ind_cunivs onib) (ind_sort idecl) -> - (#|ind_cunivs onib| <= 1) * All (fun cs => All is_propositional cs) (ind_cunivs onib). + (#|ind_cunivs onib| <= 1) * All (fun cs => All Sort.is_propositional cs) (ind_cunivs onib). Proof. intros kelim isp. - unfold check_ind_sorts. simpl. - destruct Universe.is_prop eqn:isp'. + unfold check_ind_sorts. + destruct ind_sort => //=. + induction (ind_cunivs onib); simpl; auto; try discriminate. destruct l; simpl. intros; split; eauto. constructor; [|constructor]. destruct forallb eqn:fo. eapply forallb_All in fo. @@ -471,11 +450,8 @@ Proof. destruct (ind_kelim idecl); intuition cbn in H; try congruence. intros leb. destruct (ind_kelim idecl); simpl in *; intuition congruence. - + destruct Universe.is_sprop eqn:issp. - induction (ind_cunivs onib); simpl; auto; try discriminate. + + induction (ind_cunivs onib); simpl; auto; try discriminate. destruct (ind_kelim idecl); simpl in *; intuition congruence. - unfold is_propositional in isp. - now rewrite isp' issp in isp. Qed. Lemma sorts_local_ctx_All_local_assum_impl {cf:checker_flags} Σ @@ -489,9 +465,9 @@ Proof. destruct a as [na [b|] ty]; constructor; intuition auto. destruct cs => //; eauto. destruct cs => //; eauto. destruct X. - eapply IHΔ. intros. apply (H Γ' t1 s0). right; eauto. all:auto. + eapply IHΔ. intros. apply (H Γ' t0 s0). right; eauto. all:auto. destruct cs => //. destruct X. - eapply H. left; eauto. eauto. + eapply H. left; eauto. now destruct l as (_ & ? & ? & <-). Qed. Lemma In_map {A B} (f : A -> B) (l : list A) x : @@ -510,14 +486,13 @@ Qed. that elimination to any type is allowed. *) Lemma Is_proof_mkApps_tConstruct `{cf : checker_flags} (Σ : global_env_ext) Γ ind n u mdecl idecl args : - check_univs -> wf_ext Σ -> declared_inductive (fst Σ) ind mdecl idecl -> (ind_kelim idecl <> IntoPropSProp /\ ind_kelim idecl <> IntoSProp) -> Is_proof Σ Γ (mkApps (tConstruct ind n u) args) -> #|ind_ctors idecl| <= 1 /\ ∥ All (Is_proof Σ Γ) (skipn (ind_npars mdecl) args) ∥. Proof. - intros checkunivs HΣ decli kelim [tyc [tycs [hc [hty hp]]]]. + intros HΣ decli kelim [tyc [tycs [hc [hty hp]]]]. assert (wfΣ : wf Σ) by apply HΣ. eapply inversion_mkApps in hc as [? [hc hsp]]; auto. eapply inversion_Construct in hc as [mdecl' [idecl' [cdecl' [wfΓ [declc [cu cum']]]]]]; auto. @@ -556,8 +531,8 @@ pose proof (declared_inductive_inj decli_ declc_) as [-> ->]. pose proof (onc.(on_cargs)). pose proof (onib.(ind_sorts)). - assert (Universe.is_prop (ind_sort idecl) || Universe.is_sprop (ind_sort idecl)). - { rewrite -(is_prop_subst_instance_univ u) -(is_sprop_subst_instance_univ u) => //. now subst tycs. } + assert (Sort.is_propositional (ind_sort idecl)). + { rewrite -(is_propositional_subst_instance_univ u) => //. now subst tycs. } apply check_ind_sorts_is_propositional in X1 as [nctors X1]; eauto. assert(#|ind_cunivs onib| = #|ind_ctors idecl|). clear X. clear -onib. pose proof (onib.(onConstructors)). @@ -585,7 +560,7 @@ pose proof (declared_inductive_inj decli_ declc_) as [-> ->]. autorewrite with len in X0. eapply (sorts_local_ctx_All_local_assum_impl Σ (fun Γ Γ' t => - ∑ s0 : Universe.t, Σ;;; Γ ,,, Γ' |- t : tSort s0 × is_propositional s0)). + ∑ s0 : sort, Σ;;; Γ ,,, Γ' |- t : tSort s0 × Sort.is_propositional s0)). 2:eauto. intros. exists s0. intuition auto. eapply In_map in H1 as [cs' [ins ->]]. @@ -612,12 +587,11 @@ pose proof (declared_inductive_inj decli_ declc_) as [-> ->]. Qed. Lemma elim_restriction_works_kelim `{cf : checker_flags} (Σ : global_env_ext) ind mind idecl : - check_univs -> wf_ext Σ -> declared_inductive (fst Σ) ind mind idecl -> (ind_kelim idecl <> IntoPropSProp /\ ind_kelim idecl <> IntoSProp) -> Subsingleton Σ ind. Proof. - intros cu HΣ H indk. + intros HΣ H indk. assert (wfΣ : wf Σ) by apply HΣ. destruct (on_declared_inductive H) as [[]]; eauto. intros ?. intros. @@ -630,13 +604,12 @@ Proof. Qed. Lemma elim_restriction_works `{cf : checker_flags} (Σ : global_env_ext) Γ T (ci : case_info) p c brs mind idecl : - check_univs -> wf_ext Σ -> declared_inductive (fst Σ) ci mind idecl -> Σ ;;; Γ |- tCase ci p c brs : T -> (Is_proof Σ Γ (tCase ci p c brs) -> False) -> Subsingleton Σ ci.(ci_ind). Proof. - intros cu wfΣ decli HT H. + intros wfΣ decli HT H. eapply elim_restriction_works_kelim1 in HT; eauto. eapply elim_restriction_works_kelim; eauto. destruct (ind_kelim idecl); intuition congruence. @@ -673,12 +646,12 @@ Proof. Qed. Lemma elim_restriction_works_proj `{cf : checker_flags} (Σ : global_env_ext) Γ p c mind idecl T : - check_univs -> wf_ext Σ -> + wf_ext Σ -> declared_inductive (fst Σ) p.(proj_ind) mind idecl -> Σ ;;; Γ |- tProj p c : T -> (Is_proof Σ Γ (tProj p c) -> False) -> Subsingleton Σ p.(proj_ind). Proof. - intros cu; intros. eapply elim_restriction_works_kelim; eauto. + intros. eapply elim_restriction_works_kelim; eauto. eapply elim_restriction_works_proj_kelim1 in H0; eauto. intuition congruence. Qed. @@ -687,89 +660,69 @@ Section no_prop_leq_type. Context `{cf : checker_flags}. Variable Hcf : prop_sub_type = false. -Variable Hcf' : check_univs. -Lemma leq_term_prop_sorted_l {Σ Γ v v' u u'} : - wf_ext Σ -> - PCUICEquality.leq_term Σ (global_ext_constraints Σ) v v' -> - Σ;;; Γ |- v : tSort u -> - Σ;;; Γ |- v' : tSort u' -> Universe.is_prop u -> - leq_universe (global_ext_constraints Σ) u' u. -Proof using Hcf Hcf'. - intros wfΣ leq hv hv' isp. - eapply typing_leq_term_prop in leq; eauto. - apply leq_prop_prop; intuition auto. - eapply cumul_prop_sym in leq. - eapply cumul_prop_props in leq; eauto. auto. apply wfΣ. -Qed. - -Lemma leq_term_prop_sorted_r {Σ Γ v v' u u'} : - wf_ext Σ -> - PCUICEquality.leq_term Σ (global_ext_constraints Σ) v v' -> - Σ;;; Γ |- v : tSort u -> - Σ;;; Γ |- v' : tSort u' -> Universe.is_prop u' -> - leq_universe (global_ext_constraints Σ) u u'. -Proof using Hcf Hcf'. - intros wfΣ leq hv hv' isp. - eapply typing_leq_term_prop in leq; eauto. - apply leq_prop_prop; intuition auto. - apply cumul_prop_props in leq; auto. apply wfΣ. +Lemma eq_univ_prop_compare_sort_propositional Σ pb (s s' : sort) : + Sort.is_propositional s' -> + eq_univ_prop s s' -> + compare_sort Σ pb s s'. +Proof using Hcf. + destruct s' => //. + all: unfold eq_univ_prop, compare_sort. + all: destruct s, pb => //=. + all: now rewrite Hcf. Qed. -Lemma leq_term_sprop_sorted_l {Σ Γ v v' u u'} : - wf_ext Σ -> - PCUICEquality.leq_term Σ (global_ext_constraints Σ) v v' -> - Σ;;; Γ |- v : tSort u -> - Σ;;; Γ |- v' : tSort u' -> Universe.is_sprop u -> - leq_universe (global_ext_constraints Σ) u' u. -Proof using Hcf Hcf'. - intros wfΣ leq hv hv' isp. - eapply typing_leq_term_prop in leq; eauto. - apply leq_sprop_sprop; intuition auto. - eapply cumul_prop_sym in leq. - eapply cumul_sprop_props in leq; auto. eauto. auto. apply wfΣ. +Lemma eq_univ_prop_compare_sort_propositional_r Σ pb (s s' : sort) : + Sort.is_propositional s -> + eq_univ_prop s s' -> + compare_sort Σ pb s s'. +Proof using Hcf. + destruct s => //. + all: unfold eq_univ_prop, compare_sort. + all: destruct s', pb => //=. + all: now rewrite Hcf. Qed. -Lemma leq_term_propositional_sorted_l {Σ Γ v v' u u'} : +Lemma leq_term_prop_sorted_l {Σ Γ v v' u u'} : wf_ext Σ -> PCUICEquality.leq_term Σ (global_ext_constraints Σ) v v' -> Σ;;; Γ |- v : tSort u -> - Σ;;; Γ |- v' : tSort u' -> is_propositional u -> - leq_universe (global_ext_constraints Σ) u' u. -Proof using Hcf Hcf'. + Σ;;; Γ |- v' : tSort u' -> Sort.is_propositional u -> + leq_sort (global_ext_constraints Σ) u' u. +Proof using Hcf. intros wfΣ leq hv hv' isp. - eapply orb_true_iff in isp as [isp | isp]. - - eapply leq_term_prop_sorted_l; eauto. - - eapply leq_term_sprop_sorted_l; eauto. + eapply typing_leq_term_prop in leq; eauto. 2: apply wfΣ. + apply cumul_prop_sort in leq. + now eapply eq_univ_prop_compare_sort_propositional with (pb := Cumul). Qed. -Lemma leq_term_sprop_sorted_r {Σ Γ v v' u u'} : +Lemma leq_term_prop_sorted_r {Σ Γ v v' u u'} : wf_ext Σ -> PCUICEquality.leq_term Σ (global_ext_constraints Σ) v v' -> Σ;;; Γ |- v : tSort u -> - Σ;;; Γ |- v' : tSort u' -> Universe.is_sprop u' -> - leq_universe (global_ext_constraints Σ) u u'. -Proof using Hcf Hcf'. + Σ;;; Γ |- v' : tSort u' -> Sort.is_propositional u' -> + leq_sort (global_ext_constraints Σ) u u'. +Proof using Hcf. intros wfΣ leq hv hv' isp. - eapply typing_leq_term_prop in leq; eauto. - apply leq_sprop_sprop; intuition auto. - apply cumul_sprop_props in leq; auto. apply wfΣ. + eapply typing_leq_term_prop in leq; eauto. 2: apply wfΣ. + apply cumul_prop_sort in leq. + now eapply eq_univ_prop_compare_sort_propositional with (pb := Cumul). Qed. Lemma cumul_prop_inv (Σ : global_env_ext) Γ A B u u' : wf_ext Σ -> - Universe.is_prop u -> + Sort.is_propositional u -> (((Σ ;;; Γ |- A : tSort u) * (Σ ;;; Γ |- B : tSort u')) + ((Σ ;;; Γ |- B : tSort u) * (Σ ;;; Γ |- A : tSort u')))%type -> Σ ;;; Γ |- A <= B -> ((Σ ;;; Γ |- A : tSort u) * (Σ ;;; Γ |- B : tSort u))%type. -Proof using Hcf Hcf'. +Proof using Hcf. intros wfΣ propu. intros [[HA HB]|[HB HA]] cum; split; auto; - apply cumul_alt in cum as [v [v' [[redv redv'] leq]]]. + apply cumul_alt in cum as [v [v' (redv & redv' & leq)]]. - eapply type_Cumul' with (tSort u'); eauto. eapply PCUICArities.isType_Sort. - now eapply PCUICWfUniverses.typing_wf_universe in HA. + now eapply PCUICWfUniverses.typing_wf_sort in HA. pcuic. eapply cumul_Sort. eapply subject_reduction in redv; eauto. eapply subject_reduction in redv'; eauto. @@ -779,129 +732,92 @@ Proof using Hcf Hcf'. eapply leq_term_prop_sorted_r in leq; eauto. eapply type_Cumul' with (tSort u'); eauto. eapply PCUICArities.isType_Sort. - now eapply PCUICWfUniverses.typing_wf_universe in HB. + now eapply PCUICWfUniverses.typing_wf_sort in HB. pcuic. eapply cumul_Sort; auto. Qed. -Lemma cumul_sprop_inv (Σ : global_env_ext) Γ A B u u' : +Lemma unique_sorting_family {pb} {Σ : global_env_ext} {Γ T U s s'} : wf_ext Σ -> - Universe.is_sprop u -> - (((Σ ;;; Γ |- A : tSort u) * (Σ ;;; Γ |- B : tSort u')) + - ((Σ ;;; Γ |- B : tSort u) * (Σ ;;; Γ |- A : tSort u')))%type -> - Σ ;;; Γ |- A <= B -> - ((Σ ;;; Γ |- A : tSort u) * (Σ ;;; Γ |- B : tSort u))%type. -Proof using Hcf Hcf'. - intros wfΣ propu. - intros [[HA HB]|[HB HA]] cum; split; auto; - apply cumul_alt in cum as [v [v' [[redv redv'] leq]]]. - - eapply type_Cumul' with (tSort u'); eauto. - eapply isType_Sort. - 1: now destruct u. - 1: pcuic. - eapply cumul_Sort. - eapply subject_reduction in redv; eauto. - eapply subject_reduction in redv'; eauto. - eapply leq_term_sprop_sorted_l; eauto. - - eapply subject_reduction in redv; eauto. - eapply subject_reduction in redv'; eauto. - eapply leq_term_sprop_sorted_r in leq; eauto. - eapply type_Cumul' with (tSort u'); eauto. - eapply PCUICArities.isType_Sort. - 1: now destruct u. - 1: now pcuic. - now eapply cumul_Sort. + Σ ;;; Γ |- T : tSort s -> + Σ ;;; Γ |- U : tSort s' -> + Σ ;;; Γ ⊢ T ≤[pb] U -> + Sort.to_family s = Sort.to_family s'. +Proof using Hcf. + intros wfΣ HT HU cum. + apply ws_cumul_pb_forget in cum. + eapply typing_cumul_term_prop in cum; tea. + apply cumul_prop_sort in cum. + move: cum. + destruct s, s' => //=; cbn. + all: now rewrite Hcf //. Qed. -Lemma unique_sorting_equality_prop_l {pb} {Σ : global_env_ext} {Γ T U s s'} : +Lemma unique_sorting_equality_prop {pb} {Σ : global_env_ext} {Γ T U s s'} : wf_ext Σ -> Σ ;;; Γ |- T : tSort s -> Σ ;;; Γ |- U : tSort s' -> Σ ;;; Γ ⊢ T ≤[pb] U -> - Universe.is_prop s -> Universe.is_prop s'. -Proof using Hcf Hcf'. - intros wfΣ HT HU cum isp. - eapply PCUICSpine.ws_cumul_pb_le_le in cum. - eapply ws_cumul_pb_alt_closed in cum as [v [v' [eqv]]]. - eapply subject_reduction_closed in HT; tea. - eapply subject_reduction_closed in HU; tea. - eapply leq_term_prop_sorted_l in c0; tea. all:eauto with pcuic. - eapply leq_universe_prop_r; tea; eauto with pcuic. + Sort.is_prop s = Sort.is_prop s'. +Proof using Hcf. + intros wfΣ HT HU cum. + eapply unique_sorting_family in cum; tea. + now destruct s, s'. Qed. -Lemma unique_sorting_equality_prop_r {pb} {Σ : global_env_ext} {Γ T U s s'} : +Lemma unique_sorting_equality_prop_l {pb} {Σ : global_env_ext} {Γ T U s s'} : wf_ext Σ -> Σ ;;; Γ |- T : tSort s -> Σ ;;; Γ |- U : tSort s' -> Σ ;;; Γ ⊢ T ≤[pb] U -> - Universe.is_prop s' -> Universe.is_prop s. -Proof using Hcf Hcf'. - intros wfΣ HT HU cum isp. - eapply PCUICSpine.ws_cumul_pb_le_le in cum. - eapply ws_cumul_pb_alt_closed in cum as [v [v' [eqv]]]. - eapply subject_reduction_closed in HT; tea. - eapply subject_reduction_closed in HU; tea. - eapply leq_term_prop_sorted_r in c0; tea. all:eauto with pcuic. - eapply leq_universe_prop_r; tea; eauto with pcuic. + Sort.is_prop s -> Sort.is_prop s'. +Proof using Hcf. + intros. + erewrite <- unique_sorting_equality_prop; eassumption. Qed. -Lemma unique_sorting_equality_prop {pb} {Σ : global_env_ext} {Γ T U s s'} : +Lemma unique_sorting_equality_prop_r {pb} {Σ : global_env_ext} {Γ T U s s'} : wf_ext Σ -> Σ ;;; Γ |- T : tSort s -> Σ ;;; Γ |- U : tSort s' -> Σ ;;; Γ ⊢ T ≤[pb] U -> - Universe.is_prop s = Universe.is_prop s'. -Proof using Hcf Hcf'. - intros wfΣ HT HU cum. - apply iff_is_true_eq_bool. - split. - now eapply unique_sorting_equality_prop_l; tea. - now eapply unique_sorting_equality_prop_r; tea. + Sort.is_prop s' -> Sort.is_prop s. +Proof using Hcf. + intros. + erewrite unique_sorting_equality_prop; eassumption. Qed. -Lemma unique_sorting_equality_sprop_l {pb} {Σ : global_env_ext} {Γ T U s s'} : +Lemma unique_sorting_equality_sprop {pb} {Σ : global_env_ext} {Γ T U s s'} : wf_ext Σ -> Σ ;;; Γ |- T : tSort s -> Σ ;;; Γ |- U : tSort s' -> Σ ;;; Γ ⊢ T ≤[pb] U -> - Universe.is_sprop s -> Universe.is_sprop s'. -Proof using Hcf Hcf'. - intros wfΣ HT HU cum isp. - eapply PCUICSpine.ws_cumul_pb_le_le in cum. - eapply ws_cumul_pb_alt_closed in cum as [v [v' [eqv]]]. - eapply subject_reduction_closed in HT; tea. - eapply subject_reduction_closed in HU; tea. - eapply leq_term_sprop_sorted_l in c0; tea. all:eauto with pcuic. - eapply leq_universe_sprop_r; tea; eauto with pcuic. + Sort.is_sprop s = Sort.is_sprop s'. +Proof using Hcf. + intros wfΣ HT HU cum. + eapply unique_sorting_family in cum; tea. + now destruct s, s'. Qed. -Lemma unique_sorting_equality_sprop_r {pb} {Σ : global_env_ext} {Γ T U s s'} : +Lemma unique_sorting_equality_sprop_l {pb} {Σ : global_env_ext} {Γ T U s s'} : wf_ext Σ -> Σ ;;; Γ |- T : tSort s -> Σ ;;; Γ |- U : tSort s' -> Σ ;;; Γ ⊢ T ≤[pb] U -> - Universe.is_sprop s' -> Universe.is_sprop s. -Proof using Hcf Hcf'. - intros wfΣ HT HU cum isp. - eapply PCUICSpine.ws_cumul_pb_le_le in cum. - eapply ws_cumul_pb_alt_closed in cum as [v [v' [eqv]]]. - eapply subject_reduction_closed in HT; tea. - eapply subject_reduction_closed in HU; tea. - eapply leq_term_sprop_sorted_r in c0; tea. all:eauto with pcuic. - eapply leq_universe_sprop_r; tea; eauto with pcuic. + Sort.is_sprop s -> Sort.is_sprop s'. +Proof using Hcf. + intros. + erewrite <- unique_sorting_equality_sprop; eassumption. Qed. -Lemma unique_sorting_equality_sprop {pb} {Σ : global_env_ext} {Γ T U s s'} : +Lemma unique_sorting_equality_sprop_r {pb} {Σ : global_env_ext} {Γ T U s s'} : wf_ext Σ -> Σ ;;; Γ |- T : tSort s -> Σ ;;; Γ |- U : tSort s' -> Σ ;;; Γ ⊢ T ≤[pb] U -> - Universe.is_sprop s = Universe.is_sprop s'. -Proof using Hcf Hcf'. - intros wfΣ HT HU cum. - apply iff_is_true_eq_bool. - split. - now eapply unique_sorting_equality_sprop_l; tea. - now eapply unique_sorting_equality_sprop_r; tea. + Sort.is_sprop s' -> Sort.is_sprop s. +Proof using Hcf. + intros. + erewrite unique_sorting_equality_sprop; eassumption. Qed. Lemma unique_sorting_equality_propositional {pb} {Σ : global_env_ext} {Γ T U s s'} : @@ -909,75 +825,67 @@ Lemma unique_sorting_equality_propositional {pb} {Σ : global_env_ext} {Γ T U s Σ ;;; Γ |- T : tSort s -> Σ ;;; Γ |- U : tSort s' -> Σ ;;; Γ ⊢ T ≤[pb] U -> - is_propositional s = is_propositional s'. -Proof using Hcf Hcf'. + Sort.is_propositional s = Sort.is_propositional s'. +Proof using Hcf. intros wfΣ HT HU cum. - unfold is_propositional. - destruct (Universe.is_prop s) eqn:isp => /=. symmetry. - - apply orb_true_intro; left. - now rewrite (unique_sorting_equality_prop wfΣ HT HU cum) in isp. - - destruct (Universe.is_sprop s) eqn:isp' => /=. symmetry. - apply orb_true_intro; right. - now rewrite (unique_sorting_equality_sprop wfΣ HT HU cum) in isp'. - rewrite (unique_sorting_equality_prop wfΣ HT HU cum) in isp. - rewrite (unique_sorting_equality_sprop wfΣ HT HU cum) in isp'. - rewrite isp isp' //. + eapply unique_sorting_family in cum; tea. + now destruct s, s'. Qed. Lemma cumul_prop1 (Σ : global_env_ext) Γ A B u : wf_ext Σ -> - Universe.is_prop u -> + Sort.is_prop u -> isType Σ Γ A -> Σ ;;; Γ |- B : tSort u -> Σ ;;; Γ |- A <= B -> Σ ;;; Γ |- A : tSort u. -Proof using Hcf Hcf'. +Proof using Hcf. intros. - destruct X0 as [s Hs]. - eapply cumul_prop_inv in H. 4:eauto. pcuicfo. auto. - right; eauto. + destruct X0 as (_ & s & Hs & _). + eapply cumul_prop_inv in X2 as []; eauto. + now apply Sort.is_prop_propositional. Qed. Lemma cumul_prop2 (Σ : global_env_ext) Γ A B u : wf_ext Σ -> - Universe.is_prop u -> + Sort.is_prop u -> isType Σ Γ B -> Σ ;;; Γ |- A <= B -> Σ ;;; Γ |- A : tSort u -> Σ ;;; Γ |- B : tSort u. -Proof using Hcf Hcf'. +Proof using Hcf. intros. - destruct X0 as [s Hs]. - eapply cumul_prop_inv in H. 4:eauto. pcuicfo. auto. - left; eauto. + destruct X0 as (_ & s & Hs & _). + eapply cumul_prop_inv in X1 as []; eauto. + now apply Sort.is_prop_propositional. Qed. Lemma cumul_sprop1 (Σ : global_env_ext) Γ A B u : wf_ext Σ -> - Universe.is_sprop u -> + Sort.is_sprop u -> isType Σ Γ A -> Σ ;;; Γ |- B : tSort u -> Σ ;;; Γ |- A <= B -> Σ ;;; Γ |- A : tSort u. -Proof using Hcf Hcf'. +Proof using Hcf. intros. - destruct X0 as [s Hs]. - eapply cumul_sprop_inv in H. 4:eauto. pcuicfo. auto. - right; eauto. + destruct X0 as (_ & s & Hs & _). + eapply cumul_prop_inv in X2 as []; eauto. + now apply Sort.is_sprop_propositional. Qed. Lemma cumul_sprop2 (Σ : global_env_ext) Γ A B u : wf_ext Σ -> - Universe.is_sprop u -> + Sort.is_sprop u -> isType Σ Γ B -> Σ ;;; Γ |- A <= B -> Σ ;;; Γ |- A : tSort u -> Σ ;;; Γ |- B : tSort u. -Proof using Hcf Hcf'. +Proof using Hcf. intros. - destruct X0 as [s Hs]. - eapply cumul_sprop_inv in H. 4:eauto. pcuicfo. auto. - left; eauto. + destruct X0 as (_ & s & Hs & _). + eapply cumul_prop_inv in X1 as []; eauto. + now apply Sort.is_sprop_propositional. Qed. End no_prop_leq_type. diff --git a/pcuic/theories/PCUICEquality.v b/pcuic/theories/PCUICEquality.v index b9db09530..494aa21b5 100644 --- a/pcuic/theories/PCUICEquality.v +++ b/pcuic/theories/PCUICEquality.v @@ -20,11 +20,11 @@ Instance All2_fold_len {A} P (Γ Δ : list A) : HasLen (All2_fold P Γ Δ) #|Γ| Implicit Types (cf : checker_flags). -Definition R_universe_instance R := - fun u u' => Forall2 R (List.map Universe.make u) (List.map Universe.make u'). +Definition cmp_universe_instance (cmp_univ : Universe.t -> Universe.t -> Prop) : Instance.t -> Instance.t -> Prop := + Forall2 (on_rel cmp_univ Universe.make'). -Definition R_universe_instance_dep R R' := - fun {u u'} (H : R_universe_instance R u u') => Forall2_dep R' H. +Definition cmp_universe_instance_dep cmp_univ P' := + fun {u u'} (H : cmp_universe_instance cmp_univ u u') => Forall2_dep P' H. (** Cumulative inductive types: @@ -33,26 +33,15 @@ Definition R_universe_instance_dep R R' := even on ill-formed terms. It corresponds to the right notion on well-formed terms. *) -Definition R_universe_variance (Re Rle : Universe.t -> Universe.t -> Prop) v u u' := +Definition cmp_universe_variance (cmp_univ : conv_pb -> Universe.t -> Universe.t -> Prop) pb v u u' := match v with | Variance.Irrelevant => True - | Variance.Covariant => Rle (Universe.make u) (Universe.make u') - | Variance.Invariant => Re (Universe.make u) (Universe.make u') + | Variance.Covariant => on_rel (cmp_univ pb) Universe.make' u u' + | Variance.Invariant => on_rel (cmp_univ Conv) Universe.make' u u' end. -Fixpoint R_universe_instance_variance Re Rle v u u' := - match u, u' return Prop with - | u :: us, u' :: us' => - match v with - | [] => R_universe_instance_variance Re Rle v us us' - (* Missing variance stands for irrelevance, we still check that the instances have - the same length. *) - | v :: vs => R_universe_variance Re Rle v u u' /\ - R_universe_instance_variance Re Rle vs us us' - end - | [], [] => True - | _, _ => False - end. +Definition cmp_universe_instance_variance cmp_univ pb v u u' := + Forall3 (cmp_universe_variance cmp_univ pb) v u u'. Definition global_variance_gen lookup gr napp := match gr with @@ -60,164 +49,274 @@ Definition global_variance_gen lookup gr napp := match lookup_inductive_gen lookup ind with | Some (mdecl, idecl) => match destArity [] idecl.(ind_type) with - | Some (ctx, _) => if (context_assumptions ctx) <=? napp then mdecl.(ind_variance) - else None - | None => None + | Some (ctx, _) => if (context_assumptions ctx) <=? napp then + match mdecl.(ind_variance) with + | Some var => Variance var + | None => AllEqual + end + else AllEqual + | None => AllEqual end - | None => None + | None => AllEqual end | ConstructRef ind k => match lookup_constructor_gen lookup ind k with | Some (mdecl, idecl, cdecl) => if (cdecl.(cstr_arity) + mdecl.(ind_npars))%nat <=? napp then (** Fully applied constructors are always compared at the same supertype, - which implies that no universe ws_cumul_pb needs to be checked here. *) - Some [] - else None - | _ => None + which implies that no universe ws_cumul_pb needs to be checked here. + We will still check that the instances have the same length. *) + AllIrrelevant + else AllEqual + | _ => AllEqual end - | _ => None + | _ => AllEqual end. Notation global_variance Σ := (global_variance_gen (lookup_env Σ)). -Definition R_opt_variance Re Rle v := +Definition cmp_opt_variance cmp_univ pb v := match v with - | Some v => R_universe_instance_variance Re Rle v - | None => R_universe_instance Re + | AllEqual => cmp_universe_instance (cmp_univ Conv) + | AllIrrelevant => fun l l' => #|l| = #|l'| + | Variance v => fun u u' => cmp_universe_instance (cmp_univ Conv) u u' \/ cmp_universe_instance_variance cmp_univ pb v u u' end. -Definition R_global_instance_gen Σ Re Rle gr napp := - R_opt_variance Re Rle (global_variance_gen Σ gr napp). +Lemma cmp_universe_universe_variance (cmp_univ : conv_pb -> Universe.t -> Universe.t -> Prop) pb v u u' : + RelationClasses.subrelation (cmp_univ Conv) (cmp_univ pb) -> + on_rel (cmp_univ Conv) Universe.make' u u' -> cmp_universe_variance cmp_univ pb v u u'. +Proof. + destruct v => //=. + intros H H1; apply H, H1. +Qed. + +Lemma cmp_instance_variance cmp_univ pb v u u' : + RelationClasses.subrelation (cmp_univ Conv) (cmp_univ pb) -> + #|v| = #|u| -> + cmp_universe_instance (cmp_univ Conv) u u' -> cmp_universe_instance_variance cmp_univ pb v u u'. +Proof. + intros Hsub Hlen Hu. + apply Forall2_triv in Hlen. + eapply Forall2_Forall2_Forall3 in Hu; tea. + apply Forall3_impl with (1 := Hu) => v1 u1 u1' [] _ H1. + now apply cmp_universe_universe_variance. +Qed. + +Lemma cmp_instance_opt_variance cmp_univ pb v : + RelationClasses.subrelation (cmp_opt_variance cmp_univ pb AllEqual) (cmp_opt_variance cmp_univ pb v). +Proof. + intros u u' H. + destruct v as [| |v] => //=. + 1: now apply Forall2_length in H. + now left. +Qed. + +Lemma cmp_opt_variance_var_dec cmp_univ pb v ui ui' : + RelationClasses.subrelation (cmp_univ Conv) (cmp_univ pb) -> + cmp_universe_instance (cmp_univ Conv) ui ui' \/ cmp_universe_instance_variance cmp_univ pb v ui ui' -> + { cmp_universe_instance (cmp_univ Conv) ui ui' } + { cmp_universe_instance_variance cmp_univ pb v ui ui' }. +Proof. + intros subr H. + elim:(eq_dec #|v| #|ui|). + - right. + destruct H as [|]; tas. + now eapply cmp_instance_variance. + - left. + destruct H as [|]; tas. + eapply @Forall3_Forall2_left with (Q := fun _ _ => True) in H => //. + apply Forall2_length in H. + now exfalso. +Qed. + +Lemma cmp_opt_variance_length cmp_univ pb v ui ui' : + cmp_opt_variance cmp_univ pb v ui ui' -> #|ui| = #|ui'|. +Proof. + destruct v => //=. + 1: apply Forall2_length. + move => [|]. + 1: apply Forall2_length. + intro H. + eapply @Forall2_length with (P := fun _ _ => True). + now eapply Forall3_Forall2_right => //. +Qed. + +Lemma cmp_opt_variance_or_impl cmp_universe1 cmp_universe2 cmp_universe3 pb1 pb2 pb3 v u1 u1' u2 u2' u3 u3' : + RelationClasses.subrelation (cmp_universe1 Conv) (cmp_universe1 pb1) -> + RelationClasses.subrelation (cmp_universe2 Conv) (cmp_universe2 pb2) -> + (cmp_universe_instance (cmp_universe1 Conv) u1 u1' -> cmp_universe_instance (cmp_universe2 Conv) u2 u2' -> cmp_universe_instance (cmp_universe3 Conv) u3 u3') -> + (cmp_universe_instance_variance cmp_universe1 pb1 v u1 u1' -> cmp_universe_instance_variance cmp_universe2 pb2 v u2 u2' -> cmp_universe_instance_variance cmp_universe3 pb3 v u3 u3') -> + (#|u1| = #|u1'| -> #|u2| = #|u2'| -> #|u1| = #|u2|) -> + cmp_universe_instance (cmp_universe1 Conv) u1 u1' \/ cmp_universe_instance_variance cmp_universe1 pb1 v u1 u1' -> + cmp_universe_instance (cmp_universe2 Conv) u2 u2' \/ cmp_universe_instance_variance cmp_universe2 pb2 v u2 u2' -> + cmp_universe_instance (cmp_universe3 Conv) u3 u3' \/ cmp_universe_instance_variance cmp_universe3 pb3 v u3 u3'. +Proof. + intros Hsub1 Hsub2 Hl Hr e [H|H] [H'|H']; [left; apply Hl| right; apply Hr ..]; auto. + all: apply cmp_instance_variance; tas. + - rewrite e. + all: eapply Forall2_length; tea. + + eapply @Forall3_Forall2_right with (Q := fun _ _ => True); eauto. + + eapply @Forall3_Forall2_left with (Q := fun _ _ => True); eauto. + - rewrite -e. + all: eapply Forall2_length; tea. + + eapply @Forall3_Forall2_right with (Q := fun _ _ => True); eauto. + + eapply @Forall3_Forall2_left with (Q := fun _ _ => True); eauto. +Qed. + +Definition cmp_global_instance_gen Σ cmp_universe pb gr napp := + cmp_opt_variance cmp_universe pb (global_variance_gen Σ gr napp). -Notation R_global_instance Σ := (R_global_instance_gen (lookup_env Σ)). +Notation cmp_global_instance Σ := (cmp_global_instance_gen (lookup_env Σ)). -Definition R_ind_universes {cf:checker_flags} (Σ : global_env_ext) ind n i i' := - R_global_instance Σ (eq_universe (global_ext_constraints Σ)) - (leq_universe (global_ext_constraints Σ)) (IndRef ind) n i i'. +Definition cmp_ind_universes {cf:checker_flags} (Σ : global_env_ext) ind n i i' := + cmp_global_instance Σ (compare_universe (global_ext_constraints Σ)) Cumul (IndRef ind) n i i'. -Lemma R_universe_instance_impl R R' : +Lemma cmp_universe_instance_impl R R' : RelationClasses.subrelation R R' -> - RelationClasses.subrelation (R_universe_instance R) (R_universe_instance R'). + RelationClasses.subrelation (cmp_universe_instance R) (cmp_universe_instance R'). Proof. - intros H x y xy. eapply Forall2_impl ; tea. + intros H x y xy. eapply Forall2_impl; tea; unfold on_rel; auto. Qed. -Lemma R_universe_instance_impl' R R' : +Lemma cmp_universe_instance_impl' R R' : RelationClasses.subrelation R R' -> - forall u u', R_universe_instance R u u' -> R_universe_instance R' u u'. + forall u u', cmp_universe_instance R u u' -> cmp_universe_instance R' u u'. Proof. - intros H x y xy. eapply Forall2_impl ; tea. + intros H x y xy. eapply Forall2_impl; tea; unfold on_rel; auto. Qed. -Section compare_decls. - (* leq_term compares types, eq_term bodies: - - For conversion they are both equality - - For cumulativity only the type are compared using leq. - *) - Context {eq_term leq_term : term -> term -> Type}. - Inductive compare_decls : context_decl -> context_decl -> Type := - | compare_vass {na T na' T'} : - eq_binder_annot na na' -> - leq_term T T' -> - compare_decls (vass na T) (vass na' T') - | compare_vdef {na b T na' b' T'} : - eq_binder_annot na na' -> - eq_term b b' -> - leq_term T T' -> - compare_decls (vdef na b T) (vdef na' b' T'). +Lemma cmp_universe_variance_impl R R' pb pb' v : + RelationClasses.subrelation (R Conv) (R' Conv) -> + RelationClasses.subrelation (R pb) (R' pb') -> + RelationClasses.subrelation (cmp_universe_variance R pb v) (cmp_universe_variance R' pb' v). +Proof. + intros HConv Hpb x y. + destruct v => //=. + all: unfold on_rel; now auto. +Qed. + +Lemma cmp_universe_instance_variance_impl R R' pb pb' v : + RelationClasses.subrelation (R Conv) (R' Conv) -> + RelationClasses.subrelation (R pb) (R' pb') -> + RelationClasses.subrelation (cmp_universe_instance_variance R pb v) (cmp_universe_instance_variance R' pb' v). +Proof. + intros HConv Hpb x y xy. + eapply Forall3_impl; tea. + intros ???. + now apply cmp_universe_variance_impl. +Qed. - Derive Signature NoConfusion for compare_decls. -End compare_decls. -Arguments compare_decls : clear implicits. +Inductive eq_decl_upto_names : context_decl -> context_decl -> Type := + | compare_vass {na na' T} : + eq_binder_annot na na' -> eq_decl_upto_names (vass na T) (vass na' T) + | compare_vdef {na na' b T} : + eq_binder_annot na na' -> eq_decl_upto_names (vdef na b T) (vdef na' b T). +Derive Signature NoConfusion for eq_decl_upto_names. +#[global] Hint Constructors eq_decl_upto_names : pcuic. -Notation eq_context_upto_names := (All2 (compare_decls eq eq)). +Definition compare_decls cmp_term pb := PCUICConversion.All_decls_alpha_pb pb cmp_term. -Notation eq_context_gen eq_term leq_term := - (All2_fold (fun _ _ => compare_decls eq_term leq_term)). +Notation eq_context_upto_names := (All2 eq_decl_upto_names). -Lemma eq_context_upto_names_gen Γ Γ' : eq_context_upto_names Γ Γ' <~> eq_context_gen eq eq Γ Γ'. +Notation eq_context_gen cmp_term pb := + (All2_fold (fun _ _ => compare_decls cmp_term pb)). + +Lemma eq_decl_upto_names_gen decl decl' pb : eq_decl_upto_names decl decl' <~> compare_decls (fun _ => eq) pb decl decl'. +Proof. + split; intros e; depind e; subst; constructor; auto. +Qed. + +Lemma eq_context_upto_names_gen Γ Γ' pb : eq_context_upto_names Γ Γ' <~> eq_context_gen (fun _ => eq) pb Γ Γ'. Proof. - split; intros e; depind e; constructor; auto. + split; intros e; depind e; constructor; tas; now eapply eq_decl_upto_names_gen. Qed. -Lemma compare_decls_impl eq_term leq_term eq_term' leq_term' : - subrelation eq_term eq_term' -> - subrelation leq_term leq_term' -> - subrelation (compare_decls eq_term leq_term) - (compare_decls eq_term' leq_term'). +Lemma compare_decls_impl cmp_term cmp_term' pb pb' : + subrelation (cmp_term Conv) (cmp_term' Conv) -> + subrelation (cmp_term pb) (cmp_term' pb') -> + subrelation (compare_decls cmp_term pb) (compare_decls cmp_term' pb'). Proof. intros he hle x y []; constructor; auto. Qed. -Lemma eq_context_gen_impl eq_term leq_term eq_term' leq_term' : - subrelation eq_term eq_term' -> - subrelation leq_term leq_term' -> - subrelation (eq_context_gen eq_term leq_term) (eq_context_gen eq_term' leq_term'). +Lemma eq_context_gen_impl cmp_term cmp_term' pb pb' : + subrelation (cmp_term Conv) (cmp_term' Conv) -> + subrelation (cmp_term pb) (cmp_term' pb') -> + subrelation (eq_context_gen cmp_term pb) (eq_context_gen cmp_term' pb'). Proof. intros he hle x y eq. eapply All2_fold_impl; tea => /=. intros; eapply compare_decls_impl; tea. Qed. -Lemma compare_decl_impl_ondecl P eq_term leq_term eq_term' leq_term' d d' : +Lemma compare_decl_impl_ondecl P cmp_term cmp_term' pb pb' d d' : ondecl P d -> - (forall x y, P x -> eq_term x y -> eq_term' x y) -> - (forall x y, P x -> leq_term x y -> leq_term' x y) -> - compare_decls eq_term leq_term d d' -> - compare_decls eq_term' leq_term' d d'. + (forall x y, P x -> cmp_term Conv x y -> cmp_term' Conv x y) -> + (forall x y, P x -> cmp_term pb x y -> cmp_term' pb' x y) -> + compare_decls cmp_term pb d d' -> + compare_decls cmp_term' pb' d d'. Proof. intros ond he hle cd; depelim cd; depelim ond; constructor => //; eauto. Qed. -Lemma compare_decl_map eq_term leq_term f g d d' : - compare_decls (fun x y => eq_term (f x) (g y)) - (fun x y => leq_term (f x) (g y)) d d' -> - compare_decls eq_term leq_term (map_decl f d) (map_decl g d'). +Lemma compare_decl_map cmp_term pb f g d d' : + compare_decls (fun pb x y => cmp_term pb (f x) (g y)) pb d d' -> + compare_decls cmp_term pb (map_decl f d) (map_decl g d'). Proof. intros h; depelim h; constructor; intuition auto. Qed. -Definition bcompare_decls (eq_term leq_term : term -> term -> bool) (d d' : context_decl) : bool := - match d, d' with - | {| decl_name := na; decl_body := None; decl_type := T |}, - {| decl_name := na'; decl_body := None; decl_type := T' |} => - eqb_binder_annot na na' && leq_term T T' - | {| decl_name := na; decl_body := Some b; decl_type := T |}, - {| decl_name := na'; decl_body := Some b'; decl_type := T' |} => - eqb_binder_annot na na' && eq_term b b' && leq_term T T' - | _, _ => false - end. - #[global] -Polymorphic Instance compare_decl_refl eq_term leq_term : - CRelationClasses.Reflexive eq_term -> - CRelationClasses.Reflexive leq_term -> - CRelationClasses.Reflexive (compare_decls eq_term leq_term). +Polymorphic Instance compare_decl_refl cmp_term pb : + CRelationClasses.Reflexive (cmp_term Conv) -> + CRelationClasses.Reflexive (cmp_term pb) -> + CRelationClasses.Reflexive (compare_decls cmp_term pb). Proof. intros heq hle d. destruct d as [na [b|] ty]; constructor; auto; reflexivity. Qed. #[global] -Polymorphic Instance compare_decl_sym eq_term leq_term : - CRelationClasses.Symmetric eq_term -> - CRelationClasses.Symmetric leq_term -> - CRelationClasses.Symmetric (compare_decls eq_term leq_term). +Polymorphic Instance compare_decl_sym cmp_term pb : + CRelationClasses.Symmetric (cmp_term Conv) -> + CRelationClasses.Symmetric (cmp_term pb) -> + CRelationClasses.Symmetric (compare_decls cmp_term pb). Proof. intros heq hle d d' []; constructor; auto; now symmetry. Qed. #[global] -Polymorphic Instance compare_decl_trans eq_term leq_term : - CRelationClasses.Transitive eq_term -> - CRelationClasses.Transitive leq_term -> - CRelationClasses.Transitive (compare_decls eq_term leq_term). +Polymorphic Instance compare_decl_trans cmp_term pb : + CRelationClasses.Transitive (cmp_term Conv) -> + CRelationClasses.Transitive (cmp_term pb) -> + CRelationClasses.Transitive (compare_decls cmp_term pb). Proof. intros hle hre x y z h h'; depelim h; depelim h'; constructor; auto; etransitivity; eauto. Qed. +#[global] +Polymorphic Instance eq_decl_upto_names_refl : + CRelationClasses.Reflexive eq_decl_upto_names. +Proof. + intros d. + destruct d as [na [b|] ty]; constructor; auto; reflexivity. +Qed. + +#[global] +Polymorphic Instance eq_decl_upto_names_sym : + CRelationClasses.Symmetric eq_decl_upto_names. +Proof. + intros d d' []; constructor; auto; now symmetry. +Qed. + +#[global] +Polymorphic Instance eq_decl_upto_names_trans : + CRelationClasses.Transitive eq_decl_upto_names. +Proof. + intros x y z h h'; depelim h; depelim h'; constructor; auto; + etransitivity; eauto. +Qed. + #[global] Instance alpha_eq_reflexive : CRelationClasses.Reflexive eq_context_upto_names. Proof. @@ -237,10 +336,10 @@ Proof. Qed. #[global] -Polymorphic Instance eq_context_refl eq_term leq_term : - CRelationClasses.Reflexive eq_term -> - CRelationClasses.Reflexive leq_term -> - CRelationClasses.Reflexive (eq_context_gen eq_term leq_term). +Polymorphic Instance eq_context_refl cmp_term pb : + CRelationClasses.Reflexive (cmp_term Conv) -> + CRelationClasses.Reflexive (cmp_term pb) -> + CRelationClasses.Reflexive (eq_context_gen cmp_term pb). Proof. intros heq hle x. eapply All2_fold_refl. @@ -248,10 +347,10 @@ Proof. Qed. #[global] -Polymorphic Instance eq_context_sym eq_term leq_term : - CRelationClasses.Symmetric eq_term -> - CRelationClasses.Symmetric leq_term -> - CRelationClasses.Symmetric (eq_context_gen eq_term leq_term). +Polymorphic Instance eq_context_sym cmp_term pb : + CRelationClasses.Symmetric (cmp_term Conv) -> + CRelationClasses.Symmetric (cmp_term pb) -> + CRelationClasses.Symmetric (eq_context_gen cmp_term pb). Proof. intros heq hle x. eapply All2_fold_sym. @@ -259,21 +358,36 @@ Proof. Qed. #[global] -Polymorphic Instance eq_context_trans eq_term leq_term : - CRelationClasses.Transitive eq_term -> - CRelationClasses.Transitive leq_term -> - CRelationClasses.Transitive (eq_context_gen eq_term leq_term). +Polymorphic Instance eq_context_trans cmp_term pb : + CRelationClasses.Transitive (cmp_term Conv) -> + CRelationClasses.Transitive (cmp_term pb) -> + CRelationClasses.Transitive (eq_context_gen cmp_term pb). Proof. intros hr x y z. eapply All2_fold_trans; intros. now transitivity y0. Qed. -Definition eq_predicate (eq_term : term -> term -> Type) Re p p' := - All2 eq_term p.(pparams) p'.(pparams) * - (R_universe_instance Re p.(puinst) p'.(puinst) * - ((eq_context_gen eq eq p.(pcontext) p'.(pcontext)) * - eq_term p.(preturn) p'.(preturn))). +Definition eq_predicate (eq_term : term -> term -> Type) eq_universe p p' := + All2 eq_term p.(pparams) p'.(pparams) × + cmp_universe_instance eq_universe p.(puinst) p'.(puinst) × + eq_context_upto_names p.(pcontext) p'.(pcontext) × + eq_term p.(preturn) p'.(preturn). + +Definition eq_branch (eq_term : term -> term -> Type) br br' := + eq_context_upto_names br.(bcontext) br'.(bcontext) × + eq_term br.(bbody) br'.(bbody). + +Definition eq_branches eq_term brs brs' := All2 (eq_branch eq_term) brs brs'. + +Definition eq_mfixpoint (eq_term : term -> term -> Type) mfix mfix' := + All2 (fun d d' => + eq_term d.(dtype) d'.(dtype) × + eq_term d.(dbody) d'.(dbody) × + d.(rarg) = d'.(rarg) × + eq_binder_annot d.(dname) d'.(dname) + ) mfix mfix'. + (** ** Syntactic ws_cumul_pb up-to universes We don't look at printing annotations *) @@ -285,146 +399,235 @@ Definition eq_predicate (eq_term : term -> term -> Type) Re p p' := Reserved Notation " Σ ⊢ t <==[ Rle , napp ] u" (at level 50, t, u at next level, format "Σ ⊢ t <==[ Rle , napp ] u"). -Inductive eq_term_upto_univ_napp Σ (Re Rle : Universe.t -> Universe.t -> Prop) (napp : nat) : term -> term -> Type := +Inductive eq_term_upto_univ_napp Σ + (cmp_universe : conv_pb -> Universe.t -> Universe.t -> Prop) + (cmp_sort : conv_pb -> sort -> sort -> Prop) + (pb : conv_pb) (napp : nat) : term -> term -> Type := | eq_Rel : forall n, - Σ ⊢ tRel n <==[ Rle , napp ] tRel n + Σ ⊢ tRel n <==[ pb , napp ] tRel n | eq_Evar : forall e args args', - All2 (eq_term_upto_univ_napp Σ Re Re 0) args args' -> - Σ ⊢ tEvar e args <==[ Rle , napp ] tEvar e args' + All2 (fun arg arg' => (Σ ⊢ arg <==[ Conv , 0 ] arg')) args args' -> + Σ ⊢ tEvar e args <==[ pb , napp ] tEvar e args' | eq_Var : forall id, - Σ ⊢ tVar id <==[ Rle , napp ] tVar id + Σ ⊢ tVar id <==[ pb , napp ] tVar id | eq_Sort : forall s s', - Rle s s' -> - Σ ⊢ tSort s <==[ Rle , napp ] tSort s' + cmp_sort pb s s' -> + Σ ⊢ tSort s <==[ pb , napp ] tSort s' | eq_App : forall t t' u u', - Σ ⊢ t <==[ Rle , S napp ] t' -> - Σ ⊢ u <==[ Re , 0 ] u' -> - Σ ⊢ tApp t u <==[ Rle , napp ] tApp t' u' + Σ ⊢ t <==[ pb , S napp ] t' -> + Σ ⊢ u <==[ Conv , 0 ] u' -> + Σ ⊢ tApp t u <==[ pb , napp ] tApp t' u' | eq_Const : forall c u u', - R_universe_instance Re u u' -> - Σ ⊢ tConst c u <==[ Rle , napp ] tConst c u' + cmp_universe_instance (cmp_universe Conv) u u' -> + Σ ⊢ tConst c u <==[ pb , napp ] tConst c u' | eq_Ind : forall i u u', - R_global_instance Σ Re Rle (IndRef i) napp u u' -> - Σ ⊢ tInd i u <==[ Rle , napp ] tInd i u' + cmp_global_instance Σ cmp_universe pb (IndRef i) napp u u' -> + Σ ⊢ tInd i u <==[ pb , napp ] tInd i u' | eq_Construct : forall i k u u', - R_global_instance Σ Re Rle (ConstructRef i k) napp u u' -> - Σ ⊢ tConstruct i k u <==[ Rle , napp ] tConstruct i k u' + cmp_global_instance Σ cmp_universe pb (ConstructRef i k) napp u u' -> + Σ ⊢ tConstruct i k u <==[ pb , napp ] tConstruct i k u' | eq_Lambda : forall na na' ty ty' t t', eq_binder_annot na na' -> - Σ ⊢ ty <==[ Re , 0 ] ty' -> - Σ ⊢ t <==[ Rle , 0 ] t' -> - Σ ⊢ tLambda na ty t <==[ Rle , napp ] tLambda na' ty' t' + Σ ⊢ ty <==[ Conv , 0 ] ty' -> + Σ ⊢ t <==[ pb , 0 ] t' -> + Σ ⊢ tLambda na ty t <==[ pb , napp ] tLambda na' ty' t' | eq_Prod : forall na na' a a' b b', eq_binder_annot na na' -> - Σ ⊢ a <==[ Re , 0 ] a' -> - Σ ⊢ b <==[ Rle , 0 ] b' -> - Σ ⊢ tProd na a b <==[ Rle , napp ] tProd na' a' b' + Σ ⊢ a <==[ Conv , 0 ] a' -> + Σ ⊢ b <==[ pb , 0 ] b' -> + Σ ⊢ tProd na a b <==[ pb , napp ] tProd na' a' b' | eq_LetIn : forall na na' t t' ty ty' u u', eq_binder_annot na na' -> - Σ ⊢ t <==[ Re , 0 ] t' -> - Σ ⊢ ty <==[ Re , 0 ] ty' -> - Σ ⊢ u <==[ Rle , 0 ] u' -> - Σ ⊢ tLetIn na t ty u <==[ Rle , napp ] tLetIn na' t' ty' u' + Σ ⊢ t <==[ Conv , 0 ] t' -> + Σ ⊢ ty <==[ Conv , 0 ] ty' -> + Σ ⊢ u <==[ pb , 0 ] u' -> + Σ ⊢ tLetIn na t ty u <==[ pb , napp ] tLetIn na' t' ty' u' | eq_Case : forall indn p p' c c' brs brs', - eq_predicate (eq_term_upto_univ_napp Σ Re Re 0) Re p p' -> - Σ ⊢ c <==[ Re , 0 ] c' -> - All2 (fun x y => - eq_context_gen eq eq (bcontext x) (bcontext y) * - (Σ ⊢ x.(bbody) <==[ Re , 0 ] y.(bbody)) - ) brs brs' -> - Σ ⊢ tCase indn p c brs <==[ Rle , napp ] tCase indn p' c' brs' + eq_predicate (fun t t' => Σ ⊢ t <==[ Conv , 0 ] t') (cmp_universe Conv) p p' -> + Σ ⊢ c <==[ Conv , 0 ] c' -> + eq_branches (fun t t' => Σ ⊢ t <==[ Conv , 0 ] t') brs brs' -> + Σ ⊢ tCase indn p c brs <==[ pb , napp ] tCase indn p' c' brs' | eq_Proj : forall p c c', - Σ ⊢ c <==[ Re , 0 ] c' -> - Σ ⊢ tProj p c <==[ Rle , napp ] tProj p c' + Σ ⊢ c <==[ Conv , 0 ] c' -> + Σ ⊢ tProj p c <==[ pb , napp ] tProj p c' | eq_Fix : forall mfix mfix' idx, - All2 (fun x y => - (Σ ⊢ x.(dtype) <==[ Re , 0 ] y.(dtype)) * - (Σ ⊢ x.(dbody) <==[ Re , 0 ] y.(dbody)) * - (x.(rarg) = y.(rarg)) * - eq_binder_annot x.(dname) y.(dname) - )%type mfix mfix' -> - Σ ⊢ tFix mfix idx <==[ Rle , napp ] tFix mfix' idx + eq_mfixpoint (fun t t' => Σ ⊢ t <==[ Conv , 0 ] t') mfix mfix' -> + Σ ⊢ tFix mfix idx <==[ pb , napp ] tFix mfix' idx | eq_CoFix : forall mfix mfix' idx, - All2 (fun x y => - (Σ ⊢ x.(dtype) <==[ Re , 0 ] y.(dtype)) * - (Σ ⊢ x.(dbody) <==[ Re , 0 ] y.(dbody)) * - (x.(rarg) = y.(rarg)) * - eq_binder_annot x.(dname) y.(dname) - ) mfix mfix' -> - Σ ⊢ tCoFix mfix idx <==[ Rle , napp ] tCoFix mfix' idx + eq_mfixpoint (fun t t' => Σ ⊢ t <==[ Conv , 0 ] t') mfix mfix' -> + Σ ⊢ tCoFix mfix idx <==[ pb , napp ] tCoFix mfix' idx | eq_Prim i i' : - onPrims (eq_term_upto_univ_napp Σ Re Re 0) Re i i' -> - eq_term_upto_univ_napp Σ Re Rle napp (tPrim i) (tPrim i') -where " Σ ⊢ t <==[ Rle , napp ] u " := (eq_term_upto_univ_napp Σ _ Rle napp t u) : type_scope. + onPrims (fun t t' => Σ ⊢ t <==[ Conv , 0 ] t') (cmp_universe Conv) i i' -> + Σ ⊢ tPrim i <==[ pb , napp ] tPrim i' +where " Σ ⊢ t <==[ pb , napp ] u " := (eq_term_upto_univ_napp Σ _ _ pb napp t u) : type_scope. -Notation eq_term_upto_univ Σ Re Rle := (eq_term_upto_univ_napp Σ Re Rle 0). +Notation eq_term_upto_univ Σ cmp_universe cmp_sort pb := (eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb 0) (only parsing). (* ** Syntactic conparison up-to universes *) -Definition compare_term_napp `{checker_flags} (pb : conv_pb) Σ φ napp (t u : term) := - eq_term_upto_univ_napp Σ (eq_universe φ) (compare_universe pb φ) napp t u. +Definition compare_term_napp `{checker_flags} Σ φ (pb : conv_pb) napp (t u : term) := + eq_term_upto_univ_napp Σ (compare_universe φ) (compare_sort φ) pb napp t u. -Definition compare_term `{checker_flags} (pb : conv_pb) Σ φ (t u : term) := - eq_term_upto_univ Σ (eq_universe φ) (compare_universe pb φ) t u. +Definition compare_term `{checker_flags} Σ φ (pb : conv_pb) (t u : term) := + eq_term_upto_univ Σ (compare_universe φ) (compare_sort φ) pb t u. (* ** Syntactic conversion up-to universes *) -Notation eq_term := (compare_term Conv). +Notation eq_term Σ φ := (compare_term Σ φ Conv). (* ** Syntactic cumulativity up-to universes *) -Notation leq_term := (compare_term Cumul). +Notation leq_term Σ φ := (compare_term Σ φ Cumul). Definition compare_opt_term `{checker_flags} (pb : conv_pb) Σ φ (t u : option term) := match t, u with - | Some t, Some u => compare_term pb Σ φ t u + | Some t, Some u => compare_term Σ φ pb t u | None, None => True | _, _ => False end. -Definition compare_decl `{checker_flags} pb Σ φ (d d' : context_decl) := - compare_decls (compare_term Conv Σ φ) (compare_term pb Σ φ) d d'. +Definition compare_decl `{checker_flags} Σ φ pb (d d' : context_decl) := + compare_decls (compare_term Σ φ) pb d d'. + +Notation eq_decl Σ φ := (compare_decl Σ φ Conv). +Notation leq_decl Σ φ := (compare_decl Σ φ Cumul). + +Definition compare_context `{checker_flags} Σ φ pb (Γ Δ : context) := + eq_context_gen (compare_term Σ φ) pb Γ Δ. + +Notation eq_context Σ φ := (compare_context Σ φ Conv). +Notation leq_context Σ φ := (compare_context Σ φ Cumul). -Notation eq_decl := (compare_decl Conv). -Notation leq_decl := (compare_decl Cumul). +Notation eq_context_upto Σ cmp_universe cmp_sort pb := + (eq_context_gen (fun pb0 => eq_term_upto_univ Σ cmp_universe cmp_sort pb0) pb). -Definition compare_context `{checker_flags} pb Σ φ (Γ Δ : context) := - eq_context_gen (compare_term Conv Σ φ) (compare_term pb Σ φ) Γ Δ. +(* TODO MOVE *) +#[global] +Existing Instance All2_symP. + +(* TODO MOVE *) +#[global] +Existing Instance Forall2_symP. + +#[global] +Instance cmp_universe_instance_refl cmp_universe : + RelationClasses.Reflexive cmp_universe -> + RelationClasses.Reflexive (cmp_universe_instance cmp_universe). +Proof. + intros refl_univ u. + apply Forall2_same; reflexivity. +Qed. + +#[global] +Instance cmp_universe_instance_sym cmp_universe : + RelationClasses.Symmetric cmp_universe -> + RelationClasses.Symmetric (cmp_universe_instance cmp_universe). +Proof. intros tRe x y. now eapply Forall2_symP. Qed. + +#[global] +Instance cmp_universe_instance_trans cmp_universe : + RelationClasses.Transitive cmp_universe -> + RelationClasses.Transitive (cmp_universe_instance cmp_universe). +Proof. intros tRe x y z. eapply Forall2_trans. tc. Qed. -Notation eq_context := (compare_context Conv). -Notation leq_context := (compare_context Cumul). +#[global] +Instance cmp_universe_variance_sym cmp_universe pb v : + RelationClasses.Symmetric (cmp_universe Conv) -> + RelationClasses.Symmetric (cmp_universe pb) -> + RelationClasses.Symmetric (cmp_universe_variance cmp_universe pb v). +Proof. + intros univ_sym univ_sym' u u' u''. + destruct v as [| |] => //=. + all: now symmetry. +Qed. + +#[global] +Instance cmp_universe_variance_trans cmp_universe pb v : + RelationClasses.Transitive (cmp_universe Conv) -> + RelationClasses.Transitive (cmp_universe pb) -> + RelationClasses.Transitive (cmp_universe_variance cmp_universe pb v). +Proof. + intros univ_trans univ_trans' u u' u''. + destruct v as [| |] => //=. + all: now etransitivity. +Qed. -Notation eq_context_upto Σ Re Rle := - (eq_context_gen (eq_term_upto_univ Σ Re Re) (eq_term_upto_univ Σ Re Rle)). +#[global] +Instance cmp_universe_instance_variance_sym cmp_universe pb v : + RelationClasses.Symmetric (cmp_universe Conv) -> + RelationClasses.Symmetric (cmp_universe pb) -> + RelationClasses.Symmetric (cmp_universe_instance_variance cmp_universe pb v). +Proof. + intros univ_sym univ_sym' u u'. + apply Forall3_symP. tc. +Qed. -Lemma R_global_instance_refl Σ Re Rle gr napp u : - RelationClasses.Reflexive Re -> - RelationClasses.Reflexive Rle -> - R_global_instance Σ Re Rle gr napp u u. +#[global] +Instance cmp_universe_instance_variance_trans cmp_universe pb v : + RelationClasses.Transitive (cmp_universe Conv) -> + RelationClasses.Transitive (cmp_universe pb) -> + RelationClasses.Transitive (cmp_universe_instance_variance cmp_universe pb v). +Proof. + intros univ_trans univ_trans' u u' u''. + apply Forall3_transP. tc. +Qed. + +#[global] +Instance cmp_global_instance_refl Σ cmp_universe pb gr napp : + RelationClasses.Reflexive (cmp_universe Conv) -> + RelationClasses.Reflexive (cmp_universe pb) -> + RelationClasses.Reflexive (cmp_global_instance Σ cmp_universe pb gr napp). Proof. intros rRE rRle. - unfold R_global_instance_gen. - destruct global_variance_gen as [v|] eqn:lookup. - - induction u in v |- *; simpl; auto; - unfold R_opt_variance in IHu; destruct v; simpl; auto. - split; auto. - destruct t; simpl; auto. - - apply Forall2_same; eauto. + unfold cmp_global_instance_gen. + destruct global_variance_gen as [| |v] => //= u. + - now apply cmp_universe_instance_refl. + - left. now apply cmp_universe_instance_refl. +Qed. + +#[global] +Instance cmp_global_instance_sym Σ cmp_universe pb gr napp : + RelationClasses.Symmetric (cmp_universe Conv) -> + RelationClasses.Symmetric (cmp_universe pb) -> + RelationClasses.Symmetric (cmp_global_instance Σ cmp_universe pb gr napp). +Proof. + intros univ_sym univ_sym'. + unfold cmp_global_instance_gen. + destruct global_variance_gen as [| |v] => //= u u'. + - now symmetry. + - intros [H | H]; [left|right]. + all: now symmetry. +Qed. + +#[global] +Instance cmp_global_instance_trans Σ cmp_universe pb gr napp : + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe pb) -> + RelationClasses.Transitive (cmp_universe Conv) -> + RelationClasses.Transitive (cmp_universe pb) -> + RelationClasses.Transitive (cmp_global_instance Σ cmp_universe pb gr napp). +Proof. + intros univ_sub univ_trans univ_trans'. + unfold cmp_global_instance_gen. + destruct global_variance_gen as [| |v] => //= u u' u''. + 1,2: now etransitivity. + + apply cmp_opt_variance_or_impl; tas. + all: now etransitivity. Qed. #[global] @@ -443,49 +646,19 @@ Proof. reflexivity. Qed. #[global] Hint Resolve eq_binder_annot_refl : core. -(* TODO MOVE *) -#[global] -Existing Instance All2_symP. - -(* TODO MOVE *) -#[global] -Instance Forall2_symP : - forall A (P : A -> A -> Prop), - RelationClasses.Symmetric P -> - Symmetric (Forall2 P). -Proof. - intros A P h l l' hl. - induction hl. all: auto. -Qed. Lemma eq_binder_relevances_refl (x : list aname) : All2 (on_rel eq binder_relevance) x x. Proof. now eapply All_All2_refl, All_refl. Qed. #[global] Hint Resolve eq_binder_relevances_refl : core. -#[global] -Instance R_universe_instance_refl Re : RelationClasses.Reflexive Re -> - RelationClasses.Reflexive (R_universe_instance Re). -Proof. intros tRe x. eapply Forall2_map. - induction x; constructor; auto. -Qed. - -#[global] -Instance R_universe_instance_sym Re : RelationClasses.Symmetric Re -> - RelationClasses.Symmetric (R_universe_instance Re). -Proof. intros tRe x y. now eapply Forall2_symP. Qed. - -#[global] -Instance R_universe_instance_trans Re : RelationClasses.Transitive Re -> - RelationClasses.Transitive (R_universe_instance Re). -Proof. intros tRe x y z. now eapply Forall2_trans. Qed. - -Lemma onctx_eq_ctx P ctx eq_term : +Lemma onctx_eq_ctx P ctx cmp_term pb : onctx P ctx -> - (forall x, P x -> eq_term x x) -> - All2_fold (fun _ _ => compare_decls eq_term eq_term) ctx ctx. + (forall x, P x -> cmp_term Conv x x) -> + (forall x, P x -> cmp_term pb x x) -> + eq_context_gen cmp_term pb ctx ctx. Proof. - intros onc HP. + intros onc HP HP'. induction onc. - constructor; auto. - constructor; auto; simpl. @@ -509,26 +682,46 @@ Proof. Qed. #[global] -Polymorphic Instance eq_term_upto_univ_refl Σ Re Rle napp : - RelationClasses.Reflexive Re -> - RelationClasses.Reflexive Rle -> - Reflexive (eq_term_upto_univ_napp Σ Re Rle napp). +Polymorphic Instance eq_branches_refl Re : + CRelationClasses.Reflexive Re -> + CRelationClasses.Reflexive (eq_branches Re). Proof. - intros hRe hRle t. - induction t in napp, Rle, hRle |- * using term_forall_list_ind. + intros hre. + intros brs. unfold eq_branches, eq_branch. + apply All2_same. intuition auto; try reflexivity. +Qed. + +#[global] +Polymorphic Instance eq_mfixpoint_refl Re : + CRelationClasses.Reflexive Re -> + CRelationClasses.Reflexive (eq_mfixpoint Re). +Proof. + intros hre. + intros mfix. + apply All2_same. intuition auto; try reflexivity. +Qed. + + +#[global] +Polymorphic Instance eq_term_upto_univ_refl Σ cmp_universe cmp_sort pb napp : + RelationClasses.Reflexive (cmp_universe Conv) -> + RelationClasses.Reflexive (cmp_universe pb) -> + RelationClasses.Reflexive (cmp_sort Conv) -> + RelationClasses.Reflexive (cmp_sort pb) -> + Reflexive (eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp). +Proof. + intros refl_univ refl_univ' refl_sort refl_sort' t. + induction t in napp, pb, refl_univ', refl_sort' |- * using term_forall_list_ind. all: simpl. all: try constructor. all: eauto. + all: try reflexivity. all: try solve [eapply All_All2 ; eauto]. - all: try solve [eapply Forall2_same ; eauto]. - all: try solve [unfold eq_predicate; solve_all; eapply All_All2; eauto]. - - apply R_global_instance_refl; auto. - - apply R_global_instance_refl; auto. - - destruct X as [? [? ?]]. - unfold eq_predicate; solve_all. - eapply All_All2; eauto. reflexivity. - eapply onctx_eq_ctx in a0; eauto. - - eapply All_All2; eauto; simpl; intuition eauto. - eapply onctx_eq_ctx in a; eauto. + - unfold eq_predicate. solve_all. + 2,3: reflexivity. + eapply All_All2; eauto. + - unfold eq_branches, eq_branch. + eapply All_All2; eauto. + intros br [_ ?]; split; eauto. reflexivity. - eapply All_All2; eauto; simpl; intuition eauto. - eapply All_All2; eauto; simpl; intuition eauto. - destruct p as [? []]; constructor; cbn in X; intuition eauto. @@ -536,81 +729,59 @@ Proof. Qed. #[global] -Instance compare_term_refl {cf} pb {Σ : global_env} φ : Reflexive (compare_term pb Σ φ). +Polymorphic Instance All2_eq_refl Σ cmp_universe cmp_sort : + RelationClasses.Reflexive (cmp_universe Conv) -> + RelationClasses.Reflexive (cmp_sort Conv) -> + CRelationClasses.Reflexive (All2 (eq_term_upto_univ Σ cmp_universe cmp_sort Conv)). +Proof using Type. + intros h h' x. apply All2_same. reflexivity. +Qed. + +#[global] +Instance compare_term_refl {cf} pb {Σ : global_env} φ : Reflexive (compare_term Σ φ pb). Proof. eapply eq_term_upto_univ_refl; tc. Qed. Derive Signature for eq_term_upto_univ_napp. -Lemma R_global_instance_sym Σ Re Rle gr napp u u' : - RelationClasses.Symmetric Re -> - RelationClasses.Symmetric Rle -> - R_global_instance Σ Re Rle gr napp u' u -> - R_global_instance Σ Re Rle gr napp u u'. -Proof. - intros rRE rRle. - unfold R_global_instance_gen. - destruct global_variance_gen as [v|] eqn:lookup. - - induction u in u', v |- *; destruct u'; simpl; auto; - destruct v as [|v vs]; unfold R_opt_variance in IHu; simpl; auto. - intros [Ra Ru']. split. - destruct v; simpl; auto. - apply IHu; auto. - - apply Forall2_symP; eauto. -Qed. - -Lemma onctx_eq_ctx_sym P ctx ctx' eq_term : +Lemma onctx_eq_ctx_sym P ctx ctx' cmp_term pb : onctx P ctx -> - (forall x, P x -> forall y, eq_term x y -> eq_term y x) -> - All2_fold (fun _ _ => compare_decls eq_term eq_term) ctx ctx' -> - All2_fold (fun _ _ => compare_decls eq_term eq_term) ctx' ctx. + (forall x, P x -> forall y, cmp_term Conv x y -> cmp_term Conv y x) -> + (forall x, P x -> forall y, cmp_term pb x y -> cmp_term pb y x) -> + eq_context_gen cmp_term pb ctx ctx' -> + eq_context_gen cmp_term pb ctx' ctx. Proof. - intros onc HP H1. + intros onc HP HP' H1. induction H1; depelim onc; constructor; intuition auto; simpl in *. - depelim p; depelim o; constructor; auto; try now symmetry. + destruct p; depelim o; constructor; eauto; now symmetry. Qed. #[global] -Polymorphic Instance eq_term_upto_univ_sym Σ Re Rle napp : - RelationClasses.Symmetric Re -> - RelationClasses.Symmetric Rle -> - Symmetric (eq_term_upto_univ_napp Σ Re Rle napp). -Proof. - intros he hle u v e. +Polymorphic Instance eq_term_upto_univ_sym Σ cmp_universe cmp_sort pb napp : + RelationClasses.Symmetric (cmp_universe Conv) -> + RelationClasses.Symmetric (cmp_universe pb) -> + RelationClasses.Symmetric (cmp_sort Conv) -> + RelationClasses.Symmetric (cmp_sort pb) -> + Symmetric (eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp). +Proof. + intros sym_univ sym_univ' sym_sort sym_sort' u v e. pose proof (@RelationClasses.symmetry _ (@eq_binder_annot name name) _). - induction u in Rle, hle, v, napp, e |- * using term_forall_list_ind. + induction u in napp, pb, sym_univ', sym_sort', v, e |- * using term_forall_list_ind. all: dependent destruction e. all: try solve [ - econstructor ; eauto using R_global_instance_sym ; - try eapply Forall2_symP ; eauto + econstructor ; eauto ; + try eapply Forall2_symP ; eauto ; easy ]. - econstructor. - eapply All2_All_mix_left in X as h; eauto. - clear a X. - induction h. - + constructor. - + destruct r as [h1 h2]. eapply h1 in h2 ; auto. + apply All2_sym. solve_all. - solve_all. destruct e as (r & ? & eq & ?). - econstructor; rewrite ?e; unfold eq_predicate in *; solve_all; eauto. - eapply All2_sym; solve_all. - unfold R_universe_instance in r |- *. - eapply Forall2_symP; eauto. - eapply onctx_eq_ctx_sym in a1; eauto. - eapply All2_sym; solve_all. - eapply onctx_eq_ctx_sym in a0; eauto. - - econstructor. - eapply All2_All_mix_left in X as h; eauto. - clear a X. - induction h. - + constructor. - + destruct r as [[h1 h2] [[[h3 h4] h5] h6]]. - eapply h1 in h3; auto. constructor; auto. - - econstructor. - eapply All2_All_mix_left in X as h; eauto. - clear a X. - induction h. - + constructor. - + destruct r as [[h1 h2] [[[h3 h4] h5] h6]]. eapply h1 in h3 ; auto. - constructor; auto. + econstructor; rewrite ?e; unfold eq_predicate, eq_branches, eq_branch in *; repeat split; eauto. + 2,3: now symmetry. + all: eapply All2_sym; solve_all. + apply All2_symP; solve_all. tc. + - econstructor. unfold eq_mfixpoint in *. + apply All2_sym. solve_all. + - econstructor. unfold eq_mfixpoint in *. + apply All2_sym. solve_all. - econstructor. depelim o; cbn in X; constructor; intuition eauto. eapply All2_All_mix_left in a0 as h; eauto. cbn in h. @@ -628,38 +799,20 @@ Proof. Qed. #[global] -Instance compare_term_sym {cf} Σ φ : Symmetric (compare_term Conv Σ φ). +Instance compare_term_sym {cf} Σ φ : Symmetric (compare_term Σ φ Conv). Proof. now intros t u; unfold compare_term; cbn; symmetry. Qed. -#[global] -Instance R_global_instance_trans Σ Re Rle gr napp : - RelationClasses.Transitive Re -> - RelationClasses.Transitive Rle -> - RelationClasses.Transitive (R_global_instance Σ Re Rle gr napp). -Proof. - intros he hle x y z. - unfold R_global_instance_gen, R_opt_variance. - destruct global_variance_gen as [v|]. - clear -he hle. - induction x in y, z, v |- *; destruct y, z, v; simpl; auto => //. eauto. - intros [Ra Rxy] [Rt Ryz]. - split; eauto. - destruct t1; simpl in *; auto. - now etransitivity; eauto. - now etransitivity; eauto. - eapply Forall2_trans; auto. -Qed. - -Lemma onctx_eq_ctx_trans P ctx ctx' ctx'' eq_term : +Lemma onctx_eq_ctx_trans P ctx ctx' ctx'' cmp_term pb : onctx P ctx -> - (forall x, P x -> forall y z, eq_term x y -> eq_term y z -> eq_term x z) -> - All2_fold (fun _ _ => compare_decls eq_term eq_term) ctx ctx' -> - All2_fold (fun _ _ => compare_decls eq_term eq_term) ctx' ctx'' -> - All2_fold (fun _ _ => compare_decls eq_term eq_term) ctx ctx''. + (forall x, P x -> forall y z, cmp_term Conv x y -> cmp_term Conv y z -> cmp_term Conv x z) -> + (forall x, P x -> forall y z, cmp_term pb x y -> cmp_term pb y z -> cmp_term pb x z) -> + eq_context_gen cmp_term pb ctx ctx' -> + eq_context_gen cmp_term pb ctx' ctx'' -> + eq_context_gen cmp_term pb ctx ctx''. Proof. - intros onc HP H1; revert ctx''. + intros onc HP HP' H1; revert ctx''. induction H1; depelim onc; intros ctx'' H; depelim H; constructor; intuition auto; simpl in *. depelim o. depelim p0. - depelim c; constructor; [now etransitivity|eauto]. @@ -679,104 +832,76 @@ Proof. Qed. #[global] -Polymorphic Instance eq_term_upto_univ_trans Σ Re Rle napp : - RelationClasses.Transitive Re -> - RelationClasses.Transitive Rle -> - Transitive (eq_term_upto_univ_napp Σ Re Rle napp). -Proof. - intros he hle u v w e1 e2. +Polymorphic Instance eq_term_upto_univ_trans Σ cmp_universe cmp_sort pb napp : + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe pb) -> + RelationClasses.Transitive (cmp_universe Conv) -> + RelationClasses.Transitive (cmp_universe pb) -> + RelationClasses.Transitive (cmp_sort Conv) -> + RelationClasses.Transitive (cmp_sort pb) -> + Transitive (eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp). +Proof. + intros sub_univ trans_univ trans_univ' trans_sort trans_sort' u v w e1 e2. pose proof (@RelationClasses.transitivity _ (@eq_binder_annot name name) _). - induction u in Rle, hle, v, w, napp, e1, e2 |- * using term_forall_list_ind. + assert (RelationClasses.subrelation (cmp_universe Conv) (cmp_universe Conv)) by reflexivity. + induction u in napp, pb, sub_univ, trans_univ', trans_sort', v, w, e1, e2 |- * using term_forall_list_ind. all: dependent destruction e1. all: try solve [ eauto ]. all: try solve [ - dependent destruction e2 ; econstructor ; eauto; - try eapply Forall2_trans ; eauto + dependent destruction e2 ; econstructor ; eauto ; + try now etransitivity ]. - dependent destruction e2. econstructor. eapply All2_All_mix_left in X as h; eauto. - clear a X. - induction h in a0, args'0 |- *. - + assumption. - + dependent destruction a0. constructor ; eauto. - destruct r as [h1 h2]. eauto. - - dependent destruction e2. - constructor. - eapply R_global_instance_trans; eauto. + eapply All2_trans'; tea. + intros u v w [[IH]]. now eapply IH. - dependent destruction e2. - constructor. - eapply R_global_instance_trans; eauto. - - dependent destruction e2. - unfold eq_predicate in *. + unfold eq_predicate, eq_branches, eq_branch in *. !!solve_all. - econstructor; unfold eq_predicate in *; solve_all; eauto. - * clear -he hh1 hh2. - revert hh1 hh2. generalize (pparams p), p'.(pparams), p'0.(pparams). - intros l l' l''. - induction 1 in l'' |- *; auto. intros H; depelim H. constructor; auto. - eapply r; eauto. apply r. - * etransitivity; eauto. - * eapply onctx_eq_ctx_trans in hh; eauto. - intros ???? -> ->; eauto. - * clear -H he a a0. - induction a in a0, brs'0 |- *. - + assumption. - + depelim a0. destruct p. constructor; auto. - destruct r as [[h1 h1'] [h2 h3]]. - split. - eapply onctx_eq_ctx_trans in h1; eauto. - intros ???? -> ->;reflexivity. - eapply h1'; eauto. + econstructor; unfold eq_predicate, eq_branches, eq_branch in *; solve_all; eauto. + 2: now etransitivity. + all: eapply All2_trans'; tea; intros ??? [[IH]]; repeat split; eauto. + * etransitivity; eassumption. + * destruct p0, p1; etransitivity; eassumption. + * destruct IH, p0, p1; eauto. - dependent destruction e2. - econstructor. + econstructor. unfold eq_mfixpoint in *. eapply All2_All_mix_left in X as h; eauto. - clear a X. - induction h in a0, mfix'0 |- *. - + assumption. - + dependent destruction a0. constructor ; eauto. - intuition eauto. - transitivity (rarg y); auto. + eapply All2_trans'; tea. + intros u v w [[[IHt IHb] (?&?&?&?)] (?&?&?&?)]. repeat split; eauto. now etransitivity. - dependent destruction e2. - econstructor. + econstructor. unfold eq_mfixpoint in *. eapply All2_All_mix_left in X as h; eauto. - clear a X. - induction h in a0, mfix'0 |- *. - + assumption. - + dependent destruction a0. constructor ; eauto. - intuition eauto. - transitivity (rarg y); auto. + eapply All2_trans'; tea. + intros u v w [[[IHt IHb] (?&?&?&?)] (?&?&?&?)]. repeat split; eauto. now etransitivity. - dependent destruction e2; constructor. - depelim o; intuition eauto. depelim o0; constructor; cbn in X; intuition eauto. - eapply All2_All_mix_left in b0 as h; eauto. - clear b0 a0. clear -he hle a2 h. revert h a2. - generalize (array_value a) (array_value a') (array_value a'0). clear -he hle. - intros l l0 l1. - induction 1 in l1 |- *; intros. - + assumption. - + dependent destruction a2. constructor ; eauto. - destruct r as [h1 h2]. eauto. + depelim o; tas. depelim o0. constructor; destruct X as (hty & hdef & harr); eauto. + eapply All2_All_mix_left in harr as h; eauto. + eapply All2_trans'; tea. + intros ??? [[IH]]; repeat split; eauto. Qed. #[global] -Instance compare_term_trans {cf} pb Σ φ : Transitive (compare_term pb Σ φ). +Instance compare_term_trans {cf} pb Σ φ : Transitive (compare_term Σ φ pb). Proof. apply eq_term_upto_univ_trans; tc. Qed. #[global] -Polymorphic Instance eq_term_upto_univ_equiv Σ Re (hRe : RelationClasses.Equivalence Re) - : Equivalence (eq_term_upto_univ Σ Re Re). +Polymorphic Instance eq_term_upto_univ_equiv Σ cmp_universe cmp_sort : + RelationClasses.Equivalence (cmp_universe Conv) -> + RelationClasses.Equivalence (cmp_sort Conv) -> + Equivalence (eq_term_upto_univ Σ cmp_universe cmp_sort Conv). Proof. constructor. all: exact _. Defined. #[global] -Polymorphic Instance eq_context_equiv {cf} Σ φ : Equivalence (eq_context_gen (eq_term Σ φ) (eq_term Σ φ)). +Polymorphic Instance eq_context_equiv {cf} Σ φ : Equivalence (eq_context_gen (compare_term Σ φ) Conv). Proof. constructor; try exact _. Qed. #[global] -Polymorphic Instance leq_context_preord {cf} Σ φ : PreOrder (eq_context_gen (eq_term Σ φ) (leq_term Σ φ)). +Polymorphic Instance leq_context_preord {cf} Σ φ : PreOrder (eq_context_gen (compare_term Σ φ) Cumul). Proof. constructor; try exact _. Qed. @@ -790,18 +915,18 @@ Polymorphic Instance leq_term_preorder {cf:checker_flags} Σ φ : PreOrder (leq_ Proof. split; tc. Qed. #[global] -Instance R_universe_instance_equiv R (hR : RelationClasses.Equivalence R) - : RelationClasses.Equivalence (R_universe_instance R). +Instance cmp_universe_instance_equiv R (hR : RelationClasses.Equivalence R) + : RelationClasses.Equivalence (cmp_universe_instance R). Proof. split. - intro. apply Forall2_same. reflexivity. - intros x y xy. eapply Forall2_sym, Forall2_impl; tea. now symmetry. - - intros x y z xy yz. eapply Forall2_trans; tea. apply hR. + - intros x y z xy yz. eapply Forall2_trans; tea. apply on_rel_trans. apply hR. Qed. -Lemma R_universe_instance_antisym Re Rle (hRe : RelationClasses.Equivalence Re) : - RelationClasses.Antisymmetric _ Re Rle -> - RelationClasses.Antisymmetric _ (R_universe_instance Re) (R_universe_instance Rle). +Lemma cmp_universe_instance_antisym cmp_universe pb (hE : RelationClasses.Equivalence (cmp_universe Conv)) : + RelationClasses.Antisymmetric _ (cmp_universe Conv) (cmp_universe pb) -> + RelationClasses.Antisymmetric _ (cmp_universe_instance (cmp_universe Conv)) (cmp_universe_instance (cmp_universe pb)). Proof. intros H x y H1 H2. eapply Forall2_sym in H2. @@ -810,33 +935,59 @@ Proof. Qed. #[global] -Instance R_global_instance_equiv Σ R (hR : RelationClasses.Equivalence R) gr napp - : RelationClasses.Equivalence (R_global_instance Σ R R gr napp). +Instance cmp_global_instance_equiv Σ cmp_universe (hE : RelationClasses.Equivalence (cmp_universe Conv)) gr napp + : RelationClasses.Equivalence (cmp_global_instance Σ cmp_universe Conv gr napp). Proof. split. - - intro. apply R_global_instance_refl; typeclasses eauto. - - intros x y xy. eapply R_global_instance_sym; auto; typeclasses eauto. - - intros x y z xy yz. eapply R_global_instance_trans; eauto; typeclasses eauto. + - intro. apply cmp_global_instance_refl; typeclasses eauto. + - intros x y xy. eapply cmp_global_instance_sym; auto; typeclasses eauto. + - intros x y z xy yz. eapply cmp_global_instance_trans; eauto; typeclasses eauto. +Qed. + +Lemma cmp_universe_variance_antisym cmp_universe pb (hRe : RelationClasses.Equivalence (cmp_universe Conv)) v u u' : + RelationClasses.Antisymmetric _ (cmp_universe Conv) (cmp_universe pb) -> + cmp_universe_variance cmp_universe pb v u u' -> + cmp_universe_variance cmp_universe pb v u' u -> + cmp_universe_variance cmp_universe Conv v u u'. +Proof. + intro hAntisym. + destruct v => //=. + apply hAntisym. +Qed. + +Lemma cmp_universe_instance_variance_antisym cmp_universe pb (hRe : RelationClasses.Equivalence (cmp_universe Conv)) l u v : + RelationClasses.Antisymmetric _ (cmp_universe Conv) (cmp_universe pb) -> + cmp_universe_instance_variance cmp_universe pb l u v -> + cmp_universe_instance_variance cmp_universe pb l v u -> + cmp_universe_instance_variance cmp_universe Conv l u v. +Proof. + intro hAntisym. + apply Forall3_antisymP. intros ???. + now eapply cmp_universe_variance_antisym. Qed. #[global] -Instance R_global_instance_antisym Σ Re Rle (hRe : RelationClasses.Equivalence Re) gr napp : - RelationClasses.Antisymmetric _ Re Rle -> - RelationClasses.Antisymmetric _ (R_global_instance Σ Re Re gr napp) (R_global_instance Σ Re Rle gr napp). +Instance cmp_global_instance_antisym Σ cmp_universe pb (hRe : RelationClasses.Equivalence (cmp_universe Conv)) gr napp : + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe pb) -> + RelationClasses.Antisymmetric _ (cmp_universe Conv) (cmp_universe pb) -> + RelationClasses.Antisymmetric _ (cmp_global_instance Σ cmp_universe Conv gr napp) (cmp_global_instance Σ cmp_universe pb gr napp). Proof. - intros hR u v. - unfold R_global_instance_gen, R_opt_variance. + intros hsub hR u v. + unfold cmp_global_instance_gen, cmp_opt_variance. destruct global_variance_gen; auto. - induction u in l, v |- *; destruct v, l; simpl; auto. - intros [at' uv] [ta vu]. split; auto. - destruct t0; simpl in *; auto. + apply cmp_opt_variance_or_impl; auto. + eapply cmp_universe_instance_variance_antisym; tea. Qed. -Lemma eq_term_upto_univ_antisym Σ Re Rle (hRe : RelationClasses.Equivalence Re) : - RelationClasses.Antisymmetric _ Re Rle -> - Antisymmetric (eq_term_upto_univ Σ Re Re) (eq_term_upto_univ Σ Re Rle). +Lemma eq_term_upto_univ_antisym Σ cmp_universe cmp_sort pb + (univ_equi : RelationClasses.Equivalence (cmp_universe Conv)) + (sort_equi : RelationClasses.Equivalence (cmp_sort Conv)) : + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe pb) -> + RelationClasses.Antisymmetric _ (cmp_universe Conv) (cmp_universe pb) -> + RelationClasses.Antisymmetric _ (cmp_sort Conv) (cmp_sort pb) -> + Antisymmetric (eq_term_upto_univ Σ cmp_universe cmp_sort Conv) (eq_term_upto_univ Σ cmp_universe cmp_sort pb). Proof. - intros hR u v. generalize 0; intros n h h'. + intros univ_sub univ_anisym sort_antisym u v. generalize 0; intros n h h'. induction u in v, n, h, h' |- * using term_forall_list_ind. all: simpl ; inversion h ; subst; inversion h' ; subst ; try constructor ; auto. @@ -850,61 +1001,63 @@ Proof. eapply eq_term_upto_univ_antisym; exact _. Qed. -Lemma global_variance_napp_mon {Σ gr napp napp' v} : +Lemma global_variance_napp_mon Σ gr {napp napp'} : napp <= napp' -> - global_variance Σ gr napp = Some v -> - global_variance Σ gr napp' = Some v. + match global_variance Σ gr napp with + | Variance v => global_variance Σ gr napp' = Variance v + | AllEqual => True + | AllIrrelevant => global_variance Σ gr napp' = AllIrrelevant + end. Proof. intros hnapp. rewrite /global_variance_gen. - destruct gr; try congruence. - - destruct lookup_inductive_gen as [[mdecl idec]|] => //. - destruct destArity as [[ctx s]|] => //. - elim: Nat.leb_spec => // cass indv. + destruct gr => //=. + - destruct lookup_inductive_gen as [[mdecl idec]|] => //=. + destruct destArity as [[ctx s]|] => //=. + elim: Nat.leb_spec => // cass. + destruct ind_variance => //=. elim: Nat.leb_spec => //. lia. - - destruct lookup_constructor_gen as [[[mdecl idecl] cdecl]|] => //. - elim: Nat.leb_spec => // cass indv. + - destruct lookup_constructor_gen as [[[mdecl idecl] cdecl]|] => //=. + elim: Nat.leb_spec => // cass. elim: Nat.leb_spec => //. lia. Qed. #[global] -Instance R_global_instance_impl_same_napp Σ Re Re' Rle Rle' gr napp : - RelationClasses.subrelation Re Re' -> - RelationClasses.subrelation Rle Rle' -> - subrelation (R_global_instance Σ Re Rle gr napp) (R_global_instance Σ Re' Rle' gr napp). +Instance cmp_global_instance_impl_same_napp Σ cmp_universe cmp_universe' pb pb' gr napp : + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe' Conv) -> + RelationClasses.subrelation (cmp_universe pb) (cmp_universe' pb') -> + subrelation (cmp_global_instance Σ cmp_universe pb gr napp) (cmp_global_instance Σ cmp_universe' pb' gr napp). Proof. - intros he hle t t'. - unfold R_global_instance_gen, R_opt_variance. - destruct global_variance_gen as [v|] eqn:glob. - induction t in v, t' |- *; destruct v, t'; simpl; auto. - intros []; split; auto. - destruct t0; simpl; auto. - now eapply R_universe_instance_impl'. + intros sub_conv sub_pb u u'. + unfold cmp_global_instance_gen, cmp_opt_variance. + destruct global_variance_gen as [| |v] => //. + 1: now apply cmp_universe_instance_impl. + + intros [H | H]; [left | right]. + 1: eapply cmp_universe_instance_impl; tea. + eapply cmp_universe_instance_variance_impl; eassumption. Qed. #[global] -Instance R_global_instance_impl Σ Re Re' Rle Rle' gr napp napp' : - RelationClasses.subrelation Re Re' -> - RelationClasses.subrelation Re Rle' -> - RelationClasses.subrelation Rle Rle' -> +Instance cmp_global_instance_impl Σ cmp_universe cmp_universe' pb pb' gr napp napp' : + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe' Conv) -> + RelationClasses.subrelation (cmp_universe pb) (cmp_universe' pb') -> napp <= napp' -> - subrelation (R_global_instance Σ Re Rle gr napp) (R_global_instance Σ Re' Rle' gr napp'). + subrelation (cmp_global_instance Σ cmp_universe pb gr napp) (cmp_global_instance Σ cmp_universe' pb' gr napp'). Proof. - intros he hle hele hnapp t t'. - unfold R_global_instance_gen, R_opt_variance. - destruct global_variance_gen as [v|] eqn:glob. - rewrite (global_variance_napp_mon hnapp glob). - induction t in v, t' |- *; destruct v, t'; simpl; auto. - intros []; split; auto. - destruct t0; simpl; auto. - destruct (global_variance _ _ napp') as [v|] eqn:glob'; eauto using R_universe_instance_impl'. - induction t in v, t' |- *; destruct v, t'; simpl; auto; intros H; inv H. - eauto. - split; auto. - destruct t0; simpl; auto. + intros sub_conv sub_pb lenapp u u'. + unfold cmp_global_instance_gen. + move: (global_variance_napp_mon Σ gr lenapp). + destruct global_variance_gen as [| |v] => //. + all: [> intros _ | intros ->; cbn ..]; auto. + 1: intro H; apply: cmp_instance_opt_variance; move: H => /=. + - now apply cmp_universe_instance_impl. + - intros [H | H]; [left | right]. + 1: eapply cmp_universe_instance_impl; tea. + eapply cmp_universe_instance_variance_impl; eassumption. Qed. -Lemma global_variance_empty gr napp env : env.(declarations) = [] -> global_variance env gr napp = None. +Lemma global_variance_empty gr napp env : env.(declarations) = [] -> global_variance env gr napp = AllEqual. Proof. destruct env; cbn => ->. destruct gr; auto. Qed. @@ -912,106 +1065,104 @@ Qed. (** Pure syntactic equality, without cumulative inductive types subtyping *) #[global] -Instance R_global_instance_empty_impl Σ Re Re' Rle Rle' gr napp napp' : - RelationClasses.subrelation Re Re' -> - RelationClasses.subrelation Rle Rle' -> - RelationClasses.subrelation Re Rle' -> - subrelation (R_global_instance empty_global_env Re Rle gr napp) (R_global_instance Σ Re' Rle' gr napp'). +Instance cmp_global_instance_empty_impl Σ cmp_universe cmp_universe' pb pb' gr napp napp' : + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe' Conv) -> + subrelation (cmp_global_instance empty_global_env cmp_universe pb gr napp) (cmp_global_instance Σ cmp_universe' pb' gr napp'). Proof. - intros he hle hele t t'. - unfold R_global_instance_gen, R_opt_variance. + intros he t t'. + unfold cmp_global_instance_gen. rewrite global_variance_empty //. - destruct global_variance_gen as [v|]; eauto using R_universe_instance_impl'. - induction t in v, t' |- *; destruct v, t'; simpl; intros H; inv H; auto. - simpl. - split; auto. - destruct t0; simpl; auto. + intro H; apply: cmp_instance_opt_variance; move: H => /=. + now apply cmp_universe_instance_impl. Qed. -Lemma onctx_eq_ctx_impl P ctx ctx' eq_term eq_term' : +Lemma onctx_eq_ctx_impl P ctx ctx' cmp_term cmp_term' pb pb' : onctx P ctx -> - (forall x, P x -> forall y, eq_term x y -> eq_term' x y) -> - eq_context_gen eq_term eq_term ctx ctx' -> - eq_context_gen eq_term' eq_term' ctx ctx'. + (forall x, P x -> forall y, cmp_term Conv x y -> cmp_term' Conv x y) -> + (forall x, P x -> forall y, cmp_term pb x y -> cmp_term' pb' x y) -> + eq_context_gen cmp_term pb ctx ctx' -> + eq_context_gen cmp_term' pb' ctx ctx'. Proof. - intros onc HP H1. + intros onc HP HP' H1. induction H1; depelim onc; constructor; eauto; intuition auto; simpl in *. destruct o; depelim p; constructor; auto. Qed. #[global] -Instance eq_term_upto_univ_impl Σ Re Re' Rle Rle' napp napp' : - RelationClasses.subrelation Re Re' -> - RelationClasses.subrelation Rle Rle' -> - RelationClasses.subrelation Re Rle' -> +Instance eq_term_upto_univ_impl Σ cmp_universe cmp_universe' cmp_sort cmp_sort' pb pb' napp napp' : + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe' Conv) -> + RelationClasses.subrelation (cmp_universe pb) (cmp_universe' pb') -> + RelationClasses.subrelation (cmp_sort Conv) (cmp_sort' Conv) -> + RelationClasses.subrelation (cmp_sort pb) (cmp_sort' pb') -> napp <= napp' -> - subrelation (eq_term_upto_univ_napp Σ Re Rle napp) (eq_term_upto_univ_napp Σ Re' Rle' napp'). + subrelation (eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp) (eq_term_upto_univ_napp Σ cmp_universe' cmp_sort' pb' napp'). Proof. - intros he hle hele hnapp t t'. - induction t in napp, napp', hnapp, t', Rle, Rle', hle, hele |- * using term_forall_list_ind; + intros univ_sub_conv univ_sub_pb sort_sub_conv sort_sub_pb hnapp t t'. + induction t in napp, napp', hnapp, pb, pb', univ_sub_pb, sort_sub_pb, t' |- * using term_forall_list_ind; try (inversion 1; subst; constructor; - eauto using R_universe_instance_impl'; fail). + eauto using cmp_universe_instance_impl'; fail). - inversion 1; subst; constructor. eapply All2_impl'; tea. eapply All_impl; eauto. - inversion 1; subst; constructor. eapply IHt1. 4:eauto. all:auto with arith. eauto. - inversion 1; subst; constructor. - eapply R_global_instance_impl. 5:eauto. all:auto. + eapply cmp_global_instance_impl. 4:eauto. all:auto. - inversion 1; subst; constructor. - eapply R_global_instance_impl. 5:eauto. all:eauto. - - inversion 1; subst; constructor; unfold eq_predicate in *; eauto; solve_all. - * eapply R_universe_instance_impl'; eauto. + eapply cmp_global_instance_impl. 4:eauto. all:eauto. + - inversion 1; subst; constructor; unfold eq_predicate, eq_branches, eq_branch in *; eauto; solve_all. + * eapply cmp_universe_instance_impl'; eauto. - inversion 1; subst; constructor. eapply All2_impl'; tea. eapply All_impl; eauto. - cbn. intros x [? ?] y [[[? ?] ?] ?]. repeat split; eauto. + cbn. intros x [? ?] y (?&?&?&?). repeat split; eauto. - inversion 1; subst; constructor. eapply All2_impl'; tea. eapply All_impl; eauto. - cbn. intros x [? ?] y [[[? ?] ?] ?]. repeat split; eauto. + cbn. intros x [? ?] y (?&?&?&?). repeat split; eauto. - intros h; depelim h. depelim o; constructor; cbn in X; constructor; intuition eauto. solve_all. Qed. #[global] -Instance eq_term_upto_univ_empty_impl Σ Re Re' Rle Rle' napp napp' : - RelationClasses.subrelation Re Re' -> - RelationClasses.subrelation Rle Rle' -> - RelationClasses.subrelation Re Rle' -> - subrelation (eq_term_upto_univ_napp empty_global_env Re Rle napp) (eq_term_upto_univ_napp Σ Re' Rle' napp'). -Proof. - intros he hle hele t t'. - induction t in napp, napp', t', Rle, Rle', hle, hele |- * using term_forall_list_ind; +Instance eq_term_upto_univ_empty_impl Σ cmp_universe cmp_universe' cmp_sort cmp_sort' pb pb' napp napp' : + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe' Conv) -> + RelationClasses.subrelation (cmp_universe pb) (cmp_universe' pb') -> + RelationClasses.subrelation (cmp_sort Conv) (cmp_sort' Conv) -> + RelationClasses.subrelation (cmp_sort pb) (cmp_sort' pb') -> + subrelation (eq_term_upto_univ_napp empty_global_env cmp_universe cmp_sort pb napp) (eq_term_upto_univ_napp Σ cmp_universe' cmp_sort' pb' napp'). +Proof. + intros univ_sub_conv univ_sub_pb sort_sub_conv sort_sub_pb t t'. + induction t in napp, napp', pb, pb', univ_sub_pb, sort_sub_pb, t' |- * using term_forall_list_ind; try (inversion 1; subst; constructor; - eauto using R_universe_instance_impl'; fail). + eauto using cmp_universe_instance_impl'; fail). - inversion 1; subst; constructor. eapply All2_impl'; tea. eapply All_impl; eauto. - inversion 1; subst; constructor. - (* eapply shelf bug... fixed in unifall *) - eapply R_global_instance_empty_impl. 4:eauto. all:eauto. + eapply cmp_global_instance_empty_impl. 2:eauto. all:auto. - inversion 1; subst; constructor. - eapply R_global_instance_empty_impl. 4:eauto. all:eauto. - - inversion 1; subst; constructor; unfold eq_predicate in *; solve_all. - * eapply R_universe_instance_impl'; eauto. + eapply cmp_global_instance_empty_impl. 2:eauto. all:eauto. + - inversion 1; subst; constructor; unfold eq_predicate, eq_branches, eq_branch in *; eauto; solve_all. + * eapply cmp_universe_instance_impl'; eauto. - inversion 1; subst; constructor. eapply All2_impl'; tea. eapply All_impl; eauto. - cbn. intros x [? ?] y [[[? ?] ?] ?]. repeat split; eauto. + cbn. intros x [? ?] y (?&?&?&?). repeat split; eauto. - inversion 1; subst; constructor. eapply All2_impl'; tea. eapply All_impl; eauto. - cbn. intros x [? ?] y [[[? ?] ?] ?]. repeat split; eauto. - - intros h; depelim h. constructor. depelim o; cbn in X; constructor; intuition eauto. + cbn. intros x [? ?] y (?&?&?&?). repeat split; eauto. + - inversion 1; subst; constructor. depelim X1; cbn in X; constructor; intuition eauto. solve_all. Qed. #[global] -Instance eq_term_upto_univ_leq Σ Re Rle napp napp' : - RelationClasses.subrelation Re Rle -> +Instance eq_term_upto_univ_leq Σ cmp_universe cmp_sort pb napp napp' : + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe pb) -> + RelationClasses.subrelation (cmp_sort Conv) (cmp_sort pb) -> napp <= napp' -> - subrelation (eq_term_upto_univ_napp Σ Re Re napp) (eq_term_upto_univ_napp Σ Re Rle napp'). + subrelation (eq_term_upto_univ_napp Σ cmp_universe cmp_sort Conv napp) (eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp'). Proof. intros H. eapply eq_term_upto_univ_impl; exact _. Qed. @@ -1033,39 +1184,56 @@ Proof. Qed. #[global] -Hint Constructors compare_decls : pcuic. +Hint Constructors PCUICConversion.All_decls_alpha_pb : pcuic. Local Ltac lih := lazymatch goal with - | ih : forall Rle v n k, eq_term_upto_univ_napp _ _ _ ?u _ -> _ - |- eq_term_upto_univ _ _ (lift _ _ ?u) _ => + | ih : forall Rle v n k, eq_term_upto_univ_napp _ _ _ _ ?u _ -> _ + |- eq_term_upto_univ _ _ _ (lift _ _ ?u) _ => eapply ih end. -Lemma eq_term_upto_univ_lift Σ Re Rle n n' k : - Proper (eq_term_upto_univ_napp Σ Re Rle n' ==> eq_term_upto_univ_napp Σ Re Rle n') (lift n k). +Lemma eq_term_upto_univ_lift Σ cmp_universe cmp_sort pb n n' k : + Proper (eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb n' ==> eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb n') (lift n k). Proof. intros u v e. - induction u in n', v, n, k, e, Rle |- * using term_forall_list_ind. + induction u in n', v, n, k, e, pb |- * using term_forall_list_ind. all: dependent destruction e. all: try solve [cbn ; constructor ; try lih ; try assumption; solve_all]. - cbn. destruct e as (? & ? & e & ?). - constructor; unfold eq_predicate in *; simpl; !!solve_all. - * rewrite -?(All2_fold_length e). - eapply hh0; eauto. - * rewrite (All2_fold_length a). now eapply hh4. + constructor; unfold eq_predicate, eq_branches, eq_branch in *; simpl; !!solve_all. + * rewrite -?(All2_length e). + eapply hh; eauto. + * rewrite (All2_length hh3). now eapply hh2. - cbn. constructor. - pose proof (All2_length a). + pose proof (All2_length e). unfold eq_mfixpoint in *. solve_all. rewrite H. eauto. - cbn. constructor. - pose proof (All2_length a). + pose proof (All2_length e). unfold eq_mfixpoint in *. solve_all. rewrite H. eauto. - cbn. constructor. depelim o; cbn in X; try constructor; cbn; intuition eauto. solve_all. Qed. -Lemma lift_compare_term `{checker_flags} pb Σ ϕ n k t t' : - compare_term pb Σ ϕ t t' -> compare_term pb Σ ϕ (lift n k t) (lift n k t'). +Lemma compare_decls_lift_decl Σ cmp_universe cmp_sort pb n k : + Proper (compare_decls (fun pb => eq_term_upto_univ Σ cmp_universe cmp_sort pb) pb ==> compare_decls (fun pb => eq_term_upto_univ Σ cmp_universe cmp_sort pb) pb) (lift_decl n k). +Proof. + intros d d' []; constructor; cbn; auto. + all: now eapply eq_term_upto_univ_lift. +Qed. + +Lemma eq_context_upto_lift_context Σ cmp_universe cmp_sort pb n k : + Proper (eq_context_upto Σ cmp_universe cmp_sort pb ==> eq_context_upto Σ cmp_universe cmp_sort pb) (lift_context n k). +Proof. + intros Γ Δ. + induction 1; rewrite -> ?lift_context_snoc0. constructor. + constructor; auto. + rewrite -(All2_fold_length X). + now eapply compare_decls_lift_decl. +Qed. + +Lemma lift_compare_term `{checker_flags} Σ ϕ pb n k t t' : + compare_term Σ ϕ pb t t' -> compare_term Σ ϕ pb (lift n k t) (lift n k t'). Proof. now apply eq_term_upto_univ_lift. Qed. @@ -1073,18 +1241,14 @@ Qed. Lemma lift_compare_decls `{checker_flags} pb Σ ϕ n k d d' : compare_decl pb Σ ϕ d d' -> compare_decl pb Σ ϕ (lift_decl n k d) (lift_decl n k d'). Proof. - intros []; constructor; cbn; intuition auto using lift_compare_term. + now apply compare_decls_lift_decl. Qed. Lemma lift_compare_context `{checker_flags} pb Σ φ l l' n k : compare_context pb Σ φ l l' -> compare_context pb Σ φ (lift_context n k l) (lift_context n k l'). Proof. - unfold compare_context. - induction 1; rewrite -> ?lift_context_snoc0. constructor. - constructor; auto. - eapply lift_compare_decls in p. - now rewrite (All2_fold_length X). + now apply eq_context_upto_lift_context. Qed. Lemma lift_eq_term {cf:checker_flags} Σ φ n k T U : @@ -1101,19 +1265,20 @@ Qed. Local Ltac sih := lazymatch goal with - | ih : forall Rle v n x y, _ -> eq_term_upto_univ _ _ ?u _ -> _ -> _ - |- eq_term_upto_univ _ _ (subst _ _ ?u) _ => eapply ih + | ih : forall Rle v n x y, _ -> eq_term_upto_univ _ _ _ ?u _ -> _ -> _ + |- eq_term_upto_univ _ _ _ (subst _ _ ?u) _ => eapply ih end. -Lemma eq_term_upto_univ_substs Σ Re Rle napp : - RelationClasses.subrelation Re Rle -> +Lemma eq_term_upto_univ_substs Σ cmp_universe cmp_sort pb napp : + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe pb) -> + RelationClasses.subrelation (cmp_sort Conv) (cmp_sort pb) -> forall u v n l l', - eq_term_upto_univ_napp Σ Re Rle napp u v -> - All2 (eq_term_upto_univ Σ Re Re) l l' -> - eq_term_upto_univ_napp Σ Re Rle napp (subst l n u) (subst l' n v). + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp u v -> + All2 (eq_term_upto_univ Σ cmp_universe cmp_sort Conv) l l' -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp (subst l n u) (subst l' n v). Proof. - unfold RelationClasses.subrelation; intros hR u v n l l' hu hl. - induction u in napp, v, n, l, l', hu, hl, Rle, hR |- * using term_forall_list_ind. + unfold RelationClasses.subrelation; intros hsub hsub' u v n l l' hu hl. + induction u in napp, v, n, l, l', hu, hl, pb, hsub, hsub' |- * using term_forall_list_ind. all: dependent destruction hu. all: try solve [ cbn ; constructor ; try sih ; eauto ]. - cbn. destruct (Nat.leb_spec0 n n0). @@ -1122,7 +1287,7 @@ Proof. destruct h as [t' [e' h]]. rewrite e'. eapply eq_term_upto_univ_lift. - eapply eq_term_upto_univ_leq. 3:eauto. all:auto with arith. + eapply eq_term_upto_univ_leq. 4:eauto. all:auto with arith. * intros h. eapply All2_nth_error_None in h as hh ; eauto. rewrite hh. apply All2_length in hl as e. rewrite <- e. @@ -1131,27 +1296,28 @@ Proof. - cbn. constructor. solve_all. - cbn. destruct e as (? & ? & e & ?). - constructor ; unfold eq_predicate; simpl; try sih ; solve_all. - * rewrite -(All2_fold_length e). eapply e1; eauto. - * rewrite (All2_fold_length a). now eapply b0. + constructor ; unfold eq_predicate, eq_branches, eq_branch in *; simpl; try sih ; solve_all. + * rewrite -(All2_length e). eapply e2; eauto. + * rewrite (All2_length a0). now eapply b0. - cbn. constructor ; try sih ; eauto. - pose proof (All2_length a). + pose proof (All2_length e). unfold eq_mfixpoint in *. solve_all. now rewrite H. - cbn. constructor ; try sih ; eauto. - pose proof (All2_length a). + pose proof (All2_length e). unfold eq_mfixpoint in *. solve_all. now rewrite H. - cbn; constructor. depelim o; cbn in X |- *; constructor; cbn; intuition eauto. solve_all. Qed. -Lemma eq_term_upto_univ_subst Σ Re Rle : - RelationClasses.subrelation Re Rle -> +Lemma eq_term_upto_univ_subst Σ cmp_universe cmp_sort pb : + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe pb) -> + RelationClasses.subrelation (cmp_sort Conv) (cmp_sort pb) -> forall u v n x y, - eq_term_upto_univ Σ Re Rle u v -> - eq_term_upto_univ Σ Re Re x y -> - eq_term_upto_univ Σ Re Rle (u{n := x}) (v{n := y}). + eq_term_upto_univ Σ cmp_universe cmp_sort pb u v -> + eq_term_upto_univ Σ cmp_universe cmp_sort Conv x y -> + eq_term_upto_univ Σ cmp_universe cmp_sort pb (u{n := x}) (v{n := y}). Proof. - intros hR u v n x y e1 e2. + intros hsub hsub' u v n x y e1 e2. eapply eq_term_upto_univ_substs; eauto. Qed. @@ -1161,7 +1327,6 @@ Lemma subst_eq_term {cf:checker_flags} Σ φ l k T U : Proof. intros Hleq. eapply eq_term_upto_univ_substs; try easy. - now eapply All2_same. Qed. Lemma subst_leq_term {cf:checker_flags} Σ φ l k T U : @@ -1170,34 +1335,32 @@ Lemma subst_leq_term {cf:checker_flags} Σ φ l k T U : Proof. intros Hleq. eapply eq_term_upto_univ_substs; try easy. - intro; apply eq_universe_leq_universe. - now eapply All2_same. + - intro; apply eq_leq_universe. + - intro; apply eq_leq_sort. Qed. (** ** Behavior on mkApps and it_mkLambda_or_LetIn ** *) -Lemma eq_term_eq_term_napp Σ Re Rle napp t t' : - RelationClasses.subrelation Re Rle -> - eq_term_upto_univ Σ Re Rle t t' -> - eq_term_upto_univ_napp Σ Re Rle napp t t'. +Lemma eq_term_eq_term_napp Σ cmp_universe cmp_sort pb napp t t' : + eq_term_upto_univ Σ cmp_universe cmp_sort pb t t' -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp t t'. Proof. - intros. eapply eq_term_upto_univ_impl. 5:eauto. - 4:auto with arith. all:typeclasses eauto. + intros. eapply eq_term_upto_univ_impl. 6:eauto. + 5:auto with arith. all:typeclasses eauto. Qed. -Lemma leq_term_leq_term_napp Σ Re Rle napp t t' : - RelationClasses.subrelation Re Rle -> - eq_term_upto_univ Σ Re Rle t t' -> - eq_term_upto_univ_napp Σ Re Rle napp t t'. +Lemma leq_term_leq_term_napp Σ cmp_universe cmp_sort pb napp t t' : + eq_term_upto_univ Σ cmp_universe cmp_sort pb t t' -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp t t'. Proof. - intros. eapply eq_term_upto_univ_impl. 5:eauto. - 4:auto with arith. all:try typeclasses eauto. + intros. eapply eq_term_upto_univ_impl. 6:eauto. + 5:auto with arith. all:try typeclasses eauto. Qed. -Lemma eq_term_upto_univ_napp_mkApps Σ Re Rle u1 l1 u2 l2 napp : - eq_term_upto_univ_napp Σ Re Rle (#|l1| + napp) u1 u2 -> - All2 (eq_term_upto_univ Σ Re Re) l1 l2 -> - eq_term_upto_univ_napp Σ Re Rle napp (mkApps u1 l1) (mkApps u2 l2). +Lemma eq_term_upto_univ_napp_mkApps Σ cmp_universe cmp_sort pb u1 l1 u2 l2 napp : + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb (#|l1| + napp) u1 u2 -> + All2 (eq_term_upto_univ Σ cmp_universe cmp_sort Conv) l1 l2 -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp (mkApps u1 l1) (mkApps u2 l2). Proof. intros hu hl. induction l1 in napp, u1, u2, l2, hu, hl |- *. - inversion hl. subst. assumption. @@ -1207,14 +1370,14 @@ Proof. + assumption. Qed. -Lemma eq_term_upto_univ_napp_mkApps_l_inv Σ Re Rle napp u l t : - eq_term_upto_univ_napp Σ Re Rle napp (mkApps u l) t -> +Lemma eq_term_upto_univ_napp_mkApps_l_inv Σ cmp_universe cmp_sort pb napp u l t : + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp (mkApps u l) t -> ∑ u' l', - eq_term_upto_univ_napp Σ Re Rle (#|l| + napp) u u' * - All2 (eq_term_upto_univ Σ Re Re) l l' * + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb (#|l| + napp) u u' * + All2 (eq_term_upto_univ Σ cmp_universe cmp_sort Conv) l l' * (t = mkApps u' l'). Proof. - intros h. induction l in napp, u, t, h, Rle |- *. + intros h. induction l in napp, u, t, h, pb |- *. - cbn in h. exists t, []. split ; auto. - cbn in h. apply IHl in h as [u' [l' [[h1 h2] h3]]]. dependent destruction h1. subst. @@ -1226,14 +1389,14 @@ Proof. + cbn. reflexivity. Qed. -Lemma eq_term_upto_univ_napp_mkApps_r_inv Σ Re Rle napp u l t : - eq_term_upto_univ_napp Σ Re Rle napp t (mkApps u l) -> +Lemma eq_term_upto_univ_napp_mkApps_r_inv Σ cmp_universe cmp_sort pb napp u l t : + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp t (mkApps u l) -> ∑ u' l', - eq_term_upto_univ_napp Σ Re Rle (#|l| + napp) u' u * - All2 (eq_term_upto_univ Σ Re Re) l' l * + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb (#|l| + napp) u' u * + All2 (eq_term_upto_univ Σ cmp_universe cmp_sort Conv) l' l * (t = mkApps u' l'). Proof. - intros h. induction l in napp, u, t, h, Rle |- *. + intros h. induction l in napp, u, t, h, pb |- *. - cbn in h. exists t, []. split ; auto. - cbn in h. apply IHl in h as [u' [l' [[h1 h2] h3]]]. dependent destruction h1. subst. @@ -1245,40 +1408,42 @@ Proof. + cbn. reflexivity. Qed. -Lemma eq_term_upto_univ_mkApps Σ Re Rle u1 l1 u2 l2 : - eq_term_upto_univ_napp Σ Re Rle #|l1| u1 u2 -> - All2 (eq_term_upto_univ Σ Re Re) l1 l2 -> - eq_term_upto_univ Σ Re Rle (mkApps u1 l1) (mkApps u2 l2). +Lemma eq_term_upto_univ_mkApps Σ cmp_universe cmp_sort pb u1 l1 u2 l2 : + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb #|l1| u1 u2 -> + All2 (eq_term_upto_univ Σ cmp_universe cmp_sort Conv) l1 l2 -> + eq_term_upto_univ Σ cmp_universe cmp_sort pb (mkApps u1 l1) (mkApps u2 l2). Proof. intros; apply eq_term_upto_univ_napp_mkApps; rewrite ?Nat.add_0_r; auto. Qed. -Lemma eq_term_upto_univ_mkApps_l_inv Σ Re Rle u l t : - eq_term_upto_univ Σ Re Rle (mkApps u l) t -> +Lemma eq_term_upto_univ_mkApps_l_inv Σ cmp_universe cmp_sort pb u l t : + eq_term_upto_univ Σ cmp_universe cmp_sort pb (mkApps u l) t -> ∑ u' l', - eq_term_upto_univ_napp Σ Re Rle #|l| u u' * - All2 (eq_term_upto_univ Σ Re Re) l l' * + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb #|l| u u' * + All2 (eq_term_upto_univ Σ cmp_universe cmp_sort Conv) l l' * (t = mkApps u' l'). Proof. intros H; apply eq_term_upto_univ_napp_mkApps_l_inv in H; rewrite ?Nat.add_0_r in H; auto. Qed. -Lemma eq_term_upto_univ_mkApps_r_inv Σ Re Rle u l t : - eq_term_upto_univ Σ Re Rle t (mkApps u l) -> +Lemma eq_term_upto_univ_mkApps_r_inv Σ cmp_universe cmp_sort pb u l t : + eq_term_upto_univ Σ cmp_universe cmp_sort pb t (mkApps u l) -> ∑ u' l', - eq_term_upto_univ_napp Σ Re Rle #|l| u' u * - All2 (eq_term_upto_univ Σ Re Re) l' l * + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb #|l| u' u * + All2 (eq_term_upto_univ Σ cmp_universe cmp_sort Conv) l' l * (t = mkApps u' l'). Proof. intros H; apply eq_term_upto_univ_napp_mkApps_r_inv in H; rewrite Nat.add_0_r in H; auto. Qed. -Lemma R_universe_instance_eq {u u'} : R_universe_instance eq u u' -> u = u'. +Lemma cmp_universe_instance_eq {u u'} : cmp_universe_instance eq u u' -> u = u'. Proof. intros H. + unfold cmp_universe_instance, on_rel in H. + apply Forall2_map in H. apply Forall2_eq in H. apply map_inj in H ; revgoals. - { apply Universe.make_inj. } + { intros ?? e. now inv e. } subst. constructor ; auto. Qed. @@ -1288,23 +1453,25 @@ Proof. red. destruct check_univs => //. Qed. -Lemma upto_eq_impl Σ Re Rle : - RelationClasses.Reflexive Re -> - RelationClasses.Reflexive Rle -> - subrelation (eq_term_upto_univ Σ eq eq) (eq_term_upto_univ Σ Re Rle). +Lemma upto_eq_impl Σ cmp_universe cmp_sort pb0 pb : + RelationClasses.Reflexive (cmp_universe Conv) -> + RelationClasses.Reflexive (cmp_universe pb) -> + RelationClasses.Reflexive (cmp_sort Conv) -> + RelationClasses.Reflexive (cmp_sort pb) -> + subrelation (eq_term_upto_univ Σ (fun _ => eq) (fun _ => eq) pb0) (eq_term_upto_univ Σ cmp_universe cmp_sort pb). Proof. - intros he hle. eapply eq_term_upto_univ_impl. 4:auto. + intros univ_refl univ_refl' sort_refl sort_refl'. eapply eq_term_upto_univ_impl. 5:auto. all: intros ? ? []; eauto. Qed. (** ** Syntactic ws_cumul_pb up to printing anotations ** *) -Definition upto_names := eq_term_upto_univ empty_global_env eq eq. +Definition upto_names := eq_term_upto_univ empty_global_env (fun _ => eq) (fun _ => eq) Conv. Infix "≡" := upto_names (at level 70). -Infix "≡'" := (eq_term_upto_univ empty_global_env eq eq) (at level 70). -Notation upto_names' := (eq_term_upto_univ empty_global_env eq eq). +Infix "≡'" := (eq_term_upto_univ empty_global_env (fun _ => eq) (fun _ => eq) Conv) (at level 70). +Notation upto_names' := (eq_term_upto_univ empty_global_env (fun _ => eq) (fun _ => eq) Conv). #[global] Instance upto_names_ref : Reflexive upto_names. @@ -1335,12 +1502,15 @@ Qed. (* all: intros u u'; eapply reflect_reflectT, eqb_spec. *) (* Qed. *) -Lemma upto_names_impl Σ Re Rle : - RelationClasses.Reflexive Re -> - RelationClasses.Reflexive Rle -> - subrelation upto_names (eq_term_upto_univ Σ Re Rle). +Lemma upto_names_impl Σ cmp_universe cmp_sort pb napp : + RelationClasses.Reflexive (cmp_universe Conv) -> + RelationClasses.Reflexive (cmp_universe pb) -> + RelationClasses.Reflexive (cmp_sort Conv) -> + RelationClasses.Reflexive (cmp_sort pb) -> + subrelation upto_names (eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp). Proof. - intros he hle. eapply eq_term_upto_univ_empty_impl; auto. + intros univ_refl univ_refl' sort_refl sort_refl'. + eapply eq_term_upto_univ_empty_impl; auto. all: intros ? ? []; eauto. Qed. @@ -1356,8 +1526,8 @@ Proof. eapply upto_names_impl ; exact _. Qed. -Lemma eq_term_upto_univ_isApp Σ Re Rle napp u v : - eq_term_upto_univ_napp Σ Re Rle napp u v -> +Lemma eq_term_upto_univ_isApp Σ cmp_universe cmp_sort pb napp u v : + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp u v -> isApp u = isApp v. Proof. induction 1. @@ -1372,14 +1542,14 @@ Inductive rel_option {A B} (R : A -> B -> Type) : option A -> option B -> Type : Derive Signature NoConfusion for rel_option. -Definition eq_decl_upto_gen Σ Re Rle d d' : Type := +Definition eq_decl_upto_gen Σ cmp_universe cmp_sort pb d d' : Type := eq_binder_annot d.(decl_name) d'.(decl_name) * - rel_option (eq_term_upto_univ Σ Re Re) d.(decl_body) d'.(decl_body) * - eq_term_upto_univ Σ Re Rle d.(decl_type) d'.(decl_type). + rel_option (eq_term_upto_univ Σ cmp_universe cmp_sort Conv) d.(decl_body) d'.(decl_body) * + eq_term_upto_univ Σ cmp_universe cmp_sort pb d.(decl_type) d'.(decl_type). (* TODO perhaps should be def *) -Lemma All2_eq_context_upto Σ Re Rle : - subrelation (All2 (eq_decl_upto_gen Σ Re Rle)) (eq_context_upto Σ Re Rle). +Lemma All2_eq_context_upto Σ cmp_universe cmp_sort pb : + subrelation (All2 (eq_decl_upto_gen Σ cmp_universe cmp_sort pb)) (eq_context_upto Σ cmp_universe cmp_sort pb). Proof. intros Γ Δ h. induction h. @@ -1392,27 +1562,45 @@ Proof. + constructor ; eauto. constructor; auto. Qed. -Lemma eq_context_upto_refl Σ Re Rle : - RelationClasses.Reflexive Re -> - RelationClasses.Reflexive Rle -> - Reflexive (eq_context_upto Σ Re Rle). + +Lemma eq_context_upto_impl Σ cmp_universe cmp_universe' cmp_sort cmp_sort' pb pb' : + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe' Conv) -> + RelationClasses.subrelation (cmp_universe pb) (cmp_universe' pb') -> + RelationClasses.subrelation (cmp_sort Conv) (cmp_sort' Conv) -> + RelationClasses.subrelation (cmp_sort pb) (cmp_sort' pb') -> + subrelation (eq_context_upto Σ cmp_universe cmp_sort pb) (eq_context_upto Σ cmp_universe' cmp_sort' pb'). +Proof. + intros. + apply eq_context_gen_impl. + all: apply eq_term_upto_univ_impl; tc. + all: auto. +Qed. + +Lemma eq_context_upto_refl Σ cmp_universe cmp_sort pb : + RelationClasses.Reflexive (cmp_universe Conv) -> + RelationClasses.Reflexive (cmp_universe pb) -> + RelationClasses.Reflexive (cmp_sort Conv) -> + RelationClasses.Reflexive (cmp_sort pb) -> + Reflexive (eq_context_upto Σ cmp_universe cmp_sort pb). Proof. exact _. Qed. -Lemma eq_context_upto_sym Σ Re Rle : - RelationClasses.Symmetric Re -> - RelationClasses.Symmetric Rle -> - Symmetric (eq_context_upto Σ Re Rle). +Lemma eq_context_upto_sym Σ cmp_universe cmp_sort pb : + RelationClasses.Symmetric (cmp_universe Conv) -> + RelationClasses.Symmetric (cmp_universe pb) -> + RelationClasses.Symmetric (cmp_sort Conv) -> + RelationClasses.Symmetric (cmp_sort pb) -> + Symmetric (eq_context_upto Σ cmp_universe cmp_sort pb). Proof. exact _. Qed. -Lemma eq_context_upto_cat Σ Re Rle Γ Δ Γ' Δ' : - eq_context_upto Σ Re Rle Γ Γ' -> - eq_context_upto Σ Re Rle Δ Δ' -> - eq_context_upto Σ Re Rle (Γ ,,, Δ) (Γ' ,,, Δ'). +Lemma eq_context_upto_cat Σ cmp_universe cmp_sort pb Γ Δ Γ' Δ' : + eq_context_upto Σ cmp_universe cmp_sort pb Γ Γ' -> + eq_context_upto Σ cmp_universe cmp_sort pb Δ Δ' -> + eq_context_upto Σ cmp_universe cmp_sort pb (Γ ,,, Δ) (Γ' ,,, Δ'). Proof. intros. eapply All2_fold_app; eauto. Qed. -Lemma eq_context_upto_rev Σ Re Rle Γ Δ : - eq_context_upto Σ Re Rle Γ Δ -> - eq_context_upto Σ Re Rle (rev Γ) (rev Δ). +Lemma eq_context_upto_rev Σ cmp_universe cmp_sort pb Γ Δ : + eq_context_upto Σ cmp_universe cmp_sort pb Γ Δ -> + eq_context_upto Σ cmp_universe cmp_sort pb (rev Γ) (rev Δ). Proof. induction 1. - constructor. @@ -1421,11 +1609,11 @@ Proof. Qed. Lemma eq_context_upto_rev' : - forall Σ Γ Δ Re Rle, - eq_context_upto Σ Re Rle Γ Δ -> - eq_context_upto Σ Re Rle (List.rev Γ) (List.rev Δ). + forall Σ Γ Δ cmp_universe cmp_sort pb, + eq_context_upto Σ cmp_universe cmp_sort pb Γ Δ -> + eq_context_upto Σ cmp_universe cmp_sort pb (List.rev Γ) (List.rev Δ). Proof. - intros Σ Γ Δ Re Rle h. + intros Σ Γ Δ cmp_universe cmp_sort pb h. induction h. - constructor. - simpl. eapply eq_context_upto_cat. @@ -1433,21 +1621,20 @@ Proof. + assumption. Qed. -Lemma eq_context_upto_length : - forall {Σ Re Rle Γ Δ}, - eq_context_upto Σ Re Rle Γ Δ -> +Lemma eq_context_upto_length {Σ cmp_universe cmp_sort pb Γ Δ} : + eq_context_upto Σ cmp_universe cmp_sort pb Γ Δ -> #|Γ| = #|Δ|. Proof. - intros Σ Re Rle Γ Δ h. - induction h. all: simpl ; auto. + apply All2_fold_length. Qed. -Lemma eq_context_upto_subst_context Σ Re Rle : - RelationClasses.subrelation Re Rle -> +Lemma eq_context_upto_subst_context Σ cmp_universe cmp_sort pb : + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe pb) -> + RelationClasses.subrelation (cmp_sort Conv) (cmp_sort pb) -> forall u v n l l', - eq_context_upto Σ Re Rle u v -> - All2 (eq_term_upto_univ Σ Re Re) l l' -> - eq_context_upto Σ Re Rle (subst_context l n u) (subst_context l' n v). + eq_context_upto Σ cmp_universe cmp_sort pb u v -> + All2 (eq_term_upto_univ Σ cmp_universe cmp_sort Conv) l l' -> + eq_context_upto Σ cmp_universe cmp_sort pb (subst_context l n u) (subst_context l' n v). Proof. intros re u v n l l'. induction 1; intros Hl. @@ -1455,27 +1642,27 @@ Proof. - rewrite !subst_context_snoc; constructor; eauto. depelim p; constructor; simpl; intuition auto; rewrite -(length_of X); - apply eq_term_upto_univ_substs; auto. reflexivity. + apply eq_term_upto_univ_substs; auto. all: reflexivity. Qed. #[global] Hint Resolve All2_fold_nil : pcuic. -Lemma eq_context_upto_smash_context Σ ctx ctx' x y : - eq_context_upto Σ eq eq ctx ctx' -> eq_context_upto Σ eq eq x y -> - eq_context_upto Σ eq eq (smash_context ctx x) (smash_context ctx' y). +Lemma eq_context_upto_smash_context Σ ctx ctx' x y pb : + eq_context_upto Σ (fun _ => eq) (fun _ => eq) pb ctx ctx' -> eq_context_upto Σ (fun _ => eq) (fun _ => eq) pb x y -> + eq_context_upto Σ (fun _ => eq) (fun _ => eq) pb (smash_context ctx x) (smash_context ctx' y). Proof. induction x in ctx, ctx', y |- *; intros eqctx eqt; inv eqt; simpl; try split; auto; try constructor; auto. depelim X0 => /=. - apply IHx; auto. apply eq_context_upto_cat; auto. constructor; pcuic. - apply IHx; auto. eapply eq_context_upto_subst_context; eauto. - typeclasses eauto. + all: reflexivity. Qed. -Lemma eq_context_upto_nth_error Σ Re Rle ctx ctx' n : - eq_context_upto Σ Re Rle ctx ctx' -> - rel_option (eq_decl_upto_gen Σ Re Rle) (nth_error ctx n) (nth_error ctx' n). +Lemma eq_context_upto_nth_error Σ cmp_universe cmp_sort pb ctx ctx' n : + eq_context_upto Σ cmp_universe cmp_sort pb ctx ctx' -> + rel_option (eq_decl_upto_gen Σ cmp_universe cmp_sort pb) (nth_error ctx n) (nth_error ctx' n). Proof. induction 1 in n |- *. - rewrite nth_error_nil. constructor. @@ -1485,38 +1672,38 @@ Proof. Qed. Lemma eq_context_impl : - forall Σ Re Re' Rle Rle', - RelationClasses.subrelation Re Re' -> - RelationClasses.subrelation Rle Rle' -> - RelationClasses.subrelation Re' Rle' -> - subrelation (eq_context_upto Σ Re Rle) (eq_context_upto Σ Re' Rle'). -Proof. - intros Σ Re Re' Rle Rle' hR hR' hReRle' Γ Δ h. + forall Σ cmp_universe cmp_universe' cmp_sort cmp_sort' pb pb', + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe' Conv) -> + RelationClasses.subrelation (cmp_universe pb) (cmp_universe' pb') -> + RelationClasses.subrelation (cmp_sort Conv) (cmp_sort' Conv) -> + RelationClasses.subrelation (cmp_sort pb) (cmp_sort' pb') -> + subrelation (eq_context_upto Σ cmp_universe cmp_sort pb) (eq_context_upto Σ cmp_universe' cmp_sort' pb'). +Proof. + intros Σ cmp_universe cmp_universe' cmp_sort cmp_sort' pb pb' univ_sub_conv univ_sub_pb sort_sub_conv sort_sub_pb Γ Δ h. induction h. - constructor. - constructor; auto. depelim p; constructor; auto. - all:eapply eq_term_upto_univ_impl. 5,10,15:tea. all:eauto. - all:now transitivity Re'. + all:eapply eq_term_upto_univ_impl; [ .. | eassumption]; eauto. Qed. -Lemma eq_term_upto_univ_it_mkLambda_or_LetIn Σ Re Rle : - RelationClasses.Reflexive Re -> - Proper (eq_context_upto Σ Re Re ==> eq_term_upto_univ Σ Re Rle ==> eq_term_upto_univ Σ Re Rle) it_mkLambda_or_LetIn. +Lemma eq_term_upto_univ_it_mkLambda_or_LetIn Σ cmp_universe cmp_sort pb : + Proper (eq_context_upto Σ cmp_universe cmp_sort Conv ==> eq_term_upto_univ Σ cmp_universe cmp_sort pb ==> eq_term_upto_univ Σ cmp_universe cmp_sort pb) it_mkLambda_or_LetIn. Proof. - intros he Γ Δ eq u v h. - induction eq in u, v, h, he |- *. + intros Γ Δ eq u v h. + induction eq in u, v, h |- *. - assumption. - simpl. cbn. apply IHeq => //. destruct p; cbn; constructor ; tas; try reflexivity. Qed. -Lemma eq_term_upto_univ_it_mkLambda_or_LetIn_r Σ Re Rle Γ : - RelationClasses.Reflexive Re -> - respectful (eq_term_upto_univ Σ Re Rle) (eq_term_upto_univ Σ Re Rle) +Lemma eq_term_upto_univ_it_mkLambda_or_LetIn_r Σ cmp_universe cmp_sort pb Γ : + RelationClasses.Reflexive (cmp_universe Conv) -> + RelationClasses.Reflexive (cmp_sort Conv) -> + respectful (eq_term_upto_univ Σ cmp_universe cmp_sort pb) (eq_term_upto_univ Σ cmp_universe cmp_sort pb) (it_mkLambda_or_LetIn Γ) (it_mkLambda_or_LetIn Γ). Proof. - intros he u v h. + intros univ_refl sort_refl u v h. induction Γ as [| [na [b|] A] Γ ih ] in u, v, h |- *. - assumption. - simpl. cbn. apply ih. constructor ; try apply eq_term_upto_univ_refl. @@ -1532,12 +1719,13 @@ Proof. apply eq_term_upto_univ_it_mkLambda_or_LetIn_r; exact _. Qed. -Lemma eq_term_upto_univ_it_mkProd_or_LetIn Σ Re Rle Γ : - RelationClasses.Reflexive Re -> - respectful (eq_term_upto_univ Σ Re Rle) (eq_term_upto_univ Σ Re Rle) +Lemma eq_term_upto_univ_it_mkProd_or_LetIn Σ cmp_universe cmp_sort pb Γ : + RelationClasses.Reflexive (cmp_universe Conv) -> + RelationClasses.Reflexive (cmp_sort Conv) -> + respectful (eq_term_upto_univ Σ cmp_universe cmp_sort pb) (eq_term_upto_univ Σ cmp_universe cmp_sort pb) (it_mkProd_or_LetIn Γ) (it_mkProd_or_LetIn Γ). Proof. - intros he u v h. + intros univ_refl sort_refl u v h. induction Γ as [| [na [b|] A] Γ ih ] in u, v, h |- *. - assumption. - simpl. cbn. apply ih. constructor ; try apply eq_term_upto_univ_refl. @@ -1565,11 +1753,11 @@ Proof. assumption. Qed. -Lemma eq_term_upto_univ_mkApps_inv Σ Re Rle u l u' l' : +Lemma eq_term_upto_univ_mkApps_inv Σ cmp_universe cmp_sort pb u l u' l' : isApp u = false -> isApp u' = false -> - eq_term_upto_univ Σ Re Rle (mkApps u l) (mkApps u' l') -> - eq_term_upto_univ_napp Σ Re Rle #|l| u u' * All2 (eq_term_upto_univ Σ Re Re) l l'. + eq_term_upto_univ Σ cmp_universe cmp_sort pb (mkApps u l) (mkApps u' l') -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb #|l| u u' * All2 (eq_term_upto_univ Σ cmp_universe cmp_sort Conv) l l'. Proof. intros hu hu' h. apply eq_term_upto_univ_mkApps_l_inv in h as hh. @@ -1579,9 +1767,9 @@ Proof. destruct h3 as [? ?]. subst. split ; auto. Qed. -Lemma isLambda_eq_term_l Σ Re u v : +Lemma isLambda_eq_term_l Σ cmp_universe cmp_sort pb u v : isLambda u -> - eq_term_upto_univ Σ Re Re u v -> + eq_term_upto_univ Σ cmp_universe cmp_sort pb u v -> isLambda v. Proof. intros h e. @@ -1589,9 +1777,9 @@ Proof. depelim e. auto. Qed. -Lemma isLambda_eq_term_r Σ Re u v : +Lemma isLambda_eq_term_r Σ cmp_universe cmp_sort pb u v : isLambda v -> - eq_term_upto_univ Σ Re Re u v -> + eq_term_upto_univ Σ cmp_universe cmp_sort pb u v -> isLambda u. Proof. intros h e. @@ -1599,9 +1787,9 @@ Proof. depelim e. auto. Qed. -Lemma isConstruct_app_eq_term_l Σ Re u v : +Lemma isConstruct_app_eq_term_l Σ cmp_universe cmp_sort pb u v : isConstruct_app u -> - eq_term_upto_univ Σ Re Re u v -> + eq_term_upto_univ Σ cmp_universe cmp_sort pb u v -> isConstruct_app v. Proof. intros h e. @@ -1621,12 +1809,12 @@ Proof. Qed. Lemma isConstruct_app_eq_term_r : - forall Σ Re u v, + forall Σ cmp_universe cmp_sort pb u v, isConstruct_app v -> - eq_term_upto_univ Σ Re Re u v -> + eq_term_upto_univ Σ cmp_universe cmp_sort pb u v -> isConstruct_app u. Proof. - intros Σ Re u v h e. + intros Σ cmp_universe cmp_sort pb u v h e. case_eq (decompose_app u). intros t1 l1 e1. case_eq (decompose_app v). intros t2 l2 e2. unfold isConstruct_app in *. @@ -1642,85 +1830,119 @@ Proof. - reflexivity. Qed. -Lemma R_global_instance_flip Σ gr napp - (Re Rle Rle' : Universe.t -> Universe.t -> Prop) u v: - RelationClasses.Reflexive Re -> - RelationClasses.Reflexive Rle -> - RelationClasses.Symmetric Re -> - RelationClasses.Transitive Re -> - RelationClasses.Transitive Rle -> - RelationClasses.subrelation Re Rle -> - (forall x y, Rle x y -> Rle' y x) -> - R_global_instance Σ Re Rle gr napp u v -> - R_global_instance Σ Re Rle' gr napp v u. -Proof. - intros Rerefl Rlerefl Resym Retrans Rletrans incl incl'. - unfold R_global_instance_gen, R_opt_variance. - destruct global_variance_gen as [vs|] eqn:var. - - induction u in vs, v |- *; destruct v; simpl; auto; - destruct vs as [|v' vs]; simpl; auto. - intros [Ra Ru']. split. - destruct v'; simpl; auto. - apply IHu; auto. - - apply Forall2_symP; eauto. -Qed. - -Lemma eq_term_upto_univ_napp_flip Σ (Re Rle Rle' : Universe.t -> Universe.t -> Prop) napp u v : - RelationClasses.Reflexive Re -> - RelationClasses.Reflexive Rle -> - RelationClasses.Symmetric Re -> - RelationClasses.Transitive Re -> - RelationClasses.Transitive Rle -> - RelationClasses.subrelation Re Rle -> - (forall x y, Rle x y -> Rle' y x) -> - eq_term_upto_univ_napp Σ Re Rle napp u v -> - eq_term_upto_univ_napp Σ Re Rle' napp v u. -Proof. - intros Rerefl Rlerefl Resym Retrans Rletrans incl incl' H. - assert (Resub : RelationClasses.subrelation Re Re). - { typeclasses eauto. } - revert Rerefl Rlerefl Resym Retrans Rletrans incl incl' Resub. - revert Re Rle u v H Rle'. - induction 1; intros; constructor; intuition auto. - all:try solve [now symmetry]. - all:eauto using R_global_instance_flip. - - eapply All2_sym. solve_all. - * eapply eq_context_sym; try tc. tas. - * now eapply eq_term_upto_univ_sym. - - eapply All2_sym. solve_all. - now eapply eq_term_upto_univ_sym. - now eapply eq_term_upto_univ_sym. - now symmetry. - - eapply All2_sym. solve_all. - now eapply eq_term_upto_univ_sym. - now eapply eq_term_upto_univ_sym. - now symmetry. - - depelim o; constructor; eauto. - now eapply eq_term_upto_univ_sym. - now eapply eq_term_upto_univ_sym. - eapply All2_sym; solve_all. - now eapply eq_term_upto_univ_sym. +Lemma cmp_universe_variance_flip cmp_universe cmp_universe' pb pb' v u u' : + RelationClasses.subrelation (cmp_universe Conv) (flip (cmp_universe' Conv)) -> + RelationClasses.subrelation (cmp_universe pb) (flip (cmp_universe' pb')) -> + cmp_universe_variance cmp_universe pb v u u' -> + cmp_universe_variance cmp_universe' pb' v u' u. +Proof. + intros conv_sym pb_sym. + destruct v => //=. + all: unfold on_rel, flip in *; now auto. Qed. -Lemma eq_univ_make : - forall u u', - Forall2 eq (map Universe.make u) (map Universe.make u') -> - u = u'. +Lemma cmp_universe_instance_variance_flip cmp_universe cmp_universe' pb pb' v u u' : + RelationClasses.subrelation (cmp_universe Conv) (flip (cmp_universe' Conv)) -> + RelationClasses.subrelation (cmp_universe pb) (flip (cmp_universe' pb')) -> + cmp_universe_instance_variance cmp_universe pb v u u' -> + cmp_universe_instance_variance cmp_universe' pb' v u' u. Proof. - intros u u' H. eapply Forall2_map' in H. - eapply Forall2_eq. eapply Forall2_impl; tea. - clear. intros [] [] H; now inversion H. + intros conv_sym pb_sym H. + induction H; constructor; eauto. + now eapply cmp_universe_variance_flip. +Qed. + + +Lemma cmp_universe_instance_flip eq_universe eq_universe' u u' : + RelationClasses.subrelation eq_universe (flip eq_universe') -> + cmp_universe_instance eq_universe u u' -> + cmp_universe_instance eq_universe' u' u. +Proof. + intros Hsub H. + apply Forall2_sym, Forall2_impl with (1 := H). + intros ??. apply Hsub. +Qed. + +Lemma cmp_global_instance_flip Σ cmp_universe cmp_universe' pb pb' gr napp u u': + RelationClasses.subrelation (cmp_universe Conv) (flip (cmp_universe' Conv)) -> + RelationClasses.subrelation (cmp_universe pb) (flip (cmp_universe' pb')) -> + cmp_global_instance Σ cmp_universe pb gr napp u u' -> + cmp_global_instance Σ cmp_universe' pb' gr napp u' u. +Proof. + intros conv_sym pb_sym. + unfold cmp_global_instance_gen, cmp_opt_variance. + destruct global_variance_gen as [| |v] => //. + 2: intros [H|H]; [left|right]; move:H. + 1,2: apply cmp_universe_instance_flip; tas; reflexivity. + now apply cmp_universe_instance_variance_flip. +Qed. + +Lemma eq_term_upto_univ_napp_flip Σ cmp_universe cmp_universe' cmp_sort cmp_sort' pb pb' napp u v : + RelationClasses.subrelation (cmp_universe Conv) (flip (cmp_universe' Conv)) -> + RelationClasses.subrelation (cmp_universe pb) (flip (cmp_universe' pb')) -> + RelationClasses.subrelation (cmp_sort Conv) (flip (cmp_sort' Conv)) -> + RelationClasses.subrelation (cmp_sort pb) (flip (cmp_sort' pb')) -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp u v -> + eq_term_upto_univ_napp Σ cmp_universe' cmp_sort' pb' napp v u. +Proof. + intros univ_sub_conv univ_sub_pb sort_sub_conv sort_sub_pb. + induction u in napp, pb, pb', univ_sub_pb, sort_sub_pb, v |- * using term_forall_list_ind; + try (inversion 1; subst; constructor; + eauto using cmp_universe_instance_flip; try (symmetry; assumption); fail). + - inversion 1; subst; constructor. + eapply All2_sym, All2_impl'; tea. + eapply All_impl; eauto. + - inversion 1; subst; constructor. now eapply sort_sub_pb. + - inversion 1; subst; constructor. + eapply cmp_global_instance_flip. 3:eauto. all:auto. + - inversion 1; subst; constructor. + eapply cmp_global_instance_flip. 3:eauto. all:eauto. + - inversion 1; subst; constructor; unfold eq_predicate, eq_branches, eq_branch in *; eauto. + solve_all. + * apply All2_sym; solve_all. + * eapply cmp_universe_instance_flip; eauto. + * symmetry. solve_all. + * apply All2_sym. solve_all. symmetry. solve_all. + - inversion 1; subst; constructor. + eapply All2_sym, All2_impl'; tea. + eapply All_impl; eauto. + cbn. intros x [? ?] y (?&?&?&?). repeat split; eauto. now symmetry. + - inversion 1; subst; constructor. + eapply All2_sym, All2_impl'; tea. + eapply All_impl; eauto. + cbn. intros x [? ?] y (?&?&?&?). repeat split; eauto. now symmetry. + - inversion 1; subst; constructor. depelim X1; tas; try now constructor. + destruct X as (hty & hdef & harr). + constructor; eauto. + 1: now eapply univ_sub_conv. + apply All2_sym; solve_all. +Qed. + +Lemma eq_context_upto_flip Σ cmp_universe cmp_universe' cmp_sort cmp_sort' pb pb' Γ Δ : + RelationClasses.subrelation (cmp_universe Conv) (flip (cmp_universe' Conv)) -> + RelationClasses.subrelation (cmp_universe pb) (flip (cmp_universe' pb')) -> + RelationClasses.subrelation (cmp_sort Conv) (flip (cmp_sort' Conv)) -> + RelationClasses.subrelation (cmp_sort pb) (flip (cmp_sort' pb')) -> + eq_context_upto Σ cmp_universe cmp_sort pb Γ Δ -> + eq_context_upto Σ cmp_universe' cmp_sort' pb' Δ Γ. +Proof. + intros univ_sub_conv univ_sub_pb sort_sub_conv sort_sub_pb H. + eapply All2_fold_flip, All2_fold_impl with (1 := H). clear H. + intros _ _ d d' H. + destruct H; constructor. + all: try solve [ eapply eq_term_upto_univ_napp_flip; [ .. | eassumption]; assumption ]. + all: now symmetry. Qed. (** ws_cumul_pb of binder annotations *) Notation eq_annots Γ Δ := (Forall2 (fun na decl => eq_binder_annot na (decl_name decl)) Γ Δ). -Lemma eq_context_gen_binder_annot Γ Δ : - eq_context_gen eq eq Γ Δ -> eq_annots (forget_types Γ) Δ. +Lemma eq_context_upto_names_binder_annot Γ Δ : + eq_context_upto_names Γ Δ -> eq_annots (forget_types Γ) Δ. Proof. induction 1; constructor; auto. - destruct p; auto. + destruct r; auto. Qed. Lemma eq_annots_fold (Γ : list aname) (f : nat -> term -> term) (Δ : context) : @@ -1770,3 +1992,73 @@ Proof. symmetry; eapply (eq_annots_subst_context _ (List.rev pars) 0). reflexivity. Qed. + +Lemma eq_context_gen_map2_set_binder_name cmp_term cmp_term' pb pb' pctx pctx' Γ Δ : + eq_context_gen cmp_term pb pctx pctx' -> + eq_context_gen cmp_term' pb' Γ Δ -> + eq_context_gen cmp_term' pb' + (map2 set_binder_name (forget_types pctx) Γ) + (map2 set_binder_name (forget_types pctx') Δ). +Proof using Type. + intros eqp. + induction 1 in pctx, pctx', eqp |- *. + - destruct eqp; cbn; constructor. + - destruct eqp; simpl; constructor; auto. + destruct c, p; constructor; auto. +Qed. + +Lemma eq_context_upto_names_map2_set_binder_name cmp_term pb pctx pctx' Γ Δ : + eq_context_upto_names pctx pctx' -> + eq_context_gen cmp_term pb Γ Δ -> + eq_context_gen cmp_term pb + (map2 set_binder_name (forget_types pctx) Γ) + (map2 set_binder_name (forget_types pctx') Δ). +Proof using Type. + intro eqctx. + eapply eq_context_gen_map2_set_binder_name with (pb := Conv). + now apply eq_context_upto_names_gen. +Qed. + +Lemma eq_context_upto_names_map2_set_binder_name_eq nas Γ Δ : + eq_context_upto_names Γ Δ -> + map2 PCUICEnvironment.set_binder_name nas Γ = + map2 PCUICEnvironment.set_binder_name nas Δ. +Proof. + induction 1 in nas |- *; cbn; auto. + destruct nas; cbn; auto. + f_equal. destruct r; subst; unfold set_binder_name; f_equal. + apply IHX. +Qed. + +Lemma trans_eq_context_upto_names trans Γ Δ : + eq_context_upto_names Γ Δ -> + eq_context_upto_names (map (map_decl trans) Γ) (map (map_decl trans) Δ). +Proof. + intro. + eapply All2_map. solve_all. + destruct X; cbn; subst; constructor; auto. +Qed. + +Lemma eq_context_upto_names_fold (f : nat -> term -> term) Γ Δ : + eq_context_upto_names Γ Δ -> + eq_context_upto_names (fold_context_k f Γ) (fold_context_k f Δ). +Proof. + induction 1. + - cbn; auto. + - rewrite !fold_context_k_snoc0 /=. + constructor; auto. + rewrite -(All2_length X). + destruct r; now constructor. +Qed. + +Lemma eq_context_upto_names_smash_context {Γ Δ} : + eq_context_upto_names Γ Δ -> + eq_context_upto_names (smash_context [] Γ) (smash_context [] Δ). +Proof. + assert (eq_context_upto_names [] []) as X by constructor. move: X. + set (Γ0 := []) at 1 3. set (Δ0 := []). clearbody Γ0 Δ0. intro X. + induction 1 in Γ0, Δ0, X |- *; cbn; try constructor; tas. + destruct r; cbn; apply IHX0. + - apply All2_app; tas. repeat constructor. assumption. + - now apply eq_context_upto_names_fold. +Qed. diff --git a/pcuic/theories/PCUICEtaExpand.v b/pcuic/theories/PCUICEtaExpand.v index 940990cc7..88f900379 100644 --- a/pcuic/theories/PCUICEtaExpand.v +++ b/pcuic/theories/PCUICEtaExpand.v @@ -73,7 +73,7 @@ Inductive expanded (Γ : list nat) : term -> Prop := | expanded_tRel (n : nat) m args : nth_error Γ n = Some m -> forall Hle : m <= #|args|, Forall (expanded Γ) args -> expanded Γ (mkApps (tRel n) args) | expanded_tVar (id : ident) : expanded Γ (tVar id) | expanded_tEvar (ev : nat) (args : list term) : Forall (expanded Γ) args -> expanded Γ (tEvar ev args) -| expanded_tSort (s : Universe.t) : expanded Γ (tSort s) +| expanded_tSort (s : sort) : expanded Γ (tSort s) | expanded_tProd (na : aname) (ty : term) (body : term) : expanded Γ (tProd na ty body) | expanded_tLambda (na : aname) (ty : term) (body : term) : expanded (0 :: Γ) body -> expanded Γ (tLambda na ty body) | expanded_tLetIn (na : aname) (def : term) (def_ty : term) (body : term) : expanded Γ def -> expanded (0 :: Γ) body -> expanded Γ (tLetIn na def def_ty body) @@ -117,7 +117,7 @@ Lemma expanded_ind : (forall (Γ : list nat) (id : ident), P Γ (tVar id)) -> (forall (Γ : list nat) (ev : nat) (args : list term), Forall (expanded Σ Γ) args -> Forall (P Γ) args -> P Γ (tEvar ev args)) -> - (forall (Γ : list nat) (s : Universe.t), P Γ (tSort s)) -> + (forall (Γ : list nat) (s : sort), P Γ (tSort s)) -> (forall (Γ : list nat) (na : aname) (ty body : term), P Γ (tProd na ty body)) -> (forall (Γ : list nat) (na : aname) (ty body : term), expanded Σ (0 :: Γ) body -> P (0 :: Γ) body -> P Γ (tLambda na ty body)) -> @@ -1449,4 +1449,4 @@ Proof. move/expanded_mkApps_inv' => [expa _]. move: expa; rewrite (arguments_mkApps t [u]). move/Forall_app => [] _ hu; now depelim hu. -Qed. \ No newline at end of file +Qed. diff --git a/pcuic/theories/PCUICExpandLetsCorrectness.v b/pcuic/theories/PCUICExpandLetsCorrectness.v index e99889a7b..53a7392ff 100644 --- a/pcuic/theories/PCUICExpandLetsCorrectness.v +++ b/pcuic/theories/PCUICExpandLetsCorrectness.v @@ -58,7 +58,7 @@ Proof. reflexivity. Qed. -Lemma trans_local_app Γ Δ : trans_local (SE.app_context Γ Δ) = trans_local Γ ,,, trans_local Δ. +Lemma trans_local_app Γ Δ : trans_local (Γ ,,, Δ) = trans_local Γ ,,, trans_local Δ. Proof. now rewrite /trans_local map_app. Qed. @@ -880,7 +880,7 @@ Definition wt {cf} Σ Γ t := ∑ T, Σ ;;; Γ |- t : T. Lemma isType_wt {cf} {Σ Γ T} : isType Σ Γ T -> wt Σ Γ T. Proof. - intros [s hs]; now exists (PCUICAst.tSort s). + intros (_ & s & hs & _); now exists (PCUICAst.tSort s). Qed. Coercion isType_wt : isType >-> wt. @@ -913,7 +913,7 @@ Section wtsub. [× declared_inductive Σ ci mdecl idecl, consistent_instance_ext Σ (PCUICEnvironment.ind_universes mdecl) (PCUICAst.puinst p), wf_predicate mdecl idecl p, - All2 (PCUICEquality.compare_decls eq eq) (PCUICCases.ind_predicate_context ci mdecl idecl) (PCUICAst.pcontext p), + eq_context_upto_names (PCUICCases.ind_predicate_context ci mdecl idecl) (PCUICAst.pcontext p), wf_local_rel Σ (Γ ,,, smash_context [] (ind_params mdecl)@[p.(puinst)]) p.(pcontext)@[p.(puinst)], PCUICSpine.spine_subst Σ Γ (PCUICAst.pparams p) (List.rev (PCUICAst.pparams p)) (PCUICEnvironment.smash_context [] (PCUICEnvironment.ind_params mdecl)@[PCUICAst.puinst p]), @@ -921,9 +921,9 @@ Section wtsub. wt Γ c & All2i (fun i cdecl br => [× wf_branch cdecl br, - All2 (PCUICEquality.compare_decls eq eq) (bcontext br) (PCUICCases.cstr_branch_context ci mdecl cdecl), + eq_context_upto_names (bcontext br) (PCUICCases.cstr_branch_context ci mdecl cdecl), wf_local_rel Σ (Γ ,,, smash_context [] (ind_params mdecl)@[p.(puinst)]) br.(bcontext)@[p.(puinst)], - All2 (PCUICEquality.compare_decls eq eq) + eq_context_upto_names (Γ ,,, PCUICCases.case_branch_context ci mdecl p (forget_types br.(bcontext)) cdecl) (Γ ,,, inst_case_branch_context p br) & wt (Γ ,,, PCUICCases.case_branch_context ci mdecl p (forget_types br.(bcontext)) cdecl) br.(bbody)]) 0 idecl.(ind_ctors) brs] @@ -945,17 +945,27 @@ Section wtsub. eexists; tea. eexists; tea. Qed. + Lemma lift_wt_inv {Γ j} : lift_typing typing Σ Γ j -> lift_wf_term (wt Σ Γ) j. + Proof. + intros (? & ? & ? & _). + split. 1: destruct j_term => //. + all: eexists; eassumption. + Qed. + Lemma wt_inv {Γ t} : wt Σ Γ t -> wt_subterm Γ t. Proof. destruct t; simpl; intros [T h]; try exact tt. - now eapply typing_wf_local in h. - now eapply inversion_Evar in h. - - eapply inversion_Prod in h as (?&?&?&?&?); tea. - split; eexists; eauto. - - eapply inversion_Lambda in h as (?&?&?&?&?); tea. - split; eexists; eauto. - - eapply inversion_LetIn in h as (?&?&?&?&?&?); tea. - repeat split; eexists; eauto. + - eapply inversion_Prod in h as (?&?&h1&?&?); tea. + apply lift_wt_inv in h1 as []. + split; tas; eexists; eauto. + - eapply inversion_Lambda in h as (?&h1&?&?); tea. + apply lift_wt_inv in h1 as []. + split; tas; eexists; eauto. + - eapply inversion_LetIn in h as (?&h1&?&?); tea. + apply lift_wt_inv in h1 as []. + repeat split; tas; eexists; eauto. - eapply inversion_App in h as (?&?&?&?&?&?); tea. split; eexists; eauto. - eapply inversion_Case in h as (mdecl&idecl&decli&indices&[]&?); tea. @@ -965,7 +975,7 @@ Section wtsub. split. eapply spine_subst_wt in s. now eapply All_rev_inv in s. exists mdecl, idecl. split; auto. now symmetry. - * eapply wf_local_app_inv. eapply wf_local_alpha. + * eapply All_local_env_app_inv. eapply wf_local_alpha. eapply All2_app; [|reflexivity]. eapply alpha_eq_subst_instance. symmetry; tea. eapply wf_ind_predicate; tea. pcuic. @@ -975,7 +985,7 @@ Section wtsub. eapply All2i_All2_mix_left in brs_ty; tea. eapply All2i_nth_hyp in brs_ty. solve_all. split; auto. - { eapply wf_local_app_inv. eapply wf_local_alpha. + { eapply All_local_env_app_inv. eapply wf_local_alpha. eapply All2_app; [|reflexivity]. eapply alpha_eq_subst_instance in a2. symmetry; tea. @@ -995,12 +1005,12 @@ Section wtsub. eexists; eauto. - eapply inversion_Fix in h as (?&?&?&?&?&?&?); tea. eapply All_prod. - eapply (All_impl a). intros ? h; exact h. - eapply (All_impl a0). intros ? h; eexists; tea. + eapply (All_impl a). intros ? (_ & ? & ? & _); now eexists. + eapply (All_impl a0). intros ? (h & _); cbn in h; eexists; tea. - eapply inversion_CoFix in h as (?&?&?&?&?&?&?); tea. eapply All_prod. - eapply (All_impl a). intros ? h; exact h. - eapply (All_impl a0). intros ? h; eexists; tea. + eapply (All_impl a). intros ? (_ & ? & ? & _); now eexists. + eapply (All_impl a0). intros ? (h & _); cbn in h; eexists; tea. - eapply inversion_Prim in h as (?&?&[]). depelim p0; constructor; eauto. 1-2:now eexists. solve_all. now eexists. eauto. Qed. @@ -1337,14 +1347,12 @@ Proof. Qed. Lemma trans_local_set_binder_name_eq nas Γ Δ : - All2 (PCUICEquality.compare_decls eq eq) Γ Δ -> + eq_context_upto_names Γ Δ -> trans_local (map2 PCUICEnvironment.set_binder_name nas Γ) = trans_local (map2 PCUICEnvironment.set_binder_name nas Δ). Proof. - induction 1 in nas |- *; cbn; auto. - destruct nas; cbn; auto. - f_equal. destruct r; subst; f_equal. - apply IHX. + intro H. + f_equal. now apply eq_context_upto_names_map2_set_binder_name_eq. Qed. Lemma map2_trans l l' : @@ -1848,7 +1856,7 @@ Proof. cbn. depelim X. now depelim X0. depelim X. depelim X0. rewrite (trans_subst (shiftnP #|Γ ,,, Δ| xpred0) (shiftnP #|Γ| xpred0)). - red in l0. now eapply subject_is_open_term in l0. solve_all. now eapply wt_on_free_vars. + apply unlift_TermTyp in l. now eapply subject_is_open_term in l. solve_all. now eapply wt_on_free_vars. constructor ; auto. Qed. @@ -1902,7 +1910,7 @@ Proof. - destruct nth_error eqn:Heq => //. simpl in H. noconf H. simpl. destruct c; noconf H => //. rewrite (trans_lift _ (shiftnP #|skipn (S i) Γ| xpred0)); eauto. - eapply nth_error_All_local_env in wt0; tea. cbn in wt0. + eapply All_local_env_nth_error in wt0; tea. apply unlift_TermTyp in wt0. now eapply subject_is_open_term. do 2 constructor. now rewrite nth_error_map Heq. @@ -2210,11 +2218,11 @@ Proof. - cbn; depelim wt. eapply red_primArray_type; cbn; eauto. Qed. -Lemma trans_R_global_instance {Σ : global_env} Re Rle gref napp u u' : - R_global_instance Σ Re Rle gref napp u u' -> - R_global_instance (trans_global_env Σ) Re Rle gref napp u u'. +Lemma trans_cmp_global_instance {Σ : global_env} Re Rle gref napp u u' : + cmp_global_instance Σ Re Rle gref napp u u' -> + cmp_global_instance (trans_global_env Σ) Re Rle gref napp u u'. Proof. - unfold PCUICEquality.R_global_instance, PCUICEquality.global_variance. + unfold PCUICEquality.cmp_global_instance, PCUICEquality.global_variance. destruct gref; simpl; auto. - unfold S.lookup_inductive, S.lookup_minductive. unfold S.lookup_inductive_gen, S.lookup_minductive_gen. @@ -2235,83 +2243,45 @@ Proof. destruct nth_error => /= //. Qed. -Lemma trans_eq_context_gen_eq_binder_annot Γ Δ : - eq_context_gen eq eq Γ Δ -> +Lemma trans_eq_context_upto_names_eq_binder_annot Γ Δ : + eq_context_upto_names Γ Δ -> All2 eq_binder_annot (map decl_name (trans_local Γ)) (map decl_name (trans_local Δ)). Proof. - move/All2_fold_All2. intros; do 2 eapply All2_map. solve_all. destruct X; cbn; auto. Qed. -Lemma trans_eq_context_gen Γ Δ : - eq_context_gen eq eq Γ Δ -> - eq_context_gen eq eq (trans_local Γ) (trans_local Δ). -Proof. - move/All2_fold_All2 => e. apply All2_fold_All2. - eapply All2_map. solve_all. - destruct X; cbn; subst; constructor; auto. -Qed. - -Lemma eq_context_gen_extended_subst {Γ Δ n} : - eq_context_gen eq eq Γ Δ -> - extended_subst Γ n = extended_subst Δ n. -Proof. - induction 1 in n |- *; cbn; auto. - destruct p; subst; cbn. f_equal; auto. - rewrite IHX. - now rewrite (PCUICConfluence.eq_context_gen_context_assumptions X). -Qed. - -Lemma eq_context_gen_smash_context {Γ Δ} : - eq_context_gen eq eq Γ Δ -> - eq_context_gen eq eq (smash_context [] Γ) (smash_context [] Δ). -Proof. - induction 1; try constructor. - rewrite (smash_context_app [] [d]) smash_context_acc. - rewrite (smash_context_app [] [d']) (smash_context_acc Γ'). - cbn. destruct p; subst; cbn. eapply All2_fold_All2. - eapply All2_app. 2:now eapply All2_fold_All2. - rewrite !lift_context_snoc /=. - rewrite (All2_fold_length X). repeat constructor; cbn; auto. - rewrite (eq_context_gen_extended_subst X). - now rewrite (PCUICConfluence.eq_context_gen_context_assumptions X). - eapply All2_fold_app => //. - constructor. -Qed. - -Lemma eq_context_upto_context_assumptions {Σ : global_env} {Re Rle} {Γ Δ} : - eq_context_upto Σ Re Rle Γ Δ -> - context_assumptions Γ = context_assumptions Δ. +Lemma trans_eq_context_upto_names Γ Δ : + eq_context_upto_names Γ Δ -> + eq_context_upto_names (trans_local Γ) (trans_local Δ). Proof. - induction 1; cbn; auto. - destruct p; subst; cbn; auto. f_equal; auto. + apply PCUICEquality.trans_eq_context_upto_names. Qed. -Lemma eq_term_upto_univ_expand_lets {cf} {Σ : global_env} {Re Rle Γ Δ t u napp} : - subrelation Re Rle -> - eq_context_upto Σ Re Rle Γ Δ -> - eq_term_upto_univ_napp Σ Re Rle napp t u -> - eq_term_upto_univ_napp Σ Re Rle napp (expand_lets Γ t) (expand_lets Δ u). +Lemma eq_term_upto_univ_expand_lets {cf} {Σ : global_env} {cmp_universe cmp_sort pb Γ Δ t u napp} : + subrelation (cmp_universe Conv) (cmp_universe pb) -> + subrelation (cmp_sort Conv) (cmp_sort pb) -> + eq_context_upto Σ cmp_universe cmp_sort pb Γ Δ -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp t u -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp (expand_lets Γ t) (expand_lets Δ u). Proof. - intros subr eqctx eqt. + intros subr subr' eqctx eqt. rewrite /expand_lets /expand_lets_k. eapply eq_term_upto_univ_substs => //. - rewrite (eq_context_upto_length eqctx). - rewrite (eq_context_upto_context_assumptions eqctx). + rewrite -(eq_context_upto_length eqctx). + rewrite (eq_context_gen_context_assumptions eqctx). now eapply eq_term_upto_univ_lift. apply (PCUICConfluence.eq_context_extended_subst eqctx). Qed. -Lemma trans_eq_term_upto_univ {cf} {Σ : global_env} {Re Rle t u napp} : - Reflexive Re -> Reflexive Rle -> - Transitive Re -> SubstUnivPreserving Re -> - subrelation Re Rle -> - PCUICEquality.eq_term_upto_univ_napp Σ Re Rle napp t u -> - eq_term_upto_univ_napp (trans_global_env Σ) Re Rle napp (trans t) (trans u). +Lemma trans_eq_term_upto_univ {cf} {Σ : global_env} {cmp_universe cmp_sort pb t u napp} : + SubstUnivPreserving (cmp_universe Conv) (cmp_universe Conv) -> + SubstUnivPreserving (cmp_universe Conv) (cmp_sort Conv) -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp t u -> + eq_term_upto_univ_napp (trans_global_env Σ) cmp_universe cmp_sort pb napp (trans t) (trans u). Proof. - intros hre hrle hret hres hsub e. - induction t using term_forall_list_ind in Rle, hrle, hsub, napp, u, e |- *. + intros substu substu' e. + induction t using term_forall_list_ind in pb, napp, u, e |- *. all: invs e; cbn. all: try solve [ constructor ; auto ]. all: try solve [ @@ -2321,50 +2291,40 @@ Proof. eapply ih ; eauto using subrelation_refl end ]. - 1,2,3,4,5,6: try solve [ constructor; solve_all; eauto using subrelation_refl ]. - all: try solve [ constructor; now eapply trans_R_global_instance ]. + 1,2,3,4,5,6: try solve [ constructor; unfold eq_mfixpoint in *; solve_all; eauto using subrelation_refl ]. + all: try solve [ constructor; now eapply trans_cmp_global_instance ]. - destruct X1 as [Hpars [Huinst [Hctx Hret]]]. destruct X as [IHpars [IHctx IHret]]. - constructor; cbn; auto. solve_all. - constructor; cbn. solve_all. - 1,3:eauto using subrelation_refl. - repeat split; eauto using subrelation_refl. - rewrite !map_context_trans. - now eapply trans_eq_context_gen. - red in X0. do 2 apply All2_map. - eapply All2_All_mix_left in X3; tea. - eapply All2_impl; tea; cbv beta. - intuition auto. - rewrite !trans_bcontext. - eapply eq_context_gen_smash_context. - now eapply trans_eq_context_gen in a. - rewrite !trans_bbody. - apply eq_term_upto_univ_expand_lets; eauto; tc. - apply eq_context_upto_subst_context; eauto; tc. - rewrite /id. - eapply PCUICConfluence.eq_context_upto_univ_subst_instance'; tc; auto. - cbn. - now eapply trans_eq_context_gen. - cbn. eapply All2_rev. solve_all. eauto using subrelation_refl. - cbn. eauto using subrelation_refl. - - constructor. solve_all; eauto using subrelation_refl. - - constructor; solve_all; eauto using subrelation_refl. - - constructor; depelim X0; cbn in X; constructor; cbn; intuition eauto. - * eapply a2; eauto using subrelation_refl. - * eapply a1; eauto using subrelation_refl. - * solve_all. eauto using subrelation_refl. + constructor; cbn; auto. + all: unfold eq_predicate, eq_branches in *. + 1: repeat split; cbn; auto. + + cbn. apply All2_map. solve_all. + + now apply trans_eq_context_upto_names. + + cbn. apply All2_map. solve_all. + destruct b; split. + * rewrite !trans_bcontext. + now apply eq_context_upto_names_smash_context, trans_eq_context_upto_names. + * rewrite !trans_bbody. + apply eq_term_upto_univ_expand_lets; eauto; tc. + apply eq_context_upto_subst_context; eauto; tc. + 1: apply eq_context_upto_names_subst_instance; eauto; tc. + 1: now eapply trans_eq_context_upto_names. + eapply All2_rev, All2_map. solve_all. + - constructor; depelim X0; cbn in X; try now constructor. + constructor; cbn; intuition eauto. + solve_all. Qed. Lemma trans_compare_term {cf} {Σ : global_env} {pb ϕ T U} : - compare_term pb Σ ϕ T U -> - compare_term (H:=cf' cf) pb (trans_global_env Σ) ϕ (trans T) (trans U). + compare_term Σ ϕ pb T U -> + compare_term (H:=cf' cf) (trans_global_env Σ) ϕ pb (trans T) (trans U). Proof. eapply trans_eq_term_upto_univ ; eauto; tc. Qed. Lemma trans_leq_term {cf} {Σ : global_env} ϕ T U : PCUICEquality.leq_term Σ ϕ T U -> - @compare_term (cf' cf) Cumul (trans_global_env Σ) ϕ (trans T) (trans U). + @compare_term (cf' cf) (trans_global_env Σ) ϕ Cumul (trans T) (trans U). Proof. eapply trans_eq_term_upto_univ; eauto; tc. Qed. @@ -2389,7 +2349,7 @@ Section wtcumul. Reserved Notation " Σ ;;; Γ |-- t <=[ le ] u " (at level 50, Γ, le, t, u at next level). Inductive wt_cumul_pb (pb : conv_pb) (Σ : global_env_ext) (Γ : context) : term -> term -> Type := - | wt_cumul_refl t u : compare_term pb Σ.1 (global_ext_constraints Σ) t u -> Σ ;;; Γ |-- t <=[pb] u + | wt_cumul_refl t u : compare_term Σ.1 (global_ext_constraints Σ) pb t u -> Σ ;;; Γ |-- t <=[pb] u | wt_cumul_red_l t u v : wt_red1 Σ Γ t v -> Σ ;;; Γ |-- v <=[pb] u -> Σ ;;; Γ |-- t <=[pb] u | wt_cumul_red_r t u v : Σ ;;; Γ |-- t <=[pb] v -> wt_red1 Σ Γ u v -> Σ ;;; Γ |-- t <=[pb] u where " Σ ;;; Γ |-- t <=[ pb ] u " := (wt_cumul_pb pb Σ Γ t u) : type_scope. @@ -2484,7 +2444,7 @@ Lemma trans_wf_local' {cf} : (fun Σ0 Γ0 _ (t T : PCUICAst.term) _ => TT.typing (H:=cf' cf) (trans_global Σ0) (trans_local Γ0) (trans t) (trans T)) in - ST.All_local_env_over ST.typing P Σ Γ wfΓ -> + ST.All_local_env_over (ST.typing Σ) (P Σ) Γ wfΓ -> TTwf_local (trans_global Σ) (trans_local Γ). Proof. intros. @@ -2492,8 +2452,8 @@ Proof. - simpl. constructor. - simpl. econstructor. + eapply IHX. - + simpl. destruct tu. exists x. eapply Hs. - - simpl. constructor; auto. red. destruct tu. exists x. auto. + + simpl. destruct tu. repeat (eexists; tea). + - simpl. constructor; auto. destruct tu. repeat (eexists; tea). Qed. Lemma trans_wf_local_env {cf} Σ Γ : @@ -2512,9 +2472,8 @@ Proof. - simpl. constructor. - simpl. econstructor. + eapply IHX. - + simpl. destruct t0. exists x. now eapply p. - - simpl. constructor; auto. red. destruct t0. exists x. intuition auto. - red. red in t1. destruct t1 => //. + + simpl. destruct t0 as (_ & ? & ? & _). repeat (eexists; cbn; tea). now eapply p. + - simpl. constructor; auto. red. destruct t0 as (? & ? & ? & _); cbn in *. repeat (eexists; cbn; tea). all: intuition eauto. Qed. Lemma trans_branches {cf} Σ Γ brs btys ps: @@ -2549,7 +2508,7 @@ Lemma trans_mfix_All {cf} Σ Γ mfix idx : (Γ : SE.context) (b ty : PCUICAst.term) => ST.typing Σ Γ b ty × TT.typing (H := cf' cf) (trans_global Σ) (trans_local Γ) (trans b) (trans ty)) Σ) - (SE.app_context Γ (SE.fix_context mfix)) -> + (Γ ,,, (SE.fix_context mfix)) -> TTwf_local (trans_global Σ) (trans_local Γ ,,, fix_context (map (map_def trans trans) mfix)). Proof. @@ -2557,21 +2516,18 @@ Proof. rewrite -(trans_fix_context (shiftnP #|Γ| xpred0) _ idx) //. match goal with |- TTwf_local _ ?A => - replace A with (trans_local (SE.app_context Γ (SE.fix_context mfix))) + replace A with (trans_local (Γ ,,, (SE.fix_context mfix))) end. 2: { - unfold trans_local, SE.app_context. + unfold trans_local, app_context. now rewrite map_app. } induction X;cbn;constructor;auto;cbn in *. - - destruct t0 as (?&?&?). - exists x. - apply t1. - - destruct t0 as (?&?&?). - eexists;eassumption. - - destruct t1. - assumption. + - destruct t0 as (_ & ? & (? & ?) & _). + repeat (eexists; tea). + - destruct t0 as ((? & ?) & ? & (? & ?) & _). + repeat (eexists; tea). Qed. @@ -3396,39 +3352,23 @@ Lemma weaken_prop_gen {cf} : weaken_env_prop_gen cumulSpec0 (lift_typing typing) (λ (Σ : global_env_ext) (Γ : context) (t T : term), typing (H:=cf' cf) (trans_global Σ) (trans_local Γ) (trans t) (trans T))). Proof. - intros Σ Σ' u wf wf' ext Γ t T. - destruct T. - - intros Ht. - eapply (weakening_env (trans_global (Σ, u))), extends_trans; eauto; destruct ext as (wfΣ&wfΣ'&ext); eauto. - - - intros [s Hs]. exists s. - eapply (weakening_env (trans_global (Σ, u))), extends_trans; eauto; destruct ext as (wfΣ&wfΣ'&ext); eauto. + intros Σ Σ' u wf wf' ext Γ j Hj. + apply lift_typing_impl with (1 := Hj); intros. + eapply (weakening_env (trans_global (Σ, u))), extends_trans; eauto; destruct ext as (wfΣ&wfΣ'&ext); eauto. Defined. (* to get a version that fits in one of the non-_gen lemmas, we need a strict enough version of extension to imply [wf_trans] from extension. *) Lemma weaken_prop {cf} : weaken_env_strictly_decls_prop cumulSpec0 (lift_typing typing) - (lift_typing - (λ (Σ : global_env_ext) (Γ : context) (t T : term), - wf_trans Σ → - typing (H:=cf' cf) (trans_global Σ) (trans_local Γ) (trans t) (trans T))). + (fun Σ Γ j => wf_trans Σ -> lift_typing0 (fun t T => + typing (H:=cf' cf) (trans_global Σ) (trans_local Γ) (trans t) (trans T)) j). Proof. pose weaken_prop_gen as H. - cbv [lift_typing lift_judgment infer_sort weaken_env_strictly_on_decls_prop weaken_env_prop_gen] in *. + cbv [weaken_env_strictly_on_decls_prop weaken_env_prop_gen] in *. repeat (intro; let x := lazymatch goal with x : _ |- _ => x end in pose (H x) as H'; subst H; rename H' into H). - move => Hext Γ t T. - pose (fun v wfΣ' wfΣ => H (wfΣ, (wfΣ', ltac:(tc))) Γ t T v) as H'; subst H; rename H' into H. - destruct T; eauto. - { move => H' Ht. - pose proof (strictly_extends_decls_trans Hext). - pose proof (strictly_extends_decls_wf _ _ Ht _). - eauto using strictly_extends_decls_trans, strictly_extends_decls_wf. } - { move => [s Hs]. exists s. - move => Ht. - pose proof (fun pf wfΣ wfΣ' => (H (s; pf) wfΣ wfΣ').π2) as H'; subst H; rename H' into H. - cbv [weaken_prop_gen] in *; cbn [projT1] in *. - pose proof (strictly_extends_decls_trans Hext). - pose proof (strictly_extends_decls_wf _ _ Ht _). - eauto. } + move => Hext Γ j Hj wftΣ'. + pose (fun v wfΣ' wfΣ => H (wfΣ, (wfΣ', ltac:(tc))) Γ j v) as H'; subst H; rename H' into H. + eapply strictly_extends_decls_wf with (2 := strictly_extends_decls_trans Hext) in wftΣ' as wftΣ. + eapply H; eauto. Qed. Lemma trans_arities_context mdecl : @@ -3566,12 +3506,12 @@ Proof. Qed. Lemma All_over_All {cf} Σ Γ wfΓ : - ST.All_local_env_over ST.typing - (fun (Σ : SE.global_env_ext) (Γ : SE.context) + ST.All_local_env_over (ST.typing Σ) + (fun (Γ : SE.context) (_ : wf_local Σ Γ) (t T : PCUICAst.term) (_ : ST.typing Σ Γ t T) => wf_trans Σ -> - TT.typing (H := cf' cf) (trans_global Σ) (trans_local Γ) (trans t) (trans T)) Σ Γ wfΓ -> + TT.typing (H := cf' cf) (trans_global Σ) (trans_local Γ) (trans t) (trans T)) Γ wfΓ -> ST.All_local_env (ST.lift_typing (fun (Σ0 : SE.global_env_ext) (Γ0 : SE.context) @@ -3584,17 +3524,15 @@ Proof. induction H;cbn. - constructor. - constructor. - + apply IHAll_local_env_over_gen. + + apply IHAll_local_env_over_sorting. + cbn in *. - destruct tu. - eexists;split;auto;try assumption. + apply lift_sorting_it_impl with tu => //. + intros HT wfΣ; split; tas. auto. - constructor. - + apply IHAll_local_env_over_gen. + + apply IHAll_local_env_over_sorting. + cbn in *. - destruct tu. - eexists;split;auto;eassumption. - + cbn in *. - split;auto;eassumption. + apply lift_sorting_it_impl with tu => //=. + all: intros HT wfΣ; split; tas. all: auto. Qed. Lemma trans_prim_ty p prim_ty : trans (prim_type p prim_ty) = prim_type (map_prim trans p) prim_ty. @@ -3614,10 +3552,12 @@ Proof. wf_trans Σ -> TT.typing (H:=cf' cf) (trans_global Σ) (trans_local Γ) (trans t) (trans T) )%type + (fun Σ Γ j => wf_trans Σ -> lift_typing0 (fun t T => TT.typing (H:=cf' cf) (trans_global Σ) (trans_local Γ) (trans t) (trans T)) j) (fun Σ Γ => wf_trans Σ -> TT.All_local_env (TT.lift_typing (TT.typing (H:=cf' cf)) (trans_global Σ)) (trans_local Γ)) );intros. + - apply lift_typing_impl with (1 := X) => ?? [] _ ? //; auto. - eapply trans_wf_local_env => //. now eapply All_over_All. - rewrite (trans_lift _ (shiftnP #|skipn (S n) Γ| xpred0)). eapply closed_wf_local in wfΓ; tea. @@ -3652,14 +3592,14 @@ Proof. + now apply trans_consistent_instance_ext. + red in X. epose proof (declared_constructor_inv weaken_prop _ X isdecl) as [cs [hnth onc]]. destruct onc. red in on_ctype. - destruct on_ctype as [s Hs]. - rewrite /type_of_constructor. forward Hs. eauto. - exists (s@[u]). + unshelve eapply lift_sorting_ex_it_impl_gen with (on_ctype _) => //= Hs. + rewrite /type_of_constructor. rewrite (trans_subst (shiftnP #|ind_bodies mdecl| xpred0) (shiftnP 0 xpred0)). pose proof (declared_constructor_closed_gen_type isdecl). eapply closedn_on_free_vars. len in H0. now rewrite closedn_subst_instance. eapply (inds_is_open_terms []). - eapply (substitution (Γ := trans_local Γ) (Δ := []) (T:=tSort s@[u])). + eexists. + eapply (substitution (Γ := trans_local Γ) (Δ := []) (T:=tSort _@[u])). rewrite trans_inds. eapply (weaken_subslet (Γ := trans_local Γ) (Γ' := [])); eauto with pcuic. eapply subslet_inds. eapply (trans_declared_inductive _ _ _ _ isdecl). now eapply trans_consistent_instance_ext. @@ -3700,7 +3640,7 @@ Proof. + rewrite <- trans_global_ext_constraints. eassumption. + eassert (ctx_inst _ _ _ _) as Hctxi by (eapply ctx_inst_impl with (1 := X5); now intros ? []). - eassert (PCUICEnvTyping.ctx_inst (fun Σ _ _ _ => wf_trans Σ -> @typing (cf' _) _ _ _ _) _ _ _ _) as IHctxi. + eassert (PCUICEnvTyping.ctx_inst (fun _ _ _ => wf_trans Σ -> @typing (cf' _) _ _ _ _) _ _ _) as IHctxi. { eapply ctx_inst_impl with (1 := X5). intros ? ? [? r]; exact r. } move: Hctxi IHctxi. cbn. have wfctx : wf_local Σ (Γ ,,, (ind_params mdecl,,, ind_indices idecl)@[puinst p]). @@ -3713,7 +3653,7 @@ Proof. generalize (pparams p ++ indices). change T.PCUICEnvironment.subst_instance_context with subst_instance_context. rewrite -/context. - generalize (ind_params mdecl ,,, ind_indices idecl)@[puinst p] as Δ. + generalize (ind_params mdecl ,,, ind_indices idecl : context)@[puinst p] as Δ. intros c; revert Γ. induction c using ctx_length_rev_ind. * intros Γ l wf. intros c; depelim c. constructor. @@ -3738,7 +3678,8 @@ Proof. exact X. } { intros c; depelim c. constructor. - destruct (wf_local_app_inv wfctx) as [w _]. depelim w. + destruct (All_local_env_app_inv wfctx) as [w _]. depelim w. + apply unlift_TermTyp in l0 as Hb. unshelve epose proof (substitution_wf_local (Γ':=[vdef na b t]) _ wfctx). shelve. { now eapply subslet_def_tip. } rewrite subst_telescope_subst_context in H. @@ -3748,7 +3689,7 @@ Proof. rewrite -subst_telescope_subst_context in X. rewrite [map trans_decl _](trans_subst_telescope (shiftnP #|Δ| xpred0) (shiftnP (S #|Δ|) xpred0)) in X. - cbn. rewrite (subject_is_open_term l1) //. rewrite List.rev_involutive. + cbn. rewrite (subject_is_open_term Hb) //. rewrite List.rev_involutive. eapply wf_local_closed_context in wfctx. now move: wfctx; rewrite on_free_vars_ctx_app /= => /andP[]. exact X. } @@ -3761,10 +3702,9 @@ Proof. have declc : declared_constructor Σ (ci, i) mdecl idecl cdecl. { split; tea. } move: wfbr; rewrite /wf_branch /wf_branch_gen. cbn. - apply compare_decls_eq_context in cd. rewrite /cstr_branch_context in cd. - eapply trans_eq_context_gen in cd. - eapply eq_context_gen_smash_context in cd. + eapply trans_eq_context_upto_names in cd. + eapply eq_context_upto_names_smash_context in cd. intros eqann. eapply (eq_annots_subst_context _ (map trans (inds (inductive_mind ci) (abstract_instance (ind_universes mdecl)) (ind_bodies mdecl))) #|ind_params mdecl|). eapply (eq_annots_expand_lets_ctx _ (trans_local (ind_params mdecl))). @@ -3782,7 +3722,7 @@ Proof. apply (declared_constructor_closed_args declc). } apply (inds_is_open_terms []). rewrite !trans_bcontext. - now eapply eq_context_gen_binder_annot in cd. + now eapply eq_context_upto_names_binder_annot in cd. + eapply All2i_map. eapply All2i_map_right. eapply Forall2_All2 in H4. eapply All2i_nth_hyp in X8. @@ -3867,9 +3807,7 @@ Proof. + rewrite map_length H. now destruct mdecl. - rewrite trans_dtype /=. assert (is_open_term Γ (tFix mfix n)). - { eapply (subject_is_open_term (Σ := Σ)). econstructor; tea. solve_all. - destruct a as [s Hs]. exists s; intuition eauto. - solve_all. now destruct b. } + { eapply (subject_is_open_term (Σ := Σ)). econstructor; tea. } eapply TT.type_Fix; auto. + rewrite /trans_local map_app in X. now eapply TT.All_local_env_app_inv in X as []. @@ -3879,26 +3817,25 @@ Proof. destruct decl. unfold map_def. reflexivity. - + eapply All_map, (All_impl X0). - intuition auto. destruct X as [s [? ?]]. - exists s; intuition auto. + + eapply All_map, (All_impl X1). + intros d (_ & ? & ? & _) => //. + repeat (eexists; tea). + fold trans. subst types. eapply All_map. - eapply All_prod in X0; tea. clear X1. - eapply All_impl; tea. intros d [[Hdb IHdb] [hs [hdty ihdty]]]. - len. rewrite -(trans_fix_context _ _ _ H2). - rewrite -trans_local_app. - rewrite (trans_lift _ (shiftnP #|Γ| xpred0)) in IHdb. - eapply (subject_is_open_term (Σ := Σ)); tea. - len in IHdb. eauto. + eapply All_prod with (1 := X3) in X0; tea. clear X1 X2 X3. + eapply All_impl; tea. intros d ((IHdb & ? & IHdt & _) & (_ & ? & ? & _)); cbn in *; tas. + repeat (eexists; tea); cbn. + all: len; rewrite -(trans_fix_context _ _ 0 H2). + all: rewrite -trans_local_app. + all: rewrite (trans_lift _ (shiftnP #|Γ| xpred0)) in IHdb, IHdt. + all: try eapply (subject_is_open_term (Σ := Σ)); tea. + all: len in IHdb; len in IHdt; eauto. + rewrite (trans_wf_fixpoint _ (shiftnP #|Γ| xpred0) #|mfix|) //. - rewrite trans_dtype. simpl. assert (is_open_term Γ (tCoFix mfix n)). - { eapply (subject_is_open_term (Σ := Σ)). econstructor; tea. solve_all. - destruct a as [s Hs]. exists s; intuition eauto. - solve_all. now destruct b. } + { eapply (subject_is_open_term (Σ := Σ)). econstructor; tea. } eapply TT.type_CoFix; auto. + rewrite /trans_local map_app in X. now eapply TT.All_local_env_app_inv in X as []. @@ -3908,18 +3845,20 @@ Proof. destruct decl. unfold map_def. reflexivity. + + eapply All_map, (All_impl X1). + intros d (_ & ? & ? & _) => //. + repeat (eexists; tea). + fold trans. - eapply All_map, (All_impl X0). - intros x [s ?]; exists s; intuition auto. - + fold trans;subst types. + subst types. eapply All_map. - eapply All_prod in X0; tea. clear X1. - eapply All_impl; tea. intros d [[Hdb IHdb] [hs [hdty ihdty]]]. - len. rewrite -(trans_fix_context _ _ _ H2). exact 0. - rewrite -trans_local_app. - rewrite (trans_lift _ (shiftnP #|Γ| xpred0)) in IHdb. - eapply (subject_is_open_term (Σ := Σ)); tea. - len in IHdb. eauto. + eapply All_prod with (1 := X3) in X0; tea. clear X1 X2 X3. + eapply All_impl; tea. intros d ((IHdb & ? & IHdt & _) & (_ & ? & ? & _)); cbn in *; tas. + repeat (eexists; tea); cbn. + all: len; rewrite -(trans_fix_context _ _ 0 H2). + all: rewrite -trans_local_app. + all: rewrite (trans_lift _ (shiftnP #|Γ| xpred0)) in IHdb, IHdt. + all: try eapply (subject_is_open_term (Σ := Σ)); tea. + all: len in IHdb; len in IHdt; eauto. + rewrite trans_wf_cofixpoint //. - cbn. rewrite trans_prim_ty. econstructor; eauto. rewrite prim_val_tag_map //. * now eapply trans_declared_constant in H0. @@ -3928,12 +3867,12 @@ Proof. * depelim X0; depelim X1; constructor; intuition eauto. cbn. solve_all. - eapply (type_ws_cumul_pb (pb:=Cumul)). + eauto. - + now exists s. + + now eapply has_sort_isType. + eapply cumulSpec_cumulAlgo_curry in X4; fvs. eapply ws_cumul_pb_forget in X4. eapply cumul_decorate in X4; tea. 2:eapply validity; tea. - 2:now exists s. + 2:now eapply has_sort_isType. eapply into_ws_cumul_pb. eapply (trans_cumul (Σ := Σ)); eauto. eapply (trans_on_free_vars_ctx 0). now eapply wf_local_closed_context in wfΓ. @@ -4004,14 +3943,16 @@ Proof. Qed. Lemma ctx_inst_expand_lets {cf} {Σ : global_env_ext} {wfΣ : wf Σ} {univs} {Γ Δ} {s ctx} : - PCUICTyping.ctx_inst (λ (Σ : global_env_ext) (Γ : context) (t T : term), Σ;;; Γ |- t : T) (Σ, univs) (Γ ,,, Δ) s ctx -> - PCUICTyping.ctx_inst (λ (Σ : global_env_ext) (Γ : context) (t T : term), Σ;;; Γ |- t : T) (Σ, univs) (Γ ,,, smash_context [] Δ) (map (expand_lets Δ) s) (List.rev (expand_lets_ctx Δ (List.rev ctx))). + PCUICTyping.ctx_inst (λ Γ t T , lift_typing typing (Σ, univs) Γ (TermTyp t T)) (Γ ,,, Δ) s ctx -> + PCUICTyping.ctx_inst (λ Γ t T , lift_typing typing (Σ, univs) Γ (TermTyp t T)) (Γ ,,, smash_context [] Δ) (map (expand_lets Δ) s) (List.rev (expand_lets_ctx Δ (List.rev ctx))). Proof. induction 1. - cbn. constructor. - cbn. rewrite expand_lets_ctx_app /=. rewrite List.rev_app_distr /=. constructor. - cbn. now eapply typing_expand_lets in t0. + cbn. + apply lift_typing_f_impl with (1 := t0) => // ?? Hty. + now eapply typing_expand_lets in Hty. rewrite subst_telescope_subst_context. rewrite -subst_context_subst_telescope in IHX. now rewrite (subst_context_expand_lets_ctx _ [i]). @@ -4046,15 +3987,16 @@ Qed. Lemma trans_ctx_inst_expand_lets {cf} {Σ : global_env_ext} {wfΣ : wf Σ} {Γ Δ} {s} : wf_trans Σ -> wf_local Σ (Γ ,,, List.rev Δ) -> - PCUICTyping.ctx_inst (λ (Σ : global_env_ext) (Γ : context) (t T : term), Σ;;; Γ |- t : T) Σ Γ s Δ -> - PCUICTyping.ctx_inst (λ (Σ : global_env_ext) (Γ : context) (t T : term), - typing (H:=cf' cf) Σ Γ t T) (trans_global Σ) - (trans_local Γ) (map trans s) (trans_local Δ). + PCUICTyping.ctx_inst (fun Γ t T => lift_typing typing Σ Γ (TermTyp t T)) Γ s Δ -> + PCUICTyping.ctx_inst (fun Γ t T => lift_typing (typing (H:=cf' cf)) (trans_global Σ) Γ (TermTyp t T)) (trans_local Γ) (map trans s) (trans_local Δ). Proof. intros wf wfctx i. induction i in wfctx |- *. - cbn. constructor. - - cbn. constructor. eapply pcuic_expand_lets; tea. + - cbn. constructor. + apply lift_typing_f_impl with (1 := t0) => // ?? H. + eapply pcuic_expand_lets; tea. + apply unlift_TermTyp in t0. cbn in wfctx. rewrite -app_assoc in wfctx. unshelve epose proof (substitution_wf_local (Σ := Σ) (Γ' := [vass na t]) _ wfctx). shelve. { now eapply subslet_ass_tip. } @@ -4067,14 +4009,14 @@ Proof. - cbn. constructor. cbn in wfctx. rewrite -app_assoc in wfctx. unshelve epose proof (substitution_wf_local (Σ := Σ) (Γ' := [vdef na b t]) _ wfctx). shelve. - { eapply subslet_def_tip. eapply wf_local_app_inv in wfctx as [wf' _]. + { eapply subslet_def_tip. eapply All_local_env_app_inv in wfctx as [wf' _]. eapply unlift_TermTyp. now depelim wf'. } rewrite subst_context_subst_telescope in X. specialize (IHi X). rewrite (trans_local_subst_telescope (shiftnP (S #|Γ|) xpred0) (shiftnP #|Γ| xpred0)) in IHi. { apply wf_local_closed_context in wfctx. now move/onfvs_app: wfctx => /=. } - { eapply wf_local_app_inv in wfctx as [wf' _]. - depelim wf'. cbn. now rewrite (subject_is_open_term l0). } + { eapply All_local_env_app_inv in wfctx as [wf' _]. + depelim wf'. apply unlift_TermTyp in l. cbn. now rewrite (subject_is_open_term l). } apply IHi. Qed. @@ -4278,7 +4220,7 @@ Proof. red in l. now eapply ws_cumul_pb_forget in eqt. rewrite (All2_fold_length X). fvs. - now apply IHX. - - destruct l1; cbn in l0, l2. + - destruct l as (? & ? & ? & _), l0 as (? & ? & ? & _); cbn in *. constructor; auto. eapply convSpec_convAlgo_curry in eqb; tea. now apply ws_cumul_pb_forget in eqb. @@ -4313,10 +4255,10 @@ Proof. pose proof (trans_cumul wfΣ' cum); tea. eapply (cumulAlgo_cumulSpec _ (pb:=Cumul)). eapply into_ws_cumul_pb; tea. - destruct wtT. apply trans_is_closed_context. fvs. + destruct wtT as (_ & ? & ? & _). apply trans_is_closed_context. fvs. apply trans_on_free_vars; len; fvs. apply trans_on_free_vars; len; fvs. - destruct wtT. fvs. + destruct wtT as (_ & ? & ? & _). fvs. Qed. Lemma trans_convSpec {cf} {Σ : PCUICEnvironment.global_env_ext} {Γ T U} {wfΣ : PCUICTyping.wf Σ} : @@ -4326,7 +4268,7 @@ Lemma trans_convSpec {cf} {Σ : PCUICEnvironment.global_env_ext} {Γ T U} {wfΣ convSpec (H:=cf' cf) (trans_global Σ) (trans_local Γ) (trans T) (trans U). Proof. intros wfΣ' wtT wtU cum. - eapply convSpec_convAlgo_curry in cum; fvs. + eapply convSpec_convAlgo_curry in cum; tas. eapply ws_cumul_pb_forget in cum. eapply conv_decorate in cum; tea. pose proof (trans_conv wfΣ' cum); tea. @@ -4351,22 +4293,21 @@ Proof. induction 1; cbn; constructor; auto. eapply IHX. now depelim wfl. now depelim wfr. destruct p; constructor; cbn in *; auto. + all: depelim wfl; depelim wfr; destruct l as (? & ? & ? & _), l0 as (? & ? & ? & _); cbn in *. - rewrite -trans_local_app. - depelim wfl; depelim wfr. red in l, l0. - destruct l0 as [s Hs]. destruct l as [s' Hs']. eapply trans_cumulSpec in eqt; tea. - { now exists s'. } - { exists s. eapply context_cumulativity_spec; tea. + { now eapply has_sort_isType. } + { eapply has_sort_isType. eapply context_cumulativity_spec; tea. eapply All2_fold_app. reflexivity. apply X. } - - rewrite -trans_local_app. depelim wfl; depelim wfr. red in l, l0. + - rewrite -trans_local_app. eapply (trans_convSpec (Σ := Σ)) => //. now exists t. - { red in l2. exists t'. eapply context_cumulativity_spec; tea. + { exists t'. eapply context_cumulativity_spec; tea. eapply All2_fold_app. reflexivity. apply X. } - rewrite -trans_local_app. - depelim wfl; depelim wfr. red in l, l0. eapply (trans_cumulSpec (Σ := Σ)) => //. - { destruct l1 as [s Hs]. exists s. eapply context_cumulativity_spec; tea. + { now eapply has_sort_isType. } + { eapply has_sort_isType. eapply context_cumulativity_spec; tea. eapply All2_fold_app. reflexivity. apply X. } Qed. @@ -4408,7 +4349,7 @@ Section wtcumul'. Inductive wt_cumul_pb_hetero (pb : conv_pb) (Σ : global_env_ext) (Γ Γ' : context) : term -> term -> Type := | wt_cumul_refl' t u : wt Σ Γ t -> wt Σ Γ' u -> - compare_term pb Σ.1 (global_ext_constraints Σ) t u -> Σ ;;; Γ | Γ' |-- t <=[pb] u + compare_term Σ.1 (global_ext_constraints Σ) pb t u -> Σ ;;; Γ | Γ' |-- t <=[pb] u | wt_cumul_red_l' t u v : wt_red1 Σ Γ t v -> Σ ;;; Γ | Γ' |-- v <=[pb] u -> Σ ;;; Γ | Γ' |-- t <=[pb] u | wt_cumul_red_r' t u v : Σ ;;; Γ | Γ' |-- t <=[pb] v -> wt_red1 Σ Γ' u v -> Σ ;;; Γ | Γ' |-- t <=[pb] u where " Σ ;;; Γ | Γ' |-- t <=[ le ] u " := (wt_cumul_pb_hetero le Σ Γ Γ' t u) : type_scope. @@ -4574,15 +4515,15 @@ Proof. eapply IHX. now depelim wfl. now depelim wfr. now depelim ass. now depelim ass'. destruct p; constructor; cbn in *; auto. - rewrite -trans_local_app. - depelim wfl; depelim wfr. red in l, l0. destruct l, l0. + depelim wfl; depelim wfr. pose proof (l.2.π2.1). pose proof (l0.2.π2.1). cbn in X0, X1. eapply (cumulAlgo_cumulSpec _ (pb:=Cumul)). apply into_ws_cumul_pb. eapply (trans_cumul' (Σ := Σ) (Γ' := Γ' ,,, Γ'0)) => //. - eapply cumul_decorate_hetero; tea. now eexists. now eexists. + eapply cumul_decorate_hetero; tea. len. now rewrite (All2_fold_length X) eqlen. now depelim ass. now depelim ass'. eapply cumulSpec_cumulAlgo_curry in eqt; eauto. - now apply ws_cumul_pb_forget in eqt. fvs. eapply (subject_is_open_term t0). + now apply ws_cumul_pb_forget in eqt. fvs. fvs. len. rewrite (All2_fold_length X) eqlen. rewrite -app_length; fvs. len. eapply All2_fold_length in X. lia. @@ -4624,14 +4565,12 @@ Lemma trans_type_local_ctx {cf} {Σ Γ Δ s} : Proof. intros wf wf'. induction Δ; cbn. - unfold PCUICLookup.wf_universe, wf_universe. + unfold PCUICLookup.wf_sort, wf_sort. destruct s => //. destruct a as [? [?|] ?] => /= //; intuition auto. - destruct a0 as [s' Hs]. exists s'. + all: apply lift_typing_f_impl with (1 := b) => //; intros. all:rewrite -trans_local_app. - now eapply (pcuic_expand_lets _ _ _ (tSort _)). - now eapply (pcuic_expand_lets _ _ _ _). - now eapply (pcuic_expand_lets _ _ _ (tSort _)). + all: now eapply pcuic_expand_lets. Qed. Lemma trans_on_context {cf} {Σ Γ} : @@ -4641,9 +4580,8 @@ Lemma trans_on_context {cf} {Σ Γ} : Proof. intros wf wf'. induction 1; cbn; constructor; auto. - destruct t0 as [s Hs]. exists s. now eapply (pcuic_expand_lets _ _ _ (tSort _)). - destruct t0 as [s Hs]. exists s. now eapply (pcuic_expand_lets _ _ _ (tSort _)). - now eapply (pcuic_expand_lets _ _ _ _). + all: apply lift_typing_f_impl with (1 := t0) => //; intros. + all: now eapply pcuic_expand_lets. Qed. Lemma Alli_map {A B} (P : nat -> B -> Type) {f : A -> B} {n l} : @@ -4709,7 +4647,7 @@ Lemma wf_ind_indices {cf:checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} {mi Proof. intros []. rewrite ind_arity_eq in onArity. - destruct onArity as [s Hs]. + destruct onArity as (_ & s & Hs & _). eapply type_it_mkProd_or_LetIn_inv in Hs as [Δs [ts [_ Hs]]]. eapply type_it_mkProd_or_LetIn_inv in Hs as [Δs' [ts' []]]. eapply typing_wf_local in t. now rewrite app_context_nil_l in t. @@ -4776,17 +4714,14 @@ Proof. constructor; try constructor; auto; try apply IHX. { now apply (fresh_global_map (Σ := Σ0)). } destruct d; cbn in *. - * cbn. red. move: ond; rewrite /on_constant_decl. - destruct c as [type [body|] univs] => /=. - intros Hty; eapply (pcuic_expand_lets (Σ0, univs) [] _ _ X Hty IHX). - intros [s Hty]. exists s. - exact (pcuic_expand_lets (Σ0, univs) [] _ _ X Hty IHX). + * apply lift_typing_f_impl with (1 := ond) => //. + intros ?? Hty; eapply (pcuic_expand_lets (Σ0, _) [] _ _ X Hty IHX). * generalize ond. intros []; econstructor; eauto. + cbn. eapply (Alli_mapi _ _ _).1, Alli_impl; tea. intros n idecl onind; generalize onind; intros []; econstructor; tea. { cbn. rewrite ind_arity_eq !trans_it_mkProd_or_LetIn //. } - { cbn. destruct onArity as [s Hty]. exists s. + { cbn. apply lift_typing_f_impl with (1 := onArity) => // ?? Hty. exact (pcuic_expand_lets (Σ0, ind_universes m) [] _ _ X Hty IHX). } { instantiate (1 := ind_cunivs). red in onConstructors. @@ -4799,11 +4734,11 @@ Proof. { eapply sorts_local_ctx_wf_local in on_cargs. eapply wf_local_closed_context in on_cargs. now move/onfvs_app: on_cargs; len. - rewrite cstr_eq in on_ctype. destruct on_ctype as [s Hs]. + rewrite cstr_eq in on_ctype. destruct on_ctype as (_ & s & Hs & _). eapply typing_wf_local in Hs. eapply weaken_wf_local; tea. } have oncindices: on_free_vars_terms (shiftnP (#|cstr_args cdecl| + #|ind_params m| + #|ind_bodies m|) xpred0) (cstr_indices cdecl). - { rewrite cstr_eq in on_ctype. destruct on_ctype as [s Hs]. + { rewrite cstr_eq in on_ctype. destruct on_ctype as (_ & s & Hs & _). eapply subject_is_open_term in Hs. move: Hs. rewrite !on_free_vars_it_mkProd_or_LetIn. move/and3P => [] _ _. rewrite /cstr_concl. @@ -4813,8 +4748,8 @@ Proof. { cbn; rewrite -cstr_args_length. rewrite context_assumptions_smash_context context_assumptions_map //. } { cbn; rewrite /trans_cstr_concl /trans_cstr_concl_head /cstr_concl_head. now len. } - { destruct on_ctype as [s Hty]. - exists s. + { destruct on_ctype as (_ & s & Hty & _). + repeat (eexists; tea). epose proof (pcuic_expand_lets (Σ0, ind_universes m) _ _ _ X Hty IHX). rewrite trans_arities_context. cbn. rewrite cstr_eq in X0. rewrite !trans_it_mkProd_or_LetIn in X0. @@ -4841,7 +4776,8 @@ Proof. rewrite expand_lets_ctx_tip /=. destruct univs => //. split. cbn in IHc. apply IHc, on_cargs. - destruct on_cargs as [hs ht]. red in ht. + destruct on_cargs as [hs hj]. + apply lift_sorting_f_it_impl with (f := fun t => _ (_ t)) (tu := hj) => // ht. have fvsc : on_free_vars_ctx (shiftnP (#|ind_params m| + #|ind_bodies m|) xpred0) c. { eapply typing_wf_local, wf_local_closed_context in ht. move/onfvs_app: ht. now len. } @@ -4860,14 +4796,15 @@ Proof. rewrite -(trans_smash_context (shiftnP (#|ind_params m| + #|ind_bodies m|) xpred0) []) //. rewrite -[trans_local (ind_params m) ++ _]trans_local_app. rewrite -[_ ++ _]trans_local_app. - intros ci. eapply (ctx_inst_expand_lets (Σ := (Σ0, ind_universes m))) in ci. + intros ci. + eapply (ctx_inst_expand_lets (Σ := (Σ0, ind_universes m))) in ci; cbn in ci. rewrite List.rev_involutive in ci. rewrite expand_lets_ctx_lift_context_cancel in ci. rewrite -(trans_expand_lets_map (shiftnP (#|ind_params m| + #|ind_bodies m|) xpred0)) //. rewrite shiftnP_add Nat.add_assoc //. rewrite -(trans_lift_context (shiftnP #|ind_params m| xpred0)). { rewrite ind_arity_eq in onArity. - destruct onArity as [s Hs]. + destruct onArity as (_ & s & Hs & _). eapply subject_is_open_term in Hs. move: Hs; rewrite !on_free_vars_it_mkProd_or_LetIn /=. move/and3P => [] _. now rewrite shiftnP0. } @@ -4880,7 +4817,7 @@ Proof. eapply weaken_wf_local; tea. eapply wf_arities_context'; tea. rewrite ind_arity_eq in onArity. - destruct onArity as [s Hs]. + destruct onArity as (_ & s & Hs & _). rewrite -it_mkProd_or_LetIn_app in Hs. eapply type_it_mkProd_or_LetIn_inv in Hs as [? [? [Hs _]]]. eapply PCUICClosedConv.sorts_local_ctx_All_local_env in Hs; eauto. @@ -4898,7 +4835,7 @@ Proof. rewrite -(trans_expand_lets (shiftnP (#|ind_params m| + #|ind_bodies m|) xpred0)) //. rewrite shiftnP_add Nat.add_assoc //. { rewrite cstr_eq in on_ctype. - destruct on_ctype as [s Hs]. + destruct on_ctype as (_ & s & Hs & _). eapply subject_is_open_term in Hs. len in Hs. move: Hs; rewrite !on_free_vars_it_mkProd_or_LetIn. move/and3P => [] onpars onargs onconcl. @@ -4909,7 +4846,7 @@ Proof. eapply (positive_cstr_trans _ _ _ (shiftnP #|ind_bodies m| xpred0)) in on_ctype_positive. exact on_ctype_positive. rewrite cstr_eq in on_ctype. - destruct on_ctype as [s Hs]. + destruct on_ctype as (_ & s & Hs & _). eapply subject_is_open_term in Hs. len in Hs. move: Hs; rewrite !on_free_vars_it_mkProd_or_LetIn. @@ -4975,7 +4912,9 @@ Proof. { do 3 eapply All2_map. eapply All2_map_inv in b. eapply All2_All in b. eapply All_All2_refl. - pose proof (ctx_inst_wt (on_cindices onc)). + pose proof (on_cindices onc) as X0. + eapply ctx_inst_impl in X0. 2: intros ??; apply /unlift_TermTyp. + apply ctx_inst_wt in X0. move/forallb_All: oncindices => hi. eapply All_mix in b; tea. clear hi. eapply (All_mix X0) in b; clear X0. @@ -5038,7 +4977,7 @@ Proof. depelim onConstructors. depelim onConstructors. have wfargs : wf_local (Σ0, ind_universes m) (arities_context (ind_bodies m),,, ind_params m,,, cstr_args c). { destruct o. eapply sorts_local_ctx_wf_local in on_cargs => //. - rewrite cstr_eq in on_ctype. destruct on_ctype as [s Hs]. + rewrite cstr_eq in on_ctype. destruct on_ctype as (_ & s & Hs & _). eapply typing_wf_local in Hs. eapply weaken_wf_local; tea. } have oncargs: on_free_vars_ctx (shiftnP (#|ind_params m| + #|ind_bodies m|) xpred0) (cstr_args c). @@ -5093,8 +5032,7 @@ Proof. } { cbn. move: ind_sorts. rewrite /check_ind_sorts. - destruct Universe.is_prop => //. - destruct Universe.is_sprop => //. + destruct Sort.to_family => //. intuition auto. rewrite /PCUICLookup.global_ext_constraints. destruct indices_matter => //. @@ -5787,22 +5725,15 @@ Proof. epose proof (PCUICInstConv.on_free_vars_ctx_inst_case_context_nil _ _ (puinst p) _ p0 p2). revert X3 H3; clear. rewrite /inst_case_predicate_context. induction 1; try constructor; eauto. - 2:destruct t1. - all:cbn in * |-. + all:unfold PCUICInduction.on_local_decl in t0; cbn in * |-. all:rewrite on_free_vars_ctx_snoc => /andP[] HΓ hna; cbn in hna; rewrite -> shiftnP_add in hna; len in hna; try len in t0; try len in t1; try len in t2; specialize (IHX3 HΓ). - + specialize (t0 hna). - rewrite PCUICInductives.expand_lets_k_ctx_snoc; constructor; eauto; cbn -[expand_lets_k] in *. - specialize (t0 Δ Γ' (Δ' ,,, Γ)); rewrite app_context_assoc in t0; specialize (t0 eq_refl). - len in t0. rewrite PCUICInductives.expand_lets_ctx_app app_context_assoc in t0; eauto. - + move/andP: hna => /= [] hb ht0. - len in n; len in t0. specialize (t0 ht0). specialize (n hb). - specialize (n Δ Γ' (Δ' ,,, Γ)); rewrite app_context_assoc in n; specialize (n eq_refl). - specialize (t0 Δ Γ' (Δ' ,,, Γ)); rewrite app_context_assoc in t0; specialize (t0 eq_refl). - len in n. len in t0. - rewrite PCUICInductives.expand_lets_ctx_app app_context_assoc in n, t0; eauto. - rewrite PCUICInductives.expand_lets_k_ctx_snoc; constructor; eauto; cbn -[expand_lets_k] in *. - split; eauto. + 2: move/andP: hna => [] hna0 hna1. + all: rewrite PCUICInductives.expand_lets_k_ctx_snoc; constructor; eauto; cbn -[expand_lets_k] in *. + all: eapply lift_wf_term_it_impl with (1 := t0) => // IH. + all: forward IH by assumption. + all: specialize (IH Δ Γ' (Δ' ,,, Γ)); forward IH by now rewrite app_context_assoc. + all: rewrite PCUICInductives.expand_lets_ctx_app app_context_assoc in IH; len in IH. * cbn. inv_on_free_vars. forward H2. { rewrite shiftnP_add in p1 => //. len in p1. now len. } specialize (H2 Δ Γ' (Δ' ,,, inst_case_context (pparams p) (puinst p) (pcontext p))). @@ -5890,6 +5821,7 @@ Lemma All_local_env_on_free_vars_ctx P P' Γ : Proof. induction 1; cbn; constructor; eauto; move: H; rewrite alli_app => /andP[] hΓ /= hass; len in hass; intuition eauto; unfold on_free_vars_decl, test_decl in *; cbn in *; rtoProp; intuition auto. + all: apply lift_wf_term_it_impl with (1 := t0) => //. Qed. Lemma trans_ne_nf Σ Γ T : @@ -5939,11 +5871,9 @@ Proof. eapply (PCUICInstConv.on_free_vars_ctx_inst_case_context_nil _ _ (puinst p)) in H5. 2:{ eapply All_forallb. solve_all. } eapply All_local_env_on_free_vars_ctx in X3; tea. - eapply All_local_env_impl; tea. cbn. intros; intuition eauto. - red. red in X. destruct T => //=; intuition eauto. - rewrite shiftnP_add in b0. len in a0. now rewrite trans_local_app in a0. - rewrite shiftnP_add in b1. len in a. now rewrite trans_local_app in a. - rewrite shiftnP_add in b. len in a. now rewrite trans_local_app in a. + eapply All_local_env_impl; tea. cbn. intros. + apply lift_wf_term_f_impl with (u := j_univ j) (1 := X); intros t (Ht & ont). + rewrite shiftnP_add in ont. len in Ht. now rewrite trans_local_app in Ht. * rewrite -(trans_local_case_context Γ0 p) //. solve_all. cbn; eauto. rewrite trans_local_app in H2. eapply H2. len. now rewrite shiftnP_add in H4. - rtoProp; intuition auto. eapply nf_lam; eauto. @@ -6082,4 +6012,4 @@ Proof. cbn. rewrite /firstorder_ind /= trans_lookup. destruct lookup_env => //. destruct g => //=. eapply trans_firstorder_mutind, trans_firstorder_env. -Qed. \ No newline at end of file +Qed. diff --git a/pcuic/theories/PCUICFirstorder.v b/pcuic/theories/PCUICFirstorder.v index cd3469598..e0c4d1bed 100644 --- a/pcuic/theories/PCUICFirstorder.v +++ b/pcuic/theories/PCUICFirstorder.v @@ -56,7 +56,7 @@ Section firstorder. (List.rev (c.(cstr_args) ++ mind.(ind_params)))%list. Definition firstorder_oneind mind (ind : one_inductive_body) := - forallb (firstorder_con mind) ind.(ind_ctors) && negb (Universe.is_level (ind_sort ind)). + forallb (firstorder_con mind) ind.(ind_ctors) && negb (Sort.is_level (ind_sort ind)). Definition firstorder_mutind (mind : mutual_inductive_body) := (* if forallb (fun decl => firstorder_type decl.(decl_type)) mind.(ind_params) then *) @@ -102,7 +102,7 @@ Context {cf : config.checker_flags}. Definition isPropositionalArity ar := match destArity [] ar with - | Some (_, s) => is_propositional s + | Some (_, s) => Sort.is_propositional s | None => false end. @@ -193,7 +193,9 @@ Lemma isType_context_conversion {Σ : global_env_ext} {wfΣ : wf Σ} {Γ Δ} {T} wf_local Σ Δ -> isType Σ Δ T. Proof using Type. - intros [s Hs]. exists s. eapply context_conversion; tea. now eapply ws_cumul_ctx_pb_forget. + intros HT ??. + apply lift_typing_impl with (1 := HT); intros ?? H. + eapply context_conversion; tea. now eapply ws_cumul_ctx_pb_forget. Qed. Lemma typing_spine_arity_spine {Σ : global_env_ext} {wfΣ : wf Σ} Γ Δ args T' i u pars : @@ -332,7 +334,7 @@ Proof using Type. now eapply isType_tLetIn_red in isty; pcuic. - depelim hi. solve_discr. specialize (i1 hd). specialize (IHhsp i1). - destruct (validity t) as [s Hs]. eapply inversion_mkApps in Hs as [? [hi _]]. + destruct (validity t) as (_ & s & Hs & _). eapply inversion_mkApps in Hs as [? [hi _]]. eapply inversion_Ind in hi as [mdecl [idecl [decli [? ?]]]]. econstructor; tea. 2:{ eapply IHhsp. eapply isType_apply in isty; tea. } now eapply isType_ws_cumul_pb_refl. eauto. @@ -639,7 +641,7 @@ Proof using Type. now eapply invert_cumul_sort_ind in Hcumul. + exfalso. eapply inversion_Prod in Hty as (? & ? & ? & ? & Hcumul); eauto. now eapply invert_cumul_sort_ind in Hcumul. - + exfalso. eapply inversion_Lambda in Hty as (? & ? & ? & ? & Hcumul); eauto. + + exfalso. eapply inversion_Lambda in Hty as (? & ? & ? & Hcumul); eauto. now eapply invert_cumul_prod_ind in Hcumul. + exfalso. eapply inversion_Ind in Hty as (? & ? & ? & ? & ? & ?); eauto. eapply PCUICInductives.declared_inductive_type in d. @@ -722,7 +724,7 @@ Proof. intros Ha H. induction H in t', Ha |- using firstorder_value_inds. eapply eq_term_upto_univ_napp_mkApps_l_inv in Ha as (? & ? & [] & ->). invs e. repeat f_equal. - - now eapply eq_univ_make. + - now eapply cmp_universe_instance_eq. - revert x0 a. clear - H0. induction H0; intros; invs a; f_equal; eauto. Qed. diff --git a/pcuic/theories/PCUICGeneration.v b/pcuic/theories/PCUICGeneration.v index 0d036ea34..e844a3790 100644 --- a/pcuic/theories/PCUICGeneration.v +++ b/pcuic/theories/PCUICGeneration.v @@ -29,13 +29,12 @@ Section Generation. Proof using Type. intros Ht Hsp. revert t Ht. induction Hsp; simpl; auto. - intros t Ht. eapply type_Cumul; eauto. eapply i.π2. + intros t Ht. eapply type_Cumul; eauto. eapply i.2.π2.1. intros. specialize (IHHsp (tApp t0 hd)). apply IHHsp. - destruct i as [s Hs]. - eapply type_App; eauto. - eapply type_Cumul; eauto. + eapply type_App; eauto. eapply i.2.π2.1. + eapply type_Cumul; eauto. eapply i.2.π2.1. Qed. Lemma type_it_mkLambda_or_LetIn : @@ -48,12 +47,12 @@ Section Generation. - assumption. - simpl. cbn. eapply ih. simpl in h. pose proof (typing_wf_local h) as hc. - dependent induction hc. - econstructor; try eassumption. exact t0.π2. + apply All_local_env_tip in hc as [hc ona]. + econstructor; try eassumption. - simpl. cbn. eapply ih. pose proof (typing_wf_local h) as hc. cbn in hc. - dependent induction hc; - econstructor; try eassumption. exact t0.π2. + apply All_local_env_tip in hc as [hc ona]. + econstructor; try eassumption. Qed. End Generation. diff --git a/pcuic/theories/PCUICGlobalEnv.v b/pcuic/theories/PCUICGlobalEnv.v index 207ddc8a2..3abe5276e 100644 --- a/pcuic/theories/PCUICGlobalEnv.v +++ b/pcuic/theories/PCUICGlobalEnv.v @@ -94,13 +94,12 @@ Section DeclaredInv. ind_npars mdecl = context_assumptions mdecl.(ind_params). Proof using wfΣ. intros h. - eapply declared_minductive_to_gen in h. + unshelve eapply declared_minductive_to_gen in h; tea. unfold declared_minductive in h. eapply lookup_on_global_env in h; tea. - destruct h as [Σ' [ext wfΣ' decl']]. + destruct h as (Σ' & ext & wfΣ' & decl'). red in decl'. destruct decl' as [h ? ? ?]. now rewrite onNpars. - Unshelve. all: eauto. Qed. End DeclaredInv. diff --git a/pcuic/theories/PCUICInductiveInversion.v b/pcuic/theories/PCUICInductiveInversion.v index 31a268911..89d146a61 100644 --- a/pcuic/theories/PCUICInductiveInversion.v +++ b/pcuic/theories/PCUICInductiveInversion.v @@ -63,8 +63,8 @@ Lemma type_tFix_inv {cf:checker_flags} (Σ : global_env_ext) Γ mfix idx T : wf Proof. intros wfΣ H. depind H. - unfold unfold_fix. rewrite e. - specialize (nth_error_all e a0) as [s Hs]. - specialize (nth_error_all e a1) as Hty. + specialize (nth_error_all e a0) as (_ & s & Hs & _). + specialize (nth_error_all e a1) as (Hb & _). cbn in Hb, Hs. simpl. destruct decl as [name ty body rarg]; simpl in *. clear e. @@ -79,7 +79,7 @@ Proof. rewrite rev_mapi. unfold mapi. assert (#|mfix| >= #|List.rev mfix|) by (rewrite List.rev_length; lia). assert (He :0 = #|mfix| - #|List.rev mfix|) by (rewrite List.rev_length; auto with arith). - rewrite {3}He. clear He. revert H. + rewrite {2}He. clear He. revert H. assert (forall i, i < #|List.rev mfix| -> nth_error (List.rev mfix) i = nth_error mfix (#|List.rev mfix| - S i)). { intros. rewrite nth_error_rev. 1: auto. now rewrite List.rev_length List.rev_involutive. } @@ -116,11 +116,8 @@ Qed. Lemma subslet_cofix {cf:checker_flags} (Σ : global_env_ext) Γ mfix : wf_local Σ Γ -> cofix_guard Σ Γ mfix -> - All (fun d : def term => ∑ s : Universe.t, Σ;;; Γ |- dtype d : tSort s) mfix -> - All - (fun d : def term => - Σ;;; Γ ,,, fix_context mfix |- dbody d - : lift0 #|fix_context mfix| (dtype d)) mfix -> + All (on_def_type (lift_typing typing Σ) Γ) mfix -> + All (on_def_body (lift_typing typing Σ) (fix_context mfix) Γ) mfix -> wf_cofixpoint Σ mfix -> subslet Σ Γ (cofix_subst mfix) (fix_context mfix). Proof. intros wfΓ hguard types bodies wfcofix. @@ -130,7 +127,7 @@ Proof. rewrite rev_mapi. unfold mapi. assert (#|mfix| >= #|List.rev mfix|) by (rewrite List.rev_length; lia). assert (He :0 = #|mfix| - #|List.rev mfix|) by (rewrite List.rev_length; auto with arith). - rewrite {3}He. clear He. revert H. + rewrite {2}He. clear He. revert H. assert (forall i, i < #|List.rev mfix| -> nth_error (List.rev mfix) i = nth_error mfix (#|List.rev mfix| - S i)). { intros. rewrite nth_error_rev. 1: auto. now rewrite List.rev_length List.rev_involutive. } @@ -165,7 +162,7 @@ Lemma type_tCoFix_inv {cf:checker_flags} (Σ : global_env_ext) Γ mfix idx T : w Proof. intros wfΣ H. depind H. - exists decl. - specialize (nth_error_all e a1) as Hty. + specialize (nth_error_all e a1) as (Hty & _). cbn in Hty. destruct decl as [name ty body rarg]; simpl in *. intuition auto. * eapply (substitution (s := cofix_subst mfix) (Δ := [])) in Hty. simpl; eauto with wf. @@ -258,7 +255,7 @@ Section OnConstructor. eapply sorts_local_ctx_wf_local in X => //. clear X. eapply weaken_wf_local => //. eapply wf_arities_context; eauto. eapply declc. - now eapply onParams. + now eapply onParams in onmind. Qed. Lemma on_constructor_subst : @@ -284,13 +281,13 @@ Section OnConstructor. eapply sorts_local_ctx_wf_local in X => //. clear X. eapply weaken_wf_local => //. eapply wf_arities_context; eauto; eapply declc. - now eapply onParams. - destruct (on_ctype onc). - rewrite onc.(cstr_eq) in t. - rewrite -it_mkProd_or_LetIn_app in t. - eapply inversion_it_mkProd_or_LetIn in t => //. - unfold cstr_concl_head in t. simpl in t. - eapply inversion_mkApps in t as [A [ta sp]]. + now eapply onParams in onmind. + destruct (on_ctype onc) as (_ & s & Hs & _). + rewrite onc.(cstr_eq) in Hs. + rewrite -it_mkProd_or_LetIn_app in Hs. + eapply inversion_it_mkProd_or_LetIn in Hs => //. + unfold cstr_concl_head in Hs. simpl in Hs. + eapply inversion_mkApps in Hs as [A [ta sp]]. eapply inversion_Rel in ta as [decl [wfΓ [nth cum']]]. rewrite nth_error_app_ge in nth. len. len in nth. @@ -305,7 +302,7 @@ Section OnConstructor. rewrite Hdecl in cum'; clear Hdecl. assert(closed (ind_type idecl)). { pose proof (oib.(onArity)). rewrite (oib.(ind_arity_eq)) in X0 |- *. - destruct X0 as [s Hs]. now apply subject_closed in Hs. } + destruct X0 as (_ & s' & Hs & _). now apply subject_closed in Hs. } rewrite lift_closed in cum' => //. eapply typing_spine_strengthen in sp; simpl. 3:tea. @@ -317,7 +314,7 @@ Section OnConstructor. clear Hlen'. rewrite [_ ,,, _]app_context_assoc in Hinst. now exists inst. - apply infer_typing_sort_impl with id oib.(onArity); intros Hs. + apply lift_typing_impl with (1 := oib.(onArity)); intros ?? Hs. now eapply weaken_ctx in Hs. Qed. End OnConstructor. @@ -473,7 +470,7 @@ Lemma mkApps_ind_typing_spine {cf:checker_flags} Σ Γ Γ' ind i (mkApps (tInd ind' i') args') -> ∑ instsubst, [× make_context_subst (List.rev Γ') inst [] = Some instsubst, - #|inst| = context_assumptions Γ', ind = ind', R_ind_universes Σ ind #|args| i i', + #|inst| = context_assumptions Γ', ind = ind', cmp_ind_universes Σ ind #|args| i i', All2 (fun par par' => Σ ;;; Γ ⊢ par = par') (map (subst0 instsubst) args) args' & subslet Σ Γ instsubst Γ']. Proof. @@ -607,7 +604,7 @@ Lemma Construct_Ind_ind_eq {cf:checker_flags} {Σ} (wfΣ : wf Σ.1): declared_constructor Σ.1 (i, n) mdecl idecl cdecl -> (i = i') * (* Universe instances match *) - R_ind_universes Σ i (context_assumptions (ind_params mdecl) + #|cstr_indices cdecl|) u u' * + cmp_ind_universes Σ i (context_assumptions (ind_params mdecl) + #|cstr_indices cdecl|) u u' * consistent_instance_ext Σ (ind_universes mdecl) u * consistent_instance_ext Σ (ind_universes mdecl) u' * (#|args| = (ind_npars mdecl + context_assumptions cdecl.(cstr_args))%nat) * @@ -718,7 +715,7 @@ Proof. eapply weaken_wf_local => //. rewrite -subst_instance_app_ctx. apply a. - - exists (map (subst_instance_univ u') x). split. + - exists (map (subst_instance_sort u') x). split. * move/onParams: onmind. rewrite /on_context. pose proof (@wf_local_instantiate _ Σ (InductiveDecl mdecl) (ind_params mdecl) u'). move=> H'. eapply X in H'; eauto. @@ -789,7 +786,7 @@ Lemma Construct_Ind_ind_eq' {cf:checker_flags} {Σ} (wfΣ : wf Σ.1): declared_constructor Σ.1 (i, n) mdecl idecl cdecl × (i = i') * (* Universe instances match *) - R_ind_universes Σ i (context_assumptions (ind_params mdecl) + #|cstr_indices cdecl|) u u' * + cmp_ind_universes Σ i (context_assumptions (ind_params mdecl) + #|cstr_indices cdecl|) u u' * consistent_instance_ext Σ (ind_universes mdecl) u * consistent_instance_ext Σ (ind_universes mdecl) u' * (#|args| = (ind_npars mdecl + context_assumptions cdecl.(cstr_args))%nat) * @@ -1020,7 +1017,7 @@ Lemma isType_mkApps_Ind_smash {cf:checker_flags} {Σ Γ ind u params args} {wfΣ Proof. move=> isdecl isty hpars. pose proof (isType_wf_local isty). - destruct isty as [s Hs]. + destruct isty as (_ & s & Hs & _). eapply invert_type_mkApps_ind in Hs as [sp cu]; tea. move=> parctx argctx. erewrite ind_arity_eq in sp. @@ -1209,7 +1206,7 @@ Derive Signature for positive_cstr. Lemma positive_cstr_it_mkProd_or_LetIn mdecl i Γ Δ t : positive_cstr mdecl i Γ (it_mkProd_or_LetIn Δ t) -> - All_local_env (fun Δ ty _ => positive_cstr_arg mdecl (Γ ,,, Δ) ty) + All_local_env (fun Δ j => positive_cstr_arg mdecl (Γ ,,, Δ) (j_typ j)) (smash_context [] Δ) * positive_cstr mdecl i (Γ ,,, smash_context [] Δ) (expand_lets Δ t). Proof. @@ -1232,7 +1229,7 @@ Proof. specialize (X Δ ltac:(len; lia) _ _ H). simpl; len. destruct X; split; auto. simpl. - eapply All_local_env_app; split. + eapply All_local_env_app. constructor; auto. eapply (All_local_env_impl _ _ _ a). intros; auto. now rewrite app_context_assoc. simpl. @@ -1389,7 +1386,7 @@ Lemma ws_cumul_pb_mkApps_eq {cf} {Σ} {wfΣ : wf Σ} Γ f f' u u' : is_closed_context Γ -> is_open_term Γ f -> is_open_term Γ f' -> - eq_term_upto_univ_napp Σ (eq_universe Σ) (leq_universe Σ) #|u| f f' -> + compare_term_napp Σ Σ Cumul #|u| f f' -> ws_cumul_pb_terms Σ Γ u u' -> Σ ;;; Γ ⊢ mkApps f u ≤ mkApps f' u'. Proof. @@ -1467,7 +1464,7 @@ Qed. Lemma positive_cstr_arg_subst {cf} {Σ} {wfΣ : wf Σ} {ind mdecl idecl Γ t u u'} : declared_inductive Σ ind mdecl idecl -> consistent_instance_ext Σ (ind_universes mdecl) u -> - R_opt_variance (eq_universe Σ) (leq_universe Σ) (ind_variance mdecl) u u' -> + cmp_opt_variance (compare_universe Σ) Cumul (match ind_variance mdecl with Some var => Variance var | None => AllEqual end) u u' -> closed_ctx (ind_arities mdecl ,,, Γ)@[u] -> Σ ;;; subst_instance u (ind_arities mdecl) ,,, subst_instance u Γ ⊢ (subst_instance u t) ≤ (subst_instance u' t) -> positive_cstr_arg mdecl Γ t -> @@ -1526,7 +1523,7 @@ Proof. rewrite /ind_subst !inds_spec !rev_mapi !nth_error_mapi. unshelve epose proof (declm' := declared_minductive_to_gen declm); eauto. rewrite H2 /=. simpl. constructor. simpl. - unfold R_global_instance, R_global_instance_gen. simpl. + unfold cmp_global_instance, cmp_global_instance_gen. simpl. assert(declared_inductive Σ {| inductive_mind := inductive_mind ind; inductive_ind := Nat.pred #|ind_bodies mdecl| - (k - #|Γ|) |} mdecl i). @@ -1589,13 +1586,10 @@ Lemma positive_cstr_closed_args_subst_arities {cf} {Σ} {wfΣ : wf Σ} {u u' Γ} declared_inductive Σ ind mdecl idecl -> consistent_instance_ext Σ (ind_universes mdecl) u -> on_constructor cumulSpec0 (lift_typing typing) (Σ.1, ind_universes mdecl) mdecl i idecl ind_indices cdecl cs -> - R_opt_variance (eq_universe Σ) (leq_universe Σ) (ind_variance mdecl) u u' -> + cmp_opt_variance (compare_universe Σ) Cumul (match ind_variance mdecl with Some var => Variance var | None => AllEqual end) u u' -> closed_ctx (subst_instance u (ind_params mdecl)) -> wf_local Σ (subst_instance u (ind_arities mdecl ,,, smash_context [] (ind_params mdecl) ,,, Γ)) -> - All_local_env - (fun (Γ : PCUICEnvironment.context) (t : term) (_ : typ_or_sort) => - positive_cstr_arg mdecl ([] ,,, (smash_context [] (ind_params mdecl) ,,, Γ)) t) - Γ -> + All_local_env (fun Γ j => positive_cstr_arg mdecl ([] ,,, (smash_context [] (ind_params mdecl) ,,, Γ)) (j_typ j)) Γ -> assumption_context Γ -> ws_cumul_ctx_pb_rel Cumul Σ (subst_instance u (ind_arities mdecl) ,,, subst_instance u @@ -1618,10 +1612,7 @@ Proof. (subst_instance u (ind_arities mdecl) ,,, subst_instance u (smash_context [] (ind_params mdecl) ,,, Γ)) (subst_instance u t)). - { rewrite -subst_instance_app_ctx. - destruct l as [s Hs]. exists s. - move: Hs. cbn. - now rewrite app_context_assoc. } + { rewrite -subst_instance_app_ctx app_context_assoc //. } depelim a. all:constructor. - eapply IHcpos. auto. now depelim ass. apply cv. @@ -1645,7 +1636,7 @@ Lemma positive_cstr_closed_args {cf} {Σ} {wfΣ : wf Σ} {u u'} {ind mdecl idecl cdecl} : declared_constructor Σ ind mdecl idecl cdecl -> consistent_instance_ext Σ (ind_universes mdecl) u -> - R_opt_variance (eq_universe Σ) (leq_universe Σ) (ind_variance mdecl) u u' -> + cmp_opt_variance (compare_universe Σ) Cumul (match ind_variance mdecl with Some var => Variance var | None => AllEqual end) u u' -> ws_cumul_ctx_pb_rel Cumul Σ (subst_instance u (ind_arities mdecl) ,,, subst_instance u (smash_context [] (PCUICEnvironment.ind_params mdecl))) @@ -1729,13 +1720,13 @@ Section Betweenu. Definition betweenu_level_expr (s : LevelExpr.t) := betweenu_level (LevelExpr.get_level s). - Definition betweenu_universe0 (u : LevelAlgExpr.t) := + Definition betweenu_universe (u : Universe.t) := LevelExprSet.for_all betweenu_level_expr u. - Definition betweenu_universe (u : Universe.t) := + Definition betweenu_sort (u : sort) := match u with - | Universe.lProp | Universe.lSProp => true - | Universe.lType l => betweenu_universe0 l + | sProp | sSProp => true + | sType u => betweenu_universe u end. Definition betweenu_instance (u : Instance.t) := @@ -2039,7 +2030,7 @@ Proof. rewrite is_closed_context_subst_instance app_context_nil_l //. Qed. -(** Morally, if variance_universes l = v i i' and R_universe_instance_variance l u u' then +(** Morally, if variance_universes l = v i i' and cmp_universe_instance_variance l u u' then i and i' can be substituted respectively by u and u'. The hard part is to show that (Σ.1, v) can also be turned into Σ by instanciating i and i' by u and u'. @@ -2051,7 +2042,7 @@ Lemma ws_cumul_pb_inst_variance {cf} {le} {Σ} {wfΣ : wf Σ} {mdecl l v i i' u variance_universes (PCUICEnvironment.ind_universes mdecl) l = Some (v, i, i') -> consistent_instance_ext Σ (ind_universes mdecl) u -> consistent_instance_ext Σ (ind_universes mdecl) u' -> - R_universe_instance_variance (eq_universe Σ) (leq_universe Σ) l u u' -> + cmp_universe_instance_variance (compare_universe Σ) Cumul l u u' -> forall t t', (Σ.1, v) ;;; Γ@[i] ⊢ t@[i] ≤[le] t'@[i'] -> Σ ;;; Γ@[u] ⊢ t@[u] ≤[le] t'@[u']. @@ -2119,16 +2110,16 @@ Proof. induction l in u, u', Ru, lenu, lenlu |- *. simpl in *. destruct u, u'; intro; rewrite ConstraintSetFact.empty_iff //. destruct u, u' => //; simpl in *. - destruct Ru as [Ra Rl]. - specialize (IHl u u' Rl). do 2 forward IHl by lia. + depelim Ru. rename H into Ra. + specialize (IHl u u' Ru). do 2 forward IHl by lia. destruct a => //; intros x; rewrite ConstraintSetFact.add_iff; intros [<-|inx]; auto. - + do 7 red in Ra; rewrite checku in Ra; + + do 5 red in Ra; rewrite checku in Ra; specialize (Ra _ sat); simpl in Ra. constructor. lia. - + do 6 red in Ra. rewrite checku in Ra. + + do 4 red in Ra. rewrite checku in Ra. specialize (Ra _ sat). - constructor. now rewrite !Universes.LevelAlgExpr.val_make in Ra. + constructor. now rewrite !Universes.Universe.val_make in Ra. Qed. Lemma All2_fold_inst {cf} {le} {Σ} {wfΣ : wf Σ} mdecl l v i i' u u' Γ' Γ : @@ -2137,7 +2128,7 @@ Lemma All2_fold_inst {cf} {le} {Σ} {wfΣ : wf Σ} mdecl l v i i' u u' Γ' Γ : consistent_instance_ext Σ (ind_universes mdecl) u -> consistent_instance_ext Σ (ind_universes mdecl) u' -> variance_universes (PCUICEnvironment.ind_universes mdecl) l = Some (v, i, i') -> - R_universe_instance_variance (eq_universe Σ) (leq_universe Σ) l u u' -> + cmp_universe_instance_variance (compare_universe Σ) Cumul l u u' -> ws_cumul_ctx_pb_rel le (Σ.1, v) (subst_instance i Γ') (subst_instance i Γ) (subst_instance i' Γ) -> ws_cumul_ctx_pb_rel le Σ (subst_instance u Γ') (subst_instance u Γ) (subst_instance u' Γ). Proof. @@ -2319,12 +2310,36 @@ Proof. eapply on_free_vars_ctx_impl => //. Qed. +Lemma cmp_global_instance_ind_inv {cf} {Σ} {wfΣ : wf Σ} cmp_universe pb ind mdecl idecl napp u u' : + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe pb) -> + declared_inductive Σ ind mdecl idecl -> + cmp_global_instance Σ cmp_universe pb (IndRef ind) napp u u' -> + ((context_assumptions (ind_indices idecl) + context_assumptions (ind_params mdecl) <= napp × + ∑ l, ind_variance mdecl = Some l × + cmp_universe_instance_variance cmp_universe pb l u u') + + cmp_universe_instance (cmp_universe Conv) u u')%type. +Proof. + intros Hsub decli. + destruct (on_declared_inductive decli) as [onmind oib]. + unshelve epose proof (decli' := declared_inductive_to_gen decli); eauto. + unfold cmp_global_instance, cmp_global_instance_gen, global_variance. + rewrite (declared_inductive_lookup_gen decli'). + rewrite oib.(ind_arity_eq). + rewrite !destArity_it_mkProd_or_LetIn. simpl. + rewrite app_context_nil_l context_assumptions_app. + elim: leb_spec_Set => // comp. 2: now right. + destruct ind_variance eqn:indv => //. 2: now right. + move/cmp_opt_variance_var_dec => [|H]. 1: now right. + left. + now repeat eexists. +Qed. + Lemma inductive_cumulative_indices {cf} {Σ} {wfΣ : wf Σ} : forall {ind mdecl idecl u u' napp}, declared_inductive Σ ind mdecl idecl -> consistent_instance_ext Σ (ind_universes mdecl) u -> consistent_instance_ext Σ (ind_universes mdecl) u' -> - R_global_instance Σ (eq_universe Σ) (leq_universe Σ) (IndRef ind) napp u u' -> + cmp_ind_universes Σ ind napp u u' -> forall Γ pars pars' parsubst parsubst', spine_subst Σ Γ pars parsubst (subst_instance u (ind_params mdecl)) -> spine_subst Σ Γ pars' parsubst' (subst_instance u' (ind_params mdecl)) -> @@ -2337,11 +2352,11 @@ Lemma inductive_cumulative_indices {cf} {Σ} {wfΣ : wf Σ} : Proof. intros * decli. destruct (on_declared_inductive decli) as [onmind oib]. - intros cu cu' Ru Γ * spu spu' cpars *. move: Ru. + intros cu cu' Ru Γ * spu spu' cpars *. unshelve epose proof (decli' := declared_inductive_to_gen decli); eauto. assert (onu : on_udecl_prop Σ (ind_universes mdecl)). { eapply (weaken_lookup_on_global_env' _ _ _ wfΣ (proj1 decli')). } - unfold R_global_instance, R_global_instance_gen. + unfold cmp_global_instance, cmp_global_instance_gen. assert (closed_ctx (subst_instance u (PCUICEnvironment.ind_params mdecl))) as clpu. @@ -2373,22 +2388,12 @@ Proof. rewrite closedn_ctx_app /=. rewrite (closed_wf_local _ (spine_dom_wf _ _ _ _ _ spu)) /=. eapply closedn_ctx_upwards; tea. lia. } - destruct global_variance_gen eqn:gv. - { move:gv. - simpl. unfold lookup_inductive. - rewrite (declared_inductive_lookup_gen decli'). - rewrite oib.(ind_arity_eq). - rewrite !destArity_it_mkProd_or_LetIn. simpl. - rewrite app_context_nil_l context_assumptions_app. - elim: leb_spec_Set => // comp. - destruct ind_variance eqn:indv => //. - move=> [=] eq. subst l0. - pose proof (oib.(onIndices)) as respv. + eapply cmp_global_instance_ind_inv in Ru as [(comp & l & indv & Ru) | Ru]; tea; tc. + { pose proof (oib.(onIndices)) as respv. rewrite indv in respv. simpl in respv. unfold ind_respects_variance in respv. destruct variance_universes as [[[v i] i']|] eqn:vu => //. - simpl => Ru. pose proof (onVariance onmind) as onvari. rewrite indv in onvari. eapply cumul_ctx_relSpec_Algo in respv. @@ -2479,6 +2484,7 @@ Proof. intros. pose proof (on_declared_constructor H) as [[onmind oib] [cs [hnth onc]]]. pose proof (onc.(on_cindices)). + eapply ctx_inst_impl in X. 2: intros ??; apply unlift_TermTyp. now eapply ctx_inst_open_terms in X. Qed. @@ -2535,15 +2541,16 @@ Lemma ws_cumul_ctx_pb_wf_local {cf:checker_flags} {Σ} {wfΣ : wf Σ.1} Γ Γ' wf_local Σ (Γ' ,,, Δ). Proof. intros wf wf' e. - eapply All_local_env_app; split => //. - eapply wf_local_app_inv in wf as []. + eapply All_local_env_app => //. + eapply All_local_env_app_inv in wf as []. eapply All_local_env_impl_ind; eauto. intros. - apply lift_typing_impl with (1 := X0); intros ? Hs. + rewrite -/(All_local_rel (lift_typing1 (typing Σ)) Γ' Γ0) /= in X, X0. + apply lift_typing_impl with (1 := X0); intros ?? Hs. eapply closed_context_cumulativity; tea. - eapply All_local_env_app; split=> //. + eapply All_local_env_app=> //. eapply ws_cumul_ctx_pb_app_same; tea. 2:now symmetry. - eapply wf_local_app in X; tea. + eapply All_local_env_app in X; tea. eauto with fvs. Qed. @@ -2552,7 +2559,7 @@ Lemma constructor_cumulative_indices {cf} {Σ} {wfΣ : wf Σ} : declared_constructor Σ c mdecl idecl cdecl -> consistent_instance_ext Σ (ind_universes mdecl) u -> consistent_instance_ext Σ (ind_universes mdecl) u' -> - R_global_instance Σ (eq_universe Σ) (leq_universe Σ) (IndRef c.1) napp u u' -> + cmp_ind_universes Σ c.1 napp u u' -> forall Γ pars pars' parsubst parsubst', spine_subst Σ Γ pars parsubst (subst_instance u (ind_params mdecl)) -> spine_subst Σ Γ pars' parsubst' (subst_instance u' (ind_params mdecl)) -> @@ -2574,13 +2581,12 @@ Lemma constructor_cumulative_indices {cf} {Σ} {wfΣ : wf Σ} : Proof. intros * declc. destruct (on_declared_constructor declc) as [[onmind oib] [cs [hnth onc]]]. - intros cu cu' Ru Γ * spu spu' cpars *. move: Ru. + intros cu cu' Ru Γ * spu spu' cpars *. unshelve epose proof (declc' := declared_constructor_to_gen declc); eauto. assert (onu : on_udecl_prop Σ (ind_universes mdecl)). { eapply (weaken_lookup_on_global_env' _ _ _ wfΣ (proj1 (proj1 declc'))). } have clΓ : is_closed_context Γ. { apply spine_dom_wf in spu; eauto with fvs. } - unfold R_global_instance, R_global_instance_gen. assert (closed_ctx (subst_instance u (PCUICEnvironment.ind_params mdecl))) as clpu. @@ -2599,23 +2605,14 @@ Proof. { apply spine_codom_wf in spu; eauto with fvs. } have clΓparsu' : is_closed_context (Γ ,,, (ind_params mdecl)@[u']). { apply spine_codom_wf in spu'; eauto with fvs. } - destruct global_variance_gen eqn:gv. - { move:gv. - simpl. unfold lookup_inductive. - rewrite (declared_inductive_lookup_gen declc'.p1). - rewrite oib.(ind_arity_eq). - rewrite !destArity_it_mkProd_or_LetIn. simpl. - rewrite app_context_nil_l context_assumptions_app. - elim: leb_spec_Set => // comp. - destruct ind_variance eqn:indv => //. - move=> [=] eq. subst l0. - pose proof (onc.(on_ctype_variance)) as respv. + eapply cmp_global_instance_ind_inv in Ru as [(comp & l & indv & Ru) | Ru]; eauto; tc. + 3: apply declc.p1. + { pose proof (onc.(on_ctype_variance)) as respv. specialize (respv _ indv). simpl in respv. unfold cstr_respects_variance in respv. destruct variance_universes as [[[v i] i']|] eqn:vu => //. destruct respv as [args idx]. - simpl => Ru. pose proof (onVariance onmind) as onvari. rewrite indv in onvari. assert (wf_local Σ @@ -2797,7 +2794,7 @@ Proof. { simpl. assert (wf_local Σ Γ) by apply spu. epose proof (on_constructor_inst declc _ cu) as [wfargs spinst]. - intros Ru; split. + split. { rewrite /pargctx /pargctx' /argctx /argctx'. rewrite !(smash_context_subst []). unshelve eapply (substitution_ws_cumul_ctx_pb_subst_conv (Γ'' := []) _ _ spu spu') => //. @@ -2819,7 +2816,7 @@ Proof. rewrite !subst_instance_app_ctx in wfargs. apply is_closed_context_weaken => //. rewrite -app_context_assoc in wfargs. - apply wf_local_app_inv in wfargs as []; eauto with fvs. + apply All_local_env_app_inv in wfargs as []; eauto with fvs. apply wf_local_closed_context in a; move: a. now rewrite !is_closed_subst_inst. } 2:now eapply conv_inds. @@ -2990,7 +2987,7 @@ Lemma projection_cumulative_indices {cf} {Σ} {wfΣ : wf Σ} : on_udecl_prop Σ (ind_universes mdecl) -> consistent_instance_ext Σ (ind_universes mdecl) u -> consistent_instance_ext Σ (ind_universes mdecl) u' -> - R_global_instance Σ (eq_universe Σ) (leq_universe Σ) (IndRef p.(proj_ind)) (ind_npars mdecl) u u' -> + cmp_ind_universes Σ p.(proj_ind) (ind_npars mdecl) u u' -> Σ ;;; projection_context p.(proj_ind) mdecl idecl u ⊢ subst_instance u pdecl.(proj_type) ≤ subst_instance u' pdecl.(proj_type). Proof. @@ -3014,25 +3011,11 @@ Proof. simpl in *. destruct s as [idecl' [idecl'nth _ _ pty pty']]. rewrite -pty. - unfold R_global_instance, R_global_instance_gen in Ru. - unfold global_variance_gen, lookup_inductive, lookup_minductive in Ru. destruct declp' as [[[? ?] ?] ?]. red in H0. - unfold lookup_inductive_gen, lookup_minductive_gen in Ru. - rewrite H0 H1 in Ru. - rewrite oib.(ind_arity_eq) in Ru. - rewrite !destArity_it_mkProd_or_LetIn /= in Ru. destruct p0 as [p0 _]. - destruct (context_assumptions _ <=? _) eqn:eq. - 2:{ - rewrite app_context_nil_l context_assumptions_app in eq. - eapply Nat.leb_nle in eq. - destruct onps. - apply length_nil in on_projs_noidx. - rewrite on_projs_noidx in eq. simpl in *. - rewrite p0.(onNpars) in eq. lia. } epose proof (declared_projection_closed declp). pose proof (wf_projection_context _ _ declp cu) as wfpctx. - destruct (ind_variance mdecl) eqn:eqv; revgoals. + eapply cmp_global_instance_ind_inv in Ru as [(comp & l & eqv & Ru) | Ru]; eauto; tc; revgoals. { eapply into_ws_cumul_pb; cycle 1. { eauto with fvs. } { rewrite -is_open_term_closed. @@ -3079,7 +3062,7 @@ Proof. rewrite subst_instance_app_ctx in onctx. epose proof (positive_cstr_closed_args declp cu) as hpos. rewrite eqv in hpos; simpl in hpos. - specialize (hpos Ru). + forward hpos. 1: now right. rewrite - !(subst_instance_smash _ _ []) in hpos. rewrite - !(expand_lets_smash_context _ []) in hpos. apply hpos in onctx. clear hpos. @@ -3152,9 +3135,9 @@ Qed. Lemma wt_ind_app_variance {cf} {Σ} {wfΣ : wf Σ} {Γ ind u l}: isType Σ Γ (mkApps (tInd ind u) l) -> ∑ mdecl, (lookup_inductive Σ ind = Some mdecl) * - (global_variance Σ (IndRef ind) #|l| = ind_variance (fst mdecl)). + (global_variance Σ (IndRef ind) #|l| = match ind_variance (fst mdecl) with Some var => Variance var | None => AllEqual end). Proof. - move => [s wat]. + intros (_ & s & wat & _). eapply inversion_mkApps in wat as [ty [Hind Hargs]]; auto. eapply inversion_Ind in Hind as [mdecl [idecl [wfΓ [decli [cu cum]]]]]; auto. eapply typing_spine_strengthen in Hargs; eauto. clear cum. @@ -3188,11 +3171,12 @@ Lemma ctx_inst_app_weak `{checker_flags} Σ (wfΣ : wf Σ.1) ind mdecl idecl (is ctx_inst Σ Γ (params ++ skipn (ind_npars mdecl) args) (List.rev (subst_instance v (ind_params mdecl ,,, ind_indices idecl))). Proof. - intros [? ty_args] ? cparams cum. - pose proof (wt_ind_app_variance (x; ty_args)) as [mdecl' [idecl' gv]]. + intros isTy ? cparams cum. + pose proof (wt_ind_app_variance isTy) as [mdecl' [idecl' gv]]. unfold lookup_inductive in idecl'. unshelve epose proof (isdecl' := declared_inductive_to_gen isdecl); eauto. rewrite (declared_inductive_lookup_gen isdecl') in idecl'. noconf idecl'. + destruct isTy as (_ & ? & ty_args & _). eapply invert_type_mkApps_ind in ty_args as [ty_args ?] ; eauto. erewrite ind_arity_eq in ty_args. 2: eapply PCUICInductives.oib ; eauto. @@ -3246,12 +3230,12 @@ Proof. eapply typing_spine_ctx_inst in ty_indices as [argsi sp]; tea. - eapply ctx_inst_cumul; tea. apply (ctx_inst_smash.1 argsi). - { apply wf_local_app_inv. apply wf_local_smash_end; tea. + { apply All_local_env_app_inv. apply wf_local_smash_end; tea. eapply substitution_wf_local; tea. eapply X1. rewrite -app_context_assoc -subst_instance_app_ctx. eapply weaken_wf_local; tea. eapply (on_minductive_wf_params_indices_inst isdecl _ c). } - { apply wf_local_app_inv. apply wf_local_smash_end; tea. + { apply All_local_env_app_inv. apply wf_local_smash_end; tea. eapply substitution_wf_local; tea. eapply X0. rewrite -app_context_assoc -subst_instance_app_ctx. eapply weaken_wf_local; tea. @@ -3267,7 +3251,7 @@ Lemma wf_local_vass {cf:checker_flags} Σ {Γ na A} s : Σ ;;; Γ |- A : tSort s -> wf_local Σ (Γ ,, vass na A). Proof. intro X; apply typing_wf_local in X as Y. - constructor; tea. eexists; eassumption. + constructor; tea. repeat (eexists; tea). Qed. Lemma isType_it_mkProd_or_LetIn {cf:checker_flags} {Σ Γ Δ T} : @@ -3281,23 +3265,21 @@ Proof. - intros T Hty. rewrite /= /mkProd_or_LetIn /=. eapply IHΔ. - apply infer_sort_impl with id Hty; intros Hs. + apply lift_sorting_it_impl_gen with Hty => // Hs. have wf := typing_wf_local Hs. depelim wf. - destruct l as [s1 Hs1]. red in l0. eapply type_Cumul'. econstructor; eauto. eapply isType_Sort; eauto. - now eapply PCUICWfUniverses.typing_wf_universe in Hs. + now eapply PCUICWfUniverses.typing_wf_sort in Hs. eapply convSpec_cumulSpec, red1_cumulSpec. repeat constructor. - - intros T [s Hs]. + - intros T (_ & s & Hs & _). apply IHΔ. red. - unfold PCUICTypingDef.typing in *. have wf := typing_wf_local Hs. depelim wf. - destruct l as [s1 Hs1]. - exists (Universe.sort_of_product s1 s). + pose proof (lift_sorting_extract l). + eapply has_sort_isType with (s := Sort.sort_of_product _ s). econstructor; eauto. Qed. @@ -3307,18 +3289,19 @@ Lemma wf_set_binder_name {cf} {Σ : global_env_ext} {wfΣ : wf Σ} {Γ} {nas Δ} wf_local Σ (Γ ,,, map2 set_binder_name nas Δ). Proof. intros ha wf. - apply wf_local_app_inv in wf as []. - eapply wf_local_app => //. - induction w in nas, ha |- *; depelim ha; cbn. constructor. - all: constructor; eauto; try apply IHw; auto. - 1,2: apply infer_typing_sort_impl with id t0; intros Hs. - all: eapply context_conversion; tea. - 1,3,5: eapply wf_local_app, IHw; eauto. - all: eapply eq_binder_annots_eq_ctx in ha. - all: eapply eq_context_upto_univ_conv_context. - all: eapply eq_context_upto_cat. - 1,3,5: reflexivity. - all: symmetry; apply ha. + apply All_local_env_app_inv in wf as [wfΓ wfΔ]. + eapply All_local_env_app => //. + induction wfΔ using All_local_rel_ind1 in nas, ha |- *; depelim ha; cbn. constructor. + specialize (IHwfΔ _ ha). + apply All_local_rel_snoc; tas. + apply lift_typing_impl with (1 := X) => // ?? Hs. + eapply context_conversion; tea. + 1: now eapply All_local_env_app. + eapply eq_binder_annots_eq_ctx in ha. + eapply eq_context_upto_names_conv_context. + eapply All2_app. + 2: reflexivity. + symmetry; apply ha. Qed. Lemma WfArity_build_case_predicate_type {cf:checker_flags} {Σ Γ ci args mdecl idecl p ps} : @@ -3326,7 +3309,7 @@ Lemma WfArity_build_case_predicate_type {cf:checker_flags} {Σ Γ ci args mdecl declared_inductive Σ.1 ci.(ci_ind) mdecl idecl -> isType Σ Γ (mkApps (tInd ci p.(puinst)) (pparams p ++ args)) -> let params := firstn (ind_npars mdecl) args in - wf_universe Σ ps -> + wf_sort Σ ps -> wf_predicate mdecl idecl p -> isWfArity Σ Γ (it_mkProd_or_LetIn (case_predicate_context ci mdecl idecl p) (tSort ps)). Proof. @@ -3334,7 +3317,7 @@ Proof. split. 2:{ eexists _, _. rewrite destArity_it_mkProd_or_LetIn. reflexivity. } rewrite /case_predicate_context /case_predicate_context_gen. - have wfΓ := typing_wf_local X.π2. + have wfΓ := typing_wf_local X.2.π2.1. eapply isType_mkApps_Ind_inv in X; tea. destruct X as [parsubst [argsubst [sppars spargs parslen argslen cu]]]. epose proof (isType_case_predicate (puinst p) (pparams p) ps wfΓ isdecl cu wfps). @@ -3434,7 +3417,7 @@ Qed. Lemma isType_is_open_term {cf} {Σ} {wfΣ : wf Σ} Γ T : isType Σ Γ T -> is_open_term Γ T. Proof. - intros [s hs]. now apply subject_is_open_term in hs. + intros (_ & s & hs & _). now apply subject_is_open_term in hs. Qed. #[global] Hint Immediate isType_is_open_term : pcuic. @@ -3462,7 +3445,7 @@ Lemma isType_subst_all_rels {cf} {Σ} {wfΣ : wf Σ} {Γ Δ} {T} : isType Σ (Γ ,,, Δ) (subst0 (all_rels Δ 0 #|Δ|) (lift #|Δ| #|Δ| T)). Proof. intros Hty. - apply infer_typing_sort_impl with id Hty; intros Hs. + apply lift_sorting_it_impl_gen with Hty => // Hs. pose proof (typing_wf_local Hs). eapply weakening_typing in Hs; tea. rewrite -(app_nil_l (lift_context _ _ _)) -/(app_context _ _) app_context_assoc in Hs. @@ -3534,8 +3517,8 @@ Proof. rewrite -Nat.add_comm lift_context_add. eapply spine_subst_weakening;tea. eapply spine_subst_to_extended_list_k; tea. - now apply wf_local_app_inv in wf. - destruct (wf_local_app_inv wf) as []. + now apply All_local_env_app_inv in wf. + destruct (All_local_env_app_inv wf) as []. epose proof (@all_rels_subst_lift _ _ _ _ _ Δ' T a). eapply typing_spine_strengthen; tea; revgoals. eapply ws_cumul_pb_eq_le; symmetry. @@ -3543,7 +3526,7 @@ Proof. eapply X. 2:pcuic. eauto with fvs. clear X. - apply infer_typing_sort_impl with id isty; intros Hs. + apply lift_sorting_f_it_impl with (f := fun t => _ (_ t)) (tu := isty) => // Hs. eapply (weakening_typing) in Hs; tea. eapply (substitution (Δ := [])) in Hs. 2:{ epose proof (spine_subst_to_extended_list_k a). apply X. } @@ -3593,7 +3576,7 @@ Lemma wf_case_predicate_context {cf : checker_flags} {Σ : global_env_ext} {wfΣ Proof. intros isdecl Hc wfp predctx. epose proof (WfArity_build_case_predicate_type wfΣ isdecl Hc - (PCUICWfUniverses.wf_universe_type1 Σ) wfp). + (PCUICWfUniverses.wf_sort_type1 Σ) wfp). destruct X. eapply isType_it_mkProd_or_LetIn_inv in i; tea. now eapply isType_wf_local in i. @@ -3639,7 +3622,8 @@ Proof. rewrite subst_instance_cons. rewrite /= subst_context_snoc /=. constructor. - { constructor. red. simpl. simpl in e. red in e. simpl in e. + { rewrite /= /subst_decl /map_decl /= /set_binder_name /=. relativize (subst _ _ _). + 1: constructor. red. simpl. simpl in e. red in e. simpl in e. rewrite -e. simpl in H. noconf H. reflexivity. simpl. rewrite subst_instance_mkApps subst_mkApps /=. f_equal. f_equal. @@ -3736,7 +3720,7 @@ Lemma wf_case_branch_type {cf : checker_flags} {Σ : global_env_ext} {wfΣ : wf Proof. intros isdecl Hc wfp bc Hp ptm wfpctx. unshelve epose proof (isdecl' := declared_inductive_to_gen isdecl); eauto. - destruct (WfArity_build_case_predicate_type wfΣ isdecl Hc (PCUICWfUniverses.typing_wf_universe _ Hp) wfp) as [wfty _]. + destruct (WfArity_build_case_predicate_type wfΣ isdecl Hc (PCUICWfUniverses.typing_wf_sort _ Hp) wfp) as [wfty _]. set wfcpc := wf_case_predicate_context isdecl Hc wfp. simpl in wfcpc. clearbody wfcpc. have clipars : closed_ctx (subst_instance (puinst p) (ind_params mdecl)). { rewrite closedn_subst_instance_context. @@ -3857,6 +3841,7 @@ Proof. eapply (pre_case_branch_context_eq (ind:=ci) nth cu wfbr). reflexivity. } pose proof (on_cindices onc). + eapply ctx_inst_impl in X0. 2: intros ??; apply unlift_TermTyp. assert (spindices : spine_subst Σ (Γ,,, pre_case_branch_context ci mdecl (pparams p) (puinst p) cdecl) @@ -3973,7 +3958,7 @@ Proof. !to_extended_list_k_lift_context to_extended_list_k_subst PCUICLiftSubst.map_subst_instance_to_extended_list_k //. } rewrite /pre_case_branch_context /subst_let_expand_k. - eexists (subst_instance p.(puinst) (ind_sort idecl)). + eapply has_sort_isType with (s := subst_instance p.(puinst) (ind_sort idecl)). relativize #|cstr_args cdecl|. eapply (substitution (s := List.rev (pparams p)) (T := tSort _)); tea. eapply sppars. @@ -3992,7 +3977,7 @@ Proof. rewrite (declared_inductive_type isdecl). rewrite subst_instance_it_mkProd_or_LetIn subst_instance_app it_mkProd_or_LetIn_app. - have wfs : wf_universe Σ (subst_instance_univ (puinst p) (ind_sort idecl)). + have wfs : wf_sort Σ (subst_instance_sort (puinst p) (ind_sort idecl)). by eapply (on_inductive_sort_inst isdecl _ cu). have wfparinds : wf_local Σ (Γ,,, subst_instance (puinst p) (ind_params mdecl),,, @@ -4099,7 +4084,7 @@ Lemma wf_case_branch_type' {cf : checker_flags} {Σ : global_env_ext} {wfΣ : wf Proof. intros isdecl Hc wfp bc Hp ptm wfpctx. unshelve epose proof (isdecl' := declared_inductive_to_gen isdecl); eauto. - destruct (WfArity_build_case_predicate_type wfΣ isdecl Hc (PCUICWfUniverses.typing_wf_universe _ Hp) wfp) as [wfty _]. + destruct (WfArity_build_case_predicate_type wfΣ isdecl Hc (PCUICWfUniverses.typing_wf_sort _ Hp) wfp) as [wfty _]. set wfcpc := wf_case_predicate_context isdecl Hc wfp. simpl in wfcpc. clearbody wfcpc. have clipars : closed_ctx (subst_instance (puinst p) (ind_params mdecl)). { rewrite closedn_subst_instance_context. @@ -4218,6 +4203,7 @@ Proof. eapply (pre_case_branch_context_eq (ind:=ci) nth cu wfbr). reflexivity. } pose proof (on_cindices onc). + eapply ctx_inst_impl in X0. 2: intros ??; apply unlift_TermTyp. assert (spindices : spine_subst Σ (Γ,,, pre_case_branch_context ci mdecl (pparams p) (puinst p) cdecl) @@ -4334,7 +4320,7 @@ Proof. !to_extended_list_k_lift_context to_extended_list_k_subst PCUICLiftSubst.map_subst_instance_to_extended_list_k //. } rewrite /pre_case_branch_context /subst_let_expand_k. - eexists (subst_instance p.(puinst) (ind_sort idecl)). + eapply has_sort_isType with (s := subst_instance p.(puinst) (ind_sort idecl)). relativize #|cstr_args cdecl|. eapply (substitution (s := List.rev (pparams p)) (T := tSort _)); tea. eapply sppars. @@ -4353,7 +4339,7 @@ Proof. rewrite (declared_inductive_type isdecl). rewrite subst_instance_it_mkProd_or_LetIn subst_instance_app it_mkProd_or_LetIn_app. - have wfs : wf_universe Σ (subst_instance_univ (puinst p) (ind_sort idecl)). + have wfs : wf_sort Σ (subst_instance_sort (puinst p) (ind_sort idecl)). by eapply (on_inductive_sort_inst isdecl _ cu). have wfparinds : wf_local Σ (Γ,,, subst_instance (puinst p) (ind_params mdecl),,, @@ -4491,7 +4477,7 @@ Lemma wf_case_branches_types' {cf : checker_flags} {Σ : global_env_ext} {wfΣ : 0 (ind_ctors idecl) brs. Proof. intros isdecl Hc wfp bc Hp ptm wfbrs conv. - destruct (WfArity_build_case_predicate_type wfΣ isdecl Hc (PCUICWfUniverses.typing_wf_universe _ Hp) wfp) as [wfty _]. + destruct (WfArity_build_case_predicate_type wfΣ isdecl Hc (PCUICWfUniverses.typing_wf_sort _ Hp) wfp) as [wfty _]. set wfcpc := wf_case_predicate_context isdecl Hc wfp. simpl in wfcpc. clearbody wfcpc. have clipars : closed_ctx (subst_instance (puinst p) (ind_params mdecl)). { rewrite closedn_subst_instance_context. @@ -4518,7 +4504,7 @@ Lemma conv_betas_typed `{cf:checker_flags} (Σ : global_env_ext) (wfΣ: wf Σ) ctx_inst Σ Γ l (List.rev Δ) -> Σ ;;; Γ |- mkApps (it_mkLambda_or_LetIn Δ t) l =s subst0 (mk_ctx_subst Δ l) t. Proof. - move=> [ps tWty] instl. + intros (_ & ps & tWty & _) instl. pose proof (wfΓ := typing_wf_local tWty). pose proof (ss := ctx_inst_spine_subst wfΓ instl). pose proof (eqsubst := mk_ctx_subst_spec' instl). @@ -4528,11 +4514,11 @@ Proof. constructor. + have?: Σ ;;; Γ |- it_mkLambda_or_LetIn Δ t : it_mkProd_or_LetIn Δ (tSort ps) by apply: PCUICGeneration.type_it_mkLambda_or_LetIn; eassumption. - exists ps; apply: PCUICSpine.type_mkApps; first eassumption. + apply has_sort_isType with (s := ps); apply: PCUICSpine.type_mkApps; first eassumption. apply: typing_spine_it_mkProd_or_LetIn_close=> //=; first eassumption. apply: isType_it_mkProd_or_LetIn. exact: validity tWty. - + apply:isType_subst; last (exists ps; eassumption). + + apply:isType_subst; last (apply has_sort_isType with (s := ps); eassumption). exact: inst_subslet ss. + apply: PCUICCumulativity.red_conv. rewrite -eqsubst. apply: red_betas_typed; last eassumption. diff --git a/pcuic/theories/PCUICInductives.v b/pcuic/theories/PCUICInductives.v index 371e042f2..2ed8d9f79 100644 --- a/pcuic/theories/PCUICInductives.v +++ b/pcuic/theories/PCUICInductives.v @@ -240,7 +240,7 @@ Section OnInductives. eapply on_declared_inductive in decli as [onmind oib]. pose proof (oib.(onArity)). rewrite oib.(ind_arity_eq) in X. - destruct X as [s Hs]. + destruct X as (_ & s & Hs & _). rewrite -it_mkProd_or_LetIn_app in Hs. eapply it_mkProd_or_LetIn_wf_local in Hs. now rewrite app_context_nil_l in Hs. @@ -258,7 +258,7 @@ End OnInductives. Lemma isType_intro {cf:checker_flags} {Σ Γ T s} : Σ ;;; Γ |- T : tSort s -> isType Σ Γ T. Proof. - now intros Hty; exists s. + eapply has_sort_isType. Qed. #[global] Hint Resolve isType_intro : pcuic. @@ -282,13 +282,13 @@ Section OnInductives. wf_local Σ Γ -> consistent_instance_ext Σ (ind_universes mdecl) u -> isType Σ Γ (it_mkProd_or_LetIn (subst_instance u (ind_params mdecl ,,, idecl.(ind_indices))) - (tSort (subst_instance_univ u idecl.(ind_sort)))). + (tSort (subst_instance_sort u idecl.(ind_sort)))). Proof using decli wfΣ. move=> wfΓ cext. destruct (on_declared_inductive decli) as [onmind oib]. pose proof (oib.(onArity)) as ar. rewrite oib.(ind_arity_eq) in ar. - destruct ar as [s ar]. + destruct ar as (_ & s & ar & _). eapply isType_weaken => //. rewrite -(subst_instance_it_mkProd_or_LetIn u _ (tSort _)). rewrite -it_mkProd_or_LetIn_app in ar. @@ -305,7 +305,7 @@ Section OnInductives. move=> wfΓ cext. destruct (on_declared_inductive decli) as [onmind oib]. pose proof (oib.(onArity)) as ar. - destruct ar as [s ar]. + destruct ar as (_ & s & ar & _). eapply isType_weaken => //. eapply (typing_subst_instance_decl Σ [] _ _ _ (InductiveDecl mdecl) u wfΣ) in ar. all:pcuic. @@ -315,24 +315,24 @@ Section OnInductives. Local Definition oi := (on_declared_inductive decli).1. Local Definition oib := (on_declared_inductive decli).2. - Lemma on_inductive_sort : wf_universe (Σ.1, ind_universes mdecl) (ind_sort idecl). + Lemma on_inductive_sort : wf_sort (Σ.1, ind_universes mdecl) (ind_sort idecl). Proof using decli wfΣ. pose proof (oib.(onArity)) as ar. rewrite oib.(ind_arity_eq) in ar. - destruct ar as [s ar]. + destruct ar as (_ & s & ar & _). eapply typing_wf_universes in ar; auto. move/andP: ar => []. rewrite wf_universes_it_mkProd_or_LetIn => /andP [] _. - now rewrite wf_universes_it_mkProd_or_LetIn => /andP [] _ /= /wf_universe_reflect. + now rewrite wf_universes_it_mkProd_or_LetIn => /andP [] _ /= /wf_sort_reflect. Qed. Lemma on_inductive_sort_inst u : consistent_instance_ext Σ (ind_universes mdecl) u -> - wf_universe Σ (subst_instance u (ind_sort idecl)). + wf_sort Σ (subst_instance u (ind_sort idecl)). Proof using decli wfΣ. generalize on_inductive_sort => wf. destruct Σ. intros cu. - eapply wf_universe_instantiate; eauto. + eapply wf_sort_instantiate; eauto. now eapply consistent_instance_ext_wf. Qed. @@ -540,7 +540,7 @@ Record wf_subslet {cf} Σ Γ s Δ := Lemma wf_subslet_dom {cf} {Σ Γ s Δ} : wf_subslet Σ Γ s Δ -> wf_local Σ Γ. Proof. - intros [wf _]. now eapply wf_local_app_inv in wf. + intros [wf _]. now eapply All_local_env_app_inv in wf. Qed. Lemma subst_id s Γ t : @@ -576,7 +576,7 @@ Lemma isType_it_mkProd_or_LetIn_inv {cf:checker_flags} {Σ : global_env_ext} {wf isType Σ Γ (it_mkProd_or_LetIn Δ T) -> isType Σ (Γ ,,, Δ) T. Proof. - intros HT; apply infer_typing_sort_impl with id HT; intros Hs. + intros HT; apply lift_sorting_it_impl_gen with HT => // Hs. now eapply inversion_it_mkProd_or_LetIn in Hs. Qed. @@ -623,31 +623,27 @@ Proof. move=> wfΣ wfΔ. induction Δ as [|[na [d|] ?] ?] in wfΔ |- *; simpl; auto; depelim wfΔ. * constructor. - * rewrite extended_subst_length. rewrite lift_closed. - do 2 red in l0. autorewrite with len. fvs. + * apply unlift_TermTyp in l. + rewrite extended_subst_length. rewrite lift_closed. fvs. constructor. auto. specialize (IHΔ wfΔ). - red in l0. - eapply weaken_ctx in l0. 3:eapply wf_local_smash_context; eauto. 2:auto. - eapply (substitution (Δ := [])) in l0. eapply l0. all:auto. + eapply weaken_ctx in l. 3:eapply wf_local_smash_context; eauto. 2:auto. + eapply (substitution (Δ := [])) in l. eapply l. all:auto. * rewrite smash_context_acc. simpl. rewrite /map_decl /= /map_decl /=. simpl. - pose proof (l.π2). + assert (isType Σ (smash_context [] Δ) (subst0 (extended_subst Δ 0) decl_type)). + { apply lift_sorting_it_impl_gen with l => // Hs. + eapply weaken_ctx in Hs. 3:eapply wf_local_smash_context; eauto. 2:auto. + eapply (substitution (Δ := [])) in Hs. eapply Hs. all:auto. } rewrite lift_closed. fvs. rewrite (lift_extended_subst _ 1). rewrite -{4}(closed_ctx_lift 1 0 Δ); fvs. constructor. { eapply (subslet_lift _ [_]); eauto. - constructor. eapply wf_local_smash_context; auto. - apply infer_typing_sort_impl with id l; intros Hs. - eapply weaken_ctx in Hs. 3:eapply wf_local_smash_context; eauto. 2:auto. - eapply (substitution (Δ := [])) in Hs. eapply Hs. all:auto. } + constructor; tas. eapply wf_local_smash_context; auto. } { eapply refine_type. econstructor. - rewrite smash_context_acc /=. constructor. + rewrite smash_context_acc /=. constructor; tas. apply wf_local_smash_context; auto. - apply infer_typing_sort_impl with id l; intros Hs. - eapply weaken_ctx in Hs. 3:eapply wf_local_smash_context; eauto. 2:auto. - eapply (substitution (Δ := [])) in Hs. eapply Hs. all:auto. reflexivity. simpl. rewrite distr_lift_subst. f_equal. @@ -665,8 +661,8 @@ Proof. generalize 0 at 1 4. induction Δ as [|[na [d|] ?] ?] in wfΔ |- *; simpl; auto; depelim wfΔ. - * intros n. rewrite extended_subst_length. rewrite lift_closed. - do 2 red in l0. autorewrite with len; fvs. + * apply unlift_TermTyp in l. + intros n. rewrite extended_subst_length. rewrite lift_closed. fvs. rewrite (reln_lift 1 0). rewrite map_map_compose map_subst_lift1. autorewrite with len. @@ -710,7 +706,7 @@ Proof. have wfsmpars : wf_local (Σ.1, ind_universes mdecl) (smash_context [] (ind_params mdecl)). { now apply wf_local_smash_context. } constructor => //. - red. exists (ind_sort idecl). + eapply has_sort_isType. eapply type_mkApps. econstructor; eauto. eapply consistent_instance_ext_abstract_instance; eauto with pcuic. eapply declared_inductive_wf_global_ext; eauto with pcuic. @@ -957,7 +953,7 @@ Proof. eapply (weakening_wf_local (Γ'' := [_])). 2:eapply wf_projection_context_gen; tea. 2:apply decli. rewrite - !skipn_subst_context. - eapply wf_local_app_skipn. + eapply All_local_env_app_skipn. rewrite -(smash_context_subst []) /=. rewrite -(smash_context_subst []) /=. rewrite subst_context_nil. @@ -1080,7 +1076,7 @@ Proof. rewrite H //. split => //. apply Hnth'. eapply meta_conv. econstructor. - destruct IH as [[s isTy] _]. + destruct IH as [(_ & s & isTy & _) _]. now eapply typing_wf_local in isTy. simpl. reflexivity. simpl. rewrite lift_mkApps. simpl. destruct ind; simpl. @@ -1090,7 +1086,7 @@ Proof. assert(subst_instance (abstract_instance (ind_universes mdecl)) pdecl.(proj_type) = pdecl.(proj_type)) as ->. { eapply (isType_subst_instance_id (Σ.1, ind_universes mdecl)); eauto with pcuic. eapply declared_inductive_wf_ext_wk; eauto with pcuic. - destruct Hdecl; eauto. } + destruct Hdecl; eauto. eapply IH. } destruct IH as [isTy [decl [nthdecl _ _ eqpdecl ptyeq]]]. move ptyeq at bottom. rewrite nthdecl in Hnth. noconf Hnth. simpl in ptyeq. @@ -1195,7 +1191,7 @@ Proof. destruct isdecl as [[[] ?] ?]; eauto. } destruct (ind_cunivs oib) as [|? []] eqn:hequ => //. eapply PCUICClosedConv.sorts_local_ctx_All_local_env in wfargs. - 2:{ eapply All_local_env_app. split; auto. + 2:{ eapply All_local_env_app; auto. red in onpars. eapply (All_local_env_impl _ _ _ onpars). intros. apply lift_typing_impl with (1 := X); intros ? Hs. eapply weaken_ctx; auto. } @@ -1227,7 +1223,7 @@ Proof. assert(wf_local (Σ.1, ind_universes mdecl) (ind_params mdecl ,, vass {| binder_name := nNamed (ind_name idecl); binder_relevance := idecl.(ind_relevance) |} (mkApps (tInd ind u) (to_extended_list (ind_params mdecl))))). - { constructor. auto. red. exists (ind_sort idecl). + { constructor. auto. eapply has_sort_isType. eapply type_mkApps. econstructor; eauto. destruct isdecl as []; eauto. subst u. eapply consistent_instance_ext_abstract_instance; eauto with pcuic. @@ -1251,13 +1247,13 @@ Proof. simpl in oib. pose proof (onProjs.(on_projs_noidx _ _ _ _ _ _)). destruct (ind_indices idecl); simpl in H; try discriminate. - simpl. rewrite [subst_instance_univ _ _]sortu. + simpl. rewrite [subst_instance_sort _ _]sortu. eapply ws_cumul_pb_compare => //. now eapply wf_local_closed_context in X0. constructor. reflexivity. rewrite -subst_instance_it_mkProd_or_LetIn. pose proof (onArity oib). rewrite -(oib.(ind_arity_eq)). - apply infer_typing_sort_impl with id X1. + apply lift_typing_impl with (1 := X1) => t T HT. eapply (weaken_ctx (Γ:=[])); auto. } intros wf. generalize (weakening_wf_local (Γ'':=[_]) wf X1). @@ -1303,17 +1299,16 @@ Proof. rewrite nth_error_app_lt in Heq'. autorewrite with len. lia. set (idx := context_assumptions (cstr_args c) - S i) in *. - unshelve epose proof (nth_error_All_local_env (n:=idx) _ wfsargs). - autorewrite with len. simpl. lia. - rename X1 into onnth. - destruct (nth_error (subst_context _ 1 _) idx) as [c2|] eqn:hidx. + unshelve epose proof ((nth_error_Some' _ idx).2 _) as (c2 & hidx); cycle -1. + eapply @All_local_env_nth_error with (n := idx) (2 := hidx) in wfsargs as onnth. + 2: autorewrite with len; simpl. 2: lia. simpl in onnth. red in onnth. cbn in onnth. assert(decl_body c2 = None). { apply nth_error_assumption_context in hidx; auto. rewrite /subst_context /lift_context. apply assumption_context_fold, smash_context_assumption_context. constructor. } - rewrite H in onnth. 2:{ simpl in onnth; contradiction. } - destruct onnth as [s Hs]. + rewrite H in onnth. + destruct onnth as (_ & s & Hs & _). Ltac havefst H := match goal with |- ?T × _ => have H : T; [|split => //] @@ -1353,7 +1348,7 @@ Proof. { assert(wfargs' : wf_local (Σ.1, ind_universes mdecl) (arities_context (ind_bodies mdecl) ,,, ind_params mdecl ,,, smash_context [] (cstr_args c))). - { apply All_local_env_app; split; auto. + { apply All_local_env_app; auto. now apply All_local_env_app_inv in wfargs as [wfindpars wfargs]. apply wf_local_rel_smash_context; auto. } eapply closed_wf_local in wfargs'; auto. @@ -1408,14 +1403,14 @@ Proof. now rewrite /indsubst inds_length. } split. unfold projection_type in H0. - rewrite H0. exists s; auto. + rewrite H0. eapply has_sort_isType; eauto. rewrite -/indb in Hs. rewrite /projection_type' -/indb -/indsubst -/projsubst. rewrite Nat.add_1_r in Hs. exact Hs. exists arg. split; auto. 2:{ apply wf_local_smash_end; auto. } eapply wf_local_smash_end in wfargs. - eapply (wf_local_app_skipn (n:=context_assumptions (cstr_args c) - S i)) in wfargs. + eapply All_local_env_app_skipn with (n:=context_assumptions (cstr_args c) - S i) in wfargs. epose proof (wf_local_nth_isType (d:=arg) (n:=0) wfargs). rewrite skipn_app in X1. rewrite skipn_length in X1. len. @@ -1529,14 +1524,15 @@ Proof. Qed. Lemma type_local_ctx_cum {cf:checker_flags} {Σ Γ Δ s s'} : - wf Σ.1 -> wf_universe Σ s' -> - leq_universe (global_ext_constraints Σ) s s' -> + wf Σ.1 -> wf_sort Σ s' -> + leq_sort (global_ext_constraints Σ) s s' -> type_local_ctx (lift_typing typing) Σ Γ Δ s -> type_local_ctx (lift_typing typing) Σ Γ Δ s'. Proof. intros wfΣ wfs leqs. induction Δ as [|[na [b|] ty] Δ]; simpl; auto; intros []; split; auto. + destruct l as (_ & ss & Hs & <-). split; cbn; eexists; split; [|eauto]. eapply type_Cumul; eauto. econstructor; pcuic. now eapply cumul_Sort. @@ -1545,7 +1541,7 @@ Qed. Lemma type_local_ctx_wf {cf:checker_flags} {Σ Γ Δ s} : wf Σ.1 -> type_local_ctx (lift_typing typing) Σ Γ Δ s -> - wf_universe Σ s. + wf_sort Σ s. Proof. intros wfΣ. induction Δ as [|[na [b|] ty] Δ]; simpl; intuition auto. @@ -1555,7 +1551,7 @@ Lemma isType_it_mkProd_or_LetIn {cf:checker_flags} Σ Γ Δ s s': wf Σ.1 -> wf_local Σ Γ -> type_local_ctx (lift_typing typing) Σ Γ Δ s -> - wf_universe Σ s' -> + wf_sort Σ s' -> isType Σ Γ (it_mkProd_or_LetIn Δ (tSort s')). Proof. move=> wfΣ wfΓ wfΔ wfs. @@ -1571,28 +1567,29 @@ Lemma isType_it_mkProd_or_LetIn_smash {cf:checker_flags} Σ Γ Δ s : isType Σ Γ (it_mkProd_or_LetIn Δ (tSort s)) -> isType Σ Γ (it_mkProd_or_LetIn (smash_context [] Δ) (tSort s)). Proof. - move=> wfΣ [u Hu]. - exists u. unfold PCUICTypingDef.typing in *. revert Γ u Hu. + intros wfΣ HT. + apply lift_sorting_it_impl_gen with HT => // Hu. set (u := HT.2.π1) in *; clearbody u. clear HT. + revert Γ u Hu. induction Δ as [|[na [b|] ty] Δ] using ctx_length_rev_ind; simpl in *; auto; rewrite !it_mkProd_or_LetIn_app /= /mkProd_or_LetIn /=; rewrite smash_context_app; intros Γ u Hu. - simpl. - assert (wfu := typing_wf_universe wfΣ Hu). - eapply inversion_LetIn in Hu as [? [? [? [? [? e]]]]]; auto. - eapply substitution_let in t1; auto. - rewrite /subst1 subst_it_mkProd_or_LetIn /= in t1. + assert (wfu := typing_wf_sort wfΣ Hu). + eapply inversion_LetIn in Hu as (T & wfty & HT & hlt); auto. + eapply substitution_let in HT; auto. + rewrite /subst1 subst_it_mkProd_or_LetIn /= in HT. specialize (X (subst_context [b] 0 Δ) ltac:(autorewrite with len; lia)). - eapply ws_cumul_pb_LetIn_l_inv in e; eauto. - eapply type_ws_cumul_pb in t1. 3:eauto. - 2:{ pcuic. eexists; econstructor; pcuic. } - specialize (X _ _ t1). + eapply ws_cumul_pb_LetIn_l_inv in hlt; eauto. + eapply type_ws_cumul_pb in HT. 3:eauto. + 2:{ eapply isType_Sort. all: pcuic. } + specialize (X _ _ HT). now rewrite -smash_context_subst /= subst_context_nil. - simpl. rewrite it_mkProd_or_LetIn_app. simpl. - assert (wfu := typing_wf_universe wfΣ Hu). - eapply inversion_Prod in Hu as [? [? [? [? ?]]]]; auto. + assert (wfu := typing_wf_sort wfΣ Hu). + eapply inversion_Prod in Hu as (s1 & s2 & wfty & Ht & hlt); auto. specialize (X Δ ltac:(reflexivity)). - specialize (X _ _ t0). + specialize (X _ _ Ht). eapply type_ws_cumul_pb; eauto. - econstructor; pcuic. + econstructor; tas. apply unlift_TypUniv in wfty. eapply isType_intro. econstructor; pcuic. Qed. @@ -1603,7 +1600,7 @@ Lemma typing_spine_to_extended_list {cf:checker_flags} Σ Δ u s : typing_spine Σ (smash_context [] (subst_instance u Δ)) (subst_instance u (it_mkProd_or_LetIn Δ (tSort s))) (to_extended_list (smash_context [] (subst_instance u Δ))) - (tSort (subst_instance_univ u s)). + (tSort (subst_instance_sort u s)). Proof. move=> wfΣ wfΔ. apply wf_arity_spine_typing_spine; auto. @@ -1638,7 +1635,7 @@ Proof. assert(wfsmash : wf_local Σ (smash_context [] (subst_instance u (ind_params mdecl)))). { eapply wf_local_smash_context; auto. } constructor; auto. red. - exists (subst_instance_univ u (ind_sort idecl)). + eapply @has_sort_isType with (s := subst_instance_sort u (ind_sort idecl)). eapply type_mkApps. econstructor; eauto. eapply decli.p1. rewrite (ind_arity_eq oib0). pose proof (on_projs_noidx _ _ _ _ _ _ onps). @@ -1693,14 +1690,14 @@ Lemma isType_mkApps_Ind_inv {cf:checker_flags} {Σ Γ ind u args} (wfΣ : wf Σ. consistent_instance_ext Σ (ind_universes mdecl) u]. Proof. move=> wfΓ isType. - destruct isType as [s Hs]. + destruct isType as (_ & s & Hs & _). eapply invert_type_mkApps_ind in Hs as [tyargs cu]; eauto. set (decli' := on_declared_inductive declm). rename declm into decli. destruct decli' as [declm decli']. pose proof (decli'.(onArity)) as ar. rewrite decli'.(ind_arity_eq) in tyargs, ar. - hnf in ar. destruct ar as [s' ar]. + destruct ar as (_ & s' & ar & _). cbn in ar. rewrite !subst_instance_it_mkProd_or_LetIn in tyargs. simpl in tyargs. rewrite -it_mkProd_or_LetIn_app in tyargs. eapply arity_typing_spine in tyargs as [argslen leqs [instsubst [wfdom wfcodom cs subs]]] => //. @@ -1737,14 +1734,14 @@ Lemma isType_mkApps_Ind_inv' {cf:checker_flags} {Σ Γ ind u args} (wfΣ : wf Σ consistent_instance_ext Σ (ind_universes mdecl) u]. Proof. move=> wfΓ isType. - destruct isType as [s Hs]. + destruct isType as (_ & s & Hs & _). eapply invert_type_mkApps_ind in Hs as [tyargs cu]; eauto. set (decli' := on_declared_inductive declm). rename declm into decli. destruct decli' as [declm decli']. pose proof (decli'.(onArity)) as ar. rewrite decli'.(ind_arity_eq) in tyargs, ar. - hnf in ar. destruct ar as [s' ar]. + destruct ar as (_ & s' & ar & _). cbn in ar. rewrite !subst_instance_it_mkProd_or_LetIn in tyargs. simpl in tyargs. rewrite -it_mkProd_or_LetIn_app in tyargs. eapply arity_typing_spine in tyargs as [argslen leqs [instsubst [wfdom wfcodom cs subs]]] => //. @@ -2101,7 +2098,7 @@ Lemma wf_local_expand_lets {cf} {Σ} {wfΣ : wf Σ} {Γ Δ Γ'} : Proof. intros hwf. rewrite /expand_lets_ctx /expand_lets_k_ctx. - destruct (wf_local_app_inv hwf) as [wfΓΔ _]. + destruct (All_local_env_app_inv hwf) as [wfΓΔ _]. rewrite -app_context_assoc in hwf. eapply (weakening_wf_local (Γ'' := smash_context [] Δ)) in hwf. len in hwf. simpl in hwf. simpl. @@ -2123,10 +2120,10 @@ Proof. rewrite -app_context_assoc in Ht. eapply (weakening_typing (Γ'' := smash_context [] Δ)) in Ht. len in Ht. simpl in Ht. simpl. - 2:{ eapply wf_local_smash_end; pcuic. now apply wf_local_app_inv in X. } + 2:{ eapply wf_local_smash_end; pcuic. now apply All_local_env_app_inv in X. } rewrite lift_context_app app_context_assoc Nat.add_0_r in Ht. eapply substitution in Ht; tea. - 2:{ eapply PCUICContexts.subslet_extended_subst. now apply wf_local_app_inv in X. } + 2:{ eapply PCUICContexts.subslet_extended_subst. now apply All_local_env_app_inv in X. } now len in Ht. Qed. @@ -2304,17 +2301,17 @@ Proof. split. * rewrite -app_context_assoc. rewrite smash_context_app_expand_acc. - eapply wf_local_app_inv in spine_dom_wf as [wfΔ wfΓ']. - eapply wf_local_app. now apply wf_local_app_inv in wfΔ. + eapply All_local_env_app_inv in spine_dom_wf as [wfΔ wfΓ']. + eapply All_local_env_app. now apply All_local_env_app_inv in wfΔ. eapply wf_local_rel_smash_context_gen; tea. - * eapply wf_local_app_inv in spine_codom_wf as [wfΔ wfΓ']. + * eapply All_local_env_app_inv in spine_codom_wf as [wfΔ wfΓ']. rewrite - !app_context_assoc -expand_lets_ctx_app. - apply wf_local_app_inv in wfΔ as []. - eapply wf_local_app. - now apply wf_local_app_inv in a as []. + apply All_local_env_app_inv in wfΔ as []. + eapply All_local_env_app. + now apply All_local_env_app_inv in a as []. rewrite smash_context_app_expand_acc. apply wf_local_rel_smash_context_gen; tea. - eapply wf_local_rel_app; tea. + eapply All_local_rel_app; tea. * rewrite /expand_lets_k_ctx /expand_lets_k -map_map_compose -[map (fun t => _) s]map_map_compose. eapply subst_context_subst. now eapply lift_context_subst. @@ -2344,7 +2341,7 @@ Proof. unshelve epose proof (on_inductive_inst isdecl _ _ _ cu); tea. rewrite -/(subst_context_let_expand _ _ _). rewrite subst_instance_app_ctx in X. - destruct X as [s Hs]. + destruct X as (_ & s & Hs & _). eapply type_it_mkProd_or_LetIn_inv in Hs as (idxs & inds & [sortsidx sortind leq]). set (idxctx := smash_context [] (ind_params mdecl)@[u] ,,, expand_lets_ctx (ind_params mdecl)@[u] (ind_indices idecl)@[u]) in *. have tyass : Σ ;;; Γ ,,, idxctx |- (decl_type iass)@[u] : tSort (ind_sort idecl)@[u]. @@ -2418,7 +2415,7 @@ Proof. rewrite subst_extended_lift //; len. now rewrite closedn_subst_instance_context. rewrite (lift_extended_subst _ #|ind_indices _|) subst_map_lift_lift_context; len. } constructor => //. - 2:{ eexists; tea. rewrite /idxctx app_context_assoc in tyass; tea. } + 2:{ eapply has_sort_isType. rewrite /idxctx app_context_assoc in tyass; tea. } apply typing_wf_local in tyass. rewrite /idxctx app_context_assoc in tyass => //. Qed. @@ -2430,7 +2427,7 @@ Lemma isType_case_predicate {cf : checker_flags} {Σ : global_env_ext} {wfΣ : w wf_local Σ Γ -> declared_inductive Σ ci mdecl idecl -> consistent_instance_ext Σ (ind_universes mdecl) u -> - wf_universe Σ ps -> + wf_sort Σ ps -> spine_subst Σ Γ params (List.rev params) (smash_context [] (subst_instance u (ind_params mdecl))) -> isType Σ Γ (it_mkProd_or_LetIn @@ -2446,13 +2443,13 @@ Proof. unshelve epose proof (on_inductive_inst isdecl _ _ _ cu); tea. rewrite -/(subst_context_let_expand _ _ _). rewrite subst_instance_app_ctx in X. - destruct X as [s Hs]. + destruct X as (_ & s & Hs & _). eapply isType_it_mkProd_or_LetIn_app in Hs. 2:eapply sp. rewrite subst_let_expand_it_mkProd_or_LetIn in Hs. eapply type_it_mkProd_or_LetIn_inv in Hs as (idxs & inds & [sortsidx sortind leq]). - eexists (sort_of_products (subst_instance u (ind_sort idecl) :: idxs) - (Universe.super ps)). + eapply has_sort_isType with (s := sort_of_products (subst_instance u (ind_sort idecl) :: idxs) + (Sort.super ps)). set (idxctx := subst_context_let_expand _ _ _) in *. have tyass : Σ ;;; Γ ,,, idxctx |- subst (List.rev params) #|ind_indices idecl| (decl_type iass)@[u] : tSort (ind_sort idecl)@[u]. @@ -2518,11 +2515,11 @@ Proof. set (sdecl := subst_decl _ _ _). rewrite -(it_mkProd_or_LetIn_app [sdecl]). eapply type_it_mkProd_or_LetIn_sorts; tea. - constructor => //. + constructor => //. repeat (eexists; tea). constructor => //. simpl. constructor => //. - now eapply sorts_local_ctx_wf_local; tea. red. - eexists; tea. + now eapply sorts_local_ctx_wf_local; tea. + repeat (eexists; tea). Qed. Lemma arity_spine_case_predicate {cf: checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} {Γ} {ci : case_info} @@ -2530,7 +2527,7 @@ Lemma arity_spine_case_predicate {cf: checker_flags} {Σ : global_env_ext} {wfΣ wf_local Σ Γ -> declared_inductive Σ ci mdecl idecl -> consistent_instance_ext Σ (ind_universes mdecl) u -> - wf_universe Σ ps -> + wf_sort Σ ps -> spine_subst Σ Γ params (List.rev params) (smash_context [] (subst_instance u (ind_params mdecl))) -> spine_subst Σ Γ indices (List.rev indices) diff --git a/pcuic/theories/PCUICInversion.v b/pcuic/theories/PCUICInversion.v index 56cbd157d..ec70dbba1 100644 --- a/pcuic/theories/PCUICInversion.v +++ b/pcuic/theories/PCUICInversion.v @@ -135,8 +135,8 @@ Section Inversion. forall {Γ s T}, Σ ;;; Γ |- tSort s : T -> wf_local Σ Γ × - wf_universe Σ s × - Σ ;;; Γ ⊢ tSort (Universe.super s) ≤ T. + wf_sort Σ s × + Σ ;;; Γ ⊢ tSort (Sort.super s) ≤ T. Proof using wfΣ. intros Γ s T h. invtac h. Qed. @@ -145,9 +145,9 @@ Section Inversion. forall {Γ na A B T}, Σ ;;; Γ |- tProd na A B : T -> ∑ s1 s2, - Σ ;;; Γ |- A : tSort s1 × + lift_typing typing Σ Γ (j_vass_s na A s1) × Σ ;;; Γ ,, vass na A |- B : tSort s2 × - Σ ;;; Γ ⊢ tSort (Universe.sort_of_product s1 s2) ≤ T. + Σ ;;; Γ ⊢ tSort (Sort.sort_of_product s1 s2) ≤ T. Proof using wfΣ. intros Γ na A B T h. invtac h. Qed. @@ -155,19 +155,19 @@ Section Inversion. Lemma inversion_Prod_size : forall {Γ na A B T}, forall H : Σ ;;; Γ |- tProd na A B : T, - ∑ s1 s2 (H1 : Σ ;;; Γ |- A : tSort s1) (H2 : Σ ;;; Γ ,, vass na A |- B : tSort s2), - typing_size H1 < typing_size H × typing_size H2 < typing_size H × - Σ ;;; Γ ⊢ tSort (Universe.sort_of_product s1 s2) ≤ T. + ∑ s1 s2 (H1 : lift_typing typing Σ Γ (j_vass_s na A s1)) (H2 : Σ ;;; Γ ,, vass na A |- B : tSort s2), + lift_typing_size (@typing_size _ _ _) _ H1 < typing_size H × typing_size H2 < typing_size H × + Σ ;;; Γ ⊢ tSort (Sort.sort_of_product s1 s2) ≤ T. Proof using wfΣ. intros Γ na A B T h. unshelve invtac h; eauto. - all: unfold typing_size at 2; fold (typing_size h1); fold (typing_size h2); lia. + all: try revert l; try revert l0; simpl; cbn; lia. Qed. Lemma inversion_Lambda : forall {Γ na A t T}, Σ ;;; Γ |- tLambda na A t : T -> - ∑ s B, - Σ ;;; Γ |- A : tSort s × + ∑ B, + lift_typing typing Σ Γ (j_vass na A) × Σ ;;; Γ ,, vass na A |- t : B × Σ ;;; Γ ⊢ tProd na A B ≤ T. Proof using wfΣ. @@ -177,9 +177,8 @@ Section Inversion. Lemma inversion_LetIn : forall {Γ na b B t T}, Σ ;;; Γ |- tLetIn na b B t : T -> - ∑ s1 A, - Σ ;;; Γ |- B : tSort s1 × - Σ ;;; Γ |- b : B × + ∑ A, + lift_typing typing Σ Γ (j_vdef na b B) × Σ ;;; Γ ,, vdef na b B |- t : A × Σ ;;; Γ ⊢ tLetIn na b B A ≤ T. Proof using wfΣ. @@ -247,7 +246,7 @@ Section Inversion. Import PCUICEquality. Variant case_inversion_data Γ ci p c brs mdecl idecl indices := | case_inv - (ps : Universe.t) + (ps : sort) (eq_npars : mdecl.(ind_npars) = ci.(ci_npar)) (predctx := case_predicate_context ci.(ci_ind) mdecl idecl p) (wf_pred : wf_predicate mdecl idecl p) @@ -256,9 +255,9 @@ Section Inversion. (conv_pctx : eq_context_upto_names p.(pcontext) (ind_predicate_context ci.(ci_ind) mdecl idecl)) (pret_ty : Σ ;;; Γ ,,, predctx |- p.(preturn) : tSort ps) (allowed_elim : is_allowed_elimination Σ idecl.(ind_kelim) ps) - (ind_inst : ctx_inst typing Σ Γ (p.(pparams) ++ indices) + (ind_inst : ctx_inst (typing Σ) Γ (p.(pparams) ++ indices) (List.rev (subst_instance p.(puinst) - (ind_params mdecl ,,, ind_indices idecl)))) + (ind_params mdecl ,,, ind_indices idecl : context)))) (scrut_ty : Σ ;;; Γ |- c : mkApps (tInd ci.(ci_ind) p.(puinst)) (p.(pparams) ++ indices)) (not_cofinite : isCoFinite mdecl.(ind_finite) = false) (ptm := it_mkLambda_or_LetIn predctx p.(preturn)) @@ -310,9 +309,8 @@ Section Inversion. let types := fix_context mfix in fix_guard Σ Γ mfix × nth_error mfix n = Some decl × - All (fun d => isType Σ Γ (dtype d)) mfix × - All (fun d => - Σ ;;; Γ ,,, types |- dbody d : (lift0 #|types|) (dtype d)) mfix × + All (on_def_type (lift_typing typing Σ) Γ) mfix × + All (on_def_body (lift_typing typing Σ) (fix_context mfix) Γ) mfix × wf_fixpoint Σ mfix × Σ ;;; Γ ⊢ dtype decl ≤ T. Proof using wfΣ. @@ -326,10 +324,8 @@ Section Inversion. cofix_guard Σ Γ mfix × let types := fix_context mfix in nth_error mfix idx = Some decl × - All (fun d => isType Σ Γ (dtype d)) mfix × - All (fun d => - Σ ;;; Γ ,,, types |- d.(dbody) : lift0 #|types| d.(dtype) - ) mfix × + All (on_def_type (lift_typing typing Σ) Γ) mfix × + All (on_def_body (lift_typing typing Σ) (fix_context mfix) Γ) mfix × wf_cofixpoint Σ mfix × Σ ;;; Γ ⊢ decl.(dtype) ≤ T. Proof using wfΣ. @@ -375,7 +371,7 @@ Section Inversion. - simpl. apply ih in h. cbn in h. destruct h as [B [h c]]. apply inversion_LetIn in h as hh. - destruct hh as [s1 [A' [? [? [? ?]]]]]. + destruct hh as (A' & wfA' & Ht & hlt). exists A'. split ; eauto. cbn. etransitivity; tea. eapply ws_cumul_pb_it_mkProd_or_LetIn_codom. @@ -383,7 +379,7 @@ Section Inversion. - simpl. apply ih in h. cbn in h. destruct h as [B [h c]]. apply inversion_Lambda in h as hh. - pose proof hh as [s1 [B' [? [? ?]]]]. + pose proof hh as (B' & wfB' & Ht & hlt). exists B'. split ; eauto. cbn. etransitivity; tea. eapply ws_cumul_pb_it_mkProd_or_LetIn_codom. diff --git a/pcuic/theories/PCUICNormal.v b/pcuic/theories/PCUICNormal.v index cd0180bb1..45b8d09bd 100644 --- a/pcuic/theories/PCUICNormal.v +++ b/pcuic/theories/PCUICNormal.v @@ -1052,7 +1052,7 @@ Proof. eauto. - depelim o. 1-2: reflexivity. eapply red_primArray_congr; eauto. - now eapply Universe.make_inj in e. + now eapply Universe.make'_inj in e. Qed. #[global] @@ -1211,7 +1211,7 @@ Qed. Definition fake_params n : context := unfold n (fun x => {| decl_name := {| binder_name := nAnon; binder_relevance := Relevant |}; decl_body := None; - decl_type := tSort Universe.type0 |}). + decl_type := tSort Sort.type0 |}). Lemma fake_params_length n : #|fake_params n| = n. Proof. @@ -1549,20 +1549,20 @@ Proof. now depelim r. Qed. -Lemma whne_eq_term_upto_univ_napp f Σ Γ t Re Rle napp t' : +Lemma whne_eq_term_upto_univ_napp f Σ Γ t cmp_universe cmp_sort pb napp t' : whne f Σ Γ t -> - eq_term_upto_univ_napp Σ Re Rle napp t t' -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp t t' -> whne f Σ Γ t'. Proof. intros wh eq. - induction wh in Re, Rle, napp, t, t', eq, wh |- *; depelim eq; + induction wh in pb, napp, t, t', eq, wh |- *; depelim eq; try solve [eauto using whne; depelim wh; solve_discr; eauto using whne]. - destruct args as [|? ? _] using MCList.rev_ind; [discriminate x|]. rewrite mkApps_app in x; cbn in x; inv x. apply eq_term_upto_univ_napp_mkApps_l_inv in eq1 as (?&?&(eq_hds&?)&->). depelim eq_hds. rewrite <- mkApps_snoc. - assert (All2 (eq_term_upto_univ Σ Re Re) (args ++ [x0]) (x1 ++ [u'])) + assert (All2 (eq_term_upto_univ Σ cmp_universe cmp_sort Conv) (args ++ [x0]) (x1 ++ [u'])) by (now apply All2_app). unfold unfold_fix in *. destruct (nth_error mfix idx) eqn:nth; [|easy]. @@ -1573,18 +1573,18 @@ Proof. destruct e0 as (?&?&?). eapply whne_fixapp. + unfold unfold_fix. - rewrite e1. + rewrite e2. reflexivity. + rewrite <- e. - destruct p. rewrite e3. reflexivity. + destruct p as (?& eqrarg &?). rewrite eqrarg. reflexivity. + eapply IHwh; eauto. - destruct args using MCList.rev_ind; [|rewrite mkApps_app in x; discriminate x]. now rewrite nth_error_nil in e0. Qed. -Lemma whnf_eq_term_upto_univ_napp f Σ Γ t Re Rle napp t' : +Lemma whnf_eq_term_upto_univ_napp f Σ Γ t cmp_universe cmp_sort pb napp t' : whnf f Σ Γ t -> - eq_term_upto_univ_napp Σ Re Rle napp t t' -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp t t' -> whnf f Σ Γ t'. Proof. intros wh eq. @@ -1606,11 +1606,11 @@ Proof. - apply eq_term_upto_univ_napp_mkApps_l_inv in eq as (?&?&(?&?)&->). depelim e. apply whnf_fixapp. destruct o as [[? [? ?]] |]. - eapply All2_nth_error_Some in a; eauto. - destruct a as (?&?&((? & ?)&?)&?). + eapply All2_nth_error_Some in e; eauto. + destruct e as (? & e & ? & ? & e2 & ?). rewrite e. left; eexists; split; eauto. rewrite <- e2. eapply All2_nth_error_None; eauto. - apply All2_length in a. + apply All2_length in e. right. apply nth_error_None. apply nth_error_None in H. lia. @@ -1778,7 +1778,7 @@ Section Normal. * clear -X0 nf_ind_all; revert X0; generalize (fix_context mfix); induction 1; constructor; intuition eauto. * clear -X nf_ind_all; revert X; induction 1; constructor; intuition eauto. * clear -X0 nf_ind_all. revert X0; induction 1; constructor; intuition eauto. - * clear -X1 nf_ind_all; revert X1; induction 1; constructor; cbn in *; intuition eauto. + * clear -X1 nf_ind_all; revert X1; induction 1; constructor; unfold on_local_decl, lift_wf_term in *; cbn in *; intuition eauto. * destruct 1; [eapply hne|eapply hlam|eapply hconstruct|eapply hcofix|eapply hind|eapply hprod]; eauto. clear -X nf_ind_all. revert X; generalize (fix_context mfix); induction 1; constructor; intuition eauto. clear -X0 nf_ind_all. induction X0; constructor; eauto. diff --git a/pcuic/theories/PCUICParallelReduction.v b/pcuic/theories/PCUICParallelReduction.v index a2073ad3f..5650eb39c 100644 --- a/pcuic/theories/PCUICParallelReduction.v +++ b/pcuic/theories/PCUICParallelReduction.v @@ -788,8 +788,8 @@ Section ParallelReduction. constructor. eapply t0. apply on_contexts_app => //. - destruct t1. - constructor; [eapply p|eapply p0]; + destruct t0. + constructor; [eapply o|eapply p]; apply on_contexts_app => //. Qed. @@ -806,8 +806,8 @@ Section ParallelReduction. constructor. eapply t0. apply on_contexts_app => //. - destruct t1. - constructor; [eapply p|eapply p0]; + destruct t0. + constructor; [eapply o|eapply p]; apply on_contexts_app => //. Qed. diff --git a/pcuic/theories/PCUICParallelReductionConfluence.v b/pcuic/theories/PCUICParallelReductionConfluence.v index 87d4c9c5c..19af08874 100644 --- a/pcuic/theories/PCUICParallelReductionConfluence.v +++ b/pcuic/theories/PCUICParallelReductionConfluence.v @@ -1728,9 +1728,8 @@ Section Rho. case e: lookup_env => [[decl|decl]|] //; simp rho. case eb: cst_body => [b|] //; simp rho. rewrite rename_inst inst_closed0 //. - apply declared_decl_closed in e => //. - hnf in e. rewrite eb in e. rewrite closedn_subst_instance; auto. - now move/andP: e => [-> _]. + apply declared_decl_closed in e as [Hb Ht]%andb_and => //. + hnf in Hb. rewrite eb /= in Hb. rewrite closedn_subst_instance; auto. - (* construct/cofix iota reduction *) simpl; simp rho. @@ -3199,7 +3198,7 @@ Section Rho. Definition fake_params n : context := unfold n (fun x => {| decl_name := {| binder_name := nAnon; binder_relevance := Relevant |}; decl_body := None; - decl_type := tSort Universe.type0 |}). + decl_type := tSort Sort.type0 |}). Lemma context_assumptions_fake_params n : context_assumptions (fake_params n) = n. diff --git a/pcuic/theories/PCUICPrincipality.v b/pcuic/theories/PCUICPrincipality.v index 21581f5d9..4f3da8802 100644 --- a/pcuic/theories/PCUICPrincipality.v +++ b/pcuic/theories/PCUICPrincipality.v @@ -60,7 +60,7 @@ Section Principality. Lemma isWfArity_sort Γ u : wf_local Σ Γ -> - wf_universe Σ u -> + wf_sort Σ u -> isWfArity Σ Γ (tSort u). Proof using Type. move=> wfΓ wfu. @@ -86,64 +86,61 @@ Section Principality. - apply inversion_Evar in hA. destruct hA. - apply inversion_Sort in hA as iA. 2: auto. repeat outsum. repeat outtimes. subst. - exists (tSort (Universe.super s)). + exists (tSort (Sort.super s)). int inversion_Sort. repeat outsum. repeat outtimes. now subst. - - apply inversion_Prod in hA as [dom1 [codom1 iA]]; auto. - repeat outtimes. + - apply inversion_Prod in hA as (dom1 & codom1 & t & t0 & w); auto. + apply unlift_TypUniv in t. specialize (IHu1 _ _ t) as [dom Hdom]. specialize (IHu2 _ _ t0) as [codom Hcodom]. destruct (Hdom _ t) as [e e']. eapply ws_cumul_pb_Sort_r_inv in e as [domu [red leq]]. destruct (Hcodom _ t0) as [e e'']. eapply ws_cumul_pb_Sort_r_inv in e as [codomu [cored coleq]]. - exists (tSort (Universe.sort_of_product domu codomu)). + exists (tSort (Sort.sort_of_product domu codomu)). int inversion_Prod. - repeat outsum; repeat outtimes. + destruct hB as (x & x0 & t1 & t2 & w0). + apply unlift_TypUniv in t1. + etransitivity. 1: auto. 2:eapply w0. destruct (Hdom _ t1) as [le' u1']. eapply ws_cumul_pb_Sort_r_inv in le' as [u' [redu' leu']]. destruct (Hcodom _ t2) as [le'' u2']. eapply ws_cumul_pb_Sort_r_inv in le'' as [u'' [redu'' leu'']]. constructor => //. fvs. constructor. - apply leq_universe_product_mon; auto. + apply leq_sort_product_mon; auto. pose proof (closed_red_confluence red redu') as [v' [redl redr]]. eapply invert_red_sort in redl. eapply invert_red_sort in redr. subst. now noconf redr. pose proof (closed_red_confluence cored redu'') as [v' [redl redr]]. eapply invert_red_sort in redl. eapply invert_red_sort in redr. subst. now noconf redr. - + eapply type_reduction; eauto. eapply red. + + repeat (eexists; eauto). eapply type_reduction; eauto. eapply red. + eapply type_reduction; eauto. eapply cored. - apply inversion_Lambda in hA => //; eauto. - repeat outsum. repeat outtimes. - destruct (IHu1 _ _ t) as [? ?]. - destruct (IHu2 _ _ t0) as [? ?]. + destruct hA as (x & H & t & w). + destruct (IHu2 _ _ t) as [x0 p]. destruct (p _ t). - destruct (p0 _ t0). - exists (tProd n u1 x2). + exists (tProd n u1 x0). int inversion_Lambda. + destruct hB as (x1 & _ & t1 & w1). repeat outsum. repeat outtimes. etransitivity; eauto. - apply ws_cumul_pb_Prod_l_inv in w2 as [na' [A' [B' [redA u1eq ?]]]] => //; auto. - destruct (p0 _ t4). + apply ws_cumul_pb_Prod_l_inv in w1 as [na' [A' [B' [redA u1eq ?]]]] => //; auto. + destruct (p _ t1). eapply ws_cumul_pb_Prod => //; auto. transitivity A' => //. now symmetry. - - eapply inversion_LetIn in hA as (s1 & bty & Hu2 & Hu1 & Hu3 & Hcum); auto. - destruct (IHu1 _ _ Hu1) as [? p]. - destruct (p _ Hu1). - destruct (IHu2 _ _ Hu2) as [? p']. - destruct (p' _ Hu2). - destruct (IHu3 _ _ Hu3) as [? p'']. - destruct (p'' _ Hu3). - exists (tLetIn n u1 u2 x1). + - eapply inversion_LetIn in hA as (bty & Hu12 & Hu3 & Hcum); auto. + destruct (IHu3 _ _ Hu3) as [u3' p3]. + destruct (p3 _ Hu3). + exists (tLetIn n u1 u2 u3'). int inversion_LetIn. - destruct hB as (s1' & bty' & Hu2' & Hu1' & Hu3' & Hcum'); eauto. + destruct hB as (bty' & Hu12' & Hu3' & Hcum'); eauto. etransitivity; eauto. - eapply ws_cumul_pb_LetIn; eauto using wt_cumul_pb_refl. - now specialize (p'' _ Hu3') as [? ?]. + eapply ws_cumul_pb_LetIn; eauto. + 1,2: destruct Hu12 as (? & ? & ? & _); cbn in *; now eapply wt_cumul_pb_refl. + now specialize (p3 _ Hu3') as [? ?]. - eapply inversion_App in hA as [na [dom [codom [tydom [tyarg tycodom]]]]] => //; auto. destruct (IHu2 _ _ tyarg). @@ -280,11 +277,11 @@ Section Principality. assert (consistent_instance_ext Σ (ind_universes x6) u'). { eapply type_reduction in t2. 2:eapply redr. eapply validity in t2; eauto. - destruct t2 as [s Hs]. + destruct t2 as (_ & s & Hs & _). eapply invert_type_mkApps_ind in Hs. intuition eauto. all:auto. eapply d. } assert (consistent_instance_ext Σ (ind_universes x6) x5). { eapply validity in t1; eauto. - destruct t1 as [s Hs]. + destruct t1 as (_ & s & Hs & _). eapply invert_type_mkApps_ind in Hs. intuition eauto. all:auto. eapply d. } set (p := {| proj_ind := ind; proj_npars := k; proj_arg := pars |}). transitivity (subst0 (u :: List.rev x0') (subst_instance x5 (proj_type x3))); cycle 1. @@ -365,10 +362,8 @@ Lemma principal_type_ind {cf:checker_flags} {Σ Γ c ind u u' args args'} {wfΣ: Σ ;;; Γ |- c : mkApps (tInd ind u) args -> Σ ;;; Γ |- c : mkApps (tInd ind u') args' -> (∑ ui', - PCUICEquality.R_global_instance Σ.1 (eq_universe (global_ext_constraints Σ)) - (leq_universe (global_ext_constraints Σ)) (IndRef ind) #|args| ui' u * - PCUICEquality.R_global_instance Σ.1 (eq_universe (global_ext_constraints Σ)) - (leq_universe (global_ext_constraints Σ)) (IndRef ind) #|args'| ui' u') * + cmp_ind_universes Σ ind #|args| ui' u * + cmp_ind_universes Σ ind #|args'| ui' u') * ws_cumul_pb_terms Σ Γ args args'. Proof. intros h h'. @@ -418,46 +413,25 @@ Proof. eapply eq_term_upto_univ_empty_impl; auto; typeclasses eauto. Qed. -Lemma eq_context_empty_eq_context {cf:checker_flags} {Σ : global_env_ext} {x y} : - eq_context_upto empty_global_env (eq_universe Σ) (eq_universe Σ) x y -> - eq_context_upto Σ (eq_universe Σ) (eq_universe Σ) x y. +Lemma eq_context_empty_eq_context {cf:checker_flags} {Σ : global_env_ext} {cmp_universe cmp_sort pb} {x y} : + eq_context_upto empty_global_env cmp_universe cmp_sort pb x y -> + eq_context_upto Σ cmp_universe cmp_sort pb x y. Proof. intros. eapply All2_fold_impl; tea. - intros ???? []; constructor; eauto using eq_term_empty_eq_term. - all:now apply eq_term_empty_eq_term. -Qed. - -Notation eq_term_napp Σ n x y := - (eq_term_upto_univ_napp Σ (eq_universe Σ) (eq_universe Σ) n x y). - -Notation leq_term_napp Σ n x y := - (eq_term_upto_univ_napp Σ (eq_universe Σ) (leq_universe Σ) n x y). - -Lemma eq_term_upto_univ_napp_leq {cf:checker_flags} {Σ : global_env_ext} {n x y} : - eq_term_napp Σ n x y -> - leq_term_napp Σ n x y. -Proof. - eapply eq_term_upto_univ_impl; auto; typeclasses eauto. -Qed. - -Lemma R_global_instance_empty_universe_instance Re Rle ref napp u u' : - R_global_instance empty_global_env Re Rle ref napp u u' -> - R_universe_instance Re u u'. -Proof. - rewrite /R_global_instance_gen. - now rewrite global_variance_empty. + intros ???? []; constructor; eauto. + all: eapply eq_term_upto_univ_empty_impl; tea; tc. Qed. Lemma eq_context_upto_inst_case_context {cf : checker_flags} {Σ : global_env_ext} pars pars' puinst puinst' ctx : - All2 (eq_term_upto_univ empty_global_env (eq_universe Σ) (eq_universe Σ)) pars pars' -> - R_universe_instance (eq_universe Σ) puinst puinst' -> - eq_context_upto Σ.1 (eq_universe Σ) (eq_universe Σ) (inst_case_context pars puinst ctx) + All2 (eq_term_upto_univ empty_global_env (compare_universe Σ) (compare_sort Σ) Conv) pars pars' -> + cmp_universe_instance (eq_universe Σ) puinst puinst' -> + eq_context_upto Σ.1 (compare_universe Σ) (compare_sort Σ) Conv (inst_case_context pars puinst ctx) (inst_case_context pars' puinst' ctx). Proof. intros onps oninst. rewrite /inst_case_context. - eapply eq_context_upto_subst_context. tc. + eapply eq_context_upto_subst_context. 1,2: tc. eapply eq_context_upto_univ_subst_instance; tc; auto. eapply All2_rev. eapply All2_impl; tea. intros. now eapply eq_term_empty_eq_term. @@ -481,38 +455,47 @@ Proof. (fun Σ Γ t T => forall (onu : on_udecl Σ.1 Σ.2), forall t' T' : term, Σ ;;; Γ |- t' : T' -> leq_term empty_global_env Σ t' t -> Σ;;; Γ |- t' : T) - (fun Σ Γ => wf_local Σ Γ)); auto;intros Σ wfΣ Γ wfΓ; intros. - 1-13:match goal with + (fun Σ Γ j => + on_udecl Σ.1 Σ.2 -> + lift_typing0 (fun t T => + forall t' T' : term, Σ ;;; Γ |- t' : T' -> leq_term empty_global_env Σ t' t -> Σ;;; Γ |- t' : T) + j) + (fun Σ Γ => wf_local Σ Γ)). + { intros ???? H ?; apply lift_typing_impl with (1 := H) => ??[] HT IH //. now apply IH. } + 1: now auto. + + all: intros Σ wfΣ Γ wfΓ; intros. + + 1-13:match goal with [ H : leq_term _ _ _ _ |- _ ] => depelim H end. all:try solve [econstructor; eauto]. - eapply inversion_Sort in X0 as [wf [wfs cum]]; auto. - eapply type_Cumul' with (tSort (Universe.super s)). + eapply type_Cumul' with (tSort (Sort.super s)). constructor; auto. eapply PCUICArities.isType_Sort; pcuic. - apply cumul_Sort. now apply leq_universe_super. + apply cumul_Sort. now apply leq_sort_super. - eapply inversion_Prod in X4 as [s1' [s2' [Ha [Hb Hs]]]]; auto. - specialize (X1 onu _ _ Ha). - specialize (X1 (eq_term_empty_leq_term X5_1)). + apply eq_term_empty_leq_term in X5_1 as X5_1'. apply eq_term_empty_eq_term in X5_1. eapply context_conversion in Hb. 3:{ constructor. apply conv_ctx_refl. constructor. eassumption. constructor. eauto. } all:eauto. - 2:{ constructor; eauto. now exists s1. } + 2:{ pcuic. } specialize (X3 onu _ _ Hb X5_2). econstructor; eauto. + { specialize (X1 onu). apply lift_sorting_it_impl_gen with X1 => // IH. eapply IH; tea. 1: now eapply unlift_TypUniv in Ha. } apply leq_term_empty_leq_term in X5_2. eapply context_conversion; eauto. - constructor; pcuic. constructor; try now symmetry; now constructor. + constructor; pcuic. 1: now eapply lift_sorting_forget_univ. constructor; try now symmetry; now constructor. pcuic. constructor; pcuic. constructor. now symmetry. - - eapply inversion_Lambda in X4 as (s & B & dom & codom & cum); auto. - specialize (X1 onu _ _ dom (eq_term_empty_leq_term X5_1)). + - eapply inversion_Lambda in X4 as (B & dom & codom & cum); auto. apply eq_term_empty_eq_term in X5_1. - assert(conv_context cumulAlgo_gen Σ (Γ ,, vass na ty) (Γ ,, vass n t)). + assert(conv_context cumulAlgo_gen Σ (Γ ,, vass na0 ty) (Γ ,, vass na t)). { repeat constructor; pcuic. } specialize (X3 onu t0 B). forward X3 by eapply context_conversion; eauto; pcuic. @@ -520,7 +503,7 @@ Proof. * econstructor. eauto. instantiate (1 := bty). eapply context_conversion; eauto; pcuic. constructor; pcuic. constructor; pcuic. symmetry; constructor; auto. - * have tyl := type_Lambda _ _ _ _ _ _ _ X0 X2. + * have tyl := type_Lambda _ _ _ _ _ _ X0 X2. now eapply PCUICValidity.validity in tyl. * eapply ws_cumul_pb_Prod; eauto. constructor; auto; fvs. @@ -528,23 +511,22 @@ Proof. eapply type_closed, closedn_on_free_vars in X2. now len in X2; len. - - eapply inversion_LetIn in X6 as (s1' & A & dom & bod & codom & cum); auto. - specialize (X1 onu _ _ dom (eq_term_empty_leq_term X7_2)). - specialize (X3 onu _ _ bod (eq_term_empty_leq_term X7_1)). - apply eq_term_empty_eq_term in X7_1. - apply eq_term_empty_eq_term in X7_2. - assert(Σ ⊢ Γ ,, vdef na t ty = Γ ,, vdef n b b_ty). + - eapply inversion_LetIn in X4 as (A & dombod & codom & cum); auto. + apply unlift_TermTyp in dombod as dombod', X0 as X0'. + apply eq_term_empty_eq_term in X5_1. + apply eq_term_empty_eq_term in X5_2. + assert(Σ ⊢ Γ ,, vdef na0 t ty = Γ ,, vdef na b b_ty). { constructor. eapply ws_cumul_ctx_pb_refl. fvs. constructor => //. constructor; fvs. constructor; fvs. } - specialize (X5 onu u A). - forward X5 by eapply closed_context_conversion; eauto; pcuic. - specialize (X5 X7_3). - eapply leq_term_empty_leq_term in X7_3. - have uty : Σ ;;; Γ ,, vdef na t ty |- u : b'_ty. + specialize (X3 onu u A). + forward X3 by eapply closed_context_conversion; eauto; pcuic. + specialize (X3 X5_3). + eapply leq_term_empty_leq_term in X5_3. + have uty : Σ ;;; Γ ,, vdef na0 t ty |- u : b'_ty. { eapply closed_context_conversion; eauto. pcuic. now symmetry. } eapply type_ws_cumul_pb. - * econstructor. eauto. eauto. + * econstructor. eauto. now instantiate (1 := b'_ty). * eapply PCUICValidity.validity; eauto. econstructor; eauto. @@ -635,7 +617,7 @@ Proof. eapply PCUICEquality.subst_eq_term. eapply PCUICUnivSubstitutionConv.eq_term_upto_univ_subst_instance; eauto; typeclasses eauto. - - eassert (ctx_inst _ _ _ _ _) as Hctxi by now eapply ctx_inst_impl with (1 := X5). + - eassert (ctx_inst _ _ _ _) as Hctxi by now eapply ctx_inst_impl with (1 := X5). assert (isType Σ Γ (mkApps ptm (indices ++ [c]))). { eapply validity. econstructor; eauto. all:split; eauto. solve_all. } @@ -660,7 +642,7 @@ Proof. eapply PCUICEquality.eq_term_upto_univ_it_mkLambda_or_LetIn; tea. tc. rewrite /predctx. rewrite /case_predicate_context /case_predicate_context_gen. - eapply eq_context_upto_map2_set_binder_name. tea. + eapply eq_context_upto_names_map2_set_binder_name. tea. rewrite /pre_case_predicate_context_gen. eapply eq_context_upto_inst_case_context => //. eapply All2_app. 2:constructor; pcuic. @@ -715,36 +697,28 @@ Proof. eapply closedn_on_free_vars in H0. now rewrite -plus_Sn_m -shiftnP_add. - - eapply inversion_Fix in X2 as (decl' & fixguard' & Hnth & types' & bodies & wffix & cum); auto. + - eapply inversion_Fix in X4 as (decl' & fixguard' & Hnth & types' & bodies & wffix & cum); auto. eapply type_Cumul_alt. econstructor; eauto. - eapply PCUICValidity.validity; eauto. - econstructor. 3:eapply H0. all:eauto. - eapply (All_impl X0); pcuicfo. - apply infer_typing_sort_impl with id X2; now intros []. - eapply (All_impl X1); pcuicfo; now destruct X2. - eapply All2_nth_error in a; eauto. - destruct a as [[[eqty _] _] _]. + now eapply All_nth_error in X0. + eapply All2_nth_error in e; eauto. + destruct e as (eqty & _). constructor. eapply eq_term_empty_leq_term in eqty. now eapply leq_term_empty_leq_term. - - eapply inversion_CoFix in X2 as (decl' & fixguard' & Hnth & types' & bodies & wfcofix & cum); auto. + - eapply inversion_CoFix in X4 as (decl' & fixguard' & Hnth & types' & bodies & wfcofix & cum); auto. eapply type_Cumul_alt. econstructor; eauto. - eapply PCUICValidity.validity; eauto. - eapply type_CoFix. 3:eapply H0. all:eauto. - eapply (All_impl X0); pcuicfo. - apply infer_typing_sort_impl with id X2; now intros []. - eapply (All_impl X1); pcuicfo; now destruct X2. - eapply All2_nth_error in a; eauto. - destruct a as [[[eqty _] _] _]. + now eapply All_nth_error in X0. + eapply All2_nth_error in e; eauto. + destruct e as (eqty & _). constructor. apply eq_term_empty_leq_term in eqty. now eapply leq_term_empty_leq_term. - epose proof (type_Prim _ _ _ _ _ X H H0 H1 X0). eapply validity in X4. depelim X1; depelim X3; depelim o. 1-2:econstructor; tea. - depelim X0. destruct X4 as [s ?]. + depelim X0. destruct X4 as (_ & s & ? & _). econstructor; tea. eapply inversion_Prim in X2 as [prim_ty' [cdecl' []]]; eauto. rewrite e2 in H. noconf H. @@ -753,16 +727,16 @@ Proof. simp prim_type in w |- *. eapply (ws_cumul_pb_Axiom_l_inv (args := [_])) in w as [u' [args' []]]; tea. 2:eapply declared_constant_from_gen, H0. 2:eapply p. eapply cumulAlgo_cumulSpec. etransitivity. now eapply red_ws_cumul_pb. - eapply All2_tip_l in a3 as [y' [-> Heq]]. red in r. eapply Forall2_map_inv in r. - eapply Forall2_tip_l in r. cbn. eapply ws_cumul_pb_eq_le. + eapply All2_tip_l in a3 as [y' [-> Heq]]. red in c0. + eapply Forall2_tip_l in c0. cbn. eapply ws_cumul_pb_eq_le. eapply (ws_cumul_pb_mkApps (args := [_]) (args' := [_])). - * constructor; fvs. constructor. red. eapply Forall2_map. destruct r as [? [-> eq]]. constructor. symmetry. + * constructor; fvs. constructor. red. destruct c0 as [? [-> eq]]. constructor. symmetry. etransitivity; tea. now symmetry. constructor. * constructor; [|constructor]. symmetry. etransitivity; tea. constructor; fvs. symmetry. now eapply eq_term_empty_eq_term. - eapply type_Cumul'. - eapply X1; eauto. now exists s. + eapply X1; eauto. now eapply has_sort_isType. auto. Qed. diff --git a/pcuic/theories/PCUICProgress.v b/pcuic/theories/PCUICProgress.v index 19abc554c..3e19bd568 100644 --- a/pcuic/theories/PCUICProgress.v +++ b/pcuic/theories/PCUICProgress.v @@ -102,7 +102,7 @@ Proof. revert f T. induction u; intros f T. simpl. intros. { exists T, H, s, HT. intuition pcuic. - econstructor. eexists; eauto. eexists; eauto. eapply isType_ws_cumul_pb_refl. eexists; eauto. } + econstructor. 3: eapply isType_ws_cumul_pb_refl. all: now eapply has_sort_isType. } intros Hf Ht. simpl in Hf. specialize (IHu (tApp f a) T). epose proof (IHu Hf) as (T' & H' & s' & H1 & H2 & H3 & H4); tea. @@ -113,11 +113,11 @@ Proof. unshelve econstructor. 5: eauto. 1: eauto. - 3:eapply isType_ws_cumul_pb_refl; eexists; eauto. - 1: eexists; eauto. + 3:eapply isType_ws_cumul_pb_refl; now eapply has_sort_isType. + 1: now eapply has_sort_isType. 1, 2: rewrite <- H2; lia. eapply typing_spine_pred_strengthen; tea. - eexists; eauto. clear Hs3. + now eapply has_sort_isType. clear Hs3. eapply inversion_Prod in HA as (? & ? & ? & ? & ?); tea. eapply isType_subst. econstructor. econstructor. rewrite subst_empty; eauto. econstructor; cbn; eauto. @@ -126,45 +126,45 @@ Qed. Lemma typing_ind_env_app_size `{cf : checker_flags} : forall (P : global_env_ext -> context -> term -> term -> Type) + (Pj : global_env_ext -> context -> judgment -> Type) (Pdecl := fun Σ Γ wfΓ t T tyT => P Σ Γ t T) (PΓ : global_env_ext -> context -> Type), + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (* (wfΓ : wf_local Σ Γ) *) j, + lift_typing_conj (typing Σ) (P Σ) Γ j -> Pj Σ Γ j) -> + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ), - All_local_env_over typing Pdecl Σ Γ wfΓ -> PΓ Σ Γ) -> + All_local_env_over (typing Σ) (Pdecl Σ) Γ wfΓ -> All_local_env (Pj Σ) Γ -> PΓ Σ Γ) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : nat) decl, nth_error Γ n = Some decl -> PΓ Σ Γ -> P Σ Γ (tRel n) (lift0 (S n) decl.(decl_type))) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (u : Universe.t), + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (u : sort), PΓ Σ Γ -> - wf_universe Σ u -> - P Σ Γ (tSort u) (tSort (Universe.super u))) -> + wf_sort Σ u -> + P Σ Γ (tSort u) (tSort (Sort.super u))) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : aname) (t b : term) (s1 s2 : Universe.t), + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (t b : term) (s1 s2 : sort), PΓ Σ Γ -> - Σ ;;; Γ |- t : tSort s1 -> - P Σ Γ t (tSort s1) -> - Σ ;;; Γ,, vass n t |- b : tSort s2 -> - P Σ (Γ,, vass n t) b (tSort s2) -> P Σ Γ (tProd n t b) (tSort (Universe.sort_of_product s1 s2))) -> + lift_typing typing Σ Γ (j_vass_s na t s1) -> + Pj Σ Γ (j_vass_s na t s1) -> + Σ ;;; Γ,, vass na t |- b : tSort s2 -> + P Σ (Γ,, vass na t) b (tSort s2) -> P Σ Γ (tProd na t b) (tSort (Sort.sort_of_product s1 s2))) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : aname) (t b : term) - (s1 : Universe.t) (bty : term), + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (t b bty : term), PΓ Σ Γ -> - Σ ;;; Γ |- 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)) -> + lift_typing typing Σ Γ (j_vass na t) -> + Pj Σ Γ (j_vass na t) -> + Σ ;;; Γ,, vass na t |- b : bty -> P Σ (Γ,, vass na t) b bty -> P Σ Γ (tLambda na t b) (tProd na t bty)) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : aname) (b b_ty b' : term) - (s1 : Universe.t) (b'_ty : term), + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (b b_ty b' b'_ty : term), PΓ Σ Γ -> - Σ ;;; Γ |- b_ty : tSort s1 -> - P Σ Γ b_ty (tSort s1) -> - Σ ;;; Γ |- b : b_ty -> - P Σ Γ b b_ty -> - Σ ;;; Γ,, vdef n b b_ty |- b' : b'_ty -> - P Σ (Γ,, vdef n b b_ty) b' b'_ty -> P Σ Γ (tLetIn n b b_ty b') (tLetIn n b b_ty b'_ty)) -> + lift_typing typing Σ Γ (j_vdef na b b_ty) -> + Pj Σ Γ (j_vdef na b b_ty) -> + Σ ;;; Γ,, vdef na b b_ty |- b' : b'_ty -> + P Σ (Γ,, vdef na b b_ty) b' b'_ty -> P Σ Γ (tLetIn na b b_ty b') (tLetIn na b b_ty b'_ty)) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (t : term) T B L s, PΓ Σ Γ -> @@ -178,7 +178,7 @@ forall (P : global_env_ext -> context -> term -> term -> Type) P Σ Γ (mkApps t L) B) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) cst u (decl : constant_body), - Forall_decls_typing P Σ.1 -> + on_global_env cumulSpec0 Pj Σ.1 -> PΓ Σ Γ -> declared_constant Σ.1 cst decl -> consistent_instance_ext Σ decl.(cst_universes) u -> @@ -186,14 +186,14 @@ forall (P : global_env_ext -> context -> term -> term -> Type) (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (ind : inductive) u mdecl idecl (isdecl : declared_inductive Σ.1 ind mdecl idecl), - Forall_decls_typing P Σ.1 -> + on_global_env cumulSpec0 Pj Σ.1 -> PΓ Σ Γ -> consistent_instance_ext Σ mdecl.(ind_universes) u -> P Σ Γ (tInd ind u) (subst_instance u (ind_type idecl))) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (ind : inductive) (i : nat) u mdecl idecl cdecl (isdecl : declared_constructor Σ.1 (ind, i) mdecl idecl cdecl), - Forall_decls_typing P Σ.1 -> + on_global_env cumulSpec0 Pj Σ.1 -> PΓ Σ Γ -> consistent_instance_ext Σ mdecl.(ind_universes) u -> P Σ Γ (tConstruct ind i u) (type_of_constructor mdecl cdecl (ind, i) u)) -> @@ -201,7 +201,7 @@ forall (P : global_env_ext -> context -> term -> term -> Type) (forall (Σ : global_env_ext) (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ), forall (ci : case_info) p c brs indices ps mdecl idecl (isdecl : declared_inductive Σ.1 ci.(ci_ind) mdecl idecl), - Forall_decls_typing P Σ.1 -> + on_global_env cumulSpec0 Pj Σ.1 -> PΓ Σ Γ -> mdecl.(ind_npars) = ci.(ci_npar) -> eq_context_upto_names p.(pcontext) (ind_predicate_context ci.(ci_ind) mdecl idecl) -> @@ -213,8 +213,8 @@ forall (P : global_env_ext -> context -> term -> term -> Type) wf_local Σ (Γ ,,, predctx) -> PΓ Σ (Γ ,,, predctx) -> is_allowed_elimination Σ idecl.(ind_kelim) ps -> - PCUICTyping.ctx_inst (Prop_conj typing P) Σ Γ (p.(pparams) ++ indices) - (List.rev (subst_instance p.(puinst) (mdecl.(ind_params) ,,, idecl.(ind_indices)))) -> + PCUICTyping.ctx_inst (Prop_conj typing P Σ) Γ (p.(pparams) ++ indices) + (List.rev (subst_instance p.(puinst) (mdecl.(ind_params) ,,, idecl.(ind_indices) : context))) -> Σ ;;; Γ |- c : mkApps (tInd ci.(ci_ind) p.(puinst)) (p.(pparams) ++ indices) -> P Σ Γ c (mkApps (tInd ci.(ci_ind) p.(puinst)) (p.(pparams) ++ indices)) -> isCoFinite mdecl.(ind_finite) = false -> @@ -230,7 +230,7 @@ forall (P : global_env_ext -> context -> term -> term -> Type) (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (p : projection) (c : term) u mdecl idecl cdecl pdecl (isdecl : declared_projection Σ.1 p mdecl idecl cdecl pdecl) args, - Forall_decls_typing P Σ.1 -> PΓ Σ Γ -> + on_global_env cumulSpec0 Pj Σ.1 -> PΓ Σ Γ -> Σ ;;; Γ |- c : mkApps (tInd p.(proj_ind) u) args -> P Σ Γ c (mkApps (tInd p.(proj_ind) u) args) -> #|args| = ind_npars mdecl -> @@ -241,8 +241,10 @@ forall (P : global_env_ext -> context -> term -> term -> Type) fix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> PΓ Σ (Γ ,,, types) -> - All (on_def_type (lift_typing2 typing P Σ) Γ) mfix -> - All (on_def_body (lift_typing2 typing P Σ) types Γ) mfix -> + All (on_def_type (lift_typing typing Σ) Γ) mfix -> + All (on_def_type (Pj Σ) Γ) mfix -> + All (on_def_body (lift_typing typing Σ) types Γ) mfix -> + All (on_def_body (Pj Σ) types Γ) mfix -> wf_fixpoint Σ.1 mfix -> P Σ Γ (tFix mfix n) decl.(dtype)) -> @@ -251,8 +253,10 @@ forall (P : global_env_ext -> context -> term -> term -> Type) cofix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> PΓ Σ (Γ ,,, types) -> - All (on_def_type (lift_typing2 typing P Σ) Γ) mfix -> - All (on_def_body (lift_typing2 typing P Σ) types Γ) mfix -> + All (on_def_type (lift_typing typing Σ) Γ) mfix -> + All (on_def_type (Pj Σ) Γ) mfix -> + All (on_def_body (lift_typing typing Σ) types Γ) mfix -> + All (on_def_body (Pj Σ) types Γ) mfix -> wf_cofixpoint Σ.1 mfix -> P Σ Γ (tCoFix mfix n) decl.(dtype)) -> @@ -274,10 +278,10 @@ forall (P : global_env_ext -> context -> term -> term -> Type) Σ ;;; Γ |- A <=s B -> P Σ Γ t B) -> - env_prop P PΓ. + env_prop P Pj PΓ. Proof. - intros P Pdecl PΓ. - intros XΓ X X0 X1 X2 X3 X4 X5 X6 X7 X8 X9 X10 X11 X12 Σ wfΣ Γ t T H. + intros P Pj Pdecl PΓ. + intros Xj XΓ X X0 X1 X2 X3 X4 X5 X6 X7 X8 X9 X10 X11 X12 Σ wfΣ Γ t T H. eapply typing_ind_env_app_size; eauto. clear Σ wfΣ Γ t T H. intros Σ wfΣ Γ wfΓ t na A B u s HΓ Hprod IHprod Ht IHt IH Hu IHu. pose proof (mkApps_decompose_app t). @@ -310,45 +314,45 @@ Qed. Lemma typing_ind_env `{cf : checker_flags} : forall (P : global_env_ext -> context -> term -> term -> Type) + (Pj : global_env_ext -> context -> judgment -> Type) (Pdecl := fun Σ Γ wfΓ t T tyT => P Σ Γ t T) (PΓ : global_env_ext -> context -> Type), + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (* (wfΓ : wf_local Σ Γ) *) j, + lift_typing_conj (typing Σ) (P Σ) Γ j -> Pj Σ Γ j) -> + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ), - All_local_env_over typing Pdecl Σ Γ wfΓ -> PΓ Σ Γ) -> + All_local_env_over (typing Σ) (Pdecl Σ) Γ wfΓ -> All_local_env (Pj Σ) Γ -> PΓ Σ Γ) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : nat) decl, nth_error Γ n = Some decl -> PΓ Σ Γ -> P Σ Γ (tRel n) (lift0 (S n) decl.(decl_type))) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (u : Universe.t), + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (u : sort), PΓ Σ Γ -> - wf_universe Σ u -> - P Σ Γ (tSort u) (tSort (Universe.super u))) -> + wf_sort Σ u -> + P Σ Γ (tSort u) (tSort (Sort.super u))) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : aname) (t b : term) (s1 s2 : Universe.t), + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (t b : term) (s1 s2 : sort), PΓ Σ Γ -> - Σ ;;; Γ |- t : tSort s1 -> - P Σ Γ t (tSort s1) -> - Σ ;;; Γ,, vass n t |- b : tSort s2 -> - P Σ (Γ,, vass n t) b (tSort s2) -> P Σ Γ (tProd n t b) (tSort (Universe.sort_of_product s1 s2))) -> + lift_typing typing Σ Γ (j_vass_s na t s1) -> + Pj Σ Γ (j_vass_s na t s1) -> + Σ ;;; Γ,, vass na t |- b : tSort s2 -> + P Σ (Γ,, vass na t) b (tSort s2) -> P Σ Γ (tProd na t b) (tSort (Sort.sort_of_product s1 s2))) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : aname) (t b : term) - (s1 : Universe.t) (bty : term), + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (t b bty : term), PΓ Σ Γ -> - Σ ;;; Γ |- 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)) -> + lift_typing typing Σ Γ (j_vass na t) -> + Pj Σ Γ (j_vass na t) -> + Σ ;;; Γ,, vass na t |- b : bty -> P Σ (Γ,, vass na t) b bty -> P Σ Γ (tLambda na t b) (tProd na t bty)) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : aname) (b b_ty b' : term) - (s1 : Universe.t) (b'_ty : term), + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (b b_ty b' b'_ty : term), PΓ Σ Γ -> - Σ ;;; Γ |- b_ty : tSort s1 -> - P Σ Γ b_ty (tSort s1) -> - Σ ;;; Γ |- b : b_ty -> - P Σ Γ b b_ty -> - Σ ;;; Γ,, vdef n b b_ty |- b' : b'_ty -> - P Σ (Γ,, vdef n b b_ty) b' b'_ty -> P Σ Γ (tLetIn n b b_ty b') (tLetIn n b b_ty b'_ty)) -> + lift_typing typing Σ Γ (j_vdef na b b_ty) -> + Pj Σ Γ (j_vdef na b b_ty) -> + Σ ;;; Γ,, vdef na b b_ty |- b' : b'_ty -> + P Σ (Γ,, vdef na b b_ty) b' b'_ty -> P Σ Γ (tLetIn na b b_ty b') (tLetIn na b b_ty b'_ty)) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (t : term) T B L s, PΓ Σ Γ -> @@ -361,7 +365,7 @@ Lemma typing_ind_env `{cf : checker_flags} : P Σ Γ (mkApps t L) B) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) cst u (decl : constant_body), - Forall_decls_typing P Σ.1 -> + on_global_env cumulSpec0 Pj Σ.1 -> PΓ Σ Γ -> declared_constant Σ.1 cst decl -> consistent_instance_ext Σ decl.(cst_universes) u -> @@ -369,14 +373,14 @@ Lemma typing_ind_env `{cf : checker_flags} : (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (ind : inductive) u mdecl idecl (isdecl : declared_inductive Σ.1 ind mdecl idecl), - Forall_decls_typing P Σ.1 -> + on_global_env cumulSpec0 Pj Σ.1 -> PΓ Σ Γ -> consistent_instance_ext Σ mdecl.(ind_universes) u -> P Σ Γ (tInd ind u) (subst_instance u (ind_type idecl))) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (ind : inductive) (i : nat) u mdecl idecl cdecl (isdecl : declared_constructor Σ.1 (ind, i) mdecl idecl cdecl), - Forall_decls_typing P Σ.1 -> + on_global_env cumulSpec0 Pj Σ.1 -> PΓ Σ Γ -> consistent_instance_ext Σ mdecl.(ind_universes) u -> P Σ Γ (tConstruct ind i u) (type_of_constructor mdecl cdecl (ind, i) u)) -> @@ -384,7 +388,7 @@ Lemma typing_ind_env `{cf : checker_flags} : (forall (Σ : global_env_ext) (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ), forall (ci : case_info) p c brs indices ps mdecl idecl (isdecl : declared_inductive Σ.1 ci.(ci_ind) mdecl idecl), - Forall_decls_typing P Σ.1 -> + on_global_env cumulSpec0 Pj Σ.1 -> PΓ Σ Γ -> mdecl.(ind_npars) = ci.(ci_npar) -> eq_context_upto_names p.(pcontext) (ind_predicate_context ci.(ci_ind) mdecl idecl) -> @@ -396,8 +400,8 @@ Lemma typing_ind_env `{cf : checker_flags} : wf_local Σ (Γ ,,, predctx) -> PΓ Σ (Γ ,,, predctx) -> is_allowed_elimination Σ idecl.(ind_kelim) ps -> - PCUICTyping.ctx_inst (Prop_conj typing P) Σ Γ (p.(pparams) ++ indices) - (List.rev (subst_instance p.(puinst) (mdecl.(ind_params) ,,, idecl.(ind_indices)))) -> + PCUICTyping.ctx_inst (Prop_conj typing P Σ) Γ (p.(pparams) ++ indices) + (List.rev (subst_instance p.(puinst) (mdecl.(ind_params) ,,, idecl.(ind_indices) : context))) -> Σ ;;; Γ |- c : mkApps (tInd ci.(ci_ind) p.(puinst)) (p.(pparams) ++ indices) -> P Σ Γ c (mkApps (tInd ci.(ci_ind) p.(puinst)) (p.(pparams) ++ indices)) -> isCoFinite mdecl.(ind_finite) = false -> @@ -413,7 +417,7 @@ Lemma typing_ind_env `{cf : checker_flags} : (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (p : projection) (c : term) u mdecl idecl cdecl pdecl (isdecl : declared_projection Σ.1 p mdecl idecl cdecl pdecl) args, - Forall_decls_typing P Σ.1 -> PΓ Σ Γ -> + on_global_env cumulSpec0 Pj Σ.1 -> PΓ Σ Γ -> Σ ;;; Γ |- c : mkApps (tInd p.(proj_ind) u) args -> P Σ Γ c (mkApps (tInd p.(proj_ind) u) args) -> #|args| = ind_npars mdecl -> @@ -424,8 +428,10 @@ Lemma typing_ind_env `{cf : checker_flags} : fix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> PΓ Σ (Γ ,,, types) -> - All (on_def_type (lift_typing2 typing P Σ) Γ) mfix -> - All (on_def_body (lift_typing2 typing P Σ) types Γ) mfix -> + All (on_def_type (lift_typing typing Σ) Γ) mfix -> + All (on_def_type (Pj Σ) Γ) mfix -> + All (on_def_body (lift_typing typing Σ) types Γ) mfix -> + All (on_def_body (Pj Σ) types Γ) mfix -> wf_fixpoint Σ.1 mfix -> P Σ Γ (tFix mfix n) decl.(dtype)) -> @@ -434,8 +440,10 @@ Lemma typing_ind_env `{cf : checker_flags} : cofix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> PΓ Σ (Γ ,,, types) -> - All (on_def_type (lift_typing2 typing P Σ) Γ) mfix -> - All (on_def_body (lift_typing2 typing P Σ) types Γ) mfix -> + All (on_def_type (lift_typing typing Σ) Γ) mfix -> + All (on_def_type (Pj Σ) Γ) mfix -> + All (on_def_body (lift_typing typing Σ) types Γ) mfix -> + All (on_def_body (Pj Σ) types Γ) mfix -> wf_cofixpoint Σ.1 mfix -> P Σ Γ (tCoFix mfix n) decl.(dtype)) -> @@ -459,7 +467,7 @@ Lemma typing_ind_env `{cf : checker_flags} : Σ ;;; Γ |- A <=s B -> P Σ Γ t B) -> - env_prop P PΓ. + env_prop P Pj PΓ. Proof. intros P Pdecl PΓ; unfold env_prop. intros XΓ X X0 X1 X2 X3 X4 X5 X6 X7 X8 X9 X10 X11 X12 X13 Σ wfΣ Γ t T H. @@ -687,13 +695,17 @@ Qed. Lemma progress_env_prop `{cf : checker_flags}: env_prop (fun Σ Γ t T => axiom_free Σ -> Γ = [] -> Σ ;;; Γ |- t : T -> {t' & Σ ⊢ t ⇝ᵥ t'} + (value Σ t)) + (fun Σ Γ j => axiom_free Σ -> Γ = [] -> lift_typing0 (fun t T => Σ ;;; Γ |- t : T -> {t' & Σ ⊢ t ⇝ᵥ t'} + (value Σ t)) j) (fun _ _ => True). Proof with eauto with wcbv; try congruence. eapply typing_ind_env... + - intros Σ wfΣ Γ j Hj Hax eq. + apply lift_typing_impl with (1 := Hj) => ?? [] // H IH //. now apply IH. - intros Σ wfΣ Γ wfΓ n decl Hdecl _ Hax -> Hty. destruct n; inv Hdecl. - - intros Σ wfΣ Γ _ n b b_ty b' s1 b'_ty _ Hb_ty IHb_ty Hb IHb Hb' IHb' Hax -> H. - destruct (IHb Hax eq_refl) as [ [t' IH] | IH]; eauto with wcbv. + - intros Σ wfΣ Γ _ n b b_ty b' b'_ty _ Hb IHb Hb' IHb' Hax -> H. + specialize (IHb Hax eq_refl). + apply unlift_TermTyp in Hb, IHb as [ [t' IH] | IH]; eauto with wcbv. - intros Σ wfΣ Γ _ t T B L s _ HT IHT Ht IHt HL Hax -> H. clear HT IHT. induction HL in H, t, Ht, IHt |- *. diff --git a/pcuic/theories/PCUICReduction.v b/pcuic/theories/PCUICReduction.v index dcf117dfe..8a0f0ee6b 100644 --- a/pcuic/theories/PCUICReduction.v +++ b/pcuic/theories/PCUICReduction.v @@ -153,8 +153,8 @@ Inductive red1 (Σ : global_env) (Γ : context) : term -> term -> Type := where " Σ ;;; Γ |- t ⇝ u " := (red1 Σ Γ t u). -Definition red1_ctx Σ := (OnOne2_local_env (on_one_decl (fun Δ t t' => red1 Σ Δ t t'))). -Definition red1_ctx_rel Σ Γ := (OnOne2_local_env (on_one_decl (fun Δ t t' => red1 Σ (Γ ,,, Δ) t t'))). +Definition red1_ctx Σ := (OnOne2_local_env (fun Δ => on_one_decl (fun t t' => red1 Σ Δ t t'))). +Definition red1_ctx_rel Σ Γ := (OnOne2_local_env (fun Δ => on_one_decl (fun t t' => red1 Σ (Γ ,,, Δ) t t'))). Lemma red1_ind_all : forall (Σ : global_env) (P : context -> term -> term -> Type), @@ -364,8 +364,8 @@ Definition red Σ Γ := clos_refl_trans (fun t u : term => Σ;;; Γ |- t ⇝ u). Notation " Σ ;;; Γ |- t ⇝* u " := (red Σ Γ t u) (at level 50, Γ, t, u at next level). Definition red_one_ctx_rel (Σ : global_env) (Γ : context) := - OnOne2_local_env - (on_one_decl (fun (Δ : context) (t t' : term) => red Σ (Γ,,, Δ) t t')). + OnOne2_local_env (fun Δ => + on_one_decl (fun (t t' : term) => red Σ (Γ,,, Δ) t t')). Definition red_ctx_rel Σ Γ := clos_refl_trans (red1_ctx_rel Σ Γ). @@ -726,7 +726,7 @@ Section ReductionCongruence. (@OnOne2 (context × _) (Trel_conj (on_Trel (red1_ctx_rel Σ Γ) fst) (on_Trel eq snd))). Definition red_one_context_decl_rel Σ Γ := - (OnOne2_local_env (on_one_decl (fun Δ t t' => red Σ (Γ ,,, Δ) t t'))). + (OnOne2_local_env (fun Δ => on_one_decl (fun t t' => red Σ (Γ ,,, Δ) t t'))). Notation red_one_context_decl Γ := (@OnOne2 (context × _) @@ -2165,7 +2165,7 @@ Section Stacks. Lemma red1_fill_context_hole Γ π pcontext u v : red1 Σ (Γ,,, stack_context π,,, context_hole_context pcontext) u v -> - OnOne2_local_env (on_one_decl (fun Γ' => red1 Σ (Γ,,, stack_context π,,, Γ'))) + OnOne2_local_env (fun Γ' => on_one_decl (red1 Σ (Γ,,, stack_context π,,, Γ'))) (fill_context_hole pcontext u) (fill_context_hole pcontext v). Proof using Type. diff --git a/pcuic/theories/PCUICSN.v b/pcuic/theories/PCUICSN.v index fbd8706d4..a8a2af995 100644 --- a/pcuic/theories/PCUICSN.v +++ b/pcuic/theories/PCUICSN.v @@ -7,6 +7,7 @@ From MetaCoq.PCUIC Require Import PCUICUnivSubstitutionTyp. Require Import Equations.Prop.DepElim. +Require Import ssreflect. (** We assume normalization of the reduction. We state is as well-foundedness of the reduction. @@ -61,7 +62,7 @@ Section Alpha. Context {cf : checker_flags} {no : normalizing_flags}. Context (Σ : global_env_ext) {normalization : NormalizationIn Σ}. - Notation eqt u v := (∥ upto_names u v ∥). + Notation eqt u v := (∥ eq u v ∥). (* Can be made into alpha-equality, but not using PCUICEquality.eq_term_upto_univ_napp *) Definition cored' Γ u v := exists u' v', cored Σ Γ u' v' /\ eqt u u' /\ eqt v v'. @@ -83,15 +84,9 @@ Section Alpha. + eapply cored_trans'. all: eassumption. Qed. - Local Instance substu_pres_eq : SubstUnivPreserving eq. + Local Instance substu_pres_eq {T} `{UnivSubst T} : SubstUnivPreserving eq (@eq T). Proof using Type. - red. intros s u u'. - unfold R_universe_instance. - intros f. eapply Forall2_map_inv in f. - assert (u = u') as ->. - { induction f; cbn; auto. f_equal; auto. - now eapply Universe.make_inj in H. } - reflexivity. + move => s u u' /cmp_universe_instance_eq -> //. Qed. Lemma cored'_postpone : @@ -103,17 +98,16 @@ Section Alpha. apply cored_alt in r. destruct r as [r]. induction r in u, v, hu, hv. - - eapply red1_eq_term_upto_univ_r in r. 10: eassumption. - all:tc. - destruct r as [u' [r e]]. + - (* eapply red1_eq_alpha in r as [u' [r e]]. *) + subst x. rename y into u'. exists u'. split. * constructor. assumption. * constructor. etransitivity. 1: eauto. now symmetry. - specialize (IHr1 y v). - destruct IHr1 as [y' [h1 [e1]]]; auto. reflexivity. + destruct IHr1 as [y' [h1 [e1]]]; auto. specialize (IHr2 u y'). - destruct IHr2 as [u' [h2 ?]]; auto. now symmetry. + destruct IHr2 as [u' [h2 ?]]; auto. exists u'. split. + eapply cored_trans'. all: eauto. + assumption. @@ -171,7 +165,6 @@ Section Alpha. - constructor; reflexivity. Qed. - (* TODO Maybe switch to eq_context *) Lemma cored_eq_context_upto_names : forall Γ Δ u v, eq_context_upto_names Γ Δ -> @@ -187,102 +180,6 @@ Section Alpha. + eapply IHh1. assumption. Qed. - Lemma cored_eq_term_upto : - forall Σ' Re Rle Γ u v u', - RelationClasses.Reflexive Re -> - SubstUnivPreserving Re -> - RelationClasses.Reflexive Rle -> - SubstUnivPreserving Rle -> - RelationClasses.Symmetric Re -> - RelationClasses.Transitive Re -> - RelationClasses.Transitive Rle -> - RelationClasses.subrelation Re Rle -> - eq_term_upto_univ Σ' Re Rle u u' -> - cored Σ Γ v u -> - exists v', cored Σ Γ v' u' /\ ∥ eq_term_upto_univ Σ' Re Rle v v' ∥. - Proof using Type. - intros Σ' Re Rle Γ u v u' X X0 X1 X2 X3 X4 X5 X6 e h. - apply cored_alt in h as [h]. - induction h in u', e |- *. - - eapply red1_eq_term_upto_univ_l in r. 9: eauto. all: auto. - destruct r as [? [? ?]]. - eexists. split. - + constructor. eassumption. - + constructor. assumption. - - specialize (IHh1 _ e). destruct IHh1 as [y' [r1 [e1]]]. - specialize (IHh2 _ e1). destruct IHh2 as [z' [r2 [e2]]]. - exists z'. split. - + eapply cored_trans'. all: eassumption. - + constructor. assumption. - Qed. - - Lemma cored_eq_context_upto : - forall Σ' Re Γ Δ u v, - RelationClasses.Reflexive Re -> - RelationClasses.Symmetric Re -> - RelationClasses.Transitive Re -> - SubstUnivPreserving Re -> - eq_context_upto Σ' Re Re Γ Δ -> - cored Σ Γ u v -> - exists u', cored Σ Δ u' v /\ ∥ eq_term_upto_univ Σ' Re Re u u' ∥. - Proof using Type. - intros Σ' Re Γ Δ u v hRe1 hRe2 hRe3 hRe4 e h. - apply cored_alt in h as [h]. - induction h. - - eapply red1_eq_context_upto_l in r. all: eauto. 2:tc. - destruct r as [? [? ?]]. - eexists. split. - + constructor. eassumption. - + constructor. assumption. - - destruct IHh1 as [x' [r1 [e1]]]. - destruct IHh2 as [y' [r2 [e2]]]. - eapply cored_eq_term_upto in r2. 10: exact e1. all: auto. - 2:tc. - destruct r2 as [y'' [r2' [e2']]]. - exists y''. split. - * eapply cored_trans'. all: eassumption. - * constructor. eapply eq_term_upto_univ_trans. all: eauto. - Qed. - - (* Lemma eq_context_upto_nlctx : - forall Γ, - eq_context_upto Σ eq eq Γ (nlctx Γ). - Proof. - intros Γ. - induction Γ as [| [na [b|] A] Γ ih ]. - - constructor. - - simpl. constructor; simpl; try apply binder_anonymize; tas. - + constructor; tas; auto. eapply eq_term_upto_univ_tm_nl. - all: auto. - eapply eq_term_upto_univ_tm_nl. - all: auto. - - simpl. constructor; auto. constructor. - + apply binder_anonymize. - + simpl. eapply eq_term_upto_univ_tm_nl. - all: auto. - Qed. - - Lemma cored_cored'_nl : - forall Γ u v, - cored Σ Γ u v -> - cored' (nlctx Γ) (nl u) (nl v). - Proof. - intros Γ u v h. - eapply cored_eq_context_upto in h. - 6: eapply eq_context_upto_nlctx. - all: auto. - - destruct h as [u' [r [e]]]. - eexists _, _. intuition eauto. - + constructor. eapply eq_term_trans. - * eapply eq_term_sym. eapply eq_term_tm_nl. - * eapply eq_term_upto_univ_impl; eauto. all:typeclasses eauto. - + constructor. eapply eq_term_sym. eapply eq_term_tm_nl. - - intros ? ? ? []. auto. - - intros ? ? ? r. apply Forall2_eq in r. apply map_inj in r. - + subst. reflexivity. - + apply Universe.make_inj. - Qed. *) - Lemma cored_cored' : forall Γ u v, cored Σ Γ u v -> diff --git a/pcuic/theories/PCUICSR.v b/pcuic/theories/PCUICSR.v index 150af3405..4642b5c6f 100644 --- a/pcuic/theories/PCUICSR.v +++ b/pcuic/theories/PCUICSR.v @@ -26,7 +26,7 @@ Derive Signature for OnOne2_local_env. Ltac rename_hyp h ht ::= my_rename_hyp h ht. Arguments Nat.sub : simpl nomatch. -Arguments Universe.sort_of_product : simpl nomatch. +Arguments Sort.sort_of_product : simpl nomatch. (* Preservation of wf_*fixpoint *) @@ -170,13 +170,8 @@ Qed. #[global] Hint Extern 0 (conv_context _ _ _) => constructor : pcuic. #[global] Hint Extern 0 (cumul_context _ _ _) => constructor : pcuic. -#[global] Hint Extern 4 (∑ s : Universe.t, typing _ _ ?T (tSort s)) => - match goal with - | [ H : isType _ _ T |- _ ] => exact H - end : pcuic. - Ltac unfold_pcuic := - progress (unfold PCUICTypingDef.typing, PCUICLookup.wf_universe in * ). + progress (unfold PCUICTypingDef.typing, PCUICLookup.wf_sort in * ). #[global] Hint Extern 10 => unfold_pcuic : pcuic. #[global] Hint Resolve red_conv red1_red red_cumul : pcuic. @@ -404,28 +399,27 @@ Lemma isType_weaken {cf} {Σ Γ Δ T} {wfΣ : wf Σ} : isType Σ (Δ ,,, Γ) T. Proof. intros wfΔ HT. - apply infer_typing_sort_impl with id HT; intros Hs. + apply lift_typing_impl with (1 := HT) => ?? Hs. eapply weaken_ctx => //. Qed. (** The crucial property on constructors of cumulative inductive types for type preservation: we don't need to compare their instances when fully applied. *) -Lemma R_global_instance_cstr_irrelevant {cf} {Σ} {wfΣ : wf Σ} {ci c} {mdecl idecl cdecl u u'} : +Lemma cmp_global_instance_cstr_irrelevant {cf} {Σ} {wfΣ : wf Σ} {ci c} {mdecl idecl cdecl u u'} : declared_constructor Σ (ci, c) mdecl idecl cdecl -> - R_ind_universes Σ ci (context_assumptions (ind_params mdecl) + #|cstr_indices cdecl|) u u' -> - R_global_instance Σ.1 (eq_universe Σ) (eq_universe Σ) (ConstructRef ci c) + cmp_ind_universes Σ ci (context_assumptions (ind_params mdecl) + #|cstr_indices cdecl|) u u' -> + cmp_global_instance Σ.1 (compare_universe Σ) Conv (ConstructRef ci c) (ind_npars mdecl + context_assumptions (cstr_args cdecl)) u u'. Proof. intros declc. pose proof (on_declared_constructor declc). pose proof (on_declared_constructor declc) as [[onind oib] [ctor_sorts [hnth onc]]]. - intros Hu. pose proof (R_global_instance_length _ _ _ _ _ _ _ Hu). - rewrite /R_global_instance_gen /R_opt_variance /= /lookup_constructor /lookup_constructor_gen. + intros Hu. pose proof (cmp_global_instance_length _ _ _ _ _ _ _ Hu). + rewrite /cmp_global_instance_gen /cmp_opt_variance /= /lookup_constructor /lookup_constructor_gen. unshelve epose proof (declc' := declared_constructor_to_gen declc); eauto. rewrite (declared_inductive_lookup_gen declc'.p1) (proj2 declc'). rewrite -(cstr_args_length onc). - case: leb_spec_Set; try lia. move=> _ /=; cbn. - now apply R_universe_instance_variance_irrelevant. + case: leb_spec_Set; try lia. Qed. Lemma wf_pre_case_branch_context {cf} {Σ} {wfΣ : wf Σ} {Γ ci mdecl idecl p} {br : branch term} {cdecl} : @@ -565,8 +559,8 @@ Proof. { eapply eq_context_upto_ws_cumul_ctx_pb => //. rewrite is_closed_context_set_binder_name //. apply eq_context_upto_cat. reflexivity. - apply eq_context_gen_eq_context_upto; tc. - apply eq_binder_annots_eq_context_gen_ctx => //. } + apply eq_context_upto_names_eq_context_upto; tc. + apply eq_binder_annots_eq => //. } depelim a; cbn in *; constructor; auto; eapply (ws_cumul_pb_ws_cumul_ctx (pb':=Conv)); tea. Qed. @@ -639,7 +633,7 @@ Proof. Qed. *) Lemma OnOne2_local_env_forget_types P ctx ctx' : - OnOne2_local_env (on_one_decl P) ctx ctx' -> + OnOne2_local_env (fun Γ => on_one_decl1 P Γ) ctx ctx' -> forget_types ctx = forget_types ctx'. Proof. induction 1; simpl. @@ -649,7 +643,7 @@ Proof. Qed. (* Lemma OnOne2_local_env_All Σ P Q R ctx ctx' : - OnOne2_local_env (on_one_decl P) ctx ctx' -> + OnOne2_local_env (fun Γ => on_one_decl1 P Γ) ctx ctx' -> All_local_env (lift_typing Q Σ) ctx -> (forall Γ t t' ty, All_local_env (lift_typing R Σ) Γ -> lift_typing Q Σ Γ t ty -> P Γ t t' -> lift_typing R Σ Γ t' ty) -> @@ -676,7 +670,7 @@ Qed. *) From MetaCoq.PCUIC Require Import PCUICContextReduction PCUICOnFreeVars. (* Lemma red_one_decl_conv_context {cf} {Σ} {wfΣ : wf Σ} {Γ Δ Δ'} : - OnOne2_local_env (on_one_decl (fun Δ : context => red1 Σ (Γ ,,, Δ))) Δ Δ' -> + OnOne2_local_env (fun Δ : context => on_one_decl (red1 Σ (Γ ,,, Δ))) Δ Δ' -> conv_context Σ (Γ ,,, Δ) (Γ ,,, Δ'). Proof. intros o. @@ -689,7 +683,7 @@ Qed. *) Lemma red_one_decl_red_ctx {cf} {Σ} {wfΣ : wf Σ} {Γ Δ Δ'} : is_closed_context (Γ ,,, Δ) -> - OnOne2_local_env (on_one_decl (fun Δ : context => red1 Σ (Γ ,,, Δ))) Δ Δ' -> + OnOne2_local_env (fun Δ : context => on_one_decl (red1 Σ (Γ ,,, Δ))) Δ Δ' -> red_ctx Σ (Γ ,,, Δ) (Γ ,,, Δ'). Proof. intros iscl o. @@ -721,7 +715,7 @@ Qed. Lemma red_one_decl_red_context {cf} {Σ} {wfΣ : wf Σ} {Γ Δ Δ'} : is_closed_context (Γ ,,, Δ) -> - OnOne2_local_env (on_one_decl (fun Δ : context => red1 Σ (Γ ,,, Δ))) Δ Δ' -> + OnOne2_local_env (fun Δ : context => on_one_decl (red1 Σ (Γ ,,, Δ))) Δ Δ' -> Σ ⊢ Γ ,,, Δ ⇝ Γ ,,, Δ'. Proof. intros iscl o. @@ -731,14 +725,14 @@ Qed. Lemma red_one_decl_ws_cumul_ctx_pb {cf} {Σ} {wfΣ : wf Σ} {Γ Δ Δ'} : is_closed_context (Γ ,,, Δ) -> - OnOne2_local_env (on_one_decl (fun Δ : context => red1 Σ (Γ ,,, Δ))) Δ Δ' -> + OnOne2_local_env (fun Δ : context => on_one_decl (red1 Σ (Γ ,,, Δ))) Δ Δ' -> Σ ⊢ Γ ,,, Δ = Γ ,,, Δ'. Proof. intros; now eapply red_ctx_ws_cumul_ctx_pb, red_one_decl_red_context. Qed. Lemma red1_it_mkLambda_or_LetIn_ctx {cf} {Σ} {wfΣ : wf Σ} Γ Δ Δ' u : - OnOne2_local_env (on_one_decl (fun Δ : context => red1 Σ (Γ ,,, Δ))) Δ Δ' -> + OnOne2_local_env (fun Δ : context => on_one_decl (red1 Σ (Γ ,,, Δ))) Δ Δ' -> red1 Σ Γ (it_mkLambda_or_LetIn Δ u) (it_mkLambda_or_LetIn Δ' u). Proof. @@ -771,8 +765,8 @@ Qed. *) Lemma ctx_inst_merge {cf} {Σ} {wfΣ : wf Σ} Γ inst inst' Δ : wf_local Σ (Γ ,,, (List.rev Δ)) -> PCUICTyping.ctx_inst - (fun (Σ : global_env_ext) (Γ : context) (t T : term) => - forall u : term, closed_red1 Σ Γ t u -> Σ;;; Γ |- u : T) Σ Γ inst Δ -> + (fun (Γ : context) (t T : term) => + forall u : term, closed_red1 Σ Γ t u -> Σ;;; Γ |- u : T) Γ inst Δ -> ctx_inst Σ Γ inst Δ -> OnOne2 (closed_red1 Σ Γ) inst inst' -> ctx_inst Σ Γ inst' Δ. @@ -796,9 +790,9 @@ Proof. now eapply closed_red1_red. repeat constructor. specialize (t0 _ c0). - eapply wf_local_app_inv, substitution_wf_local; tea. + eapply All_local_env_app_inv, substitution_wf_local; tea. now eapply subslet_ass_tip. - eapply wf_local_app_inv, substitution_wf_local; tea. + eapply All_local_env_app_inv, substitution_wf_local; tea. now eapply subslet_ass_tip. constructor; auto. eapply IHc. rewrite -subst_context_subst_telescope. @@ -812,14 +806,14 @@ Proof. eapply substitution_wf_local; tea. repeat (constructor; tea). eapply subslet_def. constructor. all:rewrite !subst_empty //. - eapply wf_local_app_inv in wf as [wf _]. now depelim wf. + eapply All_local_env_app_inv in wf as [wf _]. depelim wf. now eapply unlift_TermTyp. Qed. Lemma ctx_inst_merge' {cf} {Σ} {wfΣ : wf Σ} Γ inst inst' Δ : wf_local Σ (Γ ,,, Δ) -> PCUICTyping.ctx_inst - (fun (Σ : global_env_ext) (Γ : context) (t T : term) => - forall u : term, closed_red1 Σ Γ t u → Σ;;; Γ |- u : T) Σ Γ inst (List.rev Δ) -> + (fun (Γ : context) (t T : term) => + forall u : term, closed_red1 Σ Γ t u → Σ;;; Γ |- u : T) Γ inst (List.rev Δ) -> ctx_inst Σ Γ inst (List.rev Δ) -> OnOne2 (closed_red1 Σ Γ) inst inst' -> ctx_inst Σ Γ inst' (List.rev Δ). @@ -1196,6 +1190,7 @@ Proof. intros. pose proof (on_declared_constructor H) as [[onmind oib] [cs [hnth onc]]]. pose proof (onc.(on_cindices)). + eapply ctx_inst_impl in X. 2: intros ??; apply unlift_TermTyp. now eapply ctx_inst_open_terms in X. Qed. @@ -1234,6 +1229,7 @@ Proof. unshelve epose proof (H' := declared_constructor_to_gen H); eauto. pose proof (on_declared_constructor H) as [[onmind oib] [cs [hnth onc]]]. pose proof (onc.(on_cindices)). + eapply ctx_inst_impl in X. 2: intros ??; apply unlift_TermTyp. eapply ctx_inst_inst in X; tea. eapply ctx_inst_open_terms in X. eapply All_map, All_impl; tea. @@ -1426,7 +1422,7 @@ Lemma wf_subslet_skipn {cf Σ Γ s Δ n} : wf_subslet Σ Γ s Δ → wf_subslet Σ Γ (skipn n s) (skipn n Δ). Proof. intros []; split; auto using subslet_skipn. - now eapply wf_local_app_skipn. + now eapply All_local_env_app_skipn. Qed. Lemma isType_subst_arities {cf} {Σ} {wfΣ : wf Σ} {ind mdecl idecl u} {Γ T} : @@ -1456,10 +1452,10 @@ Lemma isType_expand_lets {cf} {Σ} {wfΣ : wf Σ} {Γ Δ T} : isType Σ (Γ ,,, Δ) T -> isType Σ (smash_context [] Γ ,,, expand_lets_ctx Γ Δ) (expand_lets_k Γ #|Δ| T). Proof. - move=> [] s. + intros (_ & s & Hs & _). eapply has_sort_isType. revert Hs. rewrite -[_ ,,, _]app_context_nil_l app_context_assoc. move/typing_expand_lets_gen; rewrite app_context_nil_l => hs. - now exists s. + apply hs. Qed. Lemma isType_subst_extended_subst {cf} {Σ} {wfΣ : wf Σ} {Γ Δ T} : @@ -1467,13 +1463,11 @@ Lemma isType_subst_extended_subst {cf} {Σ} {wfΣ : wf Σ} {Γ Δ T} : isType Σ (smash_context [] Γ ,,, subst_context (extended_subst Γ 0) 0 Δ) (subst (extended_subst Γ 0) #|Δ| T). Proof. - move=> [] s. - intros Hs. + intros (_ & s & Hs & _). eapply has_sort_isType. have wfctx := typing_wf_local Hs. generalize Hs. rewrite -[_ ,,, _]app_context_nil_l app_context_assoc. move/typing_expand_lets_gen; rewrite app_context_nil_l => hs. - exists s. rewrite /expand_lets_ctx /expand_lets_k_ctx in hs. rewrite closed_ctx_lift /= in hs. move/closed_wf_local: wfctx. rewrite closedn_ctx_app => /andP[] //. @@ -1515,12 +1509,31 @@ Proof. eapply into_ws_cumul_pb; fvs. Qed. +Lemma lift_judgment_SR {cf : checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} Γ na na' tm tm' ty ty' u : + lift_typing0 (fun t T : term => + Σ ;;; Γ |- t : T × + forall u : term, Σ ;;; Γ ⊢ t ⇝1 u -> Σ;;; Γ |- u : T) + (j_decl_s (mkdecl na tm ty) u) -> + on_one_decl (fun t t' => Σ ;;; Γ ⊢ t ⇝1 t') (mkdecl na tm ty) (mkdecl na' tm' ty') -> + lift_typing typing Σ Γ (j_decl_s (mkdecl na' tm' ty') u). +Proof. + cbn; intros tu Hdecl. + destruct tm as [tm|], tm' as [tm'|] => //; destruct Hdecl as (<- & h). 1: destruct h as [[h <-] | [h <-]]. + all: apply lift_sorting_it_impl_gen with tu => //= [] [] HT IHT; eauto. + eapply type_ws_cumul_pb; tea. + { eapply lift_sorting_forget_univ, lift_sorting_it_impl_gen with tu => // Ht. now apply Ht. } + eapply (red_ws_cumul_pb (pb:=Cumul)). + now eapply closed_red1_red. +Qed. + Lemma sr_red1 {cf:checker_flags} : env_prop SR_red1 + (fun Σ Γ j => + lift_typing0 (fun t T => Σ ;;; Γ |- t : T × forall u, Σ ;;; Γ ⊢ t ⇝1 u -> Σ ;;; Γ |- u : T) j) (fun Σ Γ => wf_local Σ Γ × (forall Γ' Δ' Δ, Γ = Γ' ,,, Δ' -> - OnOne2_local_env (on_one_decl (fun Δ : context => closed_red1 Σ (Γ',,, Δ))) Δ' Δ -> + OnOne2_local_env (fun Δ : context => on_one_decl (closed_red1 Σ (Γ',,, Δ))) Δ' Δ -> wf_local_rel Σ Γ' Δ)). Proof. apply typing_ind_env; intros Σ wfΣ Γ wfΓ; unfold SR_red1; intros **; rename_all_hyps; auto. @@ -1542,55 +1555,35 @@ Proof. end. - (* Contexts *) - split; auto. intros Γ' Δ' Δ ->. + split; auto. intros Γ' Δ' Δ ->. clear X0. induction 1. - * depelim p. subst. depelim X. constructor. - now eapply wf_local_app_inv. - apply infer_typing_sort_impl with id tu; intros _. - now eapply Hs. - * depelim X. - constructor. now eapply wf_local_app_inv. - depelim p. destruct s as [[red <-]|[red <-]]; subst. - apply infer_typing_sort_impl with id tu; intros _. - now eapply Hs. - exact tu. - red. depelim p. destruct s as [[red <-]|[red <-]]; subst. - specialize (Hs _ red). eapply type_ws_cumul_pb; tea. - apply infer_typing_sort_impl with id tu; intros _. - apply Hs. eapply (red_ws_cumul_pb (pb:=Cumul)). - now eapply closed_red1_red. - now eapply Hc. - - * depelim X; specialize (IHX0 _ X); pose proof (wf_local_closed_context all). - + constructor; auto. clear X. - pose proof (wf_local_closed_context all). - eapply wf_local_app_inv in all as []. - eapply wf_local_app in IHX0; tea. - apply infer_typing_sort_impl with id tu; intros Hty. - eapply closed_context_conversion; tea. - eapply red_one_decl_ws_cumul_ctx_pb => //. - eapply OnOne2_local_env_impl; tea. - intros ???. eapply on_one_decl_impl => ???; firstorder. - + constructor; auto. clear X. - { eapply wf_local_app_inv in all as []. - eapply wf_local_app in IHX0; tea. - apply infer_typing_sort_impl with id tu; intros Hty. - eapply closed_context_conversion; tea. - eapply red_one_decl_ws_cumul_ctx_pb => //. - eapply OnOne2_local_env_impl; tea. - intros ???. eapply on_one_decl_impl => ???; firstorder. } - { red. - clear X; eapply wf_local_app_inv in all as []. - eapply wf_local_app in IHX0; tea. - eapply closed_context_conversion; tea. - eapply red_one_decl_ws_cumul_ctx_pb => //. - eapply OnOne2_local_env_impl; tea. - intros ???. eapply on_one_decl_impl => ???; firstorder. } + * depelim X. apply All_local_rel_abs. + now eapply All_local_env_app_inv. + eapply lift_judgment_SR with (2 := p). + eapply lift_sorting_it_impl_gen with (tu := tu) => //. + * depelim X. apply All_local_rel_def. + now eapply All_local_env_app_inv. + eapply lift_judgment_SR with (2 := p); tea. + eapply lift_sorting_it_impl_gen with tu => //=. + + * assert (wf_local Σ (Γ' ,,, Γ'0) × lift_typing typing Σ (Γ' ,,, Γ) (j_decl d) × is_closed_context (Γ' ,,, Γ)) as (wfΓ' & ond & clΓ'). + { depelim X; specialize (IHX0 _ X); clear X; pose proof (wf_local_closed_context all). + all: split; [|now split]. + all: eapply All_local_env_app_inv in all as [wfΓ _]; eapply All_local_env_app in wfΓ; tea. } + clear wfΓ X IHX0. + apply All_local_rel_snoc. + 1: now apply All_local_env_app_inv in wfΓ'. + apply lift_typing_impl with (1 := ond) => // ?? Hty. + eapply closed_context_conversion; tea. + eapply red_one_decl_ws_cumul_ctx_pb => //. + eapply OnOne2_local_env_impl; tea. + intros ???. eapply on_one_decl_impl => ???; firstorder. - (* Rel delta reduction *) rewrite heq_nth_error in H. destruct decl as [na b ty]; noconf H. simpl. - pose proof (PCUICValidity.nth_error_All_local_env heq_nth_error wfΓ); eauto. + eapply All_local_env_nth_error in wfΓ as X0; tea. + apply unlift_TermTyp in X0. cbn in *. rewrite <- (firstn_skipn (S n) Γ). eapply weakening_length; auto. @@ -1598,70 +1591,76 @@ Proof. now unfold app_context; rewrite firstn_skipn. - (* Prod *) - constructor; eauto. intuition auto. - unshelve eapply (closed_context_conversion _ typeb); pcuics. - constructor. now eapply ws_cumul_ctx_pb_refl, wf_local_closed_context. + assert (X1' : lift_typing typing Σ Γ (j_vass_s na N1 s1)). + 1: now unshelve eapply lift_judgment_SR with (1 := X1) => //. + constructor; eauto. + unshelve eapply (closed_context_conversion _ typeb). + 1: constructor; tas. + 1: now eapply lift_sorting_forget_univ. + constructor. + now eapply ws_cumul_ctx_pb_refl, wf_local_closed_context. constructor; auto. eapply red_conv. now eapply closed_red1_red. - (* Lambda *) + assert (X1' : lift_typing typing Σ Γ (j_vass na M')). + 1: now unshelve eapply lift_judgment_SR with (1 := X1) => //. eapply type_Cumul_alt. eapply type_Lambda; eauto. - intuition auto. - unshelve eapply (closed_context_conversion _ typeb); pcuics. + unshelve eapply (closed_context_conversion _ typeb). pcuic. constructor. now eapply ws_cumul_ctx_pb_refl, wf_local_closed_context. constructor; auto. eapply red_conv. now eapply closed_red1_red. - assert (Σ ;;; Γ |- tLambda n t b : tProd n t bty). econstructor; pcuics. - now eapply validity in X0. + assert (Σ ;;; Γ |- tLambda na t b : tProd na t bty). econstructor; pcuics. + now eapply validity in X2. econstructor 3. eapply cumul_refl'. constructor. apply Hu. - (* LetIn body *) eapply type_Cumul_alt. - apply (@substitution_let _ _ _ Γ n b b_ty b' b'_ty typeb'). + apply (@substitution_let _ _ _ Γ na b b_ty b' b'_ty typeb'). specialize (typing_wf_local typeb') as wfd. - assert (Σ ;;; Γ |- tLetIn n b b_ty b' : tLetIn n b b_ty b'_ty). + assert (Σ ;;; Γ |- tLetIn na b b_ty b' : tLetIn na b b_ty b'_ty). econstructor; eauto. - eapply (validity X0). + eapply (validity X2). eapply cumul_red_r. apply cumul_refl'. constructor. - (* LetIn value *) + assert (X1' : lift_typing typing Σ Γ (j_vdef na r b_ty)). + { unshelve eapply lift_judgment_SR with (1 := X1); tea. firstorder. } eapply type_Cumul_alt. econstructor; eauto. - unshelve eapply (closed_context_conversion _ typeb'); pcuics. + unshelve eapply (closed_context_conversion _ typeb'). pcuic. constructor. eapply ws_cumul_ctx_pb_refl; eauto. constructor; auto. now eapply red_conv, closed_red1_red. apply ws_cumul_pb_refl; eauto. now repeat inv_on_free_vars. - assert (Σ ;;; Γ |- tLetIn n b b_ty b' : tLetIn n b b_ty b'_ty). econstructor; eauto. - eapply (validity X0). + assert (Σ ;;; Γ |- tLetIn na b b_ty b' : tLetIn na b b_ty b'_ty). econstructor; eauto. + eapply (validity X2). eapply cumul_red_r. apply cumul_refl'. constructor. apply Hu. - (* LetIn type annotation *) - specialize (forall_u _ Hu). + assert (X1' : lift_typing typing Σ Γ (j_vdef na b r)). + { unshelve eapply lift_judgment_SR with (1 := X1); tea. firstorder. } eapply type_Cumul_alt. econstructor; eauto. - eapply type_Cumul_alt. eauto. exists s1; auto. - apply: red_cumul Hu. - unshelve eapply (closed_context_conversion _ typeb'). - constructor; pcuic. - eapply type_Cumul_alt. eauto. pcuic. apply: red_cumul Hu. + unshelve eapply (closed_context_conversion _ typeb'). pcuic. constructor; auto. eapply ws_cumul_ctx_pb_refl; eauto. constructor; auto. eapply ws_cumul_pb_refl; eauto; repeat inv_on_free_vars =>//. apply: red_conv Hu. - assert (Σ ;;; Γ |- tLetIn n b b_ty b' : tLetIn n b b_ty b'_ty). econstructor; eauto. - eapply (validity X0). + assert (Σ ;;; Γ |- tLetIn na b b_ty b' : tLetIn na b b_ty b'_ty). econstructor; eauto. + eapply (validity X2). eapply cumul_red_r. apply cumul_refl'. constructor. apply Hu. - (* Application *) eapply substitution0; eauto. pose proof typet as typet'. - eapply inversion_Lambda in typet' as [s1 [B' [Ht [Hb HU]]]]=>//. + eapply inversion_Lambda in typet' as (B' & Ht & Hb & HU) =>//. apply cumul_Prod_inv in HU as [eqann eqA leqB] => //. pose proof (validity typet) as i. eapply isType_tProd in i as [Hdom Hcodom]; auto. eapply type_ws_cumul_pb; eauto. - unshelve eapply (closed_context_conversion _ Hb); pcuics. + destruct Hcodom as (_ & scodom & Hcodom & _). + unshelve eapply (closed_context_conversion _ Hb); pcuic. constructor. now apply ws_cumul_ctx_pb_refl. constructor; auto. - (* Fixpoint unfolding *) @@ -1697,10 +1696,10 @@ Proof. unshelve epose proof (H0' := declared_constant_to_gen H0); eauto. unshelve epose proof (declared_constant_inj decl decl0 _ _); tea; subst decl. destruct decl0 as [ty body' univs]; simpl in *; subst body'. - eapply on_declared_constant in H; tas; cbn in H. + eapply on_declared_constant in H as (Hb & s & Ht & _); tas; cbn in Hb, Ht. rewrite <- (app_context_nil_l Γ). - apply subject_closed in H as clb; tas. - apply type_closed in H as clty; tas. + apply subject_closed in Hb as clb; tas. + apply type_closed in Hb as clty; tas. replace (subst_instance u body) with (lift0 #|Γ| (subst_instance u body)). replace (subst_instance u ty) @@ -1714,9 +1713,9 @@ Proof. clear forall_u forall_u0 X X0. unshelve epose proof (isdecl' := declared_inductive_to_gen isdecl); eauto. destruct X4 as [wfcpc IHcpc]. - eassert (ctx_inst _ _ _ _) as Hctxi by (eapply ctx_inst_impl with (2 := X5); now intros ? []). - eassert (PCUICEnvTyping.ctx_inst _ _ _ _ _) as IHctxi. - { eapply ctx_inst_impl with (2 := X5). intros ? ? ? ? [? r]; exact r. } + eassert (ctx_inst _ _ _ _) as Hctxi by (eapply ctx_inst_impl with (1 := X5); now intros ? []). + eassert (PCUICEnvTyping.ctx_inst _ _ _ _) as IHctxi. + { eapply ctx_inst_impl with (1 := X5). intros t T [? r]; pattern Γ, t, T in r; exact r. } hide X8. pose proof typec as typec''. unfold iota_red. @@ -1745,9 +1744,9 @@ Proof. assert(wfparu : wf_local Σ (subst_instance (puinst p) (ind_params mdecl))). { eapply on_minductive_wf_params; eauto. destruct isdecl; eauto. } - assert (wfps : wf_universe Σ ps). + assert (wfps : wf_sort Σ ps). { eapply validity in IHp; auto. eapply PCUICWfUniverses.isType_wf_universes in IHp; tea. - now apply (ssrbool.elimT PCUICWfUniverses.wf_universe_reflect) in IHp. } + move: IHp => /= /PCUICWfUniverses.wf_sort_reflect //. } have lenpars := (wf_predicate_length_pars H0). unfold hidebody in X8. set (ptm := it_mkLambda_or_LetIn _ _). @@ -2076,9 +2075,9 @@ Proof. { rewrite on_free_vars_mkApps /= //. } eapply eq_term_upto_univ_mkApps. 2:reflexivity. - constructor. eapply R_global_instance_sym; tc. + constructor. eapply cmp_global_instance_sym; tc. rewrite eqargs. - now eapply (R_global_instance_cstr_irrelevant declc). + now eapply (cmp_global_instance_cstr_irrelevant declc). - (* Case congruence: on a cofix, impossible *) eapply inversion_mkApps in typec as [? [tcof ?]] => //. @@ -2095,9 +2094,9 @@ Proof. discriminate. - (* Case congruence on a parameter *) - eassert (ctx_inst _ _ _ _) as Hctxi by (eapply ctx_inst_impl with (2 := X5); now intros ? []). - eassert (PCUICEnvTyping.ctx_inst _ _ _ _ _) as X6. - { eapply ctx_inst_impl with (2 := X5). intros ? ? ? ? [? r]; exact r. } + eassert (ctx_inst _ _ _ _) as Hctxi by (eapply ctx_inst_impl with (1 := X5); now intros ? []). + eassert (PCUICEnvTyping.ctx_inst _ _ _ _) as X6. + { eapply ctx_inst_impl with (1 := X5). intros t T [? r]; pattern Γ, t, T in r; exact r. } clear X5; rename Hctxi into X5. destruct X0, X4. assert (isType Σ Γ (mkApps (it_mkLambda_or_LetIn (case_predicate_context ci mdecl idecl p) (preturn p)) (indices ++ [c]))). @@ -2177,7 +2176,7 @@ Proof. eapply closedn_smash_context. now eapply declared_inductive_closed_params_inst. } have isty' : isType Σ Γ (mkApps (tInd ci (puinst p)) (params' ++ indices)). - { eexists; eapply isType_mkApps_Ind; tea. } + { eapply has_sort_isType; eapply isType_mkApps_Ind; tea. } have wfcpc' : wf_local Σ (Γ ,,, case_predicate_context ci mdecl idecl (set_pparams p params')). { eapply wf_case_predicate_context; tea. } have eqindices : ws_cumul_pb_terms Σ Γ indices indices. @@ -2187,14 +2186,12 @@ Proof. eapply ws_cumul_pb_eq_le. eapply ws_cumul_pb_mkApps; pcuic. eapply All2_app => //. now apply: red_terms_ws_cumul_pb_terms. } set (pctx := (inst_case_predicate_context (set_pparams p params'))) in *. - pose proof (snd (All2_fold_All2 _) X1). symmetry in X9. move:X9. + pose proof X1 as X9. symmetry in X9. move:X9. change (pcontext p) with (pcontext (set_pparams p params')). move/(PCUICAlpha.inst_case_predicate_context_eq wfp') => eqctx. have wfpctx : wf_local Σ (Γ,,, inst_case_predicate_context (set_pparams p params')). { eapply wf_local_alpha; tea; auto. - eapply All2_app => //. - now eapply All2_fold_All2 in eqctx. - eapply All2_refl; reflexivity. } + eapply All2_app => //. reflexivity. } have eqpctx : Σ ⊢ Γ ,,, pctx = Γ ,,, case_predicate_context ci mdecl idecl (set_pparams p params'). { symmetry. rewrite /pctx. @@ -2203,7 +2200,7 @@ Proof. now eapply wf_local_closed_context. eapply upto_names_conv_context. eapply eq_context_upto_cat. apply eq_context_upto_refl; tc. - now apply eq_context_gen_upto. } + now apply eq_context_upto_names_eq_context_alpha. } epose proof (wf_case_branches_types' (p:=set_pparams p params') ps _ brs isdecl isty' wfp'). cbv zeta in X9; forward_keep X9. { eapply closed_context_conversion; tea. } @@ -2251,7 +2248,7 @@ Proof. apply weaken_wf_local => //. eapply on_minductive_wf_params; tea. eapply isdecl. } eapply (type_ws_cumul_pb (pb:=Conv)); tea. + eapply closed_context_conversion; tea. - + exists ps. exact wfcbcty'. + + eapply has_sort_isType with (s := ps). exact wfcbcty'. + eapply ws_cumul_pb_mkApps; tea. { rewrite /ptm. cbn [preturn set_pparams]. rewrite !lift_it_mkLambda_or_LetIn. @@ -2263,7 +2260,7 @@ Proof. rewrite !case_predicate_context_length // /= Nat.add_0_r. relativize #|pcontext p|. relativize #|cstr_args cdecl|. eapply weakening_ws_cumul_pb; tea. - apply isType_ws_cumul_pb_refl. now exists ps. + apply isType_ws_cumul_pb_refl. now eapply has_sort_isType. now apply wf_local_closed_context. now apply case_branch_context_length_args. now rewrite case_predicate_context_length. } @@ -2329,9 +2326,9 @@ Proof. - (* Case congruence on the return clause context *) clear IHHu. destruct X0, X4 as []. - eassert (ctx_inst _ _ _ _) as Hctxi by (eapply ctx_inst_impl with (2 := X5); now intros ? []). - eassert (PCUICEnvTyping.ctx_inst _ _ _ _ _) as X6. - { eapply ctx_inst_impl with (2 := X5). intros ? ? ? ? [? r]; exact r. } + eassert (ctx_inst _ _ _ _) as Hctxi by (eapply ctx_inst_impl with (1 := X5); now intros ? []). + eassert (PCUICEnvTyping.ctx_inst _ _ _ _) as X6. + { eapply ctx_inst_impl with (1 := X5). intros t T [? r]; pattern Γ, t, T in r; exact r. } clear X5; rename Hctxi into X5. set (ptm := it_mkLambda_or_LetIn _ _). assert (isType Σ Γ (mkApps ptm (indices ++ [c]))). @@ -2362,7 +2359,7 @@ Proof. destruct a1 as [wfparctx wfparctx' wfcbc wfcbty]. split => //. split => //. eapply type_ws_cumul_pb; tea. - now exists ps. + now eapply has_sort_isType. eapply (ws_cumul_pb_ws_cumul_ctx (pb':=Cumul)); tea. 2:eapply closed_red1_ws_cumul_pb, btyred. now apply ws_cumul_ctx_pb_refl, wf_local_closed_context. @@ -2378,9 +2375,9 @@ Proof. - (* Case congruence on discriminee *) destruct X0, X4. - eassert (ctx_inst _ _ _ _) as Hctxi by (eapply ctx_inst_impl with (2 := X5); now intros ? []). - eassert (PCUICEnvTyping.ctx_inst _ _ _ _ _) as X6. - { eapply ctx_inst_impl with (2 := X5). intros ? ? ? ? [? r]; exact r. } + eassert (ctx_inst _ _ _ _) as Hctxi by (eapply ctx_inst_impl with (1 := X5); now intros ? []). + eassert (PCUICEnvTyping.ctx_inst _ _ _ _) as X6. + { eapply ctx_inst_impl with (1 := X5). intros t T [? r]; pattern Γ, t, T in r; exact r. } clear X5; rename Hctxi into X5. set (ptm := it_mkLambda_or_LetIn _ _). assert (isType Σ Γ (mkApps ptm (indices ++ [c]))). @@ -2408,9 +2405,9 @@ Proof. - (* Case congruence on branches *) destruct X0, X4. - eassert (ctx_inst _ _ _ _) as Hctxi by (eapply ctx_inst_impl with (2 := X5); now intros ? []). - eassert (PCUICEnvTyping.ctx_inst _ _ _ _ _) as X6. - { eapply ctx_inst_impl with (2 := X5). intros ? ? ? ? [? r]; exact r. } + eassert (ctx_inst _ _ _ _) as Hctxi by (eapply ctx_inst_impl with (1 := X5); now intros ? []). + eassert (PCUICEnvTyping.ctx_inst _ _ _ _) as X6. + { eapply ctx_inst_impl with (1 := X5). intros t T [? r]; pattern Γ, t, T in r; exact r. } clear X5; rename Hctxi into X5. eapply type_Case; eauto. econstructor; eauto. econstructor; eauto. @@ -2490,7 +2487,7 @@ Proof. (subst_context (inds (inductive_mind p.(proj_ind)) u (ind_bodies mdecl)) #|ind_params mdecl| (subst_instance u (cstr_args cdecl)))))))); auto. ** eapply wf_local_closed_context. - eapply wf_local_app_skipn. + eapply All_local_env_app_skipn. apply wf_subslet_ctx in projsubs. apply projsubs. ** elim: p.(proj_arg). simpl. constructor. @@ -2603,7 +2600,7 @@ Proof. { rewrite firstn_skipn. unshelve epose proof (isdecl' := declared_projection_to_gen isdecl); eauto. eapply (isType_subst_instance_decl _ _ _ _ _ u wf isdecl'.p1.p1.p1) in projty; eauto. - apply infer_typing_sort_impl with id projty; intros Hs. + apply lift_sorting_f_it_impl with (tu := projty) => // Hs. rewrite /= /map_decl /= in Hs. eapply (weaken_ctx Γ) in Hs; auto. rewrite (subst_app_simpl [_]). @@ -2700,7 +2697,7 @@ Proof. rewrite H5 in cum. set (idx := S (context_assumptions (cstr_args cdecl) - S parg)) in *. assert (wfpargctxu1 : wf_local Σ (Γ ,,, skipn idx (smash_context [] pargctxu1))). - { simpl. apply wf_local_app_skipn. apply wf_local_smash_end; auto. + { simpl. apply All_local_env_app_skipn. apply wf_local_smash_end; auto. apply idxsubst0. } destruct cum as [[cr mapd] cumdecls]. destruct decl'' as [na [b|] ty]; simpl in mapd; try discriminate. @@ -2790,7 +2787,7 @@ Proof. simpl. eapply ws_cumul_pb_refl. rewrite -skipn_lift_context. eapply wf_local_closed_context. - eapply wf_local_app_skipn. + eapply All_local_env_app_skipn. eapply weakening_wf_local => //. eapply wf_local_smash_end, idxsubst0. move: Hty. @@ -2847,16 +2844,16 @@ Proof. simpl. destruct narg; discriminate. - (* Fix congruence: type reduction *) - assert(fixl :#|fix_context mfix| = #|fix_context mfix1|) by now (rewrite !fix_context_length; apply (OnOne2_length X2)). + assert(fixl :#|fix_context mfix| = #|fix_context mfix1|) by now (rewrite !fix_context_length; apply (OnOne2_length X4)). assert(convctx : conv_context cumulAlgo_gen Σ (Γ ,,, fix_context mfix) (Γ ,,, fix_context mfix1)). - { clear -wf wfΓ X X2 H2 fixl. + { clear -wf wfΓ X X4 H2 fixl. eapply All2_fold_app => //. apply conv_ctx_refl. clear X. eapply All2_fold_impl. eapply (conv_decls_fix_context (Γ := Γ)) => //. move: H2. assert (clΓ := wf_local_closed_context wfΓ). - clear fixl. move: #|mfix|. induction X2; constructor. + clear fixl. move: #|mfix|. induction X4; constructor. * destruct p as [[r p] e]. noconf e. split. now apply closed_red1_ws_cumul_pb. rewrite H. reflexivity. * cbn in H2. move/andP: H2 => [_ Htl]. @@ -2866,45 +2863,44 @@ Proof. eapply ws_cumul_pb_refl => //. * move/andP: H2 => [/andP [Hty _] _]. split; try reflexivity. apply ws_cumul_pb_refl => //. - * eapply IHX2. simpl in H2; now move/andP: H2 => []. + * eapply IHX4. simpl in H2; now move/andP: H2 => []. * cbn. intros ???? []; constructor; eauto; now apply ws_cumul_pb_forget_conv. } - assert (All (on_def_type (lift_typing typing Σ) Γ) mfix). - { apply (All_impl X0); intros d HT. - apply infer_typing_sort_impl with id HT; now intros [Hty _]. } assert (All (on_def_type (lift_typing typing Σ) Γ) mfix1). - { apply (OnOne2_All_All X2 X0). + { apply (OnOne2_All_All X4 X1). * intros d HT. - apply infer_typing_sort_impl with id HT; now intros [Hty _]. + apply lift_typing_impl with (1 := HT); now intros ?? [Hty _]. * intros d d' [[red _] eq] HT'. - apply infer_typing_sort_impl with id HT'; now intros [_ IH]. } + apply lift_sorting_it_impl_gen with HT' => //; now intros [_ IH]. } assert (wf_local Σ (Γ ,,, fix_context mfix1)). { apply All_mfix_wf; auto. } - destruct (OnOne2_nth_error _ _ _ decl _ X2 heq_nth_error) as [decl' [eqnth disj]]. + destruct (OnOne2_nth_error _ _ _ decl _ X4 heq_nth_error) as [decl' [eqnth disj]]. eapply type_Cumul_alt. econstructor; eauto. * eapply (fix_guard_red1 _ _ _ _ 0); eauto. constructor; eauto. eapply OnOne2_impl; tea; cbn; intuition auto. apply a2. apply a2. - * eapply (OnOne2_All_mix_left X0) in X2. - apply (OnOne2_All_All X2 X1). - + intros x [Hb IH]. + * eapply (OnOne2_All_mix_left X1) in X4. + apply (OnOne2_All_All X4 X3). + + intros d Hj. + unfold on_def_body. rewrite -fixl. + eapply lift_typing_impl with (1 := Hj) => ?? [] HT IHT. eapply context_conversion; eauto. - now rewrite -fixl. - + move=> [na ty b rarg] [na' ty' b' rarg'] [[[red Hred] eq] HT] [Hb IH']. - rewrite /= in red, Hred, eq, Hb, IH'. + + move=> [na ty b rarg] [na' ty' b' rarg'] [[[red Hred] eq] HT] HB. + rewrite /= in red, Hred, eq, HB, HT. noconf eq. - eapply context_conversion; eauto. - rewrite -fixl. + destruct HB as ((Hb & IHb) & s & (Ht & IHt) & e); cbn in *. + eassert (Σ ;;; Γ ,,, _ |- lift0 _ ty' : tSort s). + 2: repeat (eexists; cbn). + 2,3: eapply context_conversion; eauto. + all: rewrite -fixl. + { eapply IHt. eapply @weakening_closed_red1 with (Γ' := []); eauto. now eapply wf_local_closed_context. } eapply type_Cumul_alt. eapply Hb. - red. simple apply infer_typing_sort_impl with id HT; move=> /= [Hs IH]. - specialize (IH _ red). - eapply (weakening _ _ _ _ (tSort _)); auto. - apply All_mfix_wf; auto. + eapply has_sort_isType. rewrite fixl. eassumption. apply red_cumul, red1_red. eapply (weakening_red1 _ []); auto. 2:eapply red. - pose proof (Hs := HT.π2.1). + pose proof (Hs := HT.2.π2.1.1). eapply subject_closed in Hs. eapply (closedn_on_free_vars (P:=xpredT)) in Hs. now eapply on_free_vars_any_xpredT. @@ -2912,65 +2908,62 @@ Proof. * eapply wf_fixpoint_red1_type; eauto. eapply OnOne2_impl; tea; cbn; intuition auto. apply a2. apply a2. - * eapply All_nth_error in X3; eauto. + * eapply All_nth_error in X0; eauto. * apply conv_cumul, conv_sym. destruct disj as [<-|[[red Hred] eq]] => //. reflexivity. eapply PCUICCumulativity.red_conv. apply red1_red, red. - (* Fix congruence in body *) - assert(fixl :#|fix_context mfix| = #|fix_context mfix1|) by now (rewrite !fix_context_length; apply (OnOne2_length X2)). + assert(fixl :#|fix_context mfix| = #|fix_context mfix1|) by now (rewrite !fix_context_length; apply (OnOne2_length X4)). assert(convctx : fix_context mfix = fix_context mfix1). - { clear -wf X2. + { clear -wf X4. change fix_context with (fix_context_gen 0). - generalize 0. induction X2. + generalize 0. induction X4. destruct p as [_ eq]. noconf eq. simpl in H; noconf H. simpl. intros. now rewrite H H0. - simpl. intros n'; f_equal. apply IHX2. } - assert (All (on_def_type (lift_typing typing Σ) Γ) mfix). - { apply (All_impl X0); intros d HT. - apply infer_typing_sort_impl with id HT; now intros [Hty _]. } + simpl. intros n'; f_equal. apply IHX4. } assert (All (on_def_type (lift_typing typing Σ) Γ) mfix1). - { apply (OnOne2_All_All X2 X0). + { apply (OnOne2_All_All X4 X1). * intros d HT. - apply infer_typing_sort_impl with id HT; now intros [Hty _]. + apply lift_typing_impl with (1 := HT); now intros ?? [Hty _]. * intros d d' [[red _] eq] HT'. - noconf eq. - apply infer_typing_sort_impl with id HT'; intros [Hs IH]. now rewrite -H4. } + noconf eq. unfold on_def_type in HT' |- *. rewrite -H4. + apply lift_typing_impl with (1 := HT'); now intros ?? []. } assert (wf_local Σ (Γ ,,, fix_context mfix1)). { apply All_mfix_wf; auto. } - destruct (OnOne2_nth_error _ _ _ decl _ X2 heq_nth_error) as [decl' [eqnth disj]]. + destruct (OnOne2_nth_error _ _ _ decl _ X4 heq_nth_error) as [decl' [eqnth disj]]. eapply type_Cumul_alt. econstructor; eauto. * eapply (fix_guard_red1 _ _ _ _ 0); eauto. apply fix_red_body; eauto. eapply OnOne2_impl; tea; cbn; intuition auto. apply a2. apply a2. - * eapply (OnOne2_All_mix_left X0) in X2. - apply (OnOne2_All_All X2 X1). - + intros x [Hb IH]. - eapply context_conversion; eauto. - now rewrite -fixl. - rewrite convctx. apply conv_ctx_refl. - + move=> [na ty b rarg] [na' ty' b' rarg'] /= [[red eq] HT] [Hb IH']. + * eapply (OnOne2_All_mix_left X1) in X4. + apply (OnOne2_All_All X4 X3). + + intros d Hj. + rewrite -convctx. + eapply lift_typing_impl with (1 := Hj) => ?? [] HT //. + + move=> [na ty b rarg] [na' ty' b' rarg'] /= [[[red IH] eq] HT] HB. noconf eq. - now rewrite -convctx. + rewrite -convctx. + apply lift_sorting_it_impl_gen with HB => //= [] [Ht IHt] //. eauto. * eapply wf_fixpoint_red1_body; eauto. eapply OnOne2_impl; tea; cbn; intuition auto. apply a2. apply a2. - * eapply All_nth_error in X3; eauto. + * eapply All_nth_error in X0; eauto. * apply conv_cumul, conv_sym. destruct disj as [<-|[_ eq]]. reflexivity. noconf eq. rewrite H4; reflexivity. - (* CoFix congruence: type reduction *) - assert(fixl :#|fix_context mfix| = #|fix_context mfix1|) by now (rewrite !fix_context_length; apply (OnOne2_length X2)). + assert(fixl :#|fix_context mfix| = #|fix_context mfix1|) by now (rewrite !fix_context_length; apply (OnOne2_length X4)). assert(convctx : conv_context cumulAlgo_gen Σ (Γ ,,, fix_context mfix) (Γ ,,, fix_context mfix1)). - { clear -wf wfΓ X X2 H2 fixl. + { clear -wf wfΓ X X4 H2 fixl. eapply All2_fold_app => //; trea. clear X. eapply All2_fold_impl. eapply (conv_decls_fix_context (Γ := Γ)) => //. move: H2. assert (clΓ := wf_local_closed_context wfΓ). - clear fixl. move: #|mfix|. induction X2; constructor. + clear fixl. move: #|mfix|. induction X4; constructor. * destruct p as [[r p] e]. noconf e. split. now apply closed_red1_ws_cumul_pb. rewrite H. reflexivity. * cbn in H2. move/andP: H2 => [_ Htl]. @@ -2980,45 +2973,44 @@ Proof. eapply ws_cumul_pb_refl => //. * move/andP: H2 => [/andP [Hty _] _]. split; try reflexivity. apply ws_cumul_pb_refl => //. - * eapply IHX2. simpl in H2; now move/andP: H2 => []. + * eapply IHX4. simpl in H2; now move/andP: H2 => []. * cbn. intros ???? []; constructor; eauto; now apply ws_cumul_pb_forget_conv. } - assert (All (on_def_type (lift_typing typing Σ) Γ) mfix). - { apply (All_impl X0); intros d HT. - apply infer_typing_sort_impl with id HT; now intros [Hty _]. } assert (All (on_def_type (lift_typing typing Σ) Γ) mfix1). - { apply (OnOne2_All_All X2 X0). + { apply (OnOne2_All_All X4 X1). * intros d HT. - apply infer_typing_sort_impl with id HT; now intros [Hty _]. + apply lift_typing_impl with (1 := HT); now intros ?? [Hty _]. * intros d d' [[red _] eq] HT'. - apply infer_typing_sort_impl with id HT'; now intros [_ IH]. } + apply lift_sorting_it_impl_gen with HT' => //. now intros [_ IH]. } assert (wf_local Σ (Γ ,,, fix_context mfix1)). { apply All_mfix_wf; auto. } - destruct (OnOne2_nth_error _ _ _ decl _ X2 heq_nth_error) as [decl' [eqnth disj]]. + destruct (OnOne2_nth_error _ _ _ decl _ X4 heq_nth_error) as [decl' [eqnth disj]]. eapply type_Cumul_alt. econstructor; eauto. * eapply (cofix_guard_red1 _ _ _ _ 0); eauto. constructor; eauto. eapply OnOne2_impl; tea; cbn; intuition auto. apply a2. apply a2. - * eapply (OnOne2_All_mix_left X0) in X2. - apply (OnOne2_All_All X2 X1). - + intros x [Hb IH]. + * eapply (OnOne2_All_mix_left X1) in X4. + apply (OnOne2_All_All X4 X3). + + intros d Hj. + unfold on_def_body. rewrite -fixl. + eapply lift_typing_impl with (1 := Hj) => ?? [] HT IHT. eapply context_conversion; eauto. - now rewrite -fixl. - + move=> [na ty b rarg] [na' ty' b' rarg'] [[[red Hred] eq] HT] [Hb IH']. - rewrite /= in red, Hred, eq, Hb, IH'. + + move=> [na ty b rarg] [na' ty' b' rarg'] [[[red Hred] eq] HT] HB. + rewrite /= in red, Hred, eq, HB, HT. noconf eq. - eapply context_conversion; eauto. - rewrite -fixl. + destruct HB as ((Hb & IHb) & s & (Ht & IHt) & e); cbn in *. + eassert (Σ ;;; Γ ,,, _ |- lift0 _ ty' : tSort s). + 2: repeat (eexists; cbn). + 2,3: eapply context_conversion; eauto. + all: rewrite -fixl. + { eapply IHt. eapply @weakening_closed_red1 with (Γ' := []); eauto. now eapply wf_local_closed_context. } eapply type_Cumul_alt. eapply Hb. - red. simple apply infer_typing_sort_impl with id HT; move=> /= [Hs IH]. - specialize (IH _ red). - eapply (weakening _ _ _ _ (tSort _)); auto. - apply All_mfix_wf; auto. + eapply has_sort_isType. rewrite fixl. eassumption. apply red_cumul, red1_red. eapply (weakening_red1 _ []); auto. 2:eapply red. - pose proof (Hs := HT.π2.1). + pose proof (Hs := HT.2.π2.1.1). eapply subject_closed in Hs. eapply (closedn_on_free_vars (P:=xpredT)) in Hs. now eapply on_free_vars_any_xpredT. @@ -3026,52 +3018,49 @@ Proof. * eapply (wf_cofixpoint_red1_type _ Γ); eauto. eapply OnOne2_impl; tea; cbn; intuition auto; now eapply closed_red1_red. - * eapply All_nth_error in X3; eauto. + * eapply All_nth_error in X0; eauto. * apply conv_cumul, conv_sym. destruct disj as [<-|[[red Hred] eq]] => //. reflexivity. eapply PCUICCumulativity.red_conv. apply red1_red, red. - (* CoFix congruence in body *) - assert(fixl :#|fix_context mfix| = #|fix_context mfix1|) by now (rewrite !fix_context_length; apply (OnOne2_length X2)). + assert(fixl :#|fix_context mfix| = #|fix_context mfix1|) by now (rewrite !fix_context_length; apply (OnOne2_length X4)). assert(convctx : fix_context mfix = fix_context mfix1). - { clear -wf X2. + { clear -wf X4. change fix_context with (fix_context_gen 0). - generalize 0. induction X2. + generalize 0. induction X4. destruct p as [_ eq]. noconf eq. simpl in H; noconf H. simpl. intros. now rewrite H H0. - simpl. intros n'; f_equal. apply IHX2. } - assert (All (on_def_type (lift_typing typing Σ) Γ) mfix). - { apply (All_impl X0); intros d HT. - apply infer_typing_sort_impl with id HT; now intros [Hty _]. } + simpl. intros n'; f_equal. apply IHX4. } assert (All (on_def_type (lift_typing typing Σ) Γ) mfix1). - { apply (OnOne2_All_All X2 X0). + { apply (OnOne2_All_All X4 X1). * intros d HT. - apply infer_typing_sort_impl with id HT; now intros [Hty _]. + apply lift_typing_impl with (1 := HT); now intros ?? [Hty _]. * intros d d' [[red _] eq] HT'. - noconf eq. - apply infer_typing_sort_impl with id HT'; intros [Hs IH]. now rewrite -H4. } + noconf eq. unfold on_def_type in HT' |- *. rewrite -H4. + apply lift_typing_impl with (1 := HT'); now intros ?? []. } assert (wf_local Σ (Γ ,,, fix_context mfix1)). { apply All_mfix_wf; auto. } - destruct (OnOne2_nth_error _ _ _ decl _ X2 heq_nth_error) as [decl' [eqnth disj]]. + destruct (OnOne2_nth_error _ _ _ decl _ X4 heq_nth_error) as [decl' [eqnth disj]]. eapply type_Cumul_alt. econstructor; eauto. * eapply (cofix_guard_red1 _ _ _ _ 0); eauto. apply cofix_red_body; eauto. eapply OnOne2_impl; tea; cbn; intuition auto. apply a2. apply a2. - * eapply (OnOne2_All_mix_left X0) in X2. - apply (OnOne2_All_All X2 X1). - + intros x [Hb IH]. - eapply context_conversion; eauto. - now rewrite -fixl. - rewrite convctx. apply conv_ctx_refl. - + move=> [na ty b rarg] [na' ty' b' rarg'] /= [[red eq] HT] [Hb IH']. + * eapply (OnOne2_All_mix_left X1) in X4. + apply (OnOne2_All_All X4 X3). + + intros d Hj. + rewrite -convctx. + eapply lift_typing_impl with (1 := Hj) => ?? [] HT //. + + move=> [na ty b rarg] [na' ty' b' rarg'] /= [[[red IH] eq] HT] HB. noconf eq. - now rewrite -convctx. + rewrite -convctx. + apply lift_sorting_it_impl_gen with HB => //= [] [Ht IHt] //. eauto. * eapply wf_cofixpoint_red1_body; eauto. eapply OnOne2_impl; tea; cbn; intuition auto. apply a2. apply a2. - * eapply All_nth_error in X3; eauto. + * eapply All_nth_error in X0; eauto. * apply conv_cumul, conv_sym. destruct disj as [<-|[_ eq]]. reflexivity. noconf eq. rewrite H4; reflexivity. @@ -3093,10 +3082,10 @@ Proof. eapply (type_ws_cumul_pb (pb:=Conv)); tea. eapply (type_Prim _ _ (primArray; primArrayModel (set_array_type arr ty))); tea. constructor; cbn; eauto. - * eapply type_ws_cumul_pb; tea. eexists; eauto. + * eapply type_ws_cumul_pb; tea. eapply has_sort_isType; eauto. now eapply (red_ws_cumul_pb (pb:=Conv)), closed_red1_red. * solve_all. - eapply type_ws_cumul_pb; tea. eexists; eauto. + eapply type_ws_cumul_pb; tea. eapply has_sort_isType; eauto. now eapply (red_ws_cumul_pb (pb:=Conv)), closed_red1_red. * simp prim_type. eapply ws_cumul_pb_App. eapply ws_cumul_pb_refl; fvs. cbn. @@ -3167,7 +3156,7 @@ Lemma type_reduction {cf} {Σ} {wfΣ : wf Σ} {Γ t A B} : Proof. intros Ht Hr. eapply type_Cumul_alt. eassumption. - apply infer_typing_sort_impl with id (validity Ht); intros HA. + apply lift_sorting_it_impl_gen with (validity Ht) => // HA. eapply subject_reduction; eassumption. eapply conv_cumul. now apply PCUICCumulativity.red_conv. Defined. @@ -3207,7 +3196,7 @@ Section SRContext. Definition closed_red1_ctx Σ := OnOne2_local_env - (on_one_decl (fun (Δ : context) (t t' : term) => closed_red1 Σ Δ t t')). + (fun (Δ : context) => on_one_decl (fun (t t' : term) => closed_red1 Σ Δ t t')). Notation "Σ ⊢ Γ ⇝1 Δ" := (closed_red1_ctx Σ Γ Δ) (at level 50, Γ, Δ at next level, format "Σ ⊢ Γ ⇝1 Δ") : pcuic. @@ -3217,13 +3206,14 @@ Section SRContext. nth_error Γ n = Some decl -> ∑ s, Σ ;;; Γ |- lift0 (S n) (decl_type decl) : tSort s. Proof using Type. - induction n in Γ, decl |- *; intros hΓ e; destruct Γ; - cbn; inversion e; inversion hΓ; subst. - 1,2: rename X0 into H0. - 3,4: eapply IHn in H0; tas. - 3,4: rewrite simpl_lift0. - all: eapply infer_typing_sort_impl with id H0; intros Hs. - all: eapply (weakening _ _ [_] _ (tSort _)); tas. + intros wfΓ Hnth. + eapply All_local_env_nth_error in wfΓ as H; tea. + apply nth_error_Some_length in Hnth. + destruct H as (_ & s & Hs & _). + eapply weakening with (Γ' := firstn (S n) Γ) in Hs; cbn in Hs; eauto. + 2: rewrite /app_context firstn_skipn //. + rewrite firstn_length_le // /app_context firstn_skipn in Hs. + now eexists. Qed. Ltac invs H := inversion H; subst. @@ -3234,8 +3224,8 @@ Section SRContext. Σ ;;; Γ |- t : T -> Σ ;;; Γ' |- t : T. Proof using Type. assert(OnOne2_local_env - (on_one_decl - (fun (Δ : context) (t t' : term) => red1 Σ.1 Δ t t')) Γ Γ' -> + (fun (Δ : context) => on_one_decl + (fun (t t' : term) => red1 Σ.1 Δ t t')) Γ Γ' -> conv_context cumulAlgo_gen Σ Γ Γ'). { clear. induction 1. - destruct p as [<- r]. constructor; auto. @@ -3256,28 +3246,23 @@ Section SRContext. induction H in Γ', r, X |- *; depelim r. - constructor; auto. cbn in o. destruct o as [<- r]. - apply infer_typing_sort_impl with id t1; intros Hs. + apply lift_sorting_it_impl_gen with t1 => // Hs. eapply subject_reduction1 in Hs; eauto. - depelim X. constructor; auto. - apply infer_typing_sort_impl with id t1; intros Hs. + apply lift_sorting_it_impl_gen with t1 => // Hs. eapply context_conversion; eauto. - depelim X. - red in o. - simpl in t2. destruct o as [<- [[r ->]|[r <-]]]. - constructor; auto. - apply infer_typing_sort_impl with id t1; intros Hs. - eapply subject_reduction1; eauto. - eapply type_reduction; tea. pcuic. - constructor; auto. - eapply subject_reduction1; tea. + all: constructor; auto. + all: apply lift_sorting_it_impl_gen with t1 => // Hs. + + eapply type_reduction; tea. pcuic. + + eapply subject_reduction1; eauto. + + eapply subject_reduction1; eauto. - depelim X. - simpl in t2. constructor; auto. - apply infer_typing_sort_impl with id t1; intros Hs. - eapply context_conversion; eauto. - red; eapply context_conversion; eauto. + apply lift_sorting_it_impl_gen with t1 => // Hs. + all: eapply context_conversion; eauto. - eapply context_conversion; eauto. Qed. @@ -3287,18 +3272,16 @@ Section SRContext. induction 1; cbn in *. - destruct p as [-> r]. intro e. inversion e; subst; cbn in *. constructor; tas. - apply infer_typing_sort_impl with id X0; intros Hs. + apply lift_sorting_it_impl_gen with X0 => // Hs. eapply subject_reduction1; tea. - intro e. inversion e; subst; cbn in *. destruct p as [-> [[? []]|[? []]]]; constructor; cbn; tas. - + apply infer_typing_sort_impl with id X0; intros Hs; eapply subject_reduction1; tea. - + eapply type_Cumul_alt; tea. - apply infer_typing_sort_impl with id X0; intros Hs. - eapply subject_reduction1; tea. - econstructor 2. eassumption. eapply cumul_refl'. + all: apply lift_sorting_it_impl_gen with X0 => // Hs. + + eapply type_reduction; tea. pcuic. + + eapply subject_reduction1; tea. + eapply subject_reduction1; tea. - intro H; inversion H; subst; constructor; cbn in *; auto. - 1,2: apply infer_typing_sort_impl with id X1; intros Hs. + all: apply lift_sorting_it_impl_gen with X1 => // Hs. all: eapply subject_reduction_ctx; tea. Qed. @@ -3353,23 +3336,21 @@ Section SRContext. assert (subslet Σ Γ [b] [vdef na b t]). { pose proof (cons_let_def Σ Γ [] [] na b t) as XX. rewrite !subst_empty in XX. apply XX. constructor. - apply wf_local_app_l in X. inversion X; subst; cbn in *; assumption. + apply wf_local_app_l in X. inversion X; subst; cbn in *; now eapply unlift_TermTyp. } constructor; cbn; auto. - 1: apply infer_typing_sort_impl with id X0; intros Hs. - 1: change (tSort _) with (subst [b] #|Γ'| (tSort X0.π1)). - all: eapply PCUICSubstitution.substitution; tea. + apply lift_typing_f_impl with (1 := X0) => // ?? Hs. + eapply PCUICSubstitution.substitution; tea. - rewrite subst_context_snoc0. simpl. inversion HH; subst; cbn in *. change (Γ,, vdef na b t ,,, Γ') with (Γ ,,, [vdef na b t] ,,, Γ') in *. assert (subslet Σ Γ [b] [vdef na b t]). { pose proof (cons_let_def Σ Γ [] [] na b t) as XX. rewrite !subst_empty in XX. apply XX. constructor. - apply wf_local_app_l in X. inversion X; subst; cbn in *; assumption. } + apply wf_local_app_l in X. inversion X; subst; cbn in *; now eapply unlift_TermTyp. } constructor; cbn; auto. - apply infer_typing_sort_impl with id X0; intros Hs. - change (tSort _) with (subst [b] #|Γ'| (tSort X0.π1)). - all: eapply PCUICSubstitution.substitution; tea. + apply lift_typing_f_impl with (1 := X0) => // ?? Hs. + eapply PCUICSubstitution.substitution; tea. Qed. @@ -3386,7 +3367,7 @@ Section SRContext. isType Σ Γ B. Proof using Type. intros HT red. - apply infer_typing_sort_impl with id HT; intros Hs. + apply lift_sorting_it_impl_gen with HT => // Hs. eapply subject_reduction1; eauto. Qed. @@ -3445,7 +3426,7 @@ Section SRContext. Lemma isType_red {Σ} {wfΣ : wf Σ} {Γ T U} : isType Σ Γ T -> red Σ Γ T U -> isType Σ Γ U. Proof using Type. - intros HT red; apply infer_typing_sort_impl with id HT; intros Hs. + intros HT red; apply lift_sorting_it_impl_gen with HT => // Hs. eapply subject_reduction; eauto. Qed. @@ -3453,28 +3434,26 @@ End SRContext. Lemma isType_tLetIn {cf} {Σ} {HΣ' : wf Σ} {Γ} {na t A B} : isType Σ Γ (tLetIn na t A B) - <~> (isType Σ Γ A × (Σ ;;; Γ |- t : A) × isType Σ (Γ,, vdef na t A) B). + <~> (lift_typing typing Σ Γ (j_vdef na t A) × isType Σ (Γ,, vdef na t A) B). Proof. split; intro HH. - - destruct HH as [s H]. - apply inversion_LetIn in H; tas. destruct H as [s1 [A' [HA [Ht [HB H]]]]]. - repeat split; tas. 1: eexists; eassumption. - apply ws_cumul_pb_Sort_r_inv in H. - destruct H as [s' [H H']]. - exists s'. eapply type_reduction; tea. - apply invert_red_letin in H; tas. - destruct H as [[? [? [? [H ? ? ?]]]]|H]. + - destruct HH as (_ & s & H & _). + apply inversion_LetIn in H as (A' & HtA & HB & H); tas. + split; tas. + apply ws_cumul_pb_Sort_r_inv in H as [s' [H H']]. + eapply has_sort_isType. + eapply type_reduction; tea. + apply invert_red_letin in H as [[? [? [? [H ? ? ?]]]]|H]; tas. * discriminate. * etransitivity. + exact (@red_rel_all _ (Γ ,, vdef na t A) 0 t A' eq_refl). + eapply (weakening_red_0 (Γ' := [_]) (N := tSort _)); tea; [reflexivity|..]. erewrite -> on_free_vars_ctx_on_ctx_free_vars. apply H. apply H. apply H. - - destruct HH as (HA & Ht & HB). - apply infer_typing_sort_impl with id HB; intros Hs. + - destruct HH as (HtA & HB). + apply lift_sorting_f_it_impl with HB => // Hs. eapply type_reduction; tas. * econstructor; tea. - apply HA.π2. * apply red1_red. apply red_zeta with (b':=tSort _). Defined. diff --git a/pcuic/theories/PCUICSafeLemmata.v b/pcuic/theories/PCUICSafeLemmata.v index 9d398c82a..7bc9cffb9 100644 --- a/pcuic/theories/PCUICSafeLemmata.v +++ b/pcuic/theories/PCUICSafeLemmata.v @@ -38,8 +38,8 @@ Proof. * eapply sp. * unshelve epose proof (on_minductive_wf_params_indices_inst isdecl _ cu). rewrite PCUICUnivSubstitutionConv.subst_instance_app_ctx in X. - eapply wf_local_app_inv in X as []. - eapply wf_local_app => //. + eapply All_local_env_app_inv in X as []. + eapply All_local_env_app => //. eapply wf_local_smash_end. now eapply weaken_wf_local; tea. Qed. @@ -102,24 +102,6 @@ Section Lemmata. Context {cf : checker_flags}. Context (flags : RedFlags.t). - Instance All2_eq_refl Σ Re : - RelationClasses.Reflexive Re -> - CRelationClasses.Reflexive (All2 (eq_term_upto_univ Σ Re Re)). - Proof using Type. - intros h x. apply All2_same. reflexivity. - Qed. - - Instance All2_br_eq_refl Σ Re : - RelationClasses.Reflexive Re -> - CRelationClasses.Reflexive (All2 - (fun x y : branch term => - eq_context_upto Σ Re Re (bcontext x) (bcontext y) * - eq_term_upto_univ Σ Re Re (bbody x) (bbody y))). - Proof using Type. - intros h x. - apply All2_same; split; reflexivity. - Qed. - (* red is the reflexive transitive closure of one-step reduction and thus can't be used as well order. We thus define the transitive closure, but we take the symmetric version. @@ -182,11 +164,6 @@ Section Lemmata. eapply ws_cumul_pb_Lambda_r. assumption. Qed. - Lemma snoc_app_context {Γ Δ d} : (Γ ,,, (d :: Δ)) = (Γ ,,, Δ) ,,, [d]. - Proof using Type. - reflexivity. - Qed. - Lemma conv_alt_it_mkProd_or_LetIn : forall {wfΣ : wf Σ} Δ Γ B B', Σ ;;; (Δ ,,, Γ) ⊢ B = B' -> @@ -215,8 +192,8 @@ Section Lemmata. ∥ Σ ;;; Γ |- t : T ∥ -> welltyped Σ Γ T. Proof. intros [X]. - intros. eapply validity in X; try assumption. - destruct X. now exists (tSort x). + eapply validity in X; tas. + now apply isType_welltyped. Defined. Lemma wat_welltyped {Γ T} : @@ -228,7 +205,7 @@ Section Lemmata. Lemma welltyped_alpha Γ u v : welltyped Σ Γ u -> - eq_term_upto_univ empty_global_env eq eq u v -> + u ≡α v -> welltyped Σ Γ v. Proof using hΣ. intros [A h] e. @@ -342,34 +319,10 @@ Section Lemmata. Proof using Type. intros wfl. destruct pcontext as ((?&h)&?); simpl in *. - apply wf_local_app_inv in wfl as (_&wf). - apply wf_local_rel_app_inv in wf as (wf&_). + apply All_local_env_app_inv in wfl as (_&wf). + apply All_local_rel_app_inv in wf as (wf&_). destruct h; depelim wf; simpl in *. - all: destruct l; econstructor; eauto. - Qed. - (* TODO: rename alpha_eq *) - Lemma compare_decls_conv Γ Γ' : - eq_context_upto_names Γ Γ' -> - conv_context cumulAlgo_gen Σ Γ Γ'. - Proof using Type. - intros. - induction X; constructor; auto. - destruct r; constructor; subst; auto; reflexivity. - Qed. - - Lemma compare_decls_eq_context Γ Γ' : - eq_context_upto_names Γ Γ' <~> - eq_context_gen eq eq Γ Γ'. - Proof using Type. - split; induction 1; constructor; auto. - Qed. - - Lemma alpha_eq_inst_case_context Γ Δ pars puinst : - eq_context_upto_names Γ Δ -> - eq_context_upto_names (inst_case_context pars puinst Γ) (inst_case_context pars puinst Δ). - Proof using Type. - intros. rewrite /inst_case_context. - now eapply alpha_eq_subst_context, alpha_eq_subst_instance. + all: destruct l as (Hb & s & Hs & _); cbn in *; econstructor; eauto. Qed. Lemma welltyped_context : @@ -386,8 +339,8 @@ Section Lemmata. all: try apply inversion_App in typ as (?&?&?&?&?&?); auto. all: try apply inversion_Proj in typ as (?&?&?&?&?&?&?&?&?); auto. all: try apply inversion_Prod in typ as (?&?&?&?&?); auto. - all: try apply inversion_Lambda in typ as (?&?&?&?&?); auto. - all: try apply inversion_LetIn in typ as (?&?&?&?&?&?); auto. + all: try apply inversion_Lambda in typ as (?&?&?&?); auto. + all: try apply inversion_LetIn in typ as (?&?&?&?); auto. all: try solve [econstructor; eauto]. - apply inversion_Fix in typ as (?&?&?&?&?&?&?); eauto. destruct mfix as ((?&[])&?); simpl in *. @@ -395,7 +348,7 @@ Section Lemmata. depelim a. eauto using isType_welltyped. + eapply All_app in a0 as (_&a0). - depelim a0. + depelim a0. destruct o as (t0 & _); cbn in t0. rewrite fix_context_fix_context_alt in t0. rewrite map_app in t0. simpl in t0. @@ -407,7 +360,7 @@ Section Lemmata. depelim a. eauto using isType_welltyped. + eapply All_app in a0 as (_&a0). - depelim a0. + depelim a0. destruct o as (t0 & _); cbn in t0. rewrite fix_context_fix_context_alt in t0. rewrite map_app in t0. simpl in t0. @@ -416,7 +369,7 @@ Section Lemmata. - apply inversion_Case in typ as (?&?&?&?&[]&?); auto. rewrite app_context_assoc. destruct p. - + apply validity in scrut_ty as (?&typ). + + apply validity in scrut_ty as (_ & ? & typ & _). clear brs_ty. apply inversion_mkApps in typ as (?&_&spine); auto; simpl in *. clear -spine. @@ -446,7 +399,7 @@ Section Lemmata. eapply (declared_minductive_ind_npars x1). } now eapply spine_subst_smash in sp. ** cbn in conv_pctx. - eapply wf_local_app_inv. eapply wf_local_alpha. + eapply All_local_env_app_inv. eapply wf_local_alpha. ++ instantiate (1 := (Γ,,, stack_context π,,, smash_context [] (ind_params x)@[puinst],,, (ind_predicate_context ci x x0)@[puinst])). eapply All2_app => //. { eapply alpha_eq_subst_instance. now symmetry. } @@ -458,11 +411,10 @@ Section Lemmata. rewrite /case_predicate_context /= /case_predicate_context_gen. rewrite /pre_case_predicate_context_gen. rewrite /inst_case_context. - apply compare_decls_conv. + apply eq_context_upto_names_conv_context. eapply All2_app. 2:{ reflexivity. } - eapply compare_decls_eq_context. apply (PCUICAlpha.inst_case_predicate_context_eq (ind:=ci) wf_pred). - cbn. apply compare_decls_eq_context. now symmetry. + cbn. now symmetry. - apply inversion_Case in typ as (?&?&?&?&[]&?); auto. econstructor; eauto. @@ -490,6 +442,10 @@ Section Lemmata. cbn in *. eapply closed_context_conversion; tea. now symmetry. + - apply unlift_TypUniv in l. now econstructor. + - now eapply isType_welltyped. + - apply unlift_TermTyp in l. now econstructor. + - apply lift_sorting_forget_body in l. now eapply isType_welltyped. - eapply inversion_Prim in typ as (?&?&[]); eauto. depelim p0. now eexists. - eapply inversion_Prim in typ as (?&?&[]); eauto. @@ -601,19 +557,9 @@ Section Lemmata. welltyped Σ Γ (it_mkLambda_or_LetIn Δ t) -> welltyped Σ (Γ ,,, Δ) t. Proof using hΣ. - intros Γ Δ t h. - induction Δ as [| [na [b|] A] Δ ih ] in Γ, t, h |- *. - - assumption. - - simpl. apply ih in h. cbn in h. - destruct h as [T h]. - apply inversion_LetIn in h as hh ; auto. - destruct hh as [s1 [A' [? [? [? ?]]]]]. - exists A'. assumption. - - simpl. apply ih in h. cbn in h. - destruct h as [T h]. - apply inversion_Lambda in h as hh ; auto. - pose proof hh as [s1 [B [? [? ?]]]]. - exists B. assumption. + intros Γ Δ t (T & h). + apply inversion_it_mkLambda_or_LetIn in h as (T' & h & hle); tas. + now econstructor. Qed. Lemma it_mkLambda_or_LetIn_welltyped : @@ -965,7 +911,7 @@ Section Lemmata. inversion e. reflexivity. Qed. - Hint Resolve cumul_refl conv_alt_red : core. + Hint Resolve cumul_refl cumul_alt : core. Hint Resolve cumul_refl : core. Lemma cored_red_cored : diff --git a/pcuic/theories/PCUICSpine.v b/pcuic/theories/PCUICSpine.v index 1c808d6b6..ad1208051 100644 --- a/pcuic/theories/PCUICSpine.v +++ b/pcuic/theories/PCUICSpine.v @@ -25,7 +25,7 @@ Implicit Types (cf : checker_flags) (Σ : global_env_ext). Derive Signature for ctx_inst. -Notation ctx_inst := (ctx_inst typing). +Notation ctx_inst Σ := (ctx_inst (typing Σ)). Ltac splits := repeat split. @@ -71,29 +71,23 @@ Lemma subslet_eq_context_alpha {cf} {Σ Γ s Δ Δ'} : Proof. intros eq subs. induction subs in Δ', eq |- *; depelim eq; try constructor. - * depelim c; constructor; auto. now subst. - * depelim c; subst; constructor; auto. -Qed. - -Lemma eq_context_alpha_conv {cf} {Σ} {wfΣ : wf Σ} {Γ Γ'} : - eq_context_upto_names Γ Γ' -> conv_context cumulAlgo_gen Σ Γ Γ'. -Proof. - intros a. - eapply eq_context_upto_empty_conv_context. - eapply All2_fold_All2. - eapply (All2_impl a). - intros ?? []; constructor; subst; auto; reflexivity. + * depelim e; constructor; auto. + * depelim e; subst; constructor; auto. Qed. Lemma wf_local_alpha {cf} {Σ} {wfΣ : wf Σ} Γ Γ' : eq_context_upto_names Γ Γ' -> wf_local Σ Γ -> wf_local Σ Γ'. Proof. - induction 1; intros h; depelim h; try constructor; auto. - all:depelim r; constructor; subst; auto. - all: eapply lift_typing_impl; tea; intros T Hty. - all: eapply context_conversion; eauto. - all: now apply eq_context_alpha_conv. + induction 1; intros h. + 1: constructor. + apply All_local_env_tip in h as (h & hd). + apply All_local_env_snoc; auto. + replace (j_decl y) with (j_decl x : judgment). + 2: now destruct r. + eapply lift_typing_impl; tea; intros t T Hty. + eapply context_conversion; eauto. + now apply eq_context_upto_names_conv_context. Qed. Lemma subslet_eq_context_alpha_dom {cf} {Σ} {wfΣ : wf Σ} {Γ Γ' s Δ} : @@ -106,11 +100,11 @@ Proof. * now apply IHsubs. * eapply context_conversion; tea. eapply wf_local_alpha; tea. eapply typing_wf_local in t0. exact t0. - now eapply eq_context_alpha_conv. + now eapply eq_context_upto_names_conv_context. * now eapply IHsubs. * eapply context_conversion; tea. eapply wf_local_alpha; tea. eapply typing_wf_local in t0. exact t0. - now eapply eq_context_alpha_conv. + now eapply eq_context_upto_names_conv_context. Qed. @@ -227,7 +221,7 @@ Proof. apply IHctxi. Qed. -Record spine_subst {cf:checker_flags} Σ Γ inst s Δ := mkSpineSubst { +Record spine_subst {cf:checker_flags} Σ Γ inst s (Δ : context) := mkSpineSubst { spine_dom_wf : wf_local Σ Γ; spine_codom_wf : wf_local Σ (Γ ,,, Δ); inst_ctx_subst :> context_subst Δ inst s; @@ -356,24 +350,24 @@ Section WfEnv. ∑ A s', Σ ;;; Γ |- f : A × Σ ;;; Γ |- A : tSort s' × typing_spine Σ Γ A u T. Proof using wfΣ. induction u in f, T |- *. simpl. intros. - { destruct X as [s X]. exists T, s. intuition pcuic. eapply typing_spine_refl. eexists; eauto. } + { destruct X as (_ & s & X & _). exists T, s. intuition pcuic. eapply typing_spine_refl. eapply has_sort_isType; eauto. } intros HT Hf. simpl in Hf. destruct u. simpl in Hf, IHu. - edestruct @inversion_App_size with (H := Hf) as (na' & A' & B' & s & Hf' & Ha & HA & _ & _ & _ & HA'''); tea. eexists _, _; intuition eauto. econstructor; eauto with pcuic. - eapply isType_ws_cumul_pb_refl; eexists; eauto. + eapply isType_ws_cumul_pb_refl; eapply has_sort_isType; eauto. econstructor. all:eauto with pcuic. eapply inversion_Prod in HA as (? & ? & ? & ? & ?); tea. eapply isType_subst. econstructor. econstructor. rewrite subst_empty; eauto. econstructor; cbn; eauto. - specialize (IHu (tApp f a) T). - epose proof (IHu _ Hf) as [T' [s' [H' [H''' H'']]]]. + unshelve epose proof (IHu _ Hf) as [T' [s' [H' [H''' H'']]]]; tea. edestruct @inversion_App_size with (H := H') as (na' & A' & B' & s & Hf' & Ha & HA & _ & _ & _ & HA'''); tea. exists (tProd na' A' B'). exists s. intuition; eauto. econstructor; eauto with wf. - 1,2: eexists; eauto. 1:eapply isType_ws_cumul_pb_refl; eexists; eauto. + 1,2: pcuic. 1:eapply isType_ws_cumul_pb_refl; pcuic. eapply typing_spine_strengthen; tea. @@ -381,7 +375,6 @@ Section WfEnv. eapply inversion_Prod in HA as (? & ? & ? & ? & ?); tea. eapply isType_subst. econstructor. econstructor. rewrite subst_empty; eauto. econstructor; cbn; eauto. - Unshelve. eauto. Qed. Lemma subst_type_local_ctx {Γ Γ' Δ Δ' s ctxs} : @@ -394,10 +387,8 @@ Section WfEnv. destruct a as [na [b|] ty]; rewrite subst_context_snoc /= /subst_decl /map_decl /=; simpl; intuition auto. - 1: apply infer_typing_sort_impl with id a0; intros Hs. - 1: destruct a0 as (? & t); cbn in Hs |- *; clear t. - 2: rename b1 into Hs. - 3: rename b into Hs. + 2: rename b into b0. + all: apply lift_typing_f_impl with (1 := b0) => // ?? Hs. all: rewrite -app_context_assoc in Hs. all: eapply substitution in Hs; eauto. all: rewrite subst_context_app app_context_assoc in Hs. @@ -414,24 +405,13 @@ Section WfEnv. induction Δ' in Γ', ctxs |- *; simpl; auto. destruct a as [na [b|] ty]; rewrite subst_context_snoc /= /subst_decl /map_decl /=; intuition auto. - + eapply infer_typing_sort_impl with id a0; intros Hs. - destruct a0 as (? & t); cbn in Hs |- *; clear t. - rewrite -app_context_assoc in Hs. - eapply substitution in Hs; eauto. - rewrite subst_context_app app_context_assoc in Hs. - simpl in Hs. rewrite Nat.add_0_r in Hs. - now rewrite app_context_length in Hs. - + rewrite -app_context_assoc in b1. - eapply substitution in b1; eauto. - rewrite subst_context_app app_context_assoc Nat.add_0_r in b1. - now rewrite app_context_length in b1. - + destruct ctxs => //. - intuition auto. - rewrite -app_context_assoc in b. - eapply substitution in b; eauto. - rewrite subst_context_app app_context_assoc in b. - rewrite Nat.add_0_r in b. - now rewrite app_context_length in b. + 2: destruct ctxs => //=; destruct X0 as (a & b0); split; eauto. + all: apply lift_typing_f_impl with (1 := b0) => // ?? Hs. + all: rewrite -app_context_assoc in Hs. + all: eapply substitution in Hs; eauto. + all: rewrite subst_context_app app_context_assoc in Hs. + all: simpl in Hs; rewrite Nat.add_0_r in Hs. + all: now rewrite app_context_length in Hs. Qed. Lemma wf_arity_spine_typing_spine {Γ T args T'} : @@ -455,7 +435,7 @@ Section WfEnv. Lemma arity_typing_spine {Γ Γ' s inst s'} : typing_spine Σ Γ (it_mkProd_or_LetIn Γ' (tSort s)) inst (tSort s') -> - [× (#|inst| = context_assumptions Γ'), leq_universe (global_ext_constraints Σ) s s' & + [× (#|inst| = context_assumptions Γ'), leq_sort (global_ext_constraints Σ) s s' & ∑ instsubst, spine_subst Σ Γ inst instsubst Γ']. Proof using wfΣ. revert s inst s'. @@ -507,7 +487,7 @@ Section WfEnv. * apply subslet_app; rewrite !subst_empty //. eapply subslet_def_tip. rewrite app_context_assoc /= in Hsp. - apply wf_local_app_l in Hsp. now depelim Hsp. + apply wf_local_app_l in Hsp. depelim Hsp. now eapply unlift_TermTyp. + rewrite context_assumptions_app /=. assert (Hsp' := Hsp). rewrite it_mkProd_or_LetIn_app in Hsp'. @@ -870,8 +850,8 @@ Qed.*) now apply subslet_lift. Qed. - Lemma ctx_inst_length {Γ args Δ} : - ctx_inst Σ Γ args Δ -> + Lemma ctx_inst_length {P Γ args Δ} : + PCUICTyping.ctx_inst P Γ args Δ -> #|args| = context_assumptions Δ. Proof using Type. induction 1; simpl; auto. @@ -1058,8 +1038,8 @@ Qed.*) Hint Resolve spine_subst_is_closed_context spine_subst_is_closed_context_codom : fvs. Lemma arity_spine_it_mkProd_or_LetIn_Sort {Γ ctx s s' args inst} : - wf_universe Σ s' -> - leq_universe Σ s s' -> + wf_sort Σ s' -> + leq_sort Σ s s' -> spine_subst Σ Γ args inst ctx -> arity_spine Σ Γ (it_mkProd_or_LetIn ctx (tSort s)) args (tSort s'). Proof using wfΣ. @@ -1184,7 +1164,7 @@ Qed.*) Qed. Lemma ctx_inst_ass {Γ args na t} (c : ctx_inst Σ Γ args [vass na t]) : - ∑ i, ((args = [i]) * (lift_typing typing Σ Γ i (Typ t)) * (ctx_inst_sub c = [i]))%type. + ∑ i, ((args = [i]) * (typing Σ Γ i t) * (ctx_inst_sub c = [i]))%type. Proof using Type. depelim c; simpl in *. depelim c. exists i; constructor; auto. @@ -1199,7 +1179,7 @@ Qed.*) pose proof (ctx_inst_sub_spec ci) as msub. eapply make_context_subst_spec in msub. rewrite List.rev_involutive in msub. - split; pcuic. now eapply wf_local_app_inv in wfΔ as []. + split; pcuic. now eapply All_local_env_app_inv in wfΔ as []. move: ci msub. induction Δ in wfΔ, args |- *. simpl. intros ci. depelim ci. constructor. @@ -1219,7 +1199,8 @@ Qed.*) intros. depelim wfΔ. specialize (IHΔ _ wfΔ _ subr). constructor; auto. - red in l0. eapply (substitution (Δ := [])); eauto. + apply unlift_TermTyp in l. + eapply (substitution (Δ := [])); eauto. - pose proof (ctx_inst_ass c) as [i [[Hargs Hty] Hinst]]. rewrite Hinst in msub |- *. simpl in *. destruct msub as [subl subr]. @@ -1390,11 +1371,11 @@ Lemma all_rels_subst {cf:checker_flags} Σ Δ Γ t : Proof. intros wfΣ wf. assert(forall Γ' t (wf : wf_local Σ Γ'), - ((All_local_env_over typing - (fun Σ Γ' wfΓ' t T _ => + ((All_local_env_over (typing Σ) + (fun Γ' wfΓ' t T _ => forall Γ Δ, Γ' = Γ ,,, Δ -> red Σ.1 (Γ ,,, Δ) t (subst0 (all_rels Δ 0 #|Δ|) (lift #|Δ| #|Δ| t))) - Σ Γ' wf) * + Γ' wf) * (match t with | Some t => forall Γ Δ, Γ' = Γ ,,, Δ -> red Σ.1 (Γ ,,, Δ) t (subst0 (all_rels Δ 0 #|Δ|) (lift #|Δ| #|Δ| t)) @@ -1427,14 +1408,13 @@ Proof. intros Γ0 Δ ->. rename Γ0 into Γ. specialize (IH cf Σ). - assert (All_local_env_over typing - (fun (Σ : PCUICEnvironment.global_env_ext) - (Γ'0 : PCUICEnvironment.context) (_ : wf_local Σ Γ'0) + assert (All_local_env_over (typing Σ) + (fun (Γ'0 : PCUICEnvironment.context) (_ : wf_local Σ Γ'0) (t T : term) (_ : Σ;;; Γ'0 |- t : T) => forall Γ Δ : context, Γ'0 = Γ ,,, Δ -> red Σ.1 (Γ ,,, Δ) t (subst0 (all_rels Δ 0 #|Δ|) (lift #|Δ| #|Δ| t))) - Σ _ wf). + _ wf). { specialize (IH wfΣ (Γ ,,, Δ) None). forward IH. simpl. right. lia. apply (IH wf). } @@ -1498,19 +1478,19 @@ Proof. rewrite !on_free_vars_ctx_app => /andP[] onΓ. erewrite onΓ => /=. rewrite -{1}(firstn_skipn (S i) Δ) on_free_vars_ctx_app => /andP[] //. } - { clear X; eapply (nth_error_All_local_env (n:=i)) in wf. 2:len; lia. - rewrite nth_error_app_lt // in wf. - rewrite Hnth /= in wf. + { clear X; eapply (All_local_env_nth_error (n:=i)) in wf. + 2: rewrite nth_error_app_lt; eassumption. rewrite skipn_app in wf. replace (S i - #|Δ|) with 0 in wf. 2:lia. rewrite skipn_0 in wf. - rewrite /on_local_decl /= in wf. + rewrite /= in wf. move: wf => [] /subject_closed //. rewrite is_open_term_closed //. } rewrite skipn_length; simpl. - unshelve eapply (nth_error_All_local_env_over (n:=i)) in X. - 2:{ rewrite nth_error_app_lt //. apply Hnth. } - destruct X as [_ [Xb Xt]]. + apply All_local_env_over_2 in X. + eapply (All_local_env_nth_error (n:=i)) in X. + 2: rewrite nth_error_app_lt; eassumption. + destruct X as ((Hb & Xb) & s & (Ht & Xt) & _). specialize (Xb Γ (skipn (S i) Δ)). forward Xb. rewrite skipn_app. unfold app_context. f_equal. assert(S i - #|Δ| = 0) by lia. rewrite H. apply skipn_0. @@ -1613,7 +1593,7 @@ Section WfEnv. now rewrite lift_context_app Nat.add_0_r. Qed. - Lemma spine_subst_to_extended_list_k {Δ Γ} : + Lemma spine_subst_to_extended_list_k {Δ Γ : context} : wf_local Σ (Γ ,,, Δ) -> spine_subst Σ (Γ ,,, Δ) (reln [] 0 Δ) (all_rels Δ 0 #|Δ|) (lift_context #|Δ| 0 Δ). @@ -1661,9 +1641,9 @@ Section WfEnv. rewrite -eql in l0. autorewrite with len in l0. simpl in l0. lia. eapply (substitution (Δ := []) IHc); auto. rewrite lift_context0_app !app_context_assoc in X. cbn in X. - eapply wf_local_app_inv in X as []. + eapply All_local_env_app_inv in X as []. rewrite lift_context_snoc0 Nat.add_0_r /= in a. cbn in a. - depelim a. now cbn in l1. + depelim a. now apply unlift_TermTyp in l0. * rewrite app_length /= Nat.add_1_r in IHc. intros Hwf. specialize (IHc Hwf). constructor; auto. @@ -1677,12 +1657,12 @@ Section WfEnv. constructor. auto. rewrite -eql nth_error_app_lt ?app_length /=; try lia. rewrite nth_error_app_ge // ?Nat.sub_diag //. - destruct l0. - exists x. - change (tSort x) with - (subst0 (all_rels c (S #|l|) #|Δ|) (lift #|Δ| #|c| (tSort x))). + apply lift_sorting_f_it_impl with (f := fun t => _ (_ t)) (tu := l0) => // Hs. + set (s := l0.2.π1). + change (tSort s) with + (subst0 (all_rels c (S #|l|) #|Δ|) (lift #|Δ| #|c| (tSort s))). { eapply (substitution (Γ' := lift_context #|Δ| 0 c) (Δ := [])); cbn; auto. - change (tSort x) with (lift #|Δ| #|c| (tSort x)). + change (tSort s) with (lift #|Δ| #|c| (tSort s)). eapply (weakening_typing); eauto. } eapply ws_cumul_pb_eq_le. simpl. rewrite -{1}eql. simpl. @@ -1701,7 +1681,7 @@ Section WfEnv. forward X by auto. apply X; auto. all:eauto with fvs. rewrite -app_tip_assoc app_assoc -[(l ++ _) ++ _]app_assoc eql. - eapply wf_local_app_inv in Hwf as []. eauto with fvs. + eapply All_local_env_app_inv in Hwf as []. eauto with fvs. Qed. Lemma red_expand_let {Γ na b ty t} : @@ -1722,57 +1702,57 @@ Section WfEnv. ∑ Δs ts, [× sorts_local_ctx (lift_typing typing) Σ Γ Δ Δs, Σ ;;; Γ ,,, Δ |- t : tSort ts, - wf_universe Σ s & - leq_universe Σ (sort_of_products Δs ts) s]. + wf_sort Σ s & + leq_sort Σ (sort_of_products Δs ts) s]. Proof using wfΣ. intros h. revert Γ t s h. induction Δ; intros. - - exists [], s; splits. apply h. eauto with pcuic. apply leq_universe_refl. + - exists [], s; splits. apply h. eauto with pcuic. apply leq_sort_refl. - destruct a as [na [b|] ty]; simpl in *; rewrite /mkProd_or_LetIn /= in h. * specialize (IHΔ _ _ _ h) as (Δs & ts & [sorts IHΔ leq]). exists Δs, ts. - pose proof (PCUICWfUniverses.typing_wf_universe _ IHΔ) as wfts. - eapply inversion_LetIn in IHΔ as [s' [? [? [? [? e]]]]]; auto. - splits; eauto. now eexists. - eapply (type_ws_cumul_pb (pb:=Cumul)). eapply t2. apply isType_Sort; pcuic. - eapply ws_cumul_pb_LetIn_l_inv in e; auto. - eapply ws_cumul_pb_Sort_r_inv in e as [u' [redu' cumu']]. + pose proof (PCUICWfUniverses.typing_wf_sort _ IHΔ) as wfts. + eapply inversion_LetIn in IHΔ as (T & wfty & HT & hlt); auto. + split; eauto. + eapply (type_ws_cumul_pb (pb:=Cumul)). eapply HT. apply isType_Sort; pcuic. + eapply ws_cumul_pb_LetIn_l_inv in hlt; auto. + eapply ws_cumul_pb_Sort_r_inv in hlt as [u' [redu' cumu']]. transitivity (tSort u'). 2:{ eapply ws_cumul_pb_compare; eauto with fvs. - eapply typing_wf_local in t2. eauto with fvs. + eapply typing_wf_local in HT. eauto with fvs. econstructor. eauto with fvs. } eapply ws_cumul_pb_red. exists (tSort u'), (tSort u'). split; auto. 3:now constructor. - transitivity (lift0 1 (x {0 := b})). + transitivity (lift0 1 (T {0 := b})). eapply red_expand_let. pcuic. - eapply type_closed in t2. + eapply type_closed in HT. rewrite -is_open_term_closed //. change (tSort u') with (lift0 1 (tSort u')). eapply (weakening_closed_red (Γ := Γ ,,, Δ) (Γ' := []) (Γ'' := [_])); auto with fvs. - apply typing_wf_local in t2. eauto with fvs. - eapply closed_red_refl; eauto with fvs. + apply typing_wf_local in HT. eauto with fvs. + eapply closed_red_refl; eapply typing_wf_local in HT; eauto with fvs. * specialize (IHΔ _ _ _ h) as (Δs & ts & [sorts IHΔ leq]). - eapply inversion_Prod in IHΔ as [? [? [? [? e]]]]; tea. - exists (x :: Δs), x0. splits; tea. - eapply ws_cumul_pb_Sort_inv in e. - transitivity (sort_of_products Δs ts); auto using leq_universe_product. - simpl. eapply leq_universe_sort_of_products_mon. + eapply inversion_Prod in IHΔ as (s1 & s2 & wfty & Ht & hlt); tea. + exists (s1 :: Δs), s2. split; tea. split; tas. + eapply ws_cumul_pb_Sort_inv in hlt. + transitivity (sort_of_products Δs ts); auto using leq_sort_product. + simpl. eapply leq_sort_sort_of_products_mon. eapply Forall2_same. reflexivity. - exact: e. + exact: hlt. Qed. - Lemma leq_universe_sort_of_products {u v} : - leq_universe Σ v (sort_of_products u v). + Lemma leq_sort_sort_of_products {u v} : + leq_sort Σ v (sort_of_products u v). Proof using Type. induction u; simpl; auto. - reflexivity. - etransitivity; tea. - eapply leq_universe_sort_of_products_mon => //. + eapply leq_sort_sort_of_products_mon => //. eapply Forall2_same. reflexivity. - eapply leq_universe_product. + eapply leq_sort_product. Qed. Lemma inversion_it_mkProd_or_LetIn {Γ Δ t s} : @@ -1782,8 +1762,8 @@ Section WfEnv. move/type_it_mkProd_or_LetIn_inv => [Δs [ts [hΔ ht hs leq]]]. eapply type_Cumul; tea. eapply type_Sort; pcuic. eapply cumul_Sort. - transitivity (sort_of_products Δs ts); auto using leq_universe_product. - apply leq_universe_sort_of_products. + transitivity (sort_of_products Δs ts); auto using leq_sort_product. + apply leq_sort_sort_of_products. Qed. Lemma isType_it_mkProd_or_LetIn_app {Γ Δ Δ' args T s} : @@ -1955,7 +1935,7 @@ Section WfEnv. Proof using wfΣ. intros ha hm. eapply (PCUICConversion.substitution_ws_cumul_pb (Γ' := [vdef na a ty]) (s := [a]) (Γ'':=[])); eauto with pcuic. - eapply subslet_def_tip. now depelim ha. + eapply subslet_def_tip. depelim ha. now apply unlift_TermTyp in l. Qed. Lemma subst0_it_mkProd_or_LetIn s Γ T : subst s 0 (it_mkProd_or_LetIn Γ T) = @@ -2003,7 +1983,7 @@ Section WfEnv. repeat constructor. rewrite !subst_empty. eapply All_local_env_app_inv in wfΓΔ as [_ wf]. eapply All_local_env_app_inv in wf as [wfd _]. - depelim wfd. apply l0. + depelim wfd. now eapply unlift_TermTyp. * rewrite subst_app_simpl. move: (context_subst_length sps). now autorewrite with len => <-. @@ -2040,9 +2020,9 @@ Section WfEnv. now autorewrite with len => <-. Qed. - Arguments ctx_inst_nil {typing} {Σ} {Γ}. - Arguments PCUICTyping.ctx_inst_ass {typing} {Σ} {Γ} {na t i inst Δ}. - Arguments PCUICTyping.ctx_inst_def {typing} {Σ} {Γ} {na b t inst Δ}. + Arguments ctx_inst_nil {typing} {Γ}. + Arguments PCUICTyping.ctx_inst_ass {typing} {Γ} {na t i inst Δ}. + Arguments PCUICTyping.ctx_inst_def {typing} {Γ} {na b t inst Δ}. Lemma spine_subst_ctx_inst_sub {Γ args argsub Δ} (sp : spine_subst Σ Γ args argsub Δ) : ctx_inst_sub (spine_subst_ctx_inst sp) = argsub. @@ -2166,7 +2146,7 @@ Section WfEnv. simpl. rewrite smash_context_acc. simpl. auto. auto. } split; auto. - - eapply All_local_env_app; split; auto. + - eapply All_local_env_app; auto. eapply wf_local_rel_smash_context; auto. - induction inst_subslet0 in inst, inst_ctx_subst0, spine_codom_wf0 |- *. depelim inst_ctx_subst0. @@ -2425,8 +2405,8 @@ Section WfEnv. intros. specialize (IHHsp (tApp t0 hd)). apply IHHsp. - destruct i as [s Hs]. - eapply type_App; eauto. eapply i0.π2. + destruct i as (_ & s & Hs & _). + eapply type_App; eauto. eapply i0.2.π2.1. eapply type_ws_cumul_pb; eauto. Qed. @@ -2703,7 +2683,7 @@ Section WfEnv. rewrite app_context_assoc in cum. eapply substitution_ws_cumul_pb in cum; tea. len in cum; tea. - destruct (wf_local_app_inv wf). + destruct (All_local_env_app_inv wf). simpl. len. now eapply PCUICContexts.subslet_extended_subst. @@ -2719,7 +2699,7 @@ Section WfEnv. eapply (weakening_ws_cumul_pb (Γ'' := smash_context [] Δ)) in cum; tea. rewrite /expand_lets /expand_lets_k. eapply (PCUICConversion.substitution_ws_cumul_pb (Γ'' := [])) in cum; tea. len in cum; tea. - destruct (wf_local_app_inv wf). + destruct (All_local_env_app_inv wf). simpl. len. now eapply PCUICContexts.subslet_extended_subst. @@ -2834,7 +2814,7 @@ Section WfEnv. rewrite -[context_assumptions Γ0](smash_context_length []); cbn. relativize #|Γ0|. eapply is_open_term_lift. - destruct l0 as [s Hs]. eapply subject_closed in Hs. + destruct l0 as (_ & s & Hs & _). eapply subject_closed in Hs. rewrite is_open_term_closed in Hs. move: Hs. now rewrite !app_length -(All2_fold_length cum). reflexivity. * split; auto. @@ -2866,8 +2846,8 @@ Section WfEnv. rewrite -[context_assumptions Γ0](smash_context_length []); cbn. relativize #|Γ0|. eapply is_open_term_lift. - eapply subject_closed in l2. - rewrite is_open_term_closed in l2. move: l2. + eapply unlift_TermTyp, subject_closed in l0. + rewrite is_open_term_closed in l0. move: l0. now rewrite !app_length -(All2_fold_length cum). reflexivity. Qed. @@ -2962,7 +2942,7 @@ Section WfEnv. rewrite skipn_all_app_eq // in H. noconf H. intros HΔ; depelim HΔ. intros HΔ'; depelim HΔ'. - destruct l0 as [s Hs]. simpl. + destruct l0 as (_ & s & Hs & _). simpl. rewrite (ctx_inst_sub_subst dom) in t1. rewrite firstn_app_left // in dom. specialize (IHa _ dom HΔ HΔ'). @@ -2981,10 +2961,10 @@ Section WfEnv. eapply type_ws_cumul_pb; tea. + eapply typing_expand_lets in Hs. eapply (substitution (s := List.rev i) (Δ := [])) in Hs; tea. - simpl in Hs. now exists s; rewrite subst_context_nil /= in Hs. + simpl in Hs. eapply has_sort_isType. now rewrite subst_context_nil /= in Hs. exact X. + unshelve epose proof (ctx_inst_spine_subst _ dom); tea. - eapply wf_local_app; tea. now eapply typing_wf_local. + eapply All_local_env_app; tea. now eapply typing_wf_local. pose proof (spine_codom_wf _ _ _ _ _ X0). eapply spine_subst_smash in X0; tea. eapply (PCUICConversion.substitution_ws_cumul_pb (Γ := Γ) (Γ'' := []) X0). @@ -3001,10 +2981,11 @@ Section WfEnv. rewrite H0 in H, dom. rewrite firstn_all in dom. intros HΔ; depelim HΔ. intros HΔ'; depelim HΔ'. - destruct l as [s Hs]. simpl in *. + destruct l as (Hb & s & Hs & _). simpl in *. + destruct l0 as (Hb' & s' & Hs' & _). simpl in *. specialize (IHa _ dom). - forward IHa. apply wf_local_app_inv; pcuic. - forward IHa. apply wf_local_app_inv; pcuic. + forward IHa. apply All_local_env_app_inv; pcuic. + forward IHa. apply All_local_env_app_inv; pcuic. rewrite -(app_nil_r i). eapply (ctx_inst_app IHa). rewrite (ctx_inst_sub_subst IHa) /=. @@ -3088,8 +3069,8 @@ Section WfEnv. (Γ ,,, subst_context (List.rev s') 0 Δ)) -> wf_local Σ (Γ ,,, (List.rev Δ)) -> PCUICTyping.ctx_inst - (fun (Σ : global_env_ext) (Γ : context) (t T : term) => - forall u : term, P Σ Γ t u -> Σ;;; Γ |- u : T) Σ Γ inst Δ -> + (fun (Γ : context) (t T : term) => + forall u : term, P Σ Γ t u -> Σ;;; Γ |- u : T) Γ inst Δ -> ctx_inst Σ Γ inst Δ -> OnOne2 (P Σ Γ) inst inst' -> ctx_inst Σ Γ inst' Δ. @@ -3108,9 +3089,9 @@ Section WfEnv. eapply (HP _ _ _ [i] [hd']); tea. repeat constructor. now rewrite subst_empty. repeat constructor. now rewrite subst_empty. constructor. auto. - eapply wf_local_app_inv. eapply substitution_wf_local; tea. + eapply All_local_env_app_inv. eapply substitution_wf_local; tea. repeat (constructor; tea). rewrite subst_empty; tea. - eapply wf_local_app_inv. eapply substitution_wf_local; tea. + eapply All_local_env_app_inv. eapply substitution_wf_local; tea. repeat (constructor; tea). rewrite subst_empty; tea. now eapply t0. constructor; auto. eapply IHc. rewrite -subst_context_subst_telescope. @@ -3123,7 +3104,7 @@ Section WfEnv. rewrite -subst_context_subst_telescope. eapply substitution_wf_local; tea. repeat (constructor; tea). eapply subslet_def_tip. - eapply wf_local_app_inv in wf as [wf _]. now depelim wf. + eapply All_local_env_app_inv in wf as [wf _]. depelim wf. now eapply unlift_TermTyp. Qed. Lemma All2_ctx_inst {pb} {P} {Γ inst inst' Δ} : @@ -3135,8 +3116,8 @@ Section WfEnv. (Γ ,,, subst_context (List.rev s') 0 Δ)) -> wf_local Σ (Γ ,,, (List.rev Δ)) -> PCUICTyping.ctx_inst - (fun (Σ : global_env_ext) (Γ : context) (t T : term) => - forall u : term, P Σ Γ t u -> Σ;;; Γ |- u : T) Σ Γ inst Δ -> + (fun (Γ : context) (t T : term) => + forall u : term, P Σ Γ t u -> Σ;;; Γ |- u : T) Γ inst Δ -> ctx_inst Σ Γ inst Δ -> All2 (P Σ Γ) inst inst' -> ctx_inst Σ Γ inst' Δ. @@ -3159,16 +3140,16 @@ Section WfEnv. repeat constructor. now rewrite subst_empty. now apply subslet_ass_tip. now repeat constructor. - * eapply wf_local_app_inv. eapply substitution_wf_local; tea. + * eapply All_local_env_app_inv. eapply substitution_wf_local; tea. now apply subslet_ass_tip. - * eapply wf_local_app_inv. eapply substitution_wf_local; tea. + * eapply All_local_env_app_inv. eapply substitution_wf_local; tea. now apply subslet_ass_tip. - constructor. eapply IHc; eauto. simpl in wf. rewrite - !/(app_context _ _) app_context_assoc in wf. rewrite -subst_context_subst_telescope. eapply substitution_wf_local; tea. repeat (constructor; tea). eapply subslet_def_tip. - eapply wf_local_app_inv in wf as [wf _]. now depelim wf. + eapply All_local_env_app_inv in wf as [wf _]. depelim wf. now eapply unlift_TermTyp. Qed. Lemma ctx_inst_open_terms Γ args Δ : @@ -3188,8 +3169,8 @@ Section WfEnv. Lemma ctx_inst_eq_context {Γ Δ : context} {args args'} : wf_local Σ (Γ ,,, List.rev Δ) -> PCUICTyping.ctx_inst - (fun (Σ : global_env_ext) (Γ : context) (u A : term) => - forall v : term, upto_names' u v -> Σ;;; Γ |- v : A) Σ Γ args Δ -> + (fun (Γ : context) (u A : term) => + forall v : term, upto_names' u v -> Σ;;; Γ |- v : A) Γ args Δ -> ctx_inst Σ Γ args Δ -> All2 upto_names' args args' -> ctx_inst Σ Γ args' Δ. @@ -3202,7 +3183,7 @@ Section WfEnv. now eapply subslet_untyped_subslet. now eapply subslet_untyped_subslet. eapply All2_rev. - move/wf_local_app_inv: X => [] /wf_local_app_inv[] /wf_local_closed_context clΓ0 _ _. + move/All_local_env_app_inv: X => [] /All_local_env_app_inv[] /wf_local_closed_context clΓ0 _ _. eapply subslet_open_terms, All_rev_inv in X0. eapply subslet_open_terms, All_rev_inv in X1. solve_all. eapply into_ws_cumul_pb; tea. @@ -3239,9 +3220,8 @@ Lemma wf_local_nth_isType {cf} {Σ} {Γ n d} : isType Σ (skipn (S n) Γ) d.(decl_type). Proof. intros Hwf hnth. - epose proof (nth_error_All_local_env (nth_error_Some_length hnth) Hwf). - rewrite hnth /= in X. unfold on_local_decl in X. - destruct decl_body => //. destruct X => //. + eapply All_local_env_nth_error in Hwf; tea. + eapply lift_sorting_it_impl_gen with Hwf => //. Qed. diff --git a/pcuic/theories/PCUICSubstitution.v b/pcuic/theories/PCUICSubstitution.v index 71e820a71..a363d8e7a 100644 --- a/pcuic/theories/PCUICSubstitution.v +++ b/pcuic/theories/PCUICSubstitution.v @@ -96,7 +96,7 @@ Lemma subslet_nth_error {cf:checker_flags} {Σ Γ s Δ decl n} : Proof. induction 1 in n |- *; simpl; auto; destruct n; simpl; try congruence. - intros [= <-]. exists t; split; auto. - simpl. split; auto. exact tt. + simpl. split; auto. - intros. destruct decl as [na' [b|] ty]; cbn in *. + specialize (IHX _ H) as [t' [hnth [hty heq]]]. exists t'; intuition auto. + now apply IHX. @@ -150,21 +150,19 @@ Proof. rewrite commut_lift_subst_rec. 1: lia. f_equal; lia. Qed. -Lemma All_local_env_subst {cf:checker_flags} (P Q : context -> term -> typ_or_sort -> Type) c n k : +Lemma All_local_env_subst {cf:checker_flags} (P Q : context -> judgment -> Type) c n k : All_local_env Q c -> - (forall Γ t T, - Q Γ t T -> - P (subst_context n k Γ) (subst n (#|Γ| + k) t) - (typ_or_sort_map (subst n (#|Γ| + k)) T) + (forall Γ j, + Q Γ j -> + P (subst_context n k Γ) (judgment_map (subst n (#|Γ| + k)) j) ) -> All_local_env P (subst_context n k c). Proof. intros Hq Hf. induction Hq in |- *; try econstructor; eauto; simpl; unfold snoc; rewrite subst_context_snoc; econstructor; eauto. - - simpl. eapply (Hf _ _ Sort). eauto. - - simpl. eapply (Hf _ _ Sort). eauto. - - simpl. eapply (Hf _ _ (Typ t)). eauto. + - simpl. eapply (Hf _ (Typ _)). eauto. + - simpl. eapply (Hf _ (TermTyp _ _)). eauto. Qed. Lemma subst_length {cf:checker_flags} Σ Γ s Γ' : subs Σ Γ s Γ' -> #|s| = #|Γ'|. @@ -348,12 +346,13 @@ Proof. induction 1; auto; unfold subst_context, snoc; rewrite fold_context_k_snoc0; auto; unfold snoc; f_equal; auto; unfold map_decl; simpl. - - destruct t0 as [s Hs]. unfold vass. simpl. f_equal. + - destruct t0 as (_ & s & Ht & _). unfold vass. simpl. f_equal. eapply typed_subst; eauto. lia. - unfold vdef. + destruct t0 as (Hb & s & Ht & _). cbn in Hb. f_equal. + f_equal. eapply typed_subst; eauto. lia. - + eapply typed_subst in t1 as [Ht HT]; eauto. lia. + + eapply typed_subst; eauto. lia. Qed. Lemma subst_declared_constant `{H:checker_flags} Σ cst decl n k u : @@ -363,19 +362,14 @@ Lemma subst_declared_constant `{H:checker_flags} Σ cst decl n k u : map_constant_body (subst_instance u) decl. Proof. intros. unshelve eapply declared_constant_to_gen in H0; eauto. - eapply declared_decl_closed in H0; eauto. - unfold map_constant_body. - do 2 red in H0. destruct decl as [ty [body|] univs]; simpl in *. - - rewrite -> andb_and in H0. intuition auto. - rewrite <- (closedn_subst_instance 0 body u) in H1. - rewrite <- (closedn_subst_instance 0 ty u) in H2. + eapply declared_decl_closed in H0 as [Hb Ht]%andb_and; eauto. + unfold map_constant_body. cbn in *. + f_equal. + - rewrite <- (closedn_subst_instance 0 _ u) in Ht. + apply subst_closedn; eauto using closed_upwards with arith wf. + - destruct cst_body; cbn in *; auto. f_equal. - + apply subst_closedn; eauto using closed_upwards with arith wf. - + f_equal. apply subst_closedn; eauto using closed_upwards with arith wf. - - red in H0. f_equal. - intuition. simpl in *. - rewrite <- (closedn_subst_instance 0 ty u) in H0. - rewrite andb_true_r in H0. + rewrite <- (closedn_subst_instance 0 _ u) in Hb. eapply subst_closedn; eauto using closed_upwards with arith wf. Qed. @@ -585,17 +579,18 @@ Proof. rewrite rev_map_app. simpl. apply Alli_app in Ha as [Hl Hx]. inv Hx. clear X0. - apply onArity in X as [s Hs]. + apply onArity in X as (_ & s & Hs & _). specialize (IHl Hl). econstructor; eauto. fold (arities_context l) in *. unshelve epose proof (weakening Σ [] (arities_context l) _ _ wfΣ _ Hs). 1: now rewrite app_context_nil_l. + repeat (eexists; tea). simpl in X. eapply (env_prop_typing typecheck_closed) in Hs; eauto. rewrite -> andb_and in Hs. destruct Hs as [Hs Ht]. simpl in Hs. apply (lift_closed #|arities_context l|) in Hs. - rewrite -> Hs, app_context_nil_l in X. simpl. exists s. + rewrite -> Hs, app_context_nil_l in X. apply X. Qed. @@ -611,7 +606,7 @@ Lemma on_constructor_closed_arities {cf:checker_flags} {Σ : global_env} {wfΣ : on_constructor cumulSpec0 (lift_typing typing) (Σ, ind_universes mdecl) mdecl (inductive_ind mind) indices idecl cdecl cs -> closedn #|arities_context (ind_bodies mdecl)| cdecl.(cstr_type). Proof. - intros [? ? [s Hs] _ _ _ _]. + intros [? ? (_ & s & Hs & _) _ _ _ _]. pose proof (typing_wf_local Hs). destruct cdecl as [id cty car]. apply subject_closed in Hs; eauto. @@ -623,7 +618,7 @@ Lemma on_constructor_closed {cf:checker_flags} {Σ : global_env} {wfΣ : wf Σ} (subst_instance u cdecl.(cstr_type)) in closed cty. Proof. - intros [? ? [s Hs] _ _ _ _]. + intros [? ? (_ & s & Hs & _) _ _ _ _]. pose proof (typing_wf_local Hs). destruct cdecl as [id cty car]. apply subject_closed in Hs; eauto. @@ -1228,24 +1223,24 @@ Proof. rewrite -{3}H. now rewrite simpl_subst_k. Qed. -Lemma subst_compare_term {cf:checker_flags} le Σ (φ : ConstraintSet.t) (l : list term) (k : nat) (T U : term) : - compare_term le Σ φ T U -> compare_term le Σ φ (subst l k T) (subst l k U). +Lemma subst_compare_term {cf:checker_flags} Σ (φ : ConstraintSet.t) pb (l : list term) (k : nat) (T U : term) : + compare_term Σ φ pb T U -> compare_term Σ φ pb (subst l k T) (subst l k U). Proof. - destruct le; simpl. + destruct pb; simpl. - apply subst_eq_term. - apply subst_leq_term. Qed. -Lemma subst_compare_decl `{checker_flags} {le Σ ϕ l k d d'} : - compare_decl le Σ ϕ d d' -> compare_decl le Σ ϕ (subst_decl l k d) (subst_decl l k d'). +Lemma subst_compare_decl `{checker_flags} {Σ ϕ pb l k d d'} : + compare_decl Σ ϕ pb d d' -> compare_decl Σ ϕ pb (subst_decl l k d) (subst_decl l k d'). Proof. - intros []; constructor; auto; destruct le; + intros []; constructor; auto; destruct pb; intuition eauto using subst_compare_term, subst_eq_term, subst_leq_term. Qed. -Lemma subst_compare_context `{checker_flags} le Σ φ l l' n k : - compare_context le Σ φ l l' -> - compare_context le Σ φ (subst_context n k l) (subst_context n k l'). +Lemma subst_compare_context `{checker_flags} Σ φ pb l l' n k : + compare_context Σ φ pb l l' -> + compare_context Σ φ pb (subst_context n k l) (subst_context n k l'). Proof. induction 1; rewrite ?subst_context_snoc /=; constructor; auto. erewrite (All2_fold_length X). simpl. @@ -1279,17 +1274,17 @@ Section CtxReduction. Qed. End CtxReduction. -Record wt_cumul_pb {cf} (c : conv_pb) Σ (Γ : context) T U := +Record wt_cumul_pb {cf} (pb : conv_pb) Σ (Γ : context) T U := { wt_cumul_pb_dom : isType Σ Γ T; wt_cumul_pb_codom : isType Σ Γ U; - wt_cumul_pb_eq : cumulAlgo_gen Σ Γ c T U }. + wt_cumul_pb_eq : cumulAlgo_gen Σ Γ pb T U }. -Arguments wt_cumul_pb_dom {cf c Σ Γ T U}. -Arguments wt_cumul_pb_codom {cf c Σ Γ T U}. -Arguments wt_cumul_pb_eq {cf c Σ Γ T U}. +Arguments wt_cumul_pb_dom {cf pb Σ Γ T U}. +Arguments wt_cumul_pb_codom {cf pb Σ Γ T U}. +Arguments wt_cumul_pb_eq {cf pb Σ Γ T U}. -Definition wt_cumul {cf} := wt_cumul_pb Cumul. -Definition wt_conv {cf} := wt_cumul_pb Conv. +Definition wt_cumul {cf} Σ Γ := wt_cumul_pb Cumul Σ Γ. +Definition wt_conv {cf} Σ Γ := wt_cumul_pb Conv Σ Γ. Notation " Σ ;;; Γ |- t <= u ✓" := (wt_cumul Σ Γ t u) (at level 50, Γ, t, u at next level). Notation " Σ ;;; Γ |- t = u ✓" := (wt_conv Σ Γ t u) (at level 50, Γ, t, u at next level). @@ -1318,10 +1313,10 @@ Reserved Notation " Σ ;;; Γ |-[ P ] t <=[ le ] u" (at level 50, Γ, t, u at ne format "Σ ;;; Γ |-[ P ] t <=[ le ] u"). Inductive cumulP {cf} (pb : conv_pb) (Σ : global_env_ext) (P : nat -> bool) (Γ : context) : term -> term -> Type := -| wt_cumul_refl t u : compare_term pb Σ.1 (global_ext_constraints Σ) t u -> Σ ;;; Γ |-[P] t <=[pb] u +| wt_cumul_refl t u : compare_term Σ.1 (global_ext_constraints Σ) pb t u -> Σ ;;; Γ |-[P] t <=[pb] u | wt_cumul_red_l t u v : red1P P Σ Γ t v -> Σ ;;; Γ |-[P] v <=[pb] u -> Σ ;;; Γ |-[P] t <=[pb] u | wt_cumul_red_r t u v : Σ ;;; Γ |-[P] t <=[pb] v -> red1P P Σ Γ u v -> Σ ;;; Γ |-[P] t <=[pb] u -where " Σ ;;; Γ |-[ P ] t <=[ le ] u " := (cumulP le Σ P Γ t u) : type_scope. +where " Σ ;;; Γ |-[ P ] t <=[ pb ] u " := (cumulP pb Σ P Γ t u) : type_scope. Notation " Σ ;;; Γ |-[ P ] t <= u " := (cumulP Cumul Σ P Γ t u) (at level 50, Γ, t, u at next level, format "Σ ;;; Γ |-[ P ] t <= u") : type_scope. @@ -1331,7 +1326,7 @@ Notation " Σ ;;; Γ |-[ P ] t = u " := (cumulP Conv Σ P Γ t u) (at level 50, Lemma isType_wf_local {cf:checker_flags} {Σ Γ T} : isType Σ Γ T -> wf_local Σ Γ. Proof. - move=> [s Hs]. + intros (_ & s & Hs & _). now eapply typing_wf_local. Qed. @@ -1350,34 +1345,23 @@ Section SubstitutionLemmas. now rewrite shiftnPF_closedPT. Qed. - Lemma wt_cumul_pb_equalityP {le} {Γ : context} {T U} : wt_cumul_pb le Σ Γ T U -> cumulP le Σ (closedP #|Γ| xpredT) Γ T U. + Lemma wt_cumul_pb_equalityP {pb} {Γ : context} {T U} : wt_cumul_pb pb Σ Γ T U -> cumulP pb Σ (closedP #|Γ| xpredT) Γ T U. Proof using wfΣ. move=> [] dom. move: (isType_wf_local dom) => /closed_wf_local clΓ. rewrite closed_ctx_on_ctx_free_vars in clΓ; tea. move/isType_closedPT: dom => clT. move/isType_closedPT => clU cum. - destruct le. - { induction cum. - - constructor; auto. - - econstructor 2. - * econstructor; [|split]; tea. - * eapply IHcum => //. - now eapply red1_on_free_vars in r. - - econstructor 3. - * eapply IHcum => //. - now eapply red1_on_free_vars. - * econstructor; [|split]; tea. } - { induction cum. - - constructor; auto. - - econstructor 2. - * econstructor; [|split]; tea. - * eapply IHcum => //. - now eapply red1_on_free_vars in r. - - econstructor 3. - * eapply IHcum => //. - now eapply red1_on_free_vars. - * econstructor; [|split]; tea. } + induction cum. + - constructor; auto. + - econstructor 2. + * econstructor; [|split]; tea. + * eapply IHcum => //. + now eapply red1_on_free_vars in r. + - econstructor 3. + * eapply IHcum => //. + now eapply red1_on_free_vars. + * econstructor; [|split]; tea. Qed. Lemma substitution_red1 {Γ Γ' Γ'' s M N} : @@ -1752,11 +1736,10 @@ Qed. Qed. Lemma All_local_env_inst {Γ0 Γ' Δ s} : - All_local_env - (lift_typing - (fun (Σ : global_env_ext) (Γ : context) (t T : term) => - forall (Δ : PCUICEnvironment.context) (σ : nat -> term), - wf_local Σ Δ -> well_subst Σ Γ σ Δ -> Σ;;; Δ |- t.[σ] : T.[σ]) Σ) + All_local_env (fun Γ j => + forall (Δ : PCUICEnvironment.context) (σ : nat -> term), + wf_local Σ Δ -> well_subst Σ Γ σ Δ -> + (lift_typing0 (fun t T => Σ;;; Δ |- t.[σ] : T.[σ])) j) (Γ0,,, Γ',,, Δ) -> wf_local Σ Γ0 -> subslet Σ Γ0 s Γ' -> @@ -1779,6 +1762,11 @@ Theorem substitution_prop {cf} : env_prop (sub : subslet Σ Γ s Γ') (eqΓ0 : Γ0 = Γ ,,, Γ' ,,, Δ), wf_local Σ (Γ ,,, subst_context s 0 Δ) -> Σ ;;; Γ ,,, subst_context s 0 Δ |- subst s #|Δ| t : subst s #|Δ| T) + (fun Σ Γ0 j => + forall (Γ Γ' Δ : context) (s : list term) + (sub : subslet Σ Γ s Γ') (eqΓ0 : Γ0 = Γ ,,, Γ' ,,, Δ), + wf_local Σ (Γ ,,, subst_context s 0 Δ) -> + lift_typing0 (fun t T => Σ ;;; Γ ,,, subst_context s 0 Δ |- subst s #|Δ| t : subst s #|Δ| T) j) (fun Σ Γ0 => forall (Γ Γ' Δ : context) (s : list term) (sub : subslet Σ Γ s Γ') (eqΓ0 : Γ0 = Γ ,,, Γ' ,,, Δ), @@ -1793,14 +1781,14 @@ Proof. eapply subslet_well_subst; eauto. } 2:{ subst Γ. eapply typing_wf_local in HT. - eapply wf_local_app_inv in HT as [HΓ0 _]. - eapply wf_local_app_inv in HΓ0 as [HΓ0 _]. + eapply All_local_env_app_inv in HT as [HΓ0 _]. + eapply All_local_env_app_inv in HΓ0 as [HΓ0 _]. eapply All_local_env_inst; eauto. } unshelve eapply on_wf_global_env_impl ; tea. clear. intros * HΣ HP HQ Hty. - apply lift_typing_impl with (1 := Hty); clear -HΣ HP. intros. subst Γ. - rewrite !subst_inst. eapply X => //. + unshelve eapply lift_typing_impl with (1 := Hty _ _ X _) => t T HT. + 3: rewrite !subst_inst; eapply HT => //. now unshelve eapply subslet_well_subst. Qed. @@ -1820,8 +1808,9 @@ Corollary isType_substitution {cf} {Σ} {wfΣ : wf Σ} {Γ Γ' s Δ T} : isType Σ (Γ ,,, Γ' ,,, Δ) T -> isType Σ (Γ ,,, subst_context s 0 Δ) (subst s #|Δ| T). Proof. - intros Hs [s' Ht]. - eapply substitution in Ht; tea. now eexists. + intros Hs HT. + apply lift_typing_f_impl with (1 := HT) => // ?? Ht. + eapply substitution in Ht; tea. Qed. Corollary substitution_wf_local {cf} {Σ} {wfΣ : wf Σ} {Γ Γ' s Δ} : @@ -1851,7 +1840,7 @@ Proof. intros Ht. assert ((wf_local Σ Γ) * (Σ ;;; Γ |- u : U)%type) as [wfΓ tyu]. { apply typing_wf_local in Ht; eauto with wf. - now depelim Ht; simpl in *. + depelim Ht; split; tas. now destruct l as (? & _). } eapply (substitution (Γ':=[vdef n u U]) (Δ := [])); tea. now eapply subslet_def_tip. @@ -1859,10 +1848,10 @@ Qed. (** Substitution into a *well-typed* cumulativity/conversion derivation. *) -Lemma substitution_wt_cumul_pb {cf} {Σ} {wfΣ : wf Σ} {le Γ Γ' Γ'' s M N} : +Lemma substitution_wt_cumul_pb {cf} {Σ} {wfΣ : wf Σ} {pb Γ Γ' Γ'' s M N} : subslet Σ Γ s Γ' -> - wt_cumul_pb le Σ (Γ ,,, Γ' ,,, Γ'') M N -> - wt_cumul_pb le Σ (Γ ,,, subst_context s 0 Γ'') (subst s #|Γ''| M) (subst s #|Γ''| N). + wt_cumul_pb pb Σ (Γ ,,, Γ' ,,, Γ'') M N -> + wt_cumul_pb pb Σ (Γ ,,, subst_context s 0 Γ'') (subst s #|Γ''| M) (subst s #|Γ''| N). Proof. move=> Hs wteq; split. + eapply (isType_substitution Hs), wteq. @@ -1871,13 +1860,13 @@ Proof. - intros t u cmp. constructor. now eapply subst_compare_term. - move=> t u v red cum. - destruct le. + destruct pb. * eapply red_conv_conv. eapply substitution_let_red; tea; eauto with wf; apply red. * eapply red_cumul_cumul. eapply substitution_let_red; tea; eauto with wf; apply red. - move=> t u v red cum red'. - destruct le. + destruct pb. * eapply red_conv_conv_inv; tea. eapply substitution_let_red; tea; apply red'. * eapply red_cumul_cumul_inv; tea. diff --git a/pcuic/theories/PCUICTypedAst.v b/pcuic/theories/PCUICTypedAst.v index 92bba0678..c8d97d745 100644 --- a/pcuic/theories/PCUICTypedAst.v +++ b/pcuic/theories/PCUICTypedAst.v @@ -82,7 +82,7 @@ Inductive term {k : nat} : Type := | tRel (f : Fin.t k) | tVar (i : ident) (* For free variables (e.g. in a goal) *) | tEvar (n : nat) (l : list term) -| tSort (u : Universe.t) +| tSort (u : sort) | tProd (na : aname) (A : term) (B : @term Σ (S k)) | tLambda (na : aname) (A : term) (B : @term Σ (S k)) | tLetIn (na : aname) (b B : term) (t : @term Σ (S k)) diff --git a/pcuic/theories/PCUICTyping.v b/pcuic/theories/PCUICTyping.v index 456d91eef..c951c51f5 100644 --- a/pcuic/theories/PCUICTyping.v +++ b/pcuic/theories/PCUICTyping.v @@ -154,9 +154,9 @@ Variant case_side_conditions `{checker_flags} wf_local_fun typing Σ Γ ci p ps global environment *) (conv_pctx : eq_context_upto_names p.(pcontext) (ind_predicate_context ci.(ci_ind) mdecl idecl)) (allowed_elim : is_allowed_elimination Σ idecl.(ind_kelim) ps) - (ind_inst : ctx_inst typing Σ Γ (p.(pparams) ++ indices) + (ind_inst : ctx_inst (typing Σ) Γ (p.(pparams) ++ indices) (List.rev (subst_instance p.(puinst) - (ind_params mdecl ,,, ind_indices idecl)))) + (ind_params mdecl ,,, ind_indices idecl : context)))) (not_cofinite : isCoFinite mdecl.(ind_finite) = false). Variant case_branch_typing `{checker_flags} wf_local_fun typing Σ Γ (ci:case_info) p ps mdecl idecl ptm brs := @@ -179,8 +179,8 @@ Variant primitive_typing_hyps `{checker_flags} | prim_int_hyps i : primitive_typing_hyps typing Σ Γ (primInt; primIntModel i) | prim_float_hyps f : primitive_typing_hyps typing Σ Γ (primFloat; primFloatModel f) | prim_array_hyps a - (wfl : wf_universe Σ (Universe.make a.(array_level))) - (hty : typing Σ Γ a.(array_type) (tSort (Universe.make a.(array_level)))) + (wfl : wf_universe Σ (Universe.make' a.(array_level))) + (hty : typing Σ Γ a.(array_type) (tSort (sType (Universe.make' a.(array_level))))) (hdef : typing Σ Γ a.(array_default) a.(array_type)) (hvalue : All (fun x => typing Σ Γ x a.(array_type)) a.(array_value)) : primitive_typing_hyps typing Σ Γ (primArray; primArrayModel a). @@ -200,22 +200,21 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> | type_Sort : forall s, wf_local Σ Γ -> - wf_universe Σ s -> - Σ ;;; Γ |- tSort s : tSort (Universe.super s) + wf_sort Σ s -> + Σ ;;; Γ |- tSort s : tSort (Sort.super s) | type_Prod : forall na A B s1 s2, - Σ ;;; Γ |- A : tSort s1 -> + lift_typing typing Σ Γ (j_vass_s na A s1) -> Σ ;;; Γ ,, vass na A |- B : tSort s2 -> - Σ ;;; Γ |- tProd na A B : tSort (Universe.sort_of_product s1 s2) + Σ ;;; Γ |- tProd na A B : tSort (Sort.sort_of_product s1 s2) -| type_Lambda : forall na A t s1 B, - Σ ;;; Γ |- A : tSort s1 -> +| type_Lambda : forall na A t B, + lift_typing typing Σ Γ (j_vass na A) -> Σ ;;; Γ ,, vass na A |- t : B -> Σ ;;; Γ |- tLambda na A t : tProd na A B -| type_LetIn : forall na b B t s1 A, - Σ ;;; Γ |- B : tSort s1 -> - Σ ;;; Γ |- b : B -> +| type_LetIn : forall na b B t A, + lift_typing typing Σ Γ (j_vdef na b B) -> Σ ;;; Γ ,, vdef na b B |- t : A -> Σ ;;; Γ |- tLetIn na b B t : tLetIn na b B A @@ -267,8 +266,8 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> wf_local Σ Γ -> fix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> - All (fun d => {s & Σ ;;; Γ |- d.(dtype) : tSort s}) mfix -> - All (fun d => (Σ ;;; Γ ,,, fix_context mfix |- d.(dbody) : lift0 #|fix_context mfix| d.(dtype))) mfix -> + All (on_def_type (lift_typing1 (typing Σ)) Γ) mfix -> + All (on_def_body (lift_typing1 (typing Σ)) (fix_context mfix) Γ) mfix -> wf_fixpoint Σ mfix -> Σ ;;; Γ |- tFix mfix n : decl.(dtype) @@ -276,18 +275,18 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> wf_local Σ Γ -> cofix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> - All (fun d => {s & Σ ;;; Γ |- d.(dtype) : tSort s}) mfix -> - All (fun d => Σ ;;; Γ ,,, fix_context mfix |- d.(dbody) : lift0 #|fix_context mfix| d.(dtype)) mfix -> + All (on_def_type (lift_typing1 (typing Σ)) Γ) mfix -> + All (on_def_body (lift_typing1 (typing Σ)) (fix_context mfix) Γ) mfix -> wf_cofixpoint Σ mfix -> Σ ;;; Γ |- tCoFix mfix n : decl.(dtype) | type_Prim p prim_ty cdecl : - wf_local Σ Γ -> - primitive_constant Σ (prim_val_tag p) = Some prim_ty -> - declared_constant Σ prim_ty cdecl -> - primitive_invariants (prim_val_tag p) cdecl -> - primitive_typing_hyps typing Σ Γ p -> - Σ ;;; Γ |- tPrim p : prim_type p prim_ty + wf_local Σ Γ -> + primitive_constant Σ (prim_val_tag p) = Some prim_ty -> + declared_constant Σ prim_ty cdecl -> + primitive_invariants (prim_val_tag p) cdecl -> + primitive_typing_hyps typing Σ Γ p -> + Σ ;;; Γ |- tPrim p : prim_type p prim_ty | type_Cumul : forall t A B s, Σ ;;; Γ |- t : A -> @@ -296,7 +295,7 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> Σ ;;; Γ |- t : B where " Σ ;;; Γ |- t : T " := (typing Σ Γ t T) -and "'wf_local' Σ Γ " := (All_local_env (lift_typing typing Σ) Γ). +and "'wf_local' Σ Γ " := (All_local_env (lift_typing1 (typing Σ)) Γ). Lemma meta_conv {cf} {Σ Γ t A B} : Σ ;;; Γ |- t : A -> @@ -358,10 +357,17 @@ Arguments iswelltyped {cf Σ Γ t A} h. Definition isType_welltyped {cf Σ} {Γ T} : isType Σ Γ T -> welltyped Σ Γ T. Proof. - intros []. now econstructor. + intros [_ [s []]]. now econstructor. Qed. Global Hint Resolve isType_welltyped : pcuic. +Definition has_sort_isType {cf Σ} {Γ T} s + : Σ ;;; Γ |- T : tSort s -> isType Σ Γ T. +Proof. + repeat (eexists; tea). +Defined. +Global Hint Resolve has_sort_isType : pcuic. + Definition isWfArity {cf:checker_flags} Σ (Γ : context) T := (isType Σ Γ T × { ctx & { s & (destArity [] T = Some (ctx, s)) } }). @@ -380,14 +386,14 @@ Definition branches_size {cf} {Σ Γ ci mdecl idecl p ps ptm brs} (a : tybranches Σ Γ ci mdecl idecl p ps ptm n ctors brs) : size := (all2i_size _ (fun i x y p => Nat.max - (All_local_env_size typing_size _ _ p.2.1) + (All_local_env_size (typing_size _) _ p.2.1) (Nat.max (typing_size _ _ _ _ p.2.2.1) (typing_size _ _ _ _ p.2.2.2))) a). Section CtxInstSize. Context {cf} (typing : global_env_ext -> context -> term -> term -> Type) (typing_size : forall {Σ Γ t T}, typing Σ Γ t T -> size). - Fixpoint ctx_inst_size {Σ Γ args Δ} (c : ctx_inst typing Σ Γ args Δ) : size := + Fixpoint ctx_inst_size {Σ Γ args Δ} (c : ctx_inst (typing Σ) Γ args Δ) : size := match c with | ctx_inst_nil => 0 | ctx_inst_ass na t i inst Δ ty ctxi => (typing_size _ _ _ _ ty) + (ctx_inst_size ctxi) @@ -403,12 +409,12 @@ Section PrimitiveSize. destruct h. - exact 0. - exact 0. - - exact ((typing_size _ _ _ _ hty) + (typing_size _ _ _ _ hdef) + - all_size _ (fun x p => typing_size _ _ _ _ p) hvalue). + - exact (Nat.max (typing_size _ _ _ _ hty) (Nat.max (typing_size _ _ _ _ hdef) + (all_size _ (fun x p => typing_size _ _ _ _ p) hvalue))). Defined. End PrimitiveSize. -Definition typing_size `{checker_flags} {Σ Γ t T} (d : Σ ;;; Γ |- t : T) : size. +Definition typing_size `{cf : checker_flags} {Σ Γ t T} (d : Σ ;;; Γ |- t : T) : size. Proof. revert Σ Γ t T d. fix typing_size 5. @@ -422,24 +428,30 @@ Proof. | H : All_local_env _ _ |- _ => idtac | H : All _ _ |- _ => idtac | H : _ + _ |- _ => idtac + | H : lift_typing0 _ _ |- _ => idtac | H1 : size, H2 : size, H3 : size |- _ => exact (S (Nat.max H1 (Nat.max H2 H3))) | H1 : size, H2 : size |- _ => exact (S (Nat.max H1 H2)) | H1 : size |- _ => exact (S H1) | _ => exact 1 end. - - exact (S (All_local_env_size typing_size _ _ a)). - - exact (S (All_local_env_size typing_size _ _ a)). - - exact (S (S (All_local_env_size typing_size _ _ a))). - - exact (S (S (All_local_env_size typing_size _ _ a))). - - exact (S (S (All_local_env_size typing_size _ _ a))). - - exact (S (Nat.max (All_local_env_size typing_size _ _ wf_pctx) + - exact (S (All_local_env_size (typing_size _) _ a)). + - exact (S (All_local_env_size (typing_size _) _ a)). + - exact (S (Nat.max d (lift_typing_size (typing_size _ _) _ l))). + - exact (S (Nat.max d (lift_typing_size (typing_size _ _) _ l))). + - exact (S (Nat.max d (lift_typing_size (typing_size _ _) _ l))). + - exact (S (S (All_local_env_size (typing_size _) _ a))). + - exact (S (S (All_local_env_size (typing_size _) _ a))). + - exact (S (S (All_local_env_size (typing_size _) _ a))). + - exact (S (Nat.max (All_local_env_size (typing_size _) _ wf_pctx) (Nat.max (ctx_inst_size _ typing_size ind_inst) (Nat.max d2 (Nat.max d3 (branches_size typing_size brs_ty)))))). - - exact (S (Nat.max (Nat.max (All_local_env_size typing_size _ _ a) - (all_size _ (fun x p => (infer_sort_size (typing_sort_size typing_size)) Σ _ _ p) a0)) (all_size _ (fun x p => typing_size Σ _ _ _ p) a1))). - - exact (S (Nat.max (Nat.max (All_local_env_size typing_size _ _ a) - (all_size _ (fun x p => (infer_sort_size (typing_sort_size typing_size)) Σ _ _ p) a0)) (all_size _ (fun x p => typing_size Σ _ _ _ p) a1))). - - exact (S (Nat.max (All_local_env_size typing_size _ _ a) (primitive_typing_hyps_size typing typing_size _ _ _ p1))). + - exact (S (Nat.max + (Nat.max (All_local_env_size (typing_size _) _ a) (all_size _ (fun x p => on_def_type_size (typing_size Σ) _ _ p) a0)) + (all_size (on_def_body (lift_typing typing _) _ _) (fun x p => on_def_body_size (typing_size Σ) _ _ _ p) a1))). + - exact (S (Nat.max + (Nat.max (All_local_env_size (typing_size _) _ a) (all_size _ (fun x p => on_def_type_size (typing_size Σ) _ _ p) a0)) + (all_size (on_def_body (lift_typing typing _) _ _) (fun x p => on_def_body_size (typing_size Σ) _ _ _ p) a1))). + - exact (S (Nat.max (All_local_env_size (typing_size _) _ a) (primitive_typing_hyps_size typing typing_size _ _ _ p1))). Defined. Lemma typing_size_pos `{checker_flags} {Σ Γ t T} (d : Σ ;;; Γ |- t : T) : typing_size d > 0. @@ -465,7 +477,7 @@ Arguments lexprod [A B]. (** We make these well-formedness conditions type-classes as they are genrally globally available. *) -Definition wf `{checker_flags} := Forall_decls_typing typing. +Definition wf `{checker_flags} := on_global_env cumulSpec0 (lift_typing typing). Existing Class wf. #[global] Hint Mode wf + + : typeclass_intances. @@ -500,10 +512,20 @@ Proof. Defined. #[global] Hint Immediate wf_local_app_l : wf. -Lemma typing_wf_local `{checker_flags} {Σ} {Γ t T} : +Lemma typing_wf_local `{cf : checker_flags} {Σ} {Γ t T} : Σ ;;; Γ |- t : T -> wf_local Σ Γ. Proof. - induction 1; eauto using wf_local_app_l. + revert Σ Γ t T. + fix f 5. + destruct 1; eauto. + all: exact (f _ _ _ _ l.2.π2.1). +Defined. + +Lemma lift_typing_wf_local `{cf : checker_flags} {Σ} {Γ j} : + lift_typing typing Σ Γ j -> wf_local Σ Γ. +Proof. + destruct 1 as (_ & s & Hs & _). + now apply typing_wf_local in Hs. Defined. #[global] @@ -511,39 +533,43 @@ Hint Extern 4 (wf_local _ ?Γ) => match goal with | [ H : typing _ _ _ _ |- _ ] => exact (typing_wf_local H) | [ H : PCUICTypingDef.typing _ _ _ _ _ |- _ ] => exact (typing_wf_local H) + | [ H : lift_typing0 (typing _ _) _ |- _ ] => exact (lift_typing_wf_local H) + | [ H : PCUICEnvTyping.lift_typing (typing _ _) _ |- _ ] => exact (lift_typing_wf_local H) end : pcuic. #[global] Hint Resolve typing_wf_local : wf. -Definition env_prop `{checker_flags} (P : forall Σ Γ t T, Type) (PΓ : forall Σ Γ, Type) := +#[global] +Hint Resolve lift_typing_wf_local : wf. + +Definition env_prop `{checker_flags} (P : forall Σ Γ t T, Type) (Pj : forall Σ Γ j, Type) (PΓ : forall Σ Γ, Type) := forall Σ (wfΣ : wf Σ.1) Γ t T (ty : Σ ;;; Γ |- t : T), - Forall_decls_typing P Σ.1 * - (PΓ Σ Γ * P Σ Γ t T). + on_global_env cumulSpec0 Pj Σ.1 * (PΓ Σ Γ * P Σ Γ t T). -Lemma env_prop_typing `{checker_flags} {P PΓ} : env_prop P PΓ -> +Lemma env_prop_typing `{checker_flags} {P Pj PΓ} : env_prop P Pj PΓ -> forall Σ (wfΣ : wf Σ.1) (Γ : context) (t T : term), Σ ;;; Γ |- t : T -> P Σ Γ t T. Proof. intros. now apply X. Qed. -Lemma type_Prop_wf `{checker_flags} Σ Γ : wf_local Σ Γ -> Σ ;;; Γ |- tSort Universe.lProp : tSort Universe.type1. +Lemma type_Prop_wf `{checker_flags} Σ Γ : wf_local Σ Γ -> Σ ;;; Γ |- tSort sProp : tSort Sort.type1. Proof. repeat constructor; auto. Defined. -Lemma env_prop_wf_local `{checker_flags} {P PΓ} : env_prop P PΓ -> +Lemma env_prop_wf_local `{checker_flags} {P Pj PΓ} : env_prop P Pj PΓ -> forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ), PΓ Σ Γ. Proof. intros. pose (type_Prop_wf _ _ wfΓ). now destruct (X _ wfΣ _ _ _ t) as [? [? ?]]. Qed. -Lemma type_Prop `{checker_flags} Σ : Σ ;;; [] |- tSort Universe.lProp : tSort Universe.type1. +Lemma type_Prop `{checker_flags} Σ : Σ ;;; [] |- tSort sProp : tSort Sort.type1. repeat constructor. Defined. -Lemma env_prop_sigma `{checker_flags} {P PΓ} : env_prop P PΓ -> - forall (Σ : global_env) (wfΣ : wf Σ), Forall_decls_typing P Σ. +Lemma env_prop_sigma `{checker_flags} {P Pj PΓ} : env_prop P Pj PΓ -> + forall (Σ : global_env) (wfΣ : wf Σ), on_global_env cumulSpec0 Pj Σ. Proof. intros. red in X. eapply (X (empty_ext Σ)). apply wfΣ. @@ -556,13 +582,13 @@ Lemma type_Cumul' {cf:checker_flags} {Σ Γ t} T {T'} : Σ ;;; Γ |- T <=s T' -> Σ ;;; Γ |- t : T'. Proof. - intros Ht [s Hs] cum. + intros Ht (_ & s & Hs & _) cum. eapply type_Cumul; eauto. Qed. Lemma size_wf_local_app `{checker_flags} {Σ} (Γ Γ' : context) (Hwf : wf_local Σ (Γ ,,, Γ')) : - All_local_env_size (@typing_size _) Σ _ (wf_local_app_l _ _ _ Hwf) <= - All_local_env_size (@typing_size _) Σ _ Hwf. + All_local_env_size (@typing_size _ Σ) _ (wf_local_app_l _ _ _ Hwf) <= + All_local_env_size (@typing_size _ Σ) _ Hwf. Proof. induction Γ' in Γ, Hwf |- *; try lia. simpl. lia. depelim Hwf. @@ -570,52 +596,33 @@ Proof. - specialize (IHΓ' _ Hwf). simpl. unfold eq_rect_r. cbn. lia. Qed. -Lemma typing_wf_local_size `{checker_flags} {Σ} {Γ t T} +Lemma typing_wf_local_size `{cf : checker_flags} {Σ} {Γ t T} (d :Σ ;;; Γ |- t : T) : - All_local_env_size (@typing_size _) Σ _ (typing_wf_local d) < typing_size d. + All_local_env_size (@typing_size _ _) _ (typing_wf_local d) < typing_size d. Proof. + revert Σ Γ t T d. fix f 5. induction d; try destruct c0, c1; simpl; - change (fun (x : global_env_ext) (x0 : context) (x1 x2 : term) - (x3 : x;;; x0 |- x1 : x2) => typing_size x3) with (@typing_size H); try lia. + change (fun Γ t T (Hty : Σ;;; Γ |- t : T) => typing_size Hty) with (@typing_size cf Σ); try lia. + all: pose proof (f _ _ _ _ l.2.π2.1); unfold lift_sorting_size; cbn in *; try lia. Qed. Lemma wf_local_inv `{checker_flags} {Σ Γ'} (w : wf_local Σ Γ') : forall d Γ, Γ' = d :: Γ -> - ∑ w' : wf_local Σ Γ, - match d.(decl_body) with - | Some b => - ∑ u (ty : Σ ;;; Γ |- b : d.(decl_type)), - { ty' : Σ ;;; Γ |- d.(decl_type) : tSort u | - All_local_env_size (@typing_size _) Σ _ w' < - All_local_env_size (@typing_size _) _ _ w /\ - typing_size ty <= All_local_env_size (@typing_size _) _ _ w /\ - typing_size ty' <= All_local_env_size (@typing_size _) _ _ w } - - | None => - ∑ u, - { ty : Σ ;;; Γ |- d.(decl_type) : tSort u | - All_local_env_size (@typing_size _) Σ _ w' < - All_local_env_size (@typing_size _) _ _ w /\ - typing_size ty <= All_local_env_size (@typing_size _) _ _ w } - end. + ∑ (w' : wf_local Σ Γ) u, + { ty : lift_typing1 (typing Σ) Γ (Judge d.(decl_body) d.(decl_type) (Some u)) | + All_local_env_size (@typing_size _ Σ) _ w' < + All_local_env_size (@typing_size _ _) _ w /\ + lift_typing_size (@typing_size _ _ _) _ ty <= All_local_env_size (@typing_size _ _) _ w }. Proof. - intros d Γ. - destruct w. - - simpl. congruence. - - intros [=]. subst d Γ0. - exists w. simpl. destruct l. exists x. exists t0. pose (typing_size_pos t0). - cbn. split. - + lia. - + auto with arith. - - intros [=]. subst d Γ0. - exists w. simpl. simpl in l. destruct l as [u h]. - simpl in l0. - exists u, l0, h. cbn. - pose (typing_size_pos h). - pose (typing_size_pos l0). - intuition eauto. - all: try lia. + intros d Γ ->. + depelim w; cbn. + all: exists w, l.2.π1, (lift_sorting_extract l). + all: pose proof (typing_size_pos l.2.π2.1). + 2: pose proof (typing_size_pos l.1). + all: simpl in *. + all: simpl; split. + all: try lia. Qed. Lemma All_map_size {A} {P Q : A -> Type} (sizep : forall x, P x -> size) l @@ -640,45 +647,45 @@ Qed. Lemma typing_ind_env_app_size `{cf : checker_flags} : forall (P : global_env_ext -> context -> term -> term -> Type) + (Pj : global_env_ext -> context -> judgment -> Type) (Pdecl := fun Σ Γ wfΓ t T tyT => P Σ Γ t T) (PΓ : global_env_ext -> context -> Type), + (forall Σ (wfΣ : wf Σ) (Γ : context) (* (wfΓ : wf_local Σ Γ) *) j, + lift_typing_conj (typing Σ) (P Σ) Γ j -> Pj Σ Γ j) -> + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ), - All_local_env_over typing Pdecl Σ Γ wfΓ -> PΓ Σ Γ) -> + All_local_env_over (typing Σ) (Pdecl Σ) Γ wfΓ -> All_local_env (Pj Σ) Γ -> PΓ Σ Γ) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : nat) decl, nth_error Γ n = Some decl -> PΓ Σ Γ -> P Σ Γ (tRel n) (lift0 (S n) decl.(decl_type))) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (u : Universe.t), + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (u : sort), PΓ Σ Γ -> - wf_universe Σ u -> - P Σ Γ (tSort u) (tSort (Universe.super u))) -> + wf_sort Σ u -> + P Σ Γ (tSort u) (tSort (Sort.super u))) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : aname) (t b : term) (s1 s2 : Universe.t), + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (t b : term) (s1 s2 : sort), PΓ Σ Γ -> - Σ ;;; Γ |- t : tSort s1 -> - P Σ Γ t (tSort s1) -> - Σ ;;; Γ,, vass n t |- b : tSort s2 -> - P Σ (Γ,, vass n t) b (tSort s2) -> P Σ Γ (tProd n t b) (tSort (Universe.sort_of_product s1 s2))) -> + lift_typing typing Σ Γ (j_vass_s na t s1) -> + Pj Σ Γ (j_vass_s na t s1) -> + Σ ;;; Γ,, vass na t |- b : tSort s2 -> + P Σ (Γ,, vass na t) b (tSort s2) -> P Σ Γ (tProd na t b) (tSort (Sort.sort_of_product s1 s2))) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : aname) (t b : term) - (s1 : Universe.t) (bty : term), + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (t b bty : term), PΓ Σ Γ -> - Σ ;;; Γ |- 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)) -> + lift_typing typing Σ Γ (j_vass na t) -> + Pj Σ Γ (j_vass na t) -> + Σ ;;; Γ,, vass na t |- b : bty -> P Σ (Γ,, vass na t) b bty -> P Σ Γ (tLambda na t b) (tProd na t bty)) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : aname) (b b_ty b' : term) - (s1 : Universe.t) (b'_ty : term), + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (b b_ty b' b'_ty : term), PΓ Σ Γ -> - Σ ;;; Γ |- b_ty : tSort s1 -> - P Σ Γ b_ty (tSort s1) -> - Σ ;;; Γ |- b : b_ty -> - P Σ Γ b b_ty -> - Σ ;;; Γ,, vdef n b b_ty |- b' : b'_ty -> - P Σ (Γ,, vdef n b b_ty) b' b'_ty -> P Σ Γ (tLetIn n b b_ty b') (tLetIn n b b_ty b'_ty)) -> + lift_typing typing Σ Γ (j_vdef na b b_ty) -> + Pj Σ Γ (j_vdef na b b_ty) -> + Σ ;;; Γ,, vdef na b b_ty |- b' : b'_ty -> + P Σ (Γ,, vdef na b b_ty) b' b'_ty -> P Σ Γ (tLetIn na b b_ty b') (tLetIn na b b_ty b'_ty)) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (t : term) na A B u s, PΓ Σ Γ -> @@ -692,7 +699,7 @@ Lemma typing_ind_env_app_size `{cf : checker_flags} : P Σ Γ (tApp t u) (B{0 := u})) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) cst u (decl : constant_body), - Forall_decls_typing P Σ.1 -> + on_global_env cumulSpec0 Pj Σ.1 -> PΓ Σ Γ -> declared_constant Σ.1 cst decl -> consistent_instance_ext Σ decl.(cst_universes) u -> @@ -700,14 +707,14 @@ Lemma typing_ind_env_app_size `{cf : checker_flags} : (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (ind : inductive) u mdecl idecl (isdecl : declared_inductive Σ.1 ind mdecl idecl), - Forall_decls_typing P Σ.1 -> + on_global_env cumulSpec0 Pj Σ.1 -> PΓ Σ Γ -> consistent_instance_ext Σ mdecl.(ind_universes) u -> P Σ Γ (tInd ind u) (subst_instance u (ind_type idecl))) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (ind : inductive) (i : nat) u mdecl idecl cdecl (isdecl : declared_constructor Σ.1 (ind, i) mdecl idecl cdecl), - Forall_decls_typing P Σ.1 -> + on_global_env cumulSpec0 Pj Σ.1 -> PΓ Σ Γ -> consistent_instance_ext Σ mdecl.(ind_universes) u -> P Σ Γ (tConstruct ind i u) (type_of_constructor mdecl cdecl (ind, i) u)) -> @@ -715,7 +722,7 @@ Lemma typing_ind_env_app_size `{cf : checker_flags} : (forall (Σ : global_env_ext) (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ), forall (ci : case_info) p c brs indices ps mdecl idecl (isdecl : declared_inductive Σ.1 ci.(ci_ind) mdecl idecl), - Forall_decls_typing P Σ.1 -> + on_global_env cumulSpec0 Pj Σ.1 -> PΓ Σ Γ -> mdecl.(ind_npars) = ci.(ci_npar) -> eq_context_upto_names p.(pcontext) (ind_predicate_context ci.(ci_ind) mdecl idecl) -> @@ -727,8 +734,8 @@ Lemma typing_ind_env_app_size `{cf : checker_flags} : wf_local Σ (Γ ,,, predctx) -> PΓ Σ (Γ ,,, predctx) -> is_allowed_elimination Σ idecl.(ind_kelim) ps -> - ctx_inst (Prop_conj typing P) Σ Γ (p.(pparams) ++ indices) - (List.rev (subst_instance p.(puinst) (mdecl.(ind_params) ,,, idecl.(ind_indices)))) -> + ctx_inst (Prop_conj typing P Σ) Γ (p.(pparams) ++ indices) + (List.rev (subst_instance p.(puinst) (mdecl.(ind_params) ,,, idecl.(ind_indices) : context))) -> Σ ;;; Γ |- c : mkApps (tInd ci.(ci_ind) p.(puinst)) (p.(pparams) ++ indices) -> P Σ Γ c (mkApps (tInd ci.(ci_ind) p.(puinst)) (p.(pparams) ++ indices)) -> isCoFinite mdecl.(ind_finite) = false -> @@ -744,7 +751,7 @@ Lemma typing_ind_env_app_size `{cf : checker_flags} : (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (p : projection) (c : term) u mdecl idecl cdecl pdecl (isdecl : declared_projection Σ.1 p mdecl idecl cdecl pdecl) args, - Forall_decls_typing P Σ.1 -> PΓ Σ Γ -> + on_global_env cumulSpec0 Pj Σ.1 -> PΓ Σ Γ -> Σ ;;; Γ |- c : mkApps (tInd p.(proj_ind) u) args -> P Σ Γ c (mkApps (tInd p.(proj_ind) u) args) -> #|args| = ind_npars mdecl -> @@ -755,8 +762,10 @@ Lemma typing_ind_env_app_size `{cf : checker_flags} : fix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> PΓ Σ (Γ ,,, types) -> - All (on_def_type (lift_typing2 typing P Σ) Γ) mfix -> - All (on_def_body (lift_typing2 typing P Σ) types Γ) mfix -> + All (on_def_type (lift_typing typing Σ) Γ) mfix -> + All (on_def_type (Pj Σ) Γ) mfix -> + All (on_def_body (lift_typing typing Σ) types Γ) mfix -> + All (on_def_body (Pj Σ) types Γ) mfix -> wf_fixpoint Σ mfix -> P Σ Γ (tFix mfix n) decl.(dtype)) -> @@ -765,8 +774,10 @@ Lemma typing_ind_env_app_size `{cf : checker_flags} : cofix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> PΓ Σ (Γ ,,, types) -> - All (on_def_type (lift_typing2 typing P Σ) Γ) mfix -> - All (on_def_body (lift_typing2 typing P Σ) types Γ) mfix -> + All (on_def_type (lift_typing typing Σ) Γ) mfix -> + All (on_def_type (Pj Σ) Γ) mfix -> + All (on_def_body (lift_typing typing Σ) types Γ) mfix -> + All (on_def_body (Pj Σ) types Γ) mfix -> wf_cofixpoint Σ mfix -> P Σ Γ (tCoFix mfix n) decl.(dtype)) -> @@ -788,10 +799,10 @@ Lemma typing_ind_env_app_size `{cf : checker_flags} : Σ ;;; Γ |- A <=s B -> P Σ Γ t B) -> - env_prop P PΓ. + env_prop P Pj PΓ. Proof. - intros P Pdecl PΓ. - intros XΓ X X0 X1 X2 X3 X4 X5 X6 X7 X8 X9 X10 X11 X12 X13 Σ wfΣ Γ t T H. + intros P Pj Pdecl PΓ. + intros Xj XΓ X X0 X1 X2 X3 X4 X5 X6 X7 X8 X9 X10 X11 X12 X13 Σ wfΣ Γ t T H. (* NOTE (Danil): while porting to 8.9, I had to split original "pose" into 2 pieces, otherwise it takes forever to execure the "pose", for some reason *) pose proof (@Fix_F { Σ & { wfΣ : wf Σ.1 & { Γ & { t & { T & Σ ;;; Γ |- t : T }}}}}) as p0. @@ -828,85 +839,97 @@ Proof. destruct Xg. rename on_global_decl_d into Xg. constructor; auto; try constructor; auto. - * unshelve eset (IH' := IH ((Σ', udecl); (wfΣ; []; (tSort Universe.lProp); _; _))). + * unshelve eset (IH' := IH ((Σ', udecl); (wfΣ; []; (tSort sProp); _; _))). shelve. simpl. apply type_Prop. forward IH'. constructor 1; cbn. lia. apply IH'; auto. * simpl. simpl in *. destruct d. - + destruct c; simpl in *. - destruct cst_body0; apply lift_typing_impl with (1 := Xg); intros ? Hs. - all: specialize (IH ((Σ', udecl); wfΣ; _; _; _; Hs)). - all: forward IH; [constructor 1; simpl; subst Σ' Σg; cbn; lia|]. - all: apply IH. + + apply Xj; tas. + apply lift_typing_impl with (1 := Xg); intros ?? HT. split; tas. + specialize (IH ((Σ', udecl); wfΣ; _; _; _; HT)). + forward IH; [constructor 1; simpl; subst Σ' Σg; cbn; lia|]. + apply IH. + red in Xg. - destruct Xg as [onI onP onnp]; constructor; eauto. + destruct Xg as [onI onP onNP onV]; constructor; eauto. { unshelve eset (IH' := fun p => IH ((Σ', udecl); wfΣ; p) _). constructor. cbn; subst Σ' Σg; lia. clearbody IH'. cbn in IH'. clear IH; rename IH' into IH. - eapply Alli_impl; eauto. cbn in IH. clear onI onP onnp. intros n x Xg. + eapply Alli_impl; eauto. cbn in IH. clear onI onP onNP onV. intros n x Xg. refine {| ind_arity_eq := Xg.(ind_arity_eq); ind_cunivs := Xg.(ind_cunivs) |}. - - apply onArity in Xg. - apply lift_typing_impl with (1 := Xg); intros ? Hs. - apply (IH (_; _; _; Hs)). + - apply onArity in Xg. apply Xj; tas. + apply lift_typing_impl with (1 := Xg); intros ?? HT. split; tas. + apply (IH (_; _; _; HT)). - pose proof Xg.(onConstructors) as Xg'. eapply All2_impl; eauto. intros. destruct X as [cass tyeq onctyp oncargs oncind]. unshelve econstructor; eauto. - { apply lift_typing_impl with (1 := onctyp); intros ? Hs. - apply (IH (_; _; _; Hs)). } - { apply sorts_local_ctx_impl with (1 := oncargs); intros Γ' t' T' HT'. - apply lift_typing_impl with (1 := HT'); intros ? Hs. + { apply Xj; tas. + apply lift_typing_impl with (1 := onctyp); intros ?? HT. split; tas. + apply (IH (_; _; _; HT)). } + { apply sorts_local_ctx_impl with (1 := oncargs); intros Γ' j Hj. apply Xj; tas. + apply lift_typing_impl with (1 := Hj); intros ?? HT. split; tas. + apply (IH (_; _; _; HT)). } + { eapply ctx_inst_impl with (1 := oncind); intros ?? Hj. apply Xj; tas. + apply lift_typing_impl with (1 := Hj); intros ?? Hs. split; tas. apply (IH (_; _; _; Hs)). } - { clear -IH oncind. - revert oncind. - generalize (List.rev (lift_context #|cstr_args x0| 0 (ind_indices x))). - generalize (cstr_indices x0). induction 1; constructor; auto. - do 2 red in t0 |- *. - apply (IH (_; (_; (_; t0)))). } - pose proof (onProjections Xg); auto. - - destruct Xg. simpl. unfold check_ind_sorts in *. - destruct Universe.is_prop; auto. - destruct Universe.is_sprop; auto. - split. apply ind_sorts. destruct indices_matter; auto. - eapply type_local_ctx_impl. eapply ind_sorts. intros ??? HT. - apply lift_typing_impl with (1 := HT); intros ? Hs. - apply (IH (_; _; _; Hs)). + - pose proof (ind_sorts Xg) as Xg'. unfold check_ind_sorts in *. + destruct Sort.to_family; auto. + split. apply Xg'. destruct indices_matter; auto. + eapply type_local_ctx_impl. eapply Xg'. auto. intros ?? Hj. apply Xj; tas. + apply lift_typing_impl with (1 := Hj); intros ?? HT. split; tas. + apply (IH (_; _; _; HT)). - apply (onIndices Xg). } - { red in onP |- *. - apply All_local_env_impl with (1 := onP); intros ??? HT. - apply lift_typing_impl with (1 := HT); intros ? Hs. - apply (IH ((Σ', udecl); (wfΣ; _; _; _; Hs))). + { apply All_local_env_impl with (1 := onP); intros ?? Hj. apply Xj; tas. + apply lift_typing_impl with (1 := Hj); intros ?? HT. split; tas. + apply (IH ((Σ', udecl); (wfΣ; _; _; _; HT))). constructor 1. simpl. subst Σ' Σg; cbn; lia. } - assert (forall Γ t T (Hty : Σ ;;; Γ |- t : T), typing_size Hty < typing_size H -> - Forall_decls_typing P Σ.1 * P Σ Γ t T). + on_global_env cumulSpec0 Pj Σ.1 * P Σ Γ t T). { intros. - specialize (IH (existT _ Σ (existT _ wfΣ (existT _ _ (existT _ _ (existT _ _ Hty)))))). + specialize (IH (Σ; wfΣ; _; _; _; Hty)). simpl in IH. forward IH. constructor 2. simpl. apply H0. split; apply IH. } (* rename X13 into X14. *) - assert (Hdecls: typing_size H > 1 -> Forall_decls_typing P Σ.1). + assert (Hdecls: typing_size H > 1 -> on_global_env cumulSpec0 Pj Σ.1). { specialize (X14 _ _ _ (type_Prop _)). cbn in X14. intros Hle. apply X14. lia. } + + assert (Hjudg : forall Γ j (Hj : lift_typing typing Σ Γ j), + lift_typing_size (fun t T H => @typing_size cf Σ Γ t T H) _ Hj < typing_size H -> + Pj Σ Γ j). + { intros ?? Hj Hlt. + apply Xj; tas. + eapply lift_typing_size_impl with (Psize := fun t T H => @typing_size _ _ _ t T H) (tu := Hj). + intros ?? Hty Hlt'; split; tas. + apply X14 with (Hty := Hty). set (size_mid := lift_typing_size _ _ Hj) in *. lia. } + clear Xj. + assert (Hwf : forall Γ' (Hwf : wf_local Σ Γ'), - All_local_env_size (@typing_size _) _ _ Hwf < typing_size H -> + All_local_env_size (@typing_size _ _) _ Hwf < typing_size H -> PΓ Σ Γ'). { intros. eapply (XΓ _ _ _ Hwf); auto. - clear -Pdecl wfΣ X14 H0. + + clear -Pdecl wfΣ X14 H0. revert Hwf H0. - induction Hwf; cbn in *; try destruct t2 as [u Hu]; simpl in *; constructor. + induction Hwf; cbn in *; simpl in *; constructor. - simpl in *. apply IHHwf. lia. - - red. apply (X14 _ _ _ Hu). lia. + - red. apply (X14 _ _ _ t2.2.π2.1). cbn. lia. - simpl in *. apply IHHwf. lia. - - red. apply (X14 _ _ _ t3). lia. - - red. simpl. apply (X14 _ _ _ Hu). lia. } + - red. apply (X14 _ _ _ t2.1). cbn. lia. + - red. apply (X14 _ _ _ t2.2.π2.1). cbn. lia. + + revert Hwf H0. + induction Hwf; cbn in *; simpl in *; constructor. + 1,3: apply IHHwf; lia. + all: eapply Hjudg with (Hj := t2); simpl; cbn. + all: lia. } assert (Htywf : forall Γ' t T (Hty : Σ ;;; Γ' |- t : T), typing_size Hty <= typing_size H -> @@ -926,10 +949,8 @@ Proof. end; eauto; unshelve eapply X14; simpl; auto with arith]. - -- match reverse goal with - H : _ |- _ => eapply H - end; eauto; - unshelve eapply X14; simpl; eauto with arith wf. + all: try solve [ match reverse goal with H : _ |- _ => eapply H end; + eauto; [ unshelve eapply Hjudg | unshelve eapply X14 ]; simpl; auto with arith ]. -- match reverse goal with H : _ |- _ => eapply H @@ -955,20 +976,18 @@ Proof. ++ eapply (X14 _ _ _ H); eauto. rewrite /predctx; simpl; lia. ++ eapply (X14 _ _ _ H); eauto. rewrite /predctx; simpl; lia. ++ eapply (Hwf _ wf_pctx). rewrite /predctx; simpl. - change (fun (x : global_env_ext) (x0 : context) (x1 x2 : term) - (x3 : x;;; x0 |- x1 : x2) => typing_size x3) with (@typing_size cf). + change (fun Γ t T (Hty : Σ ;;; Γ |- t : T) => typing_size Hty) with (@typing_size cf Σ). lia. ++ clear -ind_inst X14. - assert (forall (Γ' : context) (t T : term) (Hty : Σ;;; Γ' |- t : T), + assert (forall Γ' t T (Hty : Σ;;; Γ' |- t : T), typing_size Hty <= ctx_inst_size _ (@typing_size _) ind_inst -> P Σ Γ' t T). { intros. eapply (X14 _ _ _ Hty). simpl. - change (fun (x : global_env_ext) (x0 : context) (x1 x2 : term) - (x3 : x;;; x0 |- x1 : x2) => typing_size x3) with (@typing_size cf). + change (fun _ _ _ _ H => typing_size H) with (@typing_size cf). lia. } clear -X ind_inst. revert ind_inst X. - generalize (List.rev (subst_instance (puinst p) (ind_params mdecl ,,, ind_indices idecl))). + generalize (List.rev (subst_instance (puinst p) (ind_params mdecl ,,, ind_indices idecl : context))). generalize (pparams p ++ indices). intros l c ctxi IH; induction ctxi; constructor; eauto. * split; tas. eapply (IH _ _ _ t0); simpl; auto. lia. @@ -984,8 +1003,8 @@ Proof. --- split; auto. intros brctxty. repeat split. +++ eapply (Hwf _ wfcbc); eauto. simpl. - change (fun (x : global_env_ext) (x0 : context) (x1 x2 : term) - (x3 : x;;; x0 |- x1 : x2) => typing_size x3) with (@typing_size cf). + change (fun _ _ _ _ H => typing_size H) with (@typing_size cf). + change (fun _ _ _ H => typing_size H) with (@typing_size cf Σ). lia. +++ exact t. +++ unshelve eapply (X14 _ _ _ t _); eauto. @@ -1008,94 +1027,73 @@ Proof. eapply X10; eauto; clear X10. simpl in *. * assert(forall Γ0 (t T : term) (Hty : Σ;;; Γ0 |- t : T), typing_size Hty < - S - (all_size _ (fun (x : def term) p => typing_size p) a1) -> - PΓ Σ Γ0). - {intros. eapply (Htywf _ _ _ Hty); eauto. lia. } + S (all_size (on_def_body (lift_typing typing _) _ _) (fun x p => on_def_body_size (@typing_size _ Σ) _ _ _ p) a1) -> + PΓ Σ Γ0). + {intros. eapply (Htywf _ _ _ Hty); eauto. change (fun _ _ _ H => typing_size H) with (@typing_size _ Σ) at 1 3. lia. } destruct mfix. now rewrite nth_error_nil in e. - depelim a1. - eapply (X _ _ _ t). simpl. lia. - * assert(forall (t T : term) (Hty : Σ;;; Γ |- t : T), - typing_size Hty < - S (all_size (fun x : def term => - ∑ s : Universe.t, Σ;;; Γ |- dtype x : tSort s) - (fun (x : def term) - (p : ∑ s : Universe.t, Σ;;; Γ |- dtype x : tSort s) => - let 'existT s d := p in typing_size d) a0) -> - Forall_decls_typing P Σ.1 * P Σ Γ t T). - intros; eauto. eapply (X14 _ _ _ Hty); eauto. simpl. unfold infer_sort. lia. - clear Hwf Htywf X14 a pΓ Hdecls. + depelim a1. destruct o as (t & ?). + eapply (X _ _ _ t). unfold on_def_body_sorting_size at 1. simpl. lia. + * assert(forall j Hj, + lift_typing_size (fun t T H => @typing_size _ Σ Γ t T H) j Hj < + S (all_size _ (fun x p => on_def_type_size (@typing_size _ Σ) _ _ p) a0) -> + Pj Σ Γ j). + { intros. eapply Hjudg with (Hj := Hj); eauto. simpl. change (typing Σ Γ) with (fun t T => typing Σ Γ t T). lia. } clear -a0 X. induction a0; constructor. - destruct p as [s Hs]. exists s; split; auto. - apply (X (dtype x) (tSort s) Hs). simpl. lia. - apply IHa0. intros. eapply (X _ _ Hty); eauto. - simpl. lia. + ++ eapply X with (Hj := p). cbn. + unfold on_def_type_sorting_size, lift_sorting_size, option_default_size in *. + cbn in *. lia. + ++ apply IHa0. intros. eapply X with (Hj := Hj); eauto. + simpl. lia. * simpl in X14. - assert(forall Γ0 : context, - wf_local Σ Γ0 -> - forall (t T : term) (Hty : Σ;;; Γ0 |- t : T), - typing_size Hty < - S - (all_size _ (fun (x : def term) p => typing_size p) a1) -> - Forall_decls_typing P Σ.1 * P Σ Γ0 t T). - {intros. eapply (X14 _ _ _ Hty); eauto. lia. } - clear X14 Hwf Htywf. - clear e decl f a0 Hdecls i. - remember (fix_context mfix) as mfixcontext. clear Heqmfixcontext. - - induction a1; econstructor; eauto. - ++ split; auto. - eapply (X _ (typing_wf_local p) _ _ p). simpl. lia. - ++ eapply IHa1. intros. - eapply (X _ X0 _ _ Hty). simpl; lia. + assert (forall Γ, wf_local Σ Γ -> + forall j Hj, + lift_typing_size (fun t T H => @typing_size _ Σ Γ t T H) j Hj < + S (all_size (on_def_body (lift_typing typing _) _ _) (fun x p => on_def_body_size (fun Γ t T H => @typing_size cf Σ Γ t T H) _ _ _ p) a1) -> + Pj Σ Γ j). + { intros. eapply Hjudg with (Hj := Hj); eauto. simpl. change (typing Σ Γ0) with (fun t T => typing Σ Γ0 t T). lia. } + clear -a1 X. + remember (fix_context mfix) as mfixcontext eqn:e. clear e. + induction a1; constructor. + ++ eapply X with (Hj := p). 1: eapply typing_wf_local, p.1. cbn. lia. + ++ apply IHa1. intros. eapply X with (Hj := Hj); eauto. + simpl. lia. -- clear X X0 X1 X2 X3 X4 X5 X6 X7 X8 X9 X10 X12. eapply X11; eauto; clear X11. simpl in *. * assert(forall Γ0 (t T : term) (Hty : Σ;;; Γ0 |- t : T), typing_size Hty < - S - (all_size _ (fun (x : def term) p => typing_size p) a1) -> - PΓ Σ Γ0). - {intros. eapply (Htywf _ _ _ Hty); eauto. lia. } + S (all_size (on_def_body (lift_typing typing _) _ _) (fun x p => on_def_body_size (@typing_size _ Σ) _ _ _ p) a1) -> + PΓ Σ Γ0). + {intros. eapply (Htywf _ _ _ Hty); eauto. change (fun _ _ _ H => typing_size H) with (@typing_size _ Σ) at 1 3. lia. } destruct mfix. now rewrite nth_error_nil in e. - depelim a1. - eapply (X _ _ _ t). simpl. lia. - * assert(forall (t T : term) (Hty : Σ;;; Γ |- t : T), - typing_size Hty < - S (all_size (fun x : def term => - ∑ s : Universe.t, Σ;;; Γ |- dtype x : tSort s) - (fun (x : def term) - (p : ∑ s : Universe.t, Σ;;; Γ |- dtype x : tSort s) => - let 'existT s d := p in typing_size d) a0) -> - Forall_decls_typing P Σ.1 * P Σ Γ t T). - intros; eauto. eapply (X14 _ _ _ Hty); eauto. simpl. unfold infer_sort. lia. - clear Hwf Htywf X14 a pΓ Hdecls. - clear -a0 X. - induction a0; constructor. - destruct p as [s Hs]. exists s; split; auto. - apply (X (dtype x) (tSort s) Hs). simpl. lia. - apply IHa0. intros. eapply (X _ _ Hty); eauto. - simpl. lia. - * simpl in X14. - assert(forall Γ0 : context, - wf_local Σ Γ0 -> - forall (t T : term) (Hty : Σ;;; Γ0 |- t : T), - typing_size Hty < - S - (all_size (fun x : def term => (Σ;;; Γ ,,, fix_context mfix |- dbody x : lift0 #|fix_context mfix| (dtype x))%type) - (fun (x : def term) p => typing_size p) a1) -> - Forall_decls_typing P Σ.1 * P Σ Γ0 t T). - {intros. eapply (X14 _ _ _ Hty); eauto. lia. } - clear X14 Hwf Htywf. - clear e decl c a0 Hdecls i. - remember (fix_context mfix) as mfixcontext. clear Heqmfixcontext. - - induction a1; econstructor; eauto. - ++ split; auto. - eapply (X _ (typing_wf_local p) _ _ p). simpl. lia. - ++ eapply IHa1. intros. - eapply (X _ X0 _ _ Hty). simpl; lia. + depelim a1. destruct o as (t & ?). + eapply (X _ _ _ t). unfold on_def_body_sorting_size at 1. simpl. lia. + * assert(forall j Hj, + lift_typing_size (fun t T H => @typing_size _ Σ Γ t T H) j Hj < + S (all_size _ (fun x p => on_def_type_size (@typing_size _ Σ) _ _ p) a0) -> + Pj Σ Γ j). + { intros. eapply Hjudg with (Hj := Hj); eauto. simpl. change (typing Σ Γ) with (fun t T => typing Σ Γ t T). lia. } + clear -a0 X. + induction a0; constructor. + ++ eapply X with (Hj := p). cbn. + unfold on_def_type_sorting_size, lift_sorting_size, option_default_size in *. + cbn in *. lia. + ++ apply IHa0. intros. eapply X with (Hj := Hj); eauto. + simpl. lia. + * simpl in X14. + assert (forall Γ, wf_local Σ Γ -> + forall j Hj, + lift_typing_size (fun t T H => @typing_size _ Σ Γ t T H) j Hj < + S (all_size (on_def_body (lift_typing typing _) _ _) (fun x p => on_def_body_size (fun Γ t T H => @typing_size cf Σ Γ t T H) _ _ _ p) a1) -> + Pj Σ Γ j). + { intros. eapply Hjudg with (Hj := Hj); eauto. simpl. change (typing Σ Γ0) with (fun t T => typing Σ Γ0 t T). lia. } + clear -a1 X. + remember (fix_context mfix) as mfixcontext eqn:e. clear e. + induction a1; constructor. + ++ eapply X with (Hj := p). 1: eapply typing_wf_local, p.1. cbn. lia. + ++ apply IHa1. intros. eapply X with (Hj := Hj); eauto. + simpl. lia. -- eapply X12; tea. clear -p2 X14. @@ -1108,45 +1106,45 @@ Qed. Lemma typing_ind_env `{cf : checker_flags} : forall (P : global_env_ext -> context -> term -> term -> Type) + (Pj : global_env_ext -> context -> judgment -> Type) (Pdecl := fun Σ Γ wfΓ t T tyT => P Σ Γ t T) (PΓ : global_env_ext -> context -> Type), + (forall Σ (wfΣ : wf Σ) (Γ : context) (* (wfΓ : wf_local Σ Γ) *) j, + lift_typing_conj (typing Σ) (P Σ) Γ j -> Pj Σ Γ j) -> + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ), - All_local_env_over typing Pdecl Σ Γ wfΓ -> PΓ Σ Γ) -> + All_local_env_over (typing Σ) (Pdecl Σ) Γ wfΓ -> All_local_env (Pj Σ) Γ -> PΓ Σ Γ) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : nat) decl, nth_error Γ n = Some decl -> PΓ Σ Γ -> P Σ Γ (tRel n) (lift0 (S n) decl.(decl_type))) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (u : Universe.t), + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (u : sort), PΓ Σ Γ -> - wf_universe Σ u -> - P Σ Γ (tSort u) (tSort (Universe.super u))) -> + wf_sort Σ u -> + P Σ Γ (tSort u) (tSort (Sort.super u))) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : aname) (t b : term) (s1 s2 : Universe.t), + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (t b : term) (s1 s2 : sort), PΓ Σ Γ -> - Σ ;;; Γ |- t : tSort s1 -> - P Σ Γ t (tSort s1) -> - Σ ;;; Γ,, vass n t |- b : tSort s2 -> - P Σ (Γ,, vass n t) b (tSort s2) -> P Σ Γ (tProd n t b) (tSort (Universe.sort_of_product s1 s2))) -> + lift_typing typing Σ Γ (j_vass_s na t s1) -> + Pj Σ Γ (j_vass_s na t s1) -> + Σ ;;; Γ,, vass na t |- b : tSort s2 -> + P Σ (Γ,, vass na t) b (tSort s2) -> P Σ Γ (tProd na t b) (tSort (Sort.sort_of_product s1 s2))) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : aname) (t b : term) - (s1 : Universe.t) (bty : term), + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (t b bty : term), PΓ Σ Γ -> - Σ ;;; Γ |- 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)) -> + lift_typing typing Σ Γ (j_vass na t) -> + Pj Σ Γ (j_vass na t) -> + Σ ;;; Γ,, vass na t |- b : bty -> P Σ (Γ,, vass na t) b bty -> P Σ Γ (tLambda na t b) (tProd na t bty)) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : aname) (b b_ty b' : term) - (s1 : Universe.t) (b'_ty : term), + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (b b_ty b' b'_ty : term), PΓ Σ Γ -> - Σ ;;; Γ |- b_ty : tSort s1 -> - P Σ Γ b_ty (tSort s1) -> - Σ ;;; Γ |- b : b_ty -> - P Σ Γ b b_ty -> - Σ ;;; Γ,, vdef n b b_ty |- b' : b'_ty -> - P Σ (Γ,, vdef n b b_ty) b' b'_ty -> P Σ Γ (tLetIn n b b_ty b') (tLetIn n b b_ty b'_ty)) -> + lift_typing typing Σ Γ (j_vdef na b b_ty) -> + Pj Σ Γ (j_vdef na b b_ty) -> + Σ ;;; Γ,, vdef na b b_ty |- b' : b'_ty -> + P Σ (Γ,, vdef na b b_ty) b' b'_ty -> P Σ Γ (tLetIn na b b_ty b') (tLetIn na b b_ty b'_ty)) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (t : term) na A B u s, PΓ Σ Γ -> @@ -1156,7 +1154,7 @@ Lemma typing_ind_env `{cf : checker_flags} : P Σ Γ (tApp t u) (B{0 := u})) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) cst u (decl : constant_body), - Forall_decls_typing P Σ.1 -> + on_global_env cumulSpec0 Pj Σ.1 -> PΓ Σ Γ -> declared_constant Σ.1 cst decl -> consistent_instance_ext Σ decl.(cst_universes) u -> @@ -1164,14 +1162,14 @@ Lemma typing_ind_env `{cf : checker_flags} : (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (ind : inductive) u mdecl idecl (isdecl : declared_inductive Σ.1 ind mdecl idecl), - Forall_decls_typing P Σ.1 -> + on_global_env cumulSpec0 Pj Σ.1 -> PΓ Σ Γ -> consistent_instance_ext Σ mdecl.(ind_universes) u -> P Σ Γ (tInd ind u) (subst_instance u (ind_type idecl))) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (ind : inductive) (i : nat) u mdecl idecl cdecl (isdecl : declared_constructor Σ.1 (ind, i) mdecl idecl cdecl), - Forall_decls_typing P Σ.1 -> + on_global_env cumulSpec0 Pj Σ.1 -> PΓ Σ Γ -> consistent_instance_ext Σ mdecl.(ind_universes) u -> P Σ Γ (tConstruct ind i u) (type_of_constructor mdecl cdecl (ind, i) u)) -> @@ -1179,7 +1177,7 @@ Lemma typing_ind_env `{cf : checker_flags} : (forall (Σ : global_env_ext) (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ), forall (ci : case_info) p c brs indices ps mdecl idecl (isdecl : declared_inductive Σ.1 ci.(ci_ind) mdecl idecl), - Forall_decls_typing P Σ.1 -> + on_global_env cumulSpec0 Pj Σ.1 -> PΓ Σ Γ -> mdecl.(ind_npars) = ci.(ci_npar) -> eq_context_upto_names p.(pcontext) (ind_predicate_context ci.(ci_ind) mdecl idecl) -> @@ -1191,8 +1189,8 @@ Lemma typing_ind_env `{cf : checker_flags} : P Σ (Γ ,,, predctx) p.(preturn) (tSort ps) -> PΓ Σ (Γ ,,, predctx) -> is_allowed_elimination Σ idecl.(ind_kelim) ps -> - ctx_inst (Prop_conj typing P) Σ Γ (p.(pparams) ++ indices) - (List.rev (subst_instance p.(puinst) (mdecl.(ind_params) ,,, idecl.(ind_indices)))) -> + ctx_inst (Prop_conj typing P Σ) Γ (p.(pparams) ++ indices) + (List.rev (subst_instance p.(puinst) (mdecl.(ind_params) ,,, idecl.(ind_indices) : context))) -> Σ ;;; Γ |- c : mkApps (tInd ci.(ci_ind) p.(puinst)) (p.(pparams) ++ indices) -> P Σ Γ c (mkApps (tInd ci.(ci_ind) p.(puinst)) (p.(pparams) ++ indices)) -> isCoFinite mdecl.(ind_finite) = false -> @@ -1208,7 +1206,7 @@ Lemma typing_ind_env `{cf : checker_flags} : (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (p : projection) (c : term) u mdecl idecl cdecl pdecl (isdecl : declared_projection Σ.1 p mdecl idecl cdecl pdecl) args, - Forall_decls_typing P Σ.1 -> PΓ Σ Γ -> + on_global_env cumulSpec0 Pj Σ.1 -> PΓ Σ Γ -> Σ ;;; Γ |- c : mkApps (tInd p.(proj_ind) u) args -> P Σ Γ c (mkApps (tInd p.(proj_ind) u) args) -> #|args| = ind_npars mdecl -> @@ -1219,8 +1217,10 @@ Lemma typing_ind_env `{cf : checker_flags} : fix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> PΓ Σ (Γ ,,, types) -> - All (on_def_type (lift_typing2 typing P Σ) Γ) mfix -> - All (on_def_body (lift_typing2 typing P Σ) types Γ) mfix -> + All (on_def_type (lift_typing typing Σ) Γ) mfix -> + All (on_def_type (Pj Σ) Γ) mfix -> + All (on_def_body (lift_typing typing Σ) types Γ) mfix -> + All (on_def_body (Pj Σ) types Γ) mfix -> wf_fixpoint Σ mfix -> P Σ Γ (tFix mfix n) decl.(dtype)) -> @@ -1229,8 +1229,10 @@ Lemma typing_ind_env `{cf : checker_flags} : cofix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> PΓ Σ (Γ ,,, types) -> - All (on_def_type (lift_typing2 typing P Σ) Γ) mfix -> - All (on_def_body (lift_typing2 typing P Σ) types Γ) mfix -> + All (on_def_type (lift_typing typing Σ) Γ) mfix -> + All (on_def_type (Pj Σ) Γ) mfix -> + All (on_def_body (lift_typing typing Σ) types Γ) mfix -> + All (on_def_body (Pj Σ) types Γ) mfix -> wf_cofixpoint Σ mfix -> P Σ Γ (tCoFix mfix n) decl.(dtype)) -> @@ -1252,10 +1254,10 @@ Lemma typing_ind_env `{cf : checker_flags} : Σ ;;; Γ |- A <=s B -> P Σ Γ t B) -> - env_prop P PΓ. + env_prop P Pj PΓ. Proof. - intros P Pdecl PΓ; unfold env_prop. - intros XΓ X X0 X1 X2 X3 X4 X5 X6 X7 X8 X9 X10 X11 X12 X13 Σ wfΣ Γ t T H. + intros P Pj Pdecl PΓ; unfold env_prop. + intros Xj XΓ X X0 X1 X2 X3 X4 X5 X6 X7 X8 X9 X10 X11 X12 X13 Σ wfΣ Γ t T H. apply typing_ind_env_app_size; eauto. Qed. @@ -1268,253 +1270,10 @@ Ltac my_rename_hyp h th := | (typing _ _ ?t _) => fresh "type" t | (@cumulSpec _ _ _ ?t _) => fresh "cumul" t | (@convSpec _ _ ?t _) => fresh "conv" t - | (All_local_env (lift_judgment (@typing _) _ _) ?G) => fresh "wf" G - | (All_local_env (lift_judgment (@typing _) _) _) => fresh "wf" + | (All_local_env (lift_sorting (@typing _) _ _) ?G) => fresh "wf" G + | (All_local_env (lift_sorting (@typing _) _) _) => fresh "wf" | (All_local_env _ _ ?G) => fresh "H" G | context [typing _ _ (_ ?t) _] => fresh "IH" t end. Ltac rename_hyp h ht ::= my_rename_hyp h ht. - -Section All_local_env. - (** * Lemmas about All_local_env *) - - Context {cf: checker_flags}. - - Lemma nth_error_All_local_env {P Γ n} (isdecl : n < #|Γ|) : - All_local_env P Γ -> - on_some (on_local_decl P (skipn (S n) Γ)) (nth_error Γ n). - Proof using Type. - induction 1 in n, isdecl |- *. red; simpl. - - destruct n; simpl; inv isdecl. - - destruct n. red; simpl. red. simpl. apply t0. - simpl. apply IHX. simpl in isdecl. lia. - - destruct n. auto. - apply IHX. simpl in *. lia. - Qed. - - Lemma lookup_on_global_env P (Σ : global_env) c decl : - on_global_env cumulSpec0 P Σ -> - lookup_env Σ c = Some decl -> - { Σ' : global_env & [× extends_strictly_on_decls Σ' Σ, on_global_env cumulSpec0 P Σ' & - on_global_decl cumulSpec0 P (Σ', universes_decl_of_decl decl) c decl] }. - Proof using Type. - destruct Σ as [univs Σ retro]; rewrite /on_global_env /lookup_env; cbn. - intros [cu Σp]. - induction Σp; simpl. congruence. - destruct (eqb_specT c kn); subst. - - intros [= ->]. - exists ({| universes := univs; declarations := Σ; retroknowledge := retro |}). - split. - * red; cbn. split; [split;[lsets|csets]| |]. - exists [(kn, decl)] => //. - apply Retroknowledge.extends_refl. - * split => //. - * destruct o; assumption. - - intros hl. destruct (IHΣp hl) as [Σ' []]. - exists Σ'. - split=> //. - destruct e as [eu ed]. red; cbn in *. - split; [auto| |auto]. - destruct ed as [Σ'' ->]. - exists (Σ'' ,, (kn, d)) => //. - Qed. - - Lemma All_local_env_app (P : context -> term -> typ_or_sort -> Type) l l' : - All_local_env P l * All_local_env (fun Γ t T => P (l ,,, Γ) t T) l' -> - All_local_env P (l ,,, l'). - Proof using Type. - induction l'; simpl; auto. intuition. - intuition. destruct a. destruct decl_body. - inv b. econstructor; eauto. inv b; econstructor; eauto. - Qed. - - Lemma All_local_env_app_inv (P : context -> term -> typ_or_sort -> Type) l l' : - All_local_env P (l ,,, l') -> - All_local_env P l * All_local_env (fun Γ t T => P (l ,,, Γ) t T) l'. - Proof using Type. - apply All_local_app_rel. - Qed. - - Definition wf_local_rel_app_inv {Σ Γ1 Γ2 Γ3} : - wf_local_rel Σ Γ1 (Γ2 ,,, Γ3) -> - wf_local_rel Σ Γ1 Γ2 * wf_local_rel Σ (Γ1 ,,, Γ2) Γ3. - Proof. - intros h. apply All_local_env_app_inv in h as [h1 h2]. - split. - - exact h1. - - eapply All_local_env_impl. 1: exact h2. - intros Γ t [T|] h. - all: cbn in *. - all: change PCUICEnvironment.app_context with app_context in *. - all: rewrite <- app_context_assoc. - all: auto. - Defined. - - Lemma All_local_env_lookup {P Γ n} {decl} : - All_local_env P Γ -> - nth_error Γ n = Some decl -> - on_local_decl P (skipn (S n) Γ) decl. - Proof using Type. - induction 1 in n, decl |- *. simpl. destruct n; simpl; congruence. - destruct n. red. simpl. intros [= <-]. simpl. apply t0. - simpl in *. eapply IHX. - destruct n. simpl. intros [= <-]. auto. - eapply IHX. - Qed. - - Definition wf_local_rel_app {Σ Γ1 Γ2 Γ3} : - wf_local_rel Σ Γ1 Γ2 -> wf_local_rel Σ (Γ1 ,,, Γ2) Γ3 - -> wf_local_rel Σ Γ1 (Γ2 ,,, Γ3). - Proof. - intros h1 h2. eapply All_local_env_app. - split. - - assumption. - - eapply All_local_env_impl. - + eassumption. - + change PCUICEnvironment.app_context with app_context. - intros Γ t []; cbn; - now rewrite app_context_assoc. - Defined. - - Definition wf_local_app {Σ Γ1 Γ2} : - wf_local Σ Γ1 -> - wf_local_rel Σ Γ1 Γ2 -> - wf_local Σ (Γ1 ,,, Γ2). - Proof using Type. - intros H1 H2. apply All_local_env_app. - now split. - Qed. - - Definition wf_local_app_inv {Σ Γ1 Γ2} : - wf_local Σ (Γ1 ,,, Γ2) -> - wf_local Σ Γ1 * wf_local_rel Σ Γ1 Γ2. - Proof using Type. - apply All_local_env_app_inv. - Qed. - - Lemma wf_local_app_ind {Σ Γ1 Γ2} : - wf_local Σ Γ1 -> - (wf_local Σ Γ1 -> wf_local_rel Σ Γ1 Γ2) -> - wf_local Σ (Γ1 ,,, Γ2). - Proof using Type. - intros wf IH. - apply wf_local_app; auto. - Qed. - - Lemma All_local_env_map (P : context -> term -> typ_or_sort -> Type) f l : - (forall u, f (tSort u) = tSort u) -> - All_local_env (fun Γ t T => P (map (map_decl f) Γ) (f t) (typ_or_sort_map f T)) l - -> All_local_env P (map (map_decl f) l). - Proof using Type. - intros Hf. induction 1; econstructor; eauto. - Qed. - - Definition property := - forall (Σ : global_env_ext) (Γ : context), - wf_local Σ Γ -> forall t T : term, typing Σ Γ t T -> Type. - - Definition lookup_wf_local {Γ P} (wfΓ : All_local_env P Γ) (n : nat) - (isdecl : n < #|Γ|) : - All_local_env P (skipn (S n) Γ). - Proof. - induction wfΓ in n, isdecl |- *; simpl. constructor. - cbn -[skipn] in *. destruct n. - simpl. exact wfΓ. - apply IHwfΓ. auto with arith. - destruct n. exact wfΓ. - apply IHwfΓ. auto with arith. - Defined. - - Lemma wf_local_app_skipn {Σ Γ Γ' n} : - wf_local Σ (Γ ,,, Γ') -> - wf_local Σ (Γ ,,, skipn n Γ'). - Proof using Type. - intros wf. - destruct (le_dec n #|Γ'|). - unfold app_context. - replace Γ with (skipn (n - #|Γ'|) Γ). - rewrite -skipn_app. now apply All_local_env_skipn. - replace (n - #|Γ'|) with 0 by lia. now rewrite skipn_0. - rewrite List.skipn_all2. lia. - now eapply wf_local_app_l in wf. - Qed. - - Definition on_local_decl_glob (P : term -> typ_or_sort -> Type) d := - match d.(decl_body) with - | Some b => P b (Typ d.(decl_type)) × P d.(decl_type) Sort - | None => P d.(decl_type) Sort - end. - - Definition lookup_wf_local_decl {Γ P} (wfΓ : All_local_env P Γ) (n : nat) - {decl} (eq : nth_error Γ n = Some decl) : - ∑ Pskip : All_local_env P (skipn (S n) Γ), - on_local_decl_glob (P (skipn (S n) Γ)) decl. - Proof. - induction wfΓ in n, decl, eq |- *; simpl. - - elimtype False. destruct n; depelim eq. - - destruct n. - + simpl. exists wfΓ. injection eq; intros <-. apply t0. - + apply IHwfΓ. auto with arith. - - destruct n. - + exists wfΓ. injection eq; intros <-. - simpl. split; auto. - + apply IHwfΓ. apply eq. - Defined. - - Definition on_wf_local_decl {Σ Γ} - (P : forall Σ Γ (wfΓ : wf_local Σ Γ) t T, Σ ;;; Γ |- t : T -> Type) - (wfΓ : wf_local Σ Γ) {d} (H : on_local_decl_glob (lift_typing typing Σ Γ) d) := - match d as d' return (on_local_decl_glob (lift_typing typing Σ Γ) d') -> Type with - | {| decl_name := na; decl_body := Some b; decl_type := ty |} => - fun H => (P Σ Γ wfΓ b ty H.1 * P Σ Γ wfΓ _ _ (projT2 (snd H)))%type - | {| decl_name := na; decl_body := None; decl_type := ty |} => fun H => P Σ Γ wfΓ _ _ (projT2 H) - end H. - - Lemma nth_error_All_local_env_over {P Σ Γ n decl} (eq : nth_error Γ n = Some decl) {wfΓ : wf_local Σ Γ} : - All_local_env_over typing P Σ Γ wfΓ -> - let Γ' := skipn (S n) Γ in - let p := lookup_wf_local_decl wfΓ n eq in - (All_local_env_over typing P Σ Γ' (projT1 p) * on_wf_local_decl P (projT1 p) (projT2 p))%type. - Proof. - induction 1 in n, decl, eq |- *. simpl. - - destruct n; simpl; elimtype False; discriminate eq. - - destruct n. cbn [skipn]. noconf eq. split. apply X. simpl. apply Hs. - simpl. apply IHX. - - destruct n. noconf eq. simpl. split; auto. - apply IHX. - Defined. - - Lemma All_local_env_prod_inv : - forall P Q Γ, - All_local_env (fun Δ A t => P Δ A t × Q Δ A t) Γ -> - All_local_env P Γ × All_local_env Q Γ. - Proof using Type. - intros P Q Γ h. - induction h. - - split ; constructor. - - destruct IHh, t0. - split ; constructor ; auto. - - destruct IHh, t0, t1. - split ; constructor ; auto. - Qed. - - Lemma All_local_env_lift_prod_inv : - forall Σ P Q Δ, - All_local_env (lift_typing (fun Σ Γ => Trel_conj (P Σ Γ) (Q Σ Γ)) Σ) Δ -> - All_local_env (lift_typing P Σ) Δ × All_local_env (lift_typing Q Σ) Δ. - Proof using Type. - intros Σ P Q Δ h. - induction h. - - split ; constructor. - - destruct IHh. destruct t0 as [? [? ?]]. - split ; constructor ; auto. - + cbn. eexists. eassumption. - + cbn. eexists. eassumption. - - destruct IHh. destruct t0 as [? [? ?]]. destruct t1. - split ; constructor ; auto. - + cbn. eexists. eassumption. - + cbn. eexists. eassumption. - Qed. - -End All_local_env. diff --git a/pcuic/theories/PCUICValidity.v b/pcuic/theories/PCUICValidity.v index 6eb870fa4..ca255c23c 100644 --- a/pcuic/theories/PCUICValidity.v +++ b/pcuic/theories/PCUICValidity.v @@ -27,24 +27,13 @@ Section Validity. Lemma isType_weaken_full : weaken_env_prop_full cumulSpec0 (lift_typing typing) (fun Σ Γ t T => isType Σ Γ T). Proof using Type. - do 2 red. intros. - apply infer_typing_sort_impl with id X2; intros Hs. - unshelve eapply (weaken_env_prop_typing _ _ _ _ _ X1 _ _ (Typ (tSort _))); eauto with pcuic. - red. simpl. destruct Σ. eapply Hs. + do 2 red. intros (Σ & φ) **. + unshelve eapply weaken_env_prop_typing; cbn. 5: apply X2. + all: assumption. Qed. Hint Resolve isType_weaken_full : pcuic. - Lemma isType_weaken : - weaken_env_prop cumulSpec0 (lift_typing typing) - (lift_typing (fun Σ Γ (_ T : term) => isType Σ Γ T)). - Proof using Type. - do 2 red. intros. - apply lift_typing_impl with (1 := X2); intros ? Hs. - now eapply (isType_weaken_full (Σ, _)). - Qed. - Hint Resolve isType_weaken : pcuic. - Lemma isType_extends (Σ : global_env) (Σ' : global_env) (φ : universes_decl) : wf Σ -> wf Σ' -> extends Σ Σ' -> @@ -53,25 +42,12 @@ Section Validity. isType (Σ, φ) Γ t0 -> isType (Σ', φ) Γ t0. Proof using Type. intros wfΣ wfΣ' ext Γ t Hty. - apply infer_typing_sort_impl with id Hty; intros Hs. - eapply (env_prop_typing weakening_env (Σ, φ)); auto. - Qed. - - Lemma weaken_env_prop_isType : - weaken_env_prop cumulSpec0 (lift_typing typing) - (lift_typing - (fun (Σ0 : PCUICEnvironment.global_env_ext) - (Γ0 : PCUICEnvironment.context) (_ T : term) => - isType Σ0 Γ0 T)). - Proof using Type. - red. intros Σ Σ' ϕ wfΣ wfΣ' ext * Hty. - apply lift_typing_impl with (1 := Hty); intros ? Hs. - now eapply isType_extends with Σ. + eapply weaken_env_prop_typing; cbn; tas. 3: eassumption. all: assumption. Qed. - Lemma isType_Sort_inv {Σ : global_env_ext} {Γ s} : wf Σ -> isType Σ Γ (tSort s) -> wf_universe Σ s. + Lemma isType_Sort_inv {Σ : global_env_ext} {Γ s} : wf Σ -> isType Σ Γ (tSort s) -> wf_sort Σ s. Proof using Type. - intros wfΣ [u Hu]. + intros wfΣ (_ & u & Hu & _). now eapply inversion_Sort in Hu as [? [? ?]]. Qed. @@ -83,8 +59,8 @@ Section Validity. isType Σ (subst_instance u Γ) (subst_instance u T). Proof using Type. destruct Σ as [Σ φ]. intros X X0 Hty X1. - eapply infer_typing_sort_impl with _ Hty; intros Hs. - eapply (typing_subst_instance_decl _ _ _ (tSort _)); eauto. + eapply lift_typing_fu_impl with (1 := Hty) => // ?? Hs. + eapply typing_subst_instance_decl; eauto. Qed. Lemma isWfArity_subst_instance_decl {Σ Γ T c decl u} : @@ -96,7 +72,7 @@ Section Validity. Proof using Type. destruct Σ as [Σ φ]. intros X X0 [isTy [ctx [s eq]]] X1. split. eapply isType_subst_instance_decl; eauto. - exists (subst_instance u ctx), (subst_instance_univ u s). + exists (subst_instance u ctx), (subst_instance_sort u s). rewrite (subst_instance_destArity []) eq. intuition auto. Qed. @@ -107,23 +83,11 @@ Section Validity. isType Σ Γ T. Proof using Type. intros wfΣ wfΓ HT. - apply infer_typing_sort_impl with id HT; intros Hs. - eapply (weaken_ctx (Γ:=[])); eauto. - Qed. - - Lemma nth_error_All_local_env {P : context -> term -> typ_or_sort -> Type} {Γ n d} : - nth_error Γ n = Some d -> - All_local_env P Γ -> - on_local_decl P (skipn (S n) Γ) d. - Proof using Type. - intros heq hΓ. - epose proof (nth_error_Some_length heq). - eapply (nth_error_All_local_env) in H; tea. - now rewrite heq in H. + now apply isType_weaken. Qed. Notation type_ctx := (type_local_ctx (lift_typing typing)). - Lemma type_ctx_wf_univ Σ Γ Δ s : type_ctx Σ Γ Δ s -> wf_universe Σ s. + Lemma type_ctx_wf_univ Σ Γ Δ s : type_ctx Σ Γ Δ s -> wf_sort Σ s. Proof using Type. induction Δ as [|[na [b|] ty]]; simpl; auto with pcuic. Qed. @@ -131,23 +95,12 @@ Section Validity. Notation liat := ltac:(lia) (only parsing). - Lemma eq_binder_annots_eq_ctx (Σ : global_env_ext) (Δ : context) (nas : list aname) : + Lemma eq_binder_annots_eq_ctx (Δ : context) (nas : list aname) : All2 (fun x y => eq_binder_annot x y.(decl_name)) nas Δ -> - PCUICEquality.eq_context_gen (PCUICEquality.eq_term Σ Σ) (PCUICEquality.eq_term Σ Σ) - (map2 set_binder_name nas Δ) Δ. + PCUICEquality.eq_context_upto_names (map2 set_binder_name nas Δ) Δ. Proof using Type. - induction Δ in nas |- * using PCUICInduction.ctx_length_rev_ind; simpl; intros hlen. - - depelim hlen. simpl. reflexivity. - - destruct nas as [|nas na] using rev_case => //; - pose proof (All2_length hlen) as hlen';len in hlen'; simpl in hlen'; try lia. - eapply All2_app_inv_l in hlen as (l1'&l2'&heq&alnas&allna). - depelim allna. depelim allna. - rewrite map2_app => /= //; try lia. unfold aname. - eapply app_inj_tail in heq as [<- <-]. - simpl. eapply All2_fold_app; auto. - constructor. constructor. - destruct d as [na' [d|] ty]; constructor; cbn in *; auto; - try reflexivity. + induction 1; cbn; constructor; auto. + destruct y as [na [b|] t]; now constructor. Qed. Lemma eq_term_set_binder_name (Σ : global_env_ext) (Δ : context) T U (nas : list aname) : @@ -156,12 +109,12 @@ Section Validity. PCUICEquality.eq_term Σ Σ (it_mkProd_or_LetIn (map2 set_binder_name nas Δ) T) (it_mkProd_or_LetIn Δ U) . Proof using Type. intros a; unshelve eapply eq_binder_annots_eq_ctx in a; tea. - eapply All2_fold_All2 in a. induction a in T, U |- *. - auto. - rewrite /= /mkProd_or_LetIn. destruct r => /=; intros; eapply IHa; constructor; auto. + all: reflexivity. Qed. Lemma All2_eq_binder_subst_context_inst l s k i Δ Γ : @@ -205,20 +158,15 @@ Section Validity. Qed. Lemma validity_wf_local {Σ} Γ Δ: - All_local_env - (fun (Γ0 : context) (t : term) (T : typ_or_sort) => - match T with - | Typ T0 => isType Σ (Γ,,, Γ0) T0 × Σ ;;; (Γ ,,, Γ0) |- t : T0 - | Sort => isType Σ (Γ,,, Γ0) t - end) Δ -> + All_local_rel (lift_typing typing Σ) Γ Δ -> ∑ xs, sorts_local_ctx (lift_typing typing) Σ Γ Δ xs. Proof using Type. induction 1. - - exists []; cbn; auto. exact tt. + - exists []; cbn; auto. + - destruct IHX as [xs Hxs]. + eexists (_ :: xs). cbn. split; auto. + unshelve eapply lift_sorting_extract; tea. - destruct IHX as [xs Hxs]. - destruct t0 as [s Hs]. - exists (s :: xs). cbn. split; auto. - - destruct IHX as [xs Hxs]. destruct t0 as [s Hs]. exists xs; cbn. split; auto. Qed. @@ -226,72 +174,66 @@ Section Validity. Theorem validity_env : env_prop (fun Σ Γ t T => isType Σ Γ T) - (fun Σ Γ => wf_local Σ Γ × All_local_env - (fun Γ t T => match T with Typ T => (isType Σ Γ T × Σ ;;; Γ |- t : T) | Sort => isType Σ Γ t end) Γ). + (fun Σ Γ j => lift_sorting (fun _ T => isType Σ Γ T) (fun _ s => wf_sort Σ s) j) + (fun Σ Γ => wf_local Σ Γ). Proof using Type. apply typing_ind_env; intros; rename_all_hyps. - - split => //. induction X; constructor; auto. + - apply lift_sorting_impl with (1 := X) => ?? [] HT //. + apply/isType_Sort_inv. + + - assumption. - - destruct X as [_ X]. - have hd := (nth_error_All_local_env heq_nth_error X). - destruct decl as [na [b|] ty]; cbn -[skipn] in *; destruct hd. - + eapply isType_lift; eauto. - now apply nth_error_Some_length in heq_nth_error. - + eapply isType_lift; eauto. - now apply nth_error_Some_length in heq_nth_error. - now exists x. + - (* Rel *) + destruct (All_local_env_nth_error X heq_nth_error) as (_ & s & Hs & _). + eapply isType_lift; eauto. + 1: now apply nth_error_Some_length in heq_nth_error. + now eapply has_sort_isType. - (* Universe *) - exists (Universe.super (Universe.super u)). - constructor; auto. - now apply wf_universe_super. + eapply has_sort_isType with (s := Sort.super (Sort.super u)). + constructor; auto. + now apply wf_sort_super. - (* Product *) - eexists. - eapply isType_Sort_inv in X1; eapply isType_Sort_inv in X3; auto. + eapply has_sort_isType. + apply unlift_TypUniv in X1; eapply isType_Sort_inv in X3; auto. econstructor; eauto. - now apply wf_universe_product. + now apply wf_sort_product. - (* Lambda *) - destruct X3 as [bs tybs]. - eapply isType_Sort_inv in X1; auto. - exists (Universe.sort_of_product s1 bs). - constructor; auto. + eapply lift_sorting_ex_it_impl_gen with X3 => // Hs. + pose proof (lift_sorting_extract X0). + eexists; constructor; eauto. - (* Let *) - apply infer_typing_sort_impl with id X5; unfold id in *; intros Hs. + apply lift_sorting_it_impl_gen with X3 => // Hs. eapply type_Cumul. eapply type_LetIn; eauto. econstructor; pcuic. eapply convSpec_cumulSpec, red1_cumulSpec; constructor. - (* Application *) - apply infer_typing_sort_impl with id X3; unfold id in *; intros Hs'. - move: (typing_wf_universe wf Hs') => wfs. + apply lift_sorting_it_impl_gen with X3 => // Hs'. + move: (typing_wf_sort wf Hs') => wfs. eapply (substitution0 (n := na) (T := tSort _)); eauto. - apply inversion_Prod in Hs' as [na' [s1 [s2 Hs]]]; tas. intuition. - eapply (weakening_ws_cumul_pb (pb:=Cumul) (Γ' := []) (Γ'' := [vass na A])) in b0; pcuic. - simpl in b0. + apply inversion_Prod in Hs' as [na' [s1 [s2 [Hs Hle]]]]; tas. + eapply (weakening_ws_cumul_pb (pb:=Cumul) (Γ' := []) (Γ'' := [vass na A])) in Hle; pcuic. + simpl in Hle. eapply (type_ws_cumul_pb (pb:=Cumul)); eauto. pcuic. etransitivity; tea. eapply into_ws_cumul_pb => //. all:eauto with fvs. do 2 constructor. - apply leq_universe_product. - - - destruct decl as [ty [b|] univs]; simpl in *. - * eapply declared_constant_inv in X; eauto. - red in X. simpl in X. - eapply isType_weakening; eauto. - unshelve eapply declared_constant_to_gen in H; eauto. - eapply (isType_subst_instance_decl (Γ:=[])); eauto. simpl. - exact weaken_env_prop_isType. - * have ond := on_declared_constant wf H. - do 2 red in ond. simpl in ond. - simpl in ond. - eapply isType_weakening; eauto. - unshelve eapply declared_constant_to_gen in H; eauto. - eapply (isType_subst_instance_decl (Γ:=[])); eauto. + apply leq_sort_product. + + - (* Constant *) + eapply declared_constant_inv in wf as Hc; eauto. + destruct Hc as (_ & s & Hs & _); simpl in Hs. + eapply isType_weakening; eauto. + unshelve eapply declared_constant_to_gen in H; eauto. + eapply (isType_subst_instance_decl (Γ:=[])); eauto. simpl. + now eapply has_sort_isType. + exact weaken_env_prop_typing. - (* Inductive type *) destruct (on_declared_inductive isdecl); pcuic. @@ -304,12 +246,12 @@ Section Validity. - (* Constructor type *) destruct (on_declared_constructor isdecl) as [[oni oib] [cs [declc onc]]]. unfold type_of_constructor. - eapply infer_typing_sort_impl with _ (on_ctype onc); intros Hs. + eapply lift_typing_fu_impl with (f := fun t => _ (_ t)) (1 := on_ctype onc) => //= ?? Hs. eapply instantiate_minductive in Hs; eauto. 2:(destruct isdecl as [[] ?]; eauto). simpl in Hs. eapply (weaken_ctx (Γ:=[]) Γ); eauto. - eapply (substitution (Γ := []) (s := inds _ _ _) (Δ := []) (T := tSort _)); eauto. + eapply (substitution (Γ := []) (s := inds _ _ _) (Δ := [])); eauto. eapply subslet_inds; eauto. destruct isdecl; eauto. now rewrite app_context_nil_l. @@ -324,18 +266,17 @@ Section Validity. eapply weaken_wf_local; tea. now apply (on_minductive_wf_params_indices_inst isdecl _ cu). eapply spine_subst_smash in X6; tea. - destruct X4. destruct (on_declared_inductive isdecl) as [onmind oib]. - rewrite /ptm. exists ps. + rewrite /ptm. eapply has_sort_isType with (s := ps). eapply type_mkApps; eauto. eapply type_it_mkLambda_or_LetIn; tea. have typred : isType Σ Γ (it_mkProd_or_LetIn predctx (tSort ps)). - { eapply All_local_env_app_inv in a0 as [_ onp]. + { eapply All_local_env_app_inv in X2 as [_ onp]. eapply validity_wf_local in onp as [xs Hs]. - eexists _. + eapply has_sort_isType. eapply type_it_mkProd_or_LetIn_sorts; tea. - exact X3.π2. } - have wfps : wf_universe Σ ps. + exact X3.2.π2.1. } + have wfps : wf_sort Σ ps. { pcuic. } eapply typing_spine_strengthen; tea. 2:{ rewrite /predctx /case_predicate_context /case_predicate_context_gen. @@ -369,7 +310,7 @@ Section Validity. unshelve eapply isType_mkApps_Ind_inv in X2 as [parsubst [argsubst [sppar sparg lenpars lenargs cu]]]; eauto. 2:eapply isdecl.p1. - eapply infer_typing_sort_impl with _ isdecl'; intros Hs. + eapply lift_typing_fu_impl with (f := fun t => _ (_ t)) (1 := isdecl') => // ?? Hs. unshelve epose proof (isdecl_ := declared_projection_to_gen isdecl); eauto. eapply (typing_subst_instance_decl _ _ _ _ _ _ _ wf isdecl_.p1.p1.p1) in Hs; eauto. simpl in Hs. @@ -381,7 +322,7 @@ Section Validity. eapply (substitution (Δ := [_]) sppar) in Hs. simpl in Hs. eapply (substitution (Γ' := [_]) (s := [c]) (Δ := [])) in Hs. - simpl in Hs. rewrite (subst_app_simpl [_]) /=. eassumption. + simpl in Hs. rewrite !(subst_app_simpl [c] (List.rev _)) /=. eassumption. constructor. constructor. simpl. rewrite subst_empty. rewrite subst_instance_mkApps subst_mkApps /=. @@ -392,27 +333,26 @@ Section Validity. assumption. - (* Fix *) - eapply nth_error_all in X0 as [s Hs]; eauto. - pcuic. + eapply nth_error_all in X0; eauto. - (* CoFix *) - eapply nth_error_all in X0 as [s Hs]; pcuic. + eapply nth_error_all in X0; eauto. - (* Primitive *) depelim X0; depelim X1; simp prim_type; cbn in *. - 1-2:destruct H1 as [hty hbod huniv]; exists (Universe.type0)@[[]]; change (tSort (Universe.type0)@[[]]) with (tSort Universe.type0)@[[]]; - rewrite -hty; refine (type_Const _ _ _ [] _ wfΓ H0 _); rewrite huniv //. - set (s := (Universe.make (array_level a))). + 1-2:destruct H1 as [hty hbod huniv]; eapply has_sort_isType with (s := _@[[]]); change (tSort ?s@[[]]) with (tSort s)@[[]]; + rewrite <- hty; refine (type_Const _ _ _ [] _ wfΓ H0 _); rewrite huniv //. + set (s := sType (Universe.make' (array_level a))). destruct H1 as [hty' hbod huniv]. - exists s. + eapply has_sort_isType with s. eapply (type_App _ _ _ _ _ (tSort s)); tea; cycle 1. + eapply (type_Const _ _ _ [array_level a]) in H0; tea. rewrite hty' in H0. cbn in H0. exact H0. - red. rewrite huniv. simpl. rtoProp; intuition eauto. eapply LevelSet.mem_spec. eapply (wfl (array_level a, 0)). lsets. + red. rewrite huniv. simpl. rtoProp; intuition eauto. eapply LevelSet.mem_spec. eapply (wfl (array_level a, 0)). cbn. lsets. cbn. red. destruct check_univs => //. red. red. intros v H c. csets. - + econstructor. econstructor; eauto. econstructor; eauto. + + econstructor. 2: econstructor; eauto. repeat (eexists; tea). econstructor; eauto. - (* Conv *) - now exists s. + now eapply has_sort_isType with (s := s). Qed. End Validity. @@ -429,10 +369,7 @@ Lemma wf_local_validity `{cf : checker_flags} {Σ} {wfΣ : wf Σ} Γ Δ : Proof. move=> wfΓΔ. apply: validity_wf_local. - enough (h : wf_local_rel Σ Γ Δ). - apply: All_local_env_impl; first exact h. - 1: move=> ?? [?|] //= ?; split=> //; apply: validity; eassumption. - by apply: (wf_local_app_inv _).2. + now apply: (All_local_env_app_inv _).2. Qed. @@ -478,7 +415,7 @@ Lemma type_App' {cf:checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} {Γ t na Σ;;; Γ |- u : A -> Σ;;; Γ |- tApp t u : B {0 := u}. Proof. intros Ht Hu. - have [s Hs] := validity Ht. + destruct (validity Ht) as (_ & s & Hs & _). eapply type_App; eauto. Qed. diff --git a/pcuic/theories/PCUICWcbvEval.v b/pcuic/theories/PCUICWcbvEval.v index e695cadf7..f88e4a378 100644 --- a/pcuic/theories/PCUICWcbvEval.v +++ b/pcuic/theories/PCUICWcbvEval.v @@ -782,7 +782,7 @@ Section Wcbv. move=> wfΣ Hc Hb. unshelve eapply declared_constant_to_gen in Hc; eauto. rewrite PCUICClosed.closedn_subst_instance. apply declared_decl_closed in Hc => //. simpl in Hc. red in Hc. - rewrite Hb in Hc. simpl in Hc. now move/andP: Hc. + rewrite Hb in Hc. simpl in Hc. now move/andP: Hc => []. Qed. Lemma closed_iota ci ind p c u args brs br : @@ -1445,4 +1445,4 @@ Proof. move: (skipn_nth_error array_value n). rewrite e => ->. constructor; eauto. + now cbn in i. -Qed. \ No newline at end of file +Qed. diff --git a/pcuic/theories/PCUICWeakeningConfig.v b/pcuic/theories/PCUICWeakeningConfig.v index 01ef96862..4c7255807 100644 --- a/pcuic/theories/PCUICWeakeningConfig.v +++ b/pcuic/theories/PCUICWeakeningConfig.v @@ -14,11 +14,11 @@ Set Default Goal Selector "!". (** ** Constraints *) -#[global] Instance subrel_config_impl_cmp cf1 cf2 pb cs : +#[global] Instance subrel_config_impl_cmp cf1 cf2 cs pb : config.impl cf1 cf2 -> - RelationClasses.subrelation (@compare_universe cf1 pb cs) (@compare_universe cf2 pb cs). + RelationClasses.subrelation (@compare_sort cf1 cs pb) (@compare_sort cf2 cs pb). Proof. - cbv [compare_universe eq_universe leq_universe leq_universe_n leq_universe_n_ eq_levelalg leq_levelalg_n eq_universe_ config.impl]. + cbv [compare_sort eq_sort eq_sort_ leq_sort leq_sort_n leq_sort_n_ eq_universe leq_universe_n config.impl]. destruct cf1, cf2; cbn. move => H u1 u2; move: H. repeat match goal with @@ -28,29 +28,29 @@ Proof. reflexivity. Qed. -#[global] Instance subrel_config_impl_eq_pb cf1 cf2 pb cs : +#[global] Instance subrel_config_impl_eq_pb cf1 cf2 cs pb : config.impl cf1 cf2 -> - RelationClasses.subrelation (@eq_universe cf1 cs) (@compare_universe cf2 pb cs). + RelationClasses.subrelation (@eq_sort cf1 cs) (@compare_sort cf2 cs pb). Proof. - change (@eq_universe ?cf) with (@compare_universe cf Conv). + change (@eq_sort ?cf ?φ) with (@compare_sort cf φ Conv). etransitivity; [ now eapply (@subrel_config_impl_cmp cf1 cf2) | ]. tc. Qed. #[global] Instance subrel_config_impl_eq cf1 cf2 u : config.impl cf1 cf2 -> - RelationClasses.subrelation (@eq_universe cf1 u) (@eq_universe cf2 u). -Proof. change (@eq_universe ?cf) with (@compare_universe cf Conv). tc. Qed. + RelationClasses.subrelation (@eq_sort cf1 u) (@eq_sort cf2 u). +Proof. change (@eq_sort ?cf ?φ) with (@compare_sort cf φ Conv). tc. Qed. #[global] Instance subrel_config_impl_le cf1 cf2 u : config.impl cf1 cf2 -> - RelationClasses.subrelation (@leq_universe cf1 u) (@leq_universe cf2 u). -Proof. change (@leq_universe ?cf) with (@compare_universe cf Cumul). tc. Qed. + RelationClasses.subrelation (@leq_sort cf1 u) (@leq_sort cf2 u). +Proof. change (@leq_sort ?cf ?φ) with (@compare_sort cf φ Cumul). tc. Qed. #[global] Instance subrel_config_impl_eq_le cf1 cf2 u : config.impl cf1 cf2 -> - RelationClasses.subrelation (@eq_universe cf1 u) (@leq_universe cf2 u). -Proof. change (@leq_universe ?cf) with (@compare_universe cf Cumul). tc. Qed. + RelationClasses.subrelation (@eq_sort cf1 u) (@leq_sort cf2 u). +Proof. change (@leq_sort ?cf ?φ) with (@compare_sort cf φ Cumul). tc. Qed. Lemma weakening_config_is_allowed_elimination cf1 cf2 cs u allowed : config.impl cf1 cf2 -> diff --git a/pcuic/theories/PCUICWeakeningEnv.v b/pcuic/theories/PCUICWeakeningEnv.v index 997960e6d..a7123e36f 100644 --- a/pcuic/theories/PCUICWeakeningEnv.v +++ b/pcuic/theories/PCUICWeakeningEnv.v @@ -67,66 +67,76 @@ Proof. apply global_ext_constraints_app, sub. Qed. -#[global] Instance subrel_extends_cmp {cf} pb (Σ Σ' : global_env) (ϕ : universes_decl) : +#[global] Instance subrel_extends_cmp_universe {cf} pb (Σ Σ' : global_env) (ϕ : universes_decl) : extends Σ Σ' -> - RelationClasses.subrelation (compare_universe pb (global_ext_constraints (Σ, ϕ))) - (compare_universe pb (global_ext_constraints (Σ', ϕ))). + RelationClasses.subrelation (compare_universe (global_ext_constraints (Σ, ϕ)) pb) + (compare_universe (global_ext_constraints (Σ', ϕ)) pb). Proof. intros ext u u'. apply cmp_universe_subset. apply weakening_env_global_ext_constraints, ext. Qed. +#[global] Instance subrel_extends_cmp {cf} pb (Σ Σ' : global_env) (ϕ : universes_decl) : + extends Σ Σ' -> + RelationClasses.subrelation (compare_sort (global_ext_constraints (Σ, ϕ)) pb) + (compare_sort (global_ext_constraints (Σ', ϕ)) pb). +Proof. + intros ext u u'. + apply cmp_sort_subset. + apply weakening_env_global_ext_constraints, ext. +Qed. + #[global] Instance subrel_extends_eq_pb {cf} pb (Σ Σ' : global_env) (ϕ : universes_decl) : extends Σ Σ' -> - RelationClasses.subrelation (eq_universe (global_ext_constraints (Σ, ϕ))) - (compare_universe pb (global_ext_constraints (Σ', ϕ))). + RelationClasses.subrelation (eq_sort (global_ext_constraints (Σ, ϕ))) + (compare_sort (global_ext_constraints (Σ', ϕ)) pb). Proof. - change eq_universe with (compare_universe Conv). + change (eq_sort ?φ) with (compare_sort φ Conv). intros ext. destruct pb. - tc. - - transitivity (compare_universe Conv (global_ext_constraints (Σ', ϕ))); tc. + - transitivity (compare_sort (global_ext_constraints (Σ', ϕ)) Conv); tc. Qed. #[global] Instance subrel_extends_eq {cf} (Σ Σ' : global_env) (ϕ : universes_decl) : extends Σ Σ' -> - RelationClasses.subrelation (eq_universe (global_ext_constraints (Σ, ϕ))) - (eq_universe (global_ext_constraints (Σ', ϕ))). -Proof. change eq_universe with (compare_universe Conv). tc. Qed. + RelationClasses.subrelation (eq_sort (global_ext_constraints (Σ, ϕ))) + (eq_sort (global_ext_constraints (Σ', ϕ))). +Proof. change (eq_sort ?φ) with (compare_sort φ Conv). tc. Qed. #[global] Instance subrel_extends_le {cf} (Σ Σ' : global_env) (ϕ : universes_decl) : extends Σ Σ' -> - RelationClasses.subrelation (leq_universe (global_ext_constraints (Σ, ϕ))) - (leq_universe (global_ext_constraints (Σ', ϕ))). -Proof. change leq_universe with (compare_universe Cumul). tc. Qed. + RelationClasses.subrelation (leq_sort (global_ext_constraints (Σ, ϕ))) + (leq_sort (global_ext_constraints (Σ', ϕ))). +Proof. change (leq_sort ?φ) with (compare_sort φ Cumul). tc. Qed. #[global] Instance subrel_extends_eq_le {cf} (Σ Σ' : global_env) (ϕ : universes_decl) : extends Σ Σ' -> - RelationClasses.subrelation (eq_universe (global_ext_constraints (Σ, ϕ))) - (leq_universe (global_ext_constraints (Σ', ϕ))). -Proof. change leq_universe with (compare_universe Cumul). tc. Qed. + RelationClasses.subrelation (eq_sort (global_ext_constraints (Σ, ϕ))) + (leq_sort (global_ext_constraints (Σ', ϕ))). +Proof. change (leq_sort ?φ) with (compare_sort φ Cumul). tc. Qed. Lemma subrelations_extends {cf} Σ Σ' φ : extends Σ Σ' -> - RelationClasses.subrelation (eq_universe (global_ext_constraints (Σ,φ))) (eq_universe (global_ext_constraints (Σ',φ))). + RelationClasses.subrelation (eq_sort (global_ext_constraints (Σ,φ))) (eq_sort (global_ext_constraints (Σ',φ))). Proof. typeclasses eauto. Qed. Lemma subrelations_leq_extends {cf} Σ Σ' φ : extends Σ Σ' -> - RelationClasses.subrelation (leq_universe (global_ext_constraints (Σ,φ))) (leq_universe (global_ext_constraints (Σ',φ))). + RelationClasses.subrelation (leq_sort (global_ext_constraints (Σ,φ))) (leq_sort (global_ext_constraints (Σ',φ))). Proof. typeclasses eauto. Qed. Lemma subrelations_compare_extends {cf} Σ Σ' pb φ : extends Σ Σ' -> - RelationClasses.subrelation (compare_universe pb (global_ext_constraints (Σ,φ))) - (compare_universe pb (global_ext_constraints (Σ',φ))). + RelationClasses.subrelation (compare_sort (global_ext_constraints (Σ,φ)) pb) + (compare_sort (global_ext_constraints (Σ',φ)) pb). Proof. destruct pb; typeclasses eauto. Qed. Lemma subrelations_eq_compare_extends {cf} Σ Σ' pb φ : extends Σ Σ' -> - RelationClasses.subrelation (eq_universe (global_ext_constraints (Σ,φ))) - (compare_universe pb (global_ext_constraints (Σ',φ))). + RelationClasses.subrelation (eq_sort (global_ext_constraints (Σ,φ))) + (compare_sort (global_ext_constraints (Σ',φ)) pb). Proof. destruct pb; typeclasses eauto. Qed. @@ -137,7 +147,7 @@ Lemma weakening_env_is_allowed_elimination {cf} Σ Σ' φ u allowed : Proof. destruct allowed; cbnr; trivial. intros ext [ | al]; auto. - destruct u; cbn in *; try elim al. + destruct u; cbn in *; try now right. right. unfold_univ_rel. apply al. @@ -184,7 +194,6 @@ Lemma extends_wf_universe {Σ : global_env_ext} Σ' u : extends Σ Σ' -> Proof. destruct Σ as [Σ univ]; cbn. intros [sub _]. - destruct u; simpl; auto. intros Hl. intros l inl; specialize (Hl l inl). cbn. @@ -195,6 +204,13 @@ Proof. - right. eapply global_levels_sub; tea. Qed. +Lemma extends_wf_sort {Σ : global_env_ext} Σ' s : extends Σ Σ' -> + wf_sort Σ s -> wf_sort (Σ', Σ.2) s. +Proof. + destruct s => //=. + apply extends_wf_universe. +Qed. + Definition on_udecl_prop (Σ : global_env) (udecl : universes_decl) := let levels := levels_of_udecl udecl in @@ -259,7 +275,7 @@ Qed. Section ExtendsWf. Context {cf : checker_flags}. Context {Pcmp: global_env_ext -> context -> conv_pb -> term -> term -> Type}. - Context {P: global_env_ext -> context -> term -> typ_or_sort -> Type}. + Context {P: global_env_ext -> context -> judgment -> Type}. Let wf := on_global_env Pcmp P. @@ -433,8 +449,8 @@ Definition weaken_env_prop_full_gen Definition weaken_env_prop_gen (R : global_env_ext -> global_env_ext -> Type) - (P : global_env_ext -> context -> term -> typ_or_sort -> Type) := - forall Σ Σ' φ, wf Σ -> wf Σ' -> R (Σ, φ) (Σ', φ) -> forall Γ t T, P (Σ, φ) Γ t T -> P (Σ', φ) Γ t T. + (P : global_env_ext -> context -> judgment -> Type) := + forall Σ Σ' φ, wf Σ -> wf Σ' -> R (Σ, φ) (Σ', φ) -> forall Γ j, P (Σ, φ) Γ j -> P (Σ', φ) Γ j. Definition weaken_env_prop_full := weaken_env_prop_full_gen extends. Definition weaken_env_decls_prop_full := weaken_env_prop_full_gen extends_decls. @@ -448,7 +464,7 @@ Definition weaken_env_strictly_on_decls_prop := weaken_env_prop_gen extends_stri Import CMorphisms CRelationClasses. #[global] Instance weaken_env_prop_gen_impl - : Proper (flip subrelation ==> (pointwise_relation _ (pointwise_relation _ (pointwise_relation _ (pointwise_relation _ iffT)))) ==> arrow)%signatureT weaken_env_prop_gen | 10. + : Proper (flip subrelation ==> (pointwise_relation _ (pointwise_relation _ (pointwise_relation _ iffT))) ==> arrow)%signatureT weaken_env_prop_gen | 10. Proof using Type. cbv -[weaken_env_prop_gen iffT]; cbv [weaken_env_prop_gen]; intros * H1 * H2 H3. unshelve @@ -461,7 +477,7 @@ Proof using Type. Qed. #[global] Instance Proper_weaken_env_prop_gen_respectful - : Proper (flip subrelation ==> (eq ==> eq ==> eq ==> eq ==> iffT) ==> arrow)%signatureT weaken_env_prop_gen | 10. + : Proper (flip subrelation ==> (eq ==> eq ==> eq ==> iffT) ==> arrow)%signatureT weaken_env_prop_gen | 10. Proof using Type. generalize weaken_env_prop_gen_impl; cbv -[weaken_env_prop_gen]; eauto. Qed. diff --git a/pcuic/theories/PCUICWellScopedCumulativity.v b/pcuic/theories/PCUICWellScopedCumulativity.v index a2fd67f45..24c749098 100644 --- a/pcuic/theories/PCUICWellScopedCumulativity.v +++ b/pcuic/theories/PCUICWellScopedCumulativity.v @@ -37,7 +37,7 @@ Implicit Types (cf : checker_flags) (Σ : global_env_ext). Inductive ws_cumul_pb {cf} (pb : conv_pb) (Σ : global_env_ext) (Γ : context) : term -> term -> Type := | ws_cumul_pb_compare (t u : term) : is_closed_context Γ -> is_open_term Γ t -> is_open_term Γ u -> - compare_term pb Σ.1 (global_ext_constraints Σ) t u -> Σ ;;; Γ ⊢ t ≤[pb] u + compare_term Σ.1 (global_ext_constraints Σ) pb t u -> Σ ;;; Γ ⊢ t ≤[pb] u | ws_cumul_pb_red_l (t u v : term) : is_closed_context Γ -> is_open_term Γ t -> is_open_term Γ u -> is_open_term Γ v -> @@ -125,7 +125,7 @@ Lemma ws_cumul_pb_alt `{cf : checker_flags} {pb} {Σ : global_env_ext} {wfΣ : w Σ ;;; Γ ⊢ t ≤[pb] u <~> ∑ v v', [× is_closed_context Γ, is_open_term Γ t, is_open_term Γ u, - red Σ Γ t v, red Σ Γ u v' & compare_term pb Σ (global_ext_constraints Σ) v v']. + red Σ Γ t v, red Σ Γ u v' & compare_term Σ (global_ext_constraints Σ) pb v v']. Proof. split. - induction 1. @@ -188,9 +188,9 @@ Proof. - now transitivity u'nf. } Qed. -Arguments wt_cumul_pb_dom {cf c Σ Γ T U}. -Arguments wt_cumul_pb_codom {cf c Σ Γ T U}. -Arguments wt_cumul_pb_eq {cf c Σ Γ T U}. +Arguments wt_cumul_pb_dom {cf pb Σ Γ T U}. +Arguments wt_cumul_pb_codom {cf pb Σ Γ T U}. +Arguments wt_cumul_pb_eq {cf pb Σ Γ T U}. Section EqualityLemmas. Context {cf : checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ}. @@ -299,7 +299,7 @@ Lemma ws_cumul_pb_alt_closed {cf} {pb} {Σ : global_env_ext} {wfΣ : wf Σ} Γ t Σ ;;; Γ ⊢ t ≤[pb] u <~> ∑ v v', [× closed_red Σ Γ t v, closed_red Σ Γ u v' & - compare_term pb Σ (global_ext_constraints Σ) v v']. + compare_term Σ (global_ext_constraints Σ) pb v v']. Proof. etransitivity. apply ws_cumul_pb_alt. split; intros (v & v' & cl); exists v, v'; intuition. @@ -510,9 +510,11 @@ Section WtContextConversion. Proof using Type. split. - induction 1; constructor; auto. - red in t0, t1. cbn. split; auto. + destruct t0 as (Hb & Ht). cbn in *. + split; tas. split; cbn; auto. - induction 1; [constructor|]. - destruct d as [na [b|] ty]; cbn in p; constructor; intuition auto. + destruct d as [na [b|] ty]; cbn in p; constructor; auto. + destruct p; split; tas. apply i. Qed. Lemma wt_cumul_ctx_pb_forget {pb} {Γ Γ' : context} : diff --git a/pcuic/theories/PCUICWfUniverses.v b/pcuic/theories/PCUICWfUniverses.v index 8f0afa819..63624b8ad 100644 --- a/pcuic/theories/PCUICWfUniverses.v +++ b/pcuic/theories/PCUICWfUniverses.v @@ -36,7 +36,7 @@ Qed. Section CheckerFlags. Context {cf:checker_flags}. - Lemma wf_universe_type0 Σ : wf_universe Σ Universe.type0. + Lemma wf_sort_type0 Σ : wf_sort Σ Sort.type0. Proof using Type. simpl. intros l hin%LevelExprSet.singleton_spec. @@ -44,7 +44,7 @@ Section CheckerFlags. apply global_ext_levels_InSet. Qed. - Lemma wf_universe_type1 Σ : wf_universe Σ Universe.type1. + Lemma wf_sort_type1 Σ : wf_sort Σ Sort.type1. Proof using Type. simpl. intros l hin%LevelExprSet.singleton_spec. @@ -52,10 +52,10 @@ Section CheckerFlags. apply global_ext_levels_InSet. Qed. - Lemma wf_universe_super {Σ u} : wf_universe Σ u -> wf_universe Σ (Universe.super u). + Lemma wf_sort_super {Σ u} : wf_sort Σ u -> wf_sort Σ (Sort.super u). Proof using Type. destruct u; cbn. - 1-2:intros _ l hin%LevelExprSet.singleton_spec; subst l; apply wf_universe_type1; + 1-2:intros _ l hin%LevelExprSet.singleton_spec; subst l; apply wf_sort_type1; now apply LevelExprSet.singleton_spec. intros Hl. intros l hin. @@ -63,8 +63,8 @@ Section CheckerFlags. simpl. now specialize (Hl _ int). Qed. - Lemma wf_universe_sup {Σ u u'} : wf_universe Σ u -> wf_universe Σ u' -> - wf_universe Σ (Universe.sup u u'). + Lemma wf_sort_sup {Σ u u'} : wf_sort Σ u -> wf_sort Σ u' -> + wf_sort Σ (Sort.sup u u'). Proof using Type. destruct u, u'; cbn; auto. intros Hu Hu' l [Hl|Hl]%LevelExprSet.union_spec. @@ -72,51 +72,50 @@ Section CheckerFlags. now apply (Hu' _ Hl). Qed. - Lemma wf_universe_product {Σ u u'} : wf_universe Σ u -> wf_universe Σ u' -> - wf_universe Σ (Universe.sort_of_product u u'). + Lemma wf_sort_product {Σ s s'} : wf_sort Σ s -> wf_sort Σ s' -> + wf_sort Σ (Sort.sort_of_product s s'). Proof using Type. - intros Hu Hu'. unfold Universe.sort_of_product. - destruct (Universe.is_prop u' || Universe.is_sprop u'); auto. - now apply wf_universe_sup. + intros Hu Hu'. + destruct s' => //=. + now apply wf_sort_sup. Qed. - Hint Resolve wf_universe_type1 wf_universe_super wf_universe_sup wf_universe_product : pcuic. + Hint Resolve wf_sort_type1 wf_sort_super wf_sort_sup wf_sort_product : pcuic. - Definition wf_universeb_level Σ l := + Definition wf_levelb Σ l := LevelSet.mem l (global_ext_levels Σ). - Definition wf_universe_level Σ l := + Definition wf_level Σ l := LevelSet.In l (global_ext_levels Σ). - Definition wf_universe_instance Σ u := - Forall (wf_universe_level Σ) u. + Definition wf_instance Σ u := + Forall (wf_level Σ) u. - Definition wf_universeb_instance Σ u := - forallb (wf_universeb_level Σ) u. + Definition wf_instanceb Σ u := + forallb (wf_levelb Σ) u. - Lemma wf_universe_levelP {Σ l} : reflect (wf_universe_level Σ l) (wf_universeb_level Σ l). + Lemma wf_levelP {Σ l} : reflect (wf_level Σ l) (wf_levelb Σ l). Proof using Type. - unfold wf_universe_level, wf_universeb_level. + unfold wf_level, wf_levelb. destruct LevelSet.mem eqn:ls; constructor. now apply LevelSet.mem_spec in ls. intros hin. now apply LevelSet.mem_spec in hin. Qed. - Lemma wf_universe_instanceP {Σ u} : reflect (wf_universe_instance Σ u) (wf_universeb_instance Σ u). + Lemma wf_instanceP {Σ u} : reflect (wf_instance Σ u) (wf_instanceb Σ u). Proof using Type. - unfold wf_universe_instance, wf_universeb_instance. - apply forallbP. intros x; apply wf_universe_levelP. + unfold wf_instance, wf_instanceb. + apply forallbP. intros x; apply wf_levelP. Qed. - Lemma wf_universe_subst_instance_univ (Σ : global_env_ext) univs u s : + Lemma wf_universe_subst_instance_univ (Σ : global_env_ext) univs ui u : wf Σ -> - wf_universe Σ s -> - wf_universe_instance (Σ.1, univs) u -> - wf_universe (Σ.1, univs) (subst_instance u s). + wf_universe Σ u -> + wf_instance (Σ.1, univs) ui -> + wf_universe (Σ.1, univs) (subst_instance ui u). Proof using Type. - destruct s as [| |t]; cbnr. intros wfΣ Hl Hu e [[l n] [inl ->]]%In_subst_instance. destruct l as [|s|n']; simpl; auto. - apply global_ext_levels_InSet. @@ -135,22 +134,32 @@ Section CheckerFlags. eapply nth_error_forall in Hu; eauto. apply global_ext_levels_InSet. * unfold subst_instance. simpl. - destruct (nth_error u n') eqn:hnth. + destruct (nth_error ui n') eqn:hnth. 2:{ simpl. rewrite hnth. apply global_ext_levels_InSet. } eapply nth_error_forall in Hu. 2:eauto. - change (nth_error u n') with (nth_error u n') in *. + change (nth_error ui n') with (nth_error ui n') in *. rewrite -> hnth. simpl. apply Hu. + now apply not_var_global_levels in Hl. Qed. - Lemma wf_universe_instantiate Σ univs s u φ : + Lemma wf_sort_subst_instance_sort (Σ : global_env_ext) univs u s : wf Σ -> - wf_universe (Σ, univs) s -> - wf_universe_instance (Σ, φ) u -> - wf_universe (Σ, φ) (subst_instance_univ u s). + wf_sort Σ s -> + wf_instance (Σ.1, univs) u -> + wf_sort (Σ.1, univs) (subst_instance u s). + Proof using Type. + destruct s as [| |t]; cbnr. + apply wf_universe_subst_instance_univ. + Qed. + + Lemma wf_sort_instantiate Σ univs s u φ : + wf Σ -> + wf_sort (Σ, univs) s -> + wf_instance (Σ, φ) u -> + wf_sort (Σ, φ) (subst_instance_sort u s). Proof using Type. intros wfΣ Hs. - apply (wf_universe_subst_instance_univ (Σ, univs) φ); auto. + apply (wf_sort_subst_instance_sort (Σ, univs) φ); auto. Qed. Lemma subst_instance_empty u : @@ -164,10 +173,10 @@ Section CheckerFlags. now destruct a => /= //; auto. Qed. - Lemma wf_universe_level_mono Σ u : + Lemma wf_level_mono Σ u : wf Σ -> on_udecl_prop Σ (Monomorphic_ctx) -> - Forall (wf_universe_level (Σ, Monomorphic_ctx)) u -> + Forall (wf_level (Σ, Monomorphic_ctx)) u -> forallb (fun x => ~~ Level.is_var x) u. Proof using Type. intros wf uprop. @@ -179,9 +188,9 @@ Section CheckerFlags. now pose proof (not_var_global_levels wf _ H). Qed. - Lemma wf_universe_level_sub Σ univs u : - wf_universe_level (Σ, Monomorphic_ctx) u -> - wf_universe_level (Σ, univs) u. + Lemma wf_level_sub Σ univs u : + wf_level (Σ, Monomorphic_ctx) u -> + wf_level (Σ, univs) u. Proof using cf. intros wfx. red in wfx |- *. @@ -190,14 +199,14 @@ Section CheckerFlags. eapply LevelSet.union_spec. now right. Qed. - Lemma wf_universe_instance_sub Σ univs u : - wf_universe_instance (Σ, Monomorphic_ctx) u -> - wf_universe_instance (Σ, univs) u. + Lemma wf_instance_sub Σ univs u : + wf_instance (Σ, Monomorphic_ctx) u -> + wf_instance (Σ, univs) u. Proof using cf. intros wfu. red in wfu |- *. eapply Forall_impl; eauto. - intros. red in H. cbn in H. eapply wf_universe_level_sub; eauto. + intros. red in H. cbn in H. eapply wf_level_sub; eauto. Qed. Lemma In_Level_global_ext_poly s Σ cst : @@ -217,10 +226,10 @@ Section CheckerFlags. intros x' [->|inx]; auto. Qed. - Lemma wf_universe_instance_In {Σ u} : wf_universe_instance Σ u <-> + Lemma wf_instance_In {Σ u} : wf_instance Σ u <-> (forall l, In l u -> LS.In l (global_ext_levels Σ)). Proof using Type. - unfold wf_universe_instance. + unfold wf_instance. split; intros. eapply Forall_In in H; eauto. apply In_Forall. auto. Qed. @@ -236,21 +245,21 @@ Section CheckerFlags. specialize (IHu' H). intuition auto. Qed. - Lemma wf_universe_subst_instance Σ univs u u' φ : + Lemma wf_instance_subst_instance Σ univs u u' φ : wf Σ -> on_udecl_prop Σ univs -> - wf_universe_instance (Σ, univs) u' -> - wf_universe_instance (Σ, φ) u -> - wf_universe_instance (Σ, φ) (subst_instance u u'). + wf_instance (Σ, univs) u' -> + wf_instance (Σ, φ) u -> + wf_instance (Σ, φ) (subst_instance u u'). Proof using Type. intros wfΣ onup Hs cu. destruct univs. - red in Hs |- *. - unshelve epose proof (wf_universe_level_mono _ _ _ _ Hs); eauto. + unshelve epose proof (wf_level_mono _ _ _ _ Hs); eauto. eapply forallb_Forall in H. apply Forall_map. solve_all. destruct x; simpl => //. red. apply global_ext_levels_InSet. - eapply wf_universe_level_sub; eauto. + eapply wf_level_sub; eauto. - clear onup. red in Hs |- *. eapply Forall_map, Forall_impl; eauto. @@ -274,34 +283,30 @@ Section CheckerFlags. Section WfUniverses. Context (Σ : global_env_ext). - Definition wf_universeb (s : Universe.t) : bool := - match s with - | Universe.lType l => LevelExprSet.for_all (fun l => LevelSet.mem (LevelExpr.get_level l) (global_ext_levels Σ)) l - | _ => true - end. + Definition wf_universeb (u : Universe.t) : bool := + LevelExprSet.for_all (fun l => LevelSet.mem (LevelExpr.get_level l) (global_ext_levels Σ)) u. Lemma wf_universe_reflect {u : Universe.t} : reflect (wf_universe Σ u) (wf_universeb u). Proof using Type. - destruct u; simpl; try now constructor. eapply iff_reflect. rewrite LevelExprSet.for_all_spec. split; intros. - intros l Hl; specialize (H l Hl). now eapply LS.mem_spec. - - specialize (H l H0). simpl in H. + - intros l Hl. specialize (H l Hl). now eapply LS.mem_spec in H. Qed. Fixpoint on_universes fu fc t := match t with - | tSort s => fu s + | tSort s => Sort.on_sort fu true s | tApp t u | tProd _ t u | tLambda _ t u => on_universes fu fc t && on_universes fu fc u | tCase _ p c brs => [&& - forallb fu (map Universe.make p.(puinst)) , + forallb fu (map Universe.make' p.(puinst)) , forallb (on_universes fu fc) p.(pparams) , test_context (fc #|p.(puinst)|) p.(pcontext) , on_universes fu fc p.(preturn) , @@ -313,16 +318,24 @@ Section CheckerFlags. | tFix mfix _ | tCoFix mfix _ => forallb (fun d => on_universes fu fc d.(dtype) && on_universes fu fc d.(dbody)) mfix | tConst _ u | tInd _ u | tConstruct _ _ u => - forallb fu (map Universe.make u) + forallb fu (map Universe.make' u) | tEvar _ args => forallb (on_universes fu fc) args - | tPrim p => test_primu (fun x => fu (Universe.make x)) (on_universes fu fc) p + | tPrim p => test_primu (fun x => fu (Universe.make' x)) (on_universes fu fc) p | _ => true end. Definition wf_universes t := on_universes wf_universeb closedu t. + Definition wf_sortb s := Sort.on_sort wf_universeb true s. + + Lemma wf_sort_reflect {s : sort} : + reflect (wf_sort Σ s) (wf_sortb s). + Proof using Type. + destruct s => //=; repeat constructor. + apply wf_universe_reflect. + Qed. Lemma wf_universeb_instance_forall u : - forallb wf_universeb (map Universe.make u) = wf_universeb_instance Σ u. + forallb wf_universeb (map Universe.make' u) = wf_instanceb Σ u. Proof using Type. induction u => //=. rewrite IHu. @@ -430,64 +443,68 @@ Qed. repeat match goal with | [ H: is_true (wf_universeb _ ?x) |- _ ] => apply (elimT (@wf_universe_reflect _ x)) in H | [ |- is_true (wf_universeb _ ?x) ] => apply (introT (@wf_universe_reflect _ x)) + | [ H: is_true (Sort.on_sort (wf_universeb _) _ ?x) |- _ ] => apply (elimT (@wf_sort_reflect _ x)) in H + | [ |- is_true (Sort.on_sort (wf_universeb _) _ ?x) ] => apply (introT (@wf_sort_reflect _ x)) + | [ H: is_true (wf_sortb _ ?x) |- _ ] => apply (elimT (@wf_sort_reflect _ x)) in H + | [ |- is_true (wf_sortb _ ?x) ] => apply (introT (@wf_sort_reflect _ x)) end. Lemma wf_universes_inst {Σ : global_env_ext} univs t u : wf Σ -> on_udecl_prop Σ.1 univs -> - wf_universe_instance Σ u -> + wf_instance Σ u -> wf_universes (Σ.1, univs) t -> wf_universes Σ (subst_instance u t). Proof using Type. intros wfΣ onudecl cu wft. induction t using term_forall_list_ind; simpl in *; auto; try to_prop; - try apply /andP; to_wfu; intuition eauto 4. + try apply /andP; intuition eauto 4. - all:cbn -[Universe.make] in * ; autorewrite with map; repeat (f_equal; solve_all). + all:cbn -[Universe.make'] in * ; to_wfu; autorewrite with map; repeat (f_equal; solve_all). - - to_wfu. destruct Σ as [Σ univs']. simpl in *. - eapply (wf_universe_subst_instance_univ (Σ, univs)); auto. + - destruct Σ as [Σ univs']. simpl in *. + eapply (wf_sort_subst_instance_sort (Σ, univs)); auto. - apply forallb_All. rewrite -forallb_map wf_universeb_instance_forall. apply All_forallb in wft. rewrite -forallb_map wf_universeb_instance_forall in wft. - apply/wf_universe_instanceP. - eapply wf_universe_subst_instance; eauto. + apply/wf_instanceP. + eapply wf_instance_subst_instance; eauto. destruct Σ; simpl in *. - now move/wf_universe_instanceP: wft. + now move/wf_instanceP: wft. - apply forallb_All. rewrite -forallb_map wf_universeb_instance_forall. apply All_forallb in wft. rewrite -forallb_map wf_universeb_instance_forall in wft. - apply/wf_universe_instanceP. - eapply wf_universe_subst_instance; eauto. + apply/wf_instanceP. + eapply wf_instance_subst_instance; eauto. destruct Σ; simpl in *. - now move/wf_universe_instanceP: wft. + now move/wf_instanceP: wft. - apply forallb_All. rewrite -forallb_map wf_universeb_instance_forall. apply All_forallb in wft. rewrite -forallb_map wf_universeb_instance_forall in wft. - apply/wf_universe_instanceP. - eapply wf_universe_subst_instance; eauto. + apply/wf_instanceP. + eapply wf_instance_subst_instance; eauto. destruct Σ; simpl in *. - now move/wf_universe_instanceP: wft. + now move/wf_instanceP: wft. - apply forallb_All. rewrite -forallb_map wf_universeb_instance_forall. apply All_forallb in H. rewrite -forallb_map wf_universeb_instance_forall in H. - apply/wf_universe_instanceP. - eapply wf_universe_subst_instance; eauto. + apply/wf_instanceP. + eapply wf_instance_subst_instance; eauto. destruct Σ ; simpl in *. - now move/wf_universe_instanceP: H. + now move/wf_instanceP: H. - now len. - rewrite /test_branch. rtoProp. move/andP: a => [] tctx wfu. split; auto. simpl. solve_all. now len. - - rewrite -subst_instance_univ_make. to_wfu. + - rewrite -subst_instance_universe_make. to_wfu. eapply (wf_universe_subst_instance_univ (Σ.1, univs)) => //. Qed. @@ -496,19 +513,18 @@ Qed. wf_universe (Σ', Σ.2) t. Proof using Type. intros wfΣ ext. - destruct t; simpl; auto. intros Hl l inl; specialize (Hl l inl). apply LS.union_spec. apply LS.union_spec in Hl as [Hl|Hl]; simpl. left; auto. right. now eapply global_levels_sub; [apply ext|]. Qed. - Lemma weaken_wf_universe_level {Σ : global_env_ext} Σ' t : wf Σ -> wf Σ' -> extends Σ Σ' -> - wf_universe_level Σ t -> - wf_universe_level (Σ', Σ.2) t. + Lemma weaken_wf_level {Σ : global_env_ext} Σ' t : wf Σ -> wf Σ' -> extends Σ Σ' -> + wf_level Σ t -> + wf_level (Σ', Σ.2) t. Proof using Type. intros wfΣ wfΣ' ext. - unfold wf_universe_level. + unfold wf_level. destruct t; simpl; auto using global_ext_levels_InSet; intros; apply LS.union_spec. - eapply LS.union_spec in H as [H|H]. @@ -518,30 +534,30 @@ Qed. - cbn. eapply in_var_global_ext in H; eauto. Qed. - Lemma weaken_wf_universe_instance {Σ : global_env_ext} Σ' t : wf Σ -> wf Σ' -> extends Σ.1 Σ' -> - wf_universe_instance Σ t -> - wf_universe_instance (Σ', Σ.2) t. + Lemma weaken_wf_instance {Σ : global_env_ext} Σ' t : wf Σ -> wf Σ' -> extends Σ.1 Σ' -> + wf_instance Σ t -> + wf_instance (Σ', Σ.2) t. Proof using Type. intros wfΣ wfΣ' ext. - unfold wf_universe_instance. + unfold wf_instance. intros H; eapply Forall_impl; eauto. - intros. now eapply weaken_wf_universe_level. + intros. now eapply weaken_wf_level. Qed. - Arguments Universe.make : simpl never. -Lemma test_primu_test_primu_tPrimProp {P : term -> Type} {pu put} {pu' : Level.t -> bool} {put' : term -> bool} p : - tPrimProp P p -> test_primu pu put p -> - (forall u, pu u -> pu' u) -> - (forall t, P t -> put t -> put' t) -> - test_primu pu' put' p. -Proof. - intros hp ht hf hg. - destruct p as [? []]; cbn => //. - destruct a; destruct hp; cbn in *. - rtoProp. destruct p0. intuition eauto. - eapply forallb_All in H2. eapply All_prod in a; tea. - eapply All_forallb, All_impl; tea. intuition eauto. apply hg; intuition auto. -Qed. + Arguments Universe.make' : simpl never. + Lemma test_primu_test_primu_tPrimProp {P : term -> Type} {pu put} {pu' : Level.t -> bool} {put' : term -> bool} p : + tPrimProp P p -> test_primu pu put p -> + (forall u, pu u -> pu' u) -> + (forall t, P t -> put t -> put' t) -> + test_primu pu' put' p. + Proof. + intros hp ht hf hg. + destruct p as [? []]; cbn => //. + destruct a; destruct hp; cbn in *. + rtoProp. destruct p0. intuition eauto. + eapply forallb_All in H2. eapply All_prod in a; tea. + eapply All_forallb, All_impl; tea. intuition eauto. apply hg; intuition auto. + Qed. Lemma weaken_wf_universes {Σ : global_env_ext} Σ' t : wf Σ -> wf Σ' -> extends Σ.1 Σ' -> wf_universes Σ t -> @@ -552,7 +568,8 @@ Qed. try apply /andP; to_wfu; intuition eauto 4. - solve_all. - - now eapply weaken_wf_universe. + - move: H. destruct s => //=. + now apply weaken_wf_universe. - eapply forallb_impl ; tea. now move => ? _ /wf_universe_reflect /weaken_wf_universe /wf_universe_reflect. - eapply forallb_impl ; tea. @@ -572,8 +589,9 @@ Qed. intros; to_wfu; now eapply weaken_wf_universe. Qed. - Lemma wf_universes_weaken_full : weaken_env_prop_full cumulSpec0 (lift_typing typing) (fun Σ Γ t T => - wf_universes Σ t && wf_universes Σ T). + Lemma wf_universes_weaken_full : + weaken_env_prop_full cumulSpec0 (lift_typing typing) + (fun Σ _ t T => wf_universes Σ t && wf_universes Σ T). Proof using Type. do 2 red. intros. to_prop; apply /andP; split; now apply weaken_wf_universes. @@ -581,16 +599,17 @@ Qed. Lemma wf_universes_weaken : weaken_env_prop cumulSpec0 (lift_typing typing) - (lift_typing (fun Σ Γ (t T : term) => - wf_universes Σ t && wf_universes Σ T)). + (fun Σ _ j => option_default (wf_universes Σ) (j_term j) true && wf_universes Σ (j_typ j) && option_default (wf_sortb Σ) (j_univ j) true). Proof using Type. - intros Σ Σ' φ wfΣ wfΣ' Hext Γ t T HT. - apply lift_typing_impl with (1 := HT); intros ? Hty. - now eapply (wf_universes_weaken_full (Σ, _)). + intros Σ Σ' φ wfΣ wfΣ' Hext _ j Hj. + pose proof (fun t => @weaken_wf_universes (Σ, φ) Σ' t wfΣ wfΣ' Hext). + rtoProp; repeat split; auto. + 1: destruct j_term => //; cbn in *; auto. + 1: destruct j_univ => //; cbn in *; now apply (H (tSort _)). Qed. Lemma wf_universes_inds Σ mind u bodies : - wf_universe_instance Σ u -> + wf_instance Σ u -> All (fun t : term => wf_universes Σ t) (inds mind u bodies). Proof using Type. intros wfu. @@ -600,7 +619,7 @@ Qed. constructor; auto. cbn. rewrite wf_universeb_instance_forall. - now apply /wf_universe_instanceP. + now apply /wf_instanceP. Qed. Lemma wf_universes_mkApps Σ f args : @@ -610,24 +629,16 @@ Qed. now rewrite mkApps_app forallb_app /= andb_true_r andb_assoc -IHargs. Qed. - Lemma type_local_ctx_wf Σ Γ Δ s : type_local_ctx - (lift_typing - (fun (Σ : PCUICEnvironment.global_env_ext) - (_ : PCUICEnvironment.context) (t T : term) => - wf_universes Σ t && wf_universes Σ T)) Σ Γ Δ s -> + Lemma type_local_ctx_wf Σ Γ Δ s : type_local_ctx (fun Σ _ => lift_wf_term (fun t => wf_universes Σ t)) Σ Γ Δ s -> All (fun d => option_default (wf_universes Σ) (decl_body d) true && wf_universes Σ (decl_type d)) Δ. Proof using Type. - induction Δ as [|[na [b|] ty] ?]; simpl; constructor; auto. - simpl. - destruct X as [? [? ?]]. now to_prop. - apply IHΔ. apply X. - simpl. - destruct X as [? ?]. now to_prop. - apply IHΔ. apply X. + induction Δ as [|[na [b|] ty] ?]; simpl. + 1: constructor. + all: intros [X [Hb Ht]]; constructor; cbn; auto. Qed. Lemma consistent_instance_ext_wf Σ univs u : consistent_instance_ext Σ univs u -> - wf_universe_instance Σ u. + wf_instance Σ u. Proof using Type. destruct univs; simpl. - destruct u => // /=. @@ -642,52 +653,44 @@ Qed. | [ H : on_udecl _ _, H' : on_udecl _ _ -> _ |- _ ] => specialize (H' H) end. - Local Lemma wf_sorts_local_ctx_smash (Σ : global_env_ext) mdecl args sorts : - sorts_local_ctx - (lift_typing - (fun (Σ : PCUICEnvironment.global_env_ext) - (_ : PCUICEnvironment.context) (t T : term) => - wf_universes Σ t && wf_universes Σ T)) (Σ.1, ind_universes mdecl) - (arities_context (ind_bodies mdecl),,, ind_params mdecl) - args sorts -> - sorts_local_ctx - (lift_typing - (fun (Σ : PCUICEnvironment.global_env_ext) - (_ : PCUICEnvironment.context) (t T : term) => - wf_universes Σ t && wf_universes Σ T)) (Σ.1, ind_universes mdecl) - (arities_context (ind_bodies mdecl),,, ind_params mdecl) - (smash_context [] args) sorts. + Local Lemma wf_universes_local_ctx_smash (Σ : global_env_ext) mdecl args sorts : + sorts_local_ctx (fun Σ _ => lift_wfbu_term (fun t => wf_universes Σ t) (wf_sortb Σ)) + (Σ.1, ind_universes mdecl) (arities_context (ind_bodies mdecl),,, ind_params mdecl) args sorts -> + sorts_local_ctx (fun Σ _ => lift_wfbu_term (fun t => wf_universes Σ t) (wf_sortb Σ)) + (Σ.1, ind_universes mdecl) (arities_context (ind_bodies mdecl),,, ind_params mdecl) (smash_context [] args) sorts. Proof using Type. induction args as [|[na [b|] ty] args] in sorts |- *; simpl; auto. - intros []. + intros [X Hj]. rewrite subst_context_nil. auto. - destruct sorts; auto. - intros []. + destruct sorts as [|u]; auto. + intros [X Hj]. rewrite smash_context_acc /=. split. eauto. + apply lift_wfbu_term_f_impl with (f := fun t => _ (_ t)) (1 := Hj) (fu := id) => // t Ht. rewrite wf_universes_subst. - clear -s. generalize 0. - induction args as [|[na [b|] ty] args] in sorts, s |- *; simpl in *; auto. - - destruct s as [? [[s' wf] [? ?]%andb_and]]. + clear -X. generalize 0. + induction args as [|[na [b|] ty] args] in sorts, X |- *; simpl in *; auto. + - destruct X as [? Hj]. constructor; eauto. rewrite wf_universes_subst. eapply IHargs; eauto. + move: Hj => /andP[] /andP[] /= Htm _ _. now rewrite wf_universes_lift. - - destruct sorts => //. destruct s. + - destruct sorts => //. destruct X. constructor => //. eapply IHargs; eauto. - now rewrite wf_universes_lift. Qed. - Lemma wf_sorts_local_ctx_nth_error Σ P Γ Δ s n d : - sorts_local_ctx P Σ Γ Δ s -> + Lemma wf_universes_local_ctx_nth_error Σ P Pu Γ Δ s n d : + sorts_local_ctx (fun Σ _ => lift_wfbu_term (P Σ) (Pu Σ)) Σ Γ Δ s -> nth_error Δ n = Some d -> - ∑ Γ' t, P Σ Γ' (decl_type d) t. + P Σ (decl_type d). Proof using Type. induction Δ as [|[na [b|] ty] Δ] in n, s |- *; simpl; auto. - now rewrite nth_error_nil. - - intros [h [h' h'']]. - destruct n. simpl. move=> [= <-] /=. do 2 eexists; eauto. + - move => [] h /andP[] /andP[] _ h' _. + destruct n. simpl. move=> [= <-] //=. now simpl; eapply IHΔ. - - destruct s => //. intros [h h']. - destruct n. simpl. move=> [= <-] /=. eexists; eauto. + - destruct s => //. move => [] h /andP[] /andP[] _ h' _. + destruct n. simpl. move=> [= <-] //=. now simpl; eapply IHΔ. Qed. @@ -708,7 +711,7 @@ Qed. Qed. Lemma wf_abstract_instance Σ decl : - wf_universe_instance (Σ, decl) (abstract_instance decl). + wf_instance (Σ, decl) (abstract_instance decl). Proof using Type. destruct decl as [|[u cst]]=> /=. red. constructor. @@ -747,15 +750,6 @@ Qed. Qed. - Lemma test_context_app p Γ Δ : - test_context p (Γ ,,, Δ) = test_context p Γ && test_context p Δ. - Proof using Type. - induction Δ; simpl; auto. - - now rewrite andb_true_r. - - now rewrite IHΔ andb_assoc. - Qed. - - Lemma wf_universes_it_mkLambda_or_LetIn {Σ Γ T} : wf_universes Σ (it_mkLambda_or_LetIn Γ T) = test_context (wf_universes Σ) Γ && wf_universes Σ T. Proof using Type. @@ -784,7 +778,7 @@ Qed. Qed. Lemma closedu_compare_decls k Γ Δ : - All2 (PCUICEquality.compare_decls eq eq) Γ Δ -> + PCUICEquality.eq_context_upto_names Γ Δ -> test_context (closedu k) Γ = test_context (closedu k) Δ. Proof using Type. induction 1; cbn; auto. @@ -916,12 +910,12 @@ Qed. apply closedu_smash_context_gen => //. Qed. - Lemma wf_universe_level_closed {Σ : global_env} {wfΣ : wf Σ} univs u : + Lemma wf_level_closed {Σ : global_env} {wfΣ : wf Σ} univs l : on_udecl_prop Σ univs -> - wf_universe_level (Σ, univs) u -> closedu_level #|polymorphic_instance univs| u. + wf_level (Σ, univs) l -> closedu_level #|polymorphic_instance univs| l. Proof using Type. - intros ond Ht; destruct u => //. - cbn in Ht. unfold closedu_universe, closedu_universe_levels. + intros ond Ht; destruct l => //. + cbn in Ht. unfold closedu_level. cbn. red in Ht. eapply in_var_global_ext in Ht => //. cbn in Ht. @@ -938,30 +932,38 @@ Qed. Lemma wf_universe_closed {Σ : global_env} {wfΣ : wf Σ} univs u : on_udecl_prop Σ univs -> - wf_universe (Σ, univs) u -> closedu #|polymorphic_instance univs| (tSort u). + wf_universe (Σ, univs) u -> closedu_universe #|polymorphic_instance univs| u. Proof using Type. - intros ond Ht; destruct u => //. - cbn in Ht. unfold closedu_universe, closedu_universe_levels. + intros ond Ht. + cbn in Ht. unfold closedu_universe. eapply LevelExprSet.for_all_spec. intros x y ?; subst; auto. intros i hi. specialize (Ht i hi). unfold closedu_level_expr. - apply wf_universe_level_closed => //. + apply wf_level_closed => //. Qed. - Lemma wf_universe_instance_closed {Σ : global_env} {wfΣ : wf Σ} {univs u} : + Lemma wf_sort_closed {Σ : global_env} {wfΣ : wf Σ} univs s : on_udecl_prop Σ univs -> - wf_universe_instance (Σ, univs) u -> + wf_sort (Σ, univs) s -> closedu_sort #|polymorphic_instance univs| s. + Proof using Type. + destruct s => //=. + apply wf_universe_closed. + Qed. + + Lemma wf_instance_closed {Σ : global_env} {wfΣ : wf Σ} {univs u} : + on_udecl_prop Σ univs -> + wf_instance (Σ, univs) u -> closedu_instance #|polymorphic_instance univs| u. Proof using Type. intros ond Ht. red in Ht. unfold closedu_instance. solve_all. - now eapply wf_universe_level_closed. + now eapply wf_level_closed. Qed. - Lemma wf_universe_make Σ u : wf_universe Σ (Universe.make u) -> wf_universe_level Σ u. + Lemma wf_universe_make Σ u : wf_universe Σ (Universe.make' u) -> wf_level Σ u. Proof. - rewrite /wf_universe /= => hl; rewrite /wf_universe_level. + rewrite /wf_universe /= => hl; rewrite /wf_level. apply (hl (u, 0)). lsets. Qed. @@ -970,30 +972,31 @@ Qed. wf_universes (Σ, univs) t -> closedu #|polymorphic_instance univs| t. Proof using Type. intros ond. induction t using term_forall_list_ind; cbn => //; solve_all. - - apply wf_universe_closed => //. + - apply wf_sort_closed => //. + destruct s => //=. now move/wf_universe_reflect: H. - - eapply wf_universe_instance_closed => //. + - eapply wf_instance_closed => //. apply All_forallb in H. rewrite -forallb_map wf_universeb_instance_forall in H. - now move/wf_universe_instanceP: H. - - eapply wf_universe_instance_closed => //. + now move/wf_instanceP: H. + - eapply wf_instance_closed => //. apply All_forallb in H. rewrite -forallb_map wf_universeb_instance_forall in H. - now move/wf_universe_instanceP: H. - - eapply wf_universe_instance_closed => //. + now move/wf_instanceP: H. + - eapply wf_instance_closed => //. apply All_forallb in H. rewrite -forallb_map wf_universeb_instance_forall in H. - now move/wf_universe_instanceP: H. + now move/wf_instanceP: H. - unfold test_predicate_ku in *; solve_all. - eapply wf_universe_instance_closed => //. + eapply wf_instance_closed => //. apply All_forallb in H0. rewrite -forallb_map wf_universeb_instance_forall in H0. - now move/wf_universe_instanceP: H0. + now move/wf_instanceP: H0. - unfold test_branch in *; solve_all. - unfold test_def in *; solve_all. - unfold test_def in *; solve_all. - eapply test_primu_test_primu_tPrimProp; tea; cbn; eauto. - intros. to_wfu. eapply wf_universe_level_closed; tea. + intros. to_wfu. eapply wf_level_closed; tea. now apply wf_universe_make. Qed. @@ -1063,68 +1066,51 @@ Qed. Qed. Theorem wf_types : - env_prop (fun Σ Γ t T => - wf_universes Σ t && wf_universes Σ T) - (fun Σ Γ => - All_local_env - (lift_typing (fun (Σ : global_env_ext) (Γ : context) (t T : term) => - wf_universes Σ t && wf_universes Σ T) Σ) Γ × - test_context (wf_universes Σ) Γ). + env_prop (fun Σ Γ t T => wf_universes Σ t && wf_universes Σ T) + (fun Σ _ j => option_default (wf_universes Σ) (j_term j) true && wf_universes Σ (j_typ j) && option_default (wf_sortb Σ) (j_univ j) true) + (fun Σ Γ => wf_ctx_universes Σ Γ). Proof using Type. - apply typing_ind_env; intros; rename_all_hyps; cbn; rewrite -!/(wf_universes _ _) ; + apply typing_ind_env; unfold lift_wfb_term; intros; rename_all_hyps; cbn; rewrite -!/(wf_universes _ _) ; specIH; to_prop; cbn; auto. - - split. - * induction X; constructor; auto. - destruct tu as [s tu]; exists s; simpl. - now simpl in Hs. - destruct tu as [s tu]; exists s; simpl. - now simpl in Hs. - * induction X; simpl; auto. - rewrite IHX /= /test_decl /=. now move/andP: Hs. + - destruct X as (Hb & s & (_ & (Ht & Hs)%andb_and) & e). + rewrite Ht andb_true_r. + rtoProp; split. + + destruct j_term => //. + now destruct Hb as (_ & (? & _)%andb_and). + + destruct j_univ => //. rewrite e //. + + - apply All_local_env_cst, All_forallb in X0. + apply forallb_impl with (2 := X0) => [] [na bo ty] _ //=. + rewrite andb_true_r //. - rewrite wf_universes_lift. - destruct X as [X _]. - pose proof (nth_error_Some_length heq_nth_error). - eapply nth_error_All_local_env in X; tea. - rewrite heq_nth_error /= in X. red in X. - destruct decl as [na [b|] ty]; cbn -[skipn] in *. - + now to_prop. - + destruct X as [s Hs]. now to_prop. + eapply forallb_nth_error with (n := n) in H0. rewrite heq_nth_error /= in H0. + now move/andP: H0. - apply/andP; split; to_wfu; cbn ; eauto with pcuic. - cbn in *; to_wfu ; eauto with pcuic. - rewrite wf_universes_subst. constructor. to_wfu; auto. constructor. - now move/andP: H4 => []. + now move/andP: H1 => []. - apply/andP; split. { rewrite wf_universeb_instance_forall. - apply/wf_universe_instanceP. + apply/wf_instanceP. eapply consistent_instance_ext_wf; eauto. } - pose proof (declared_constant_inv _ _ _ _ wf_universes_weaken wf X H). - red in X1; cbn in X1. - unshelve eapply declared_constant_to_gen in H; eauto. - destruct (cst_body decl). - * to_prop. - epose proof (weaken_lookup_on_global_env' Σ.1 _ _ wf H). - eapply wf_universes_inst. 2:eauto. all:eauto. - simpl in H2. - now eapply consistent_instance_ext_wf. - * move: X1 => [s /andP[Hc _]]. - to_prop. - eapply wf_universes_inst; eauto. - exact (weaken_lookup_on_global_env' Σ.1 _ _ wf H). - now eapply consistent_instance_ext_wf. + pose proof (declared_constant_inv _ _ _ _ wf_universes_weaken wf X H0) as [[_ X0]%andb_and _]%andb_and. + unshelve eapply declared_constant_to_gen in H0; eauto. + epose proof (weaken_lookup_on_global_env' Σ.1 _ _ wf H0). + eapply wf_universes_inst. 2:eauto. all:eauto. + now eapply consistent_instance_ext_wf. - apply/andP; split. { rewrite wf_universeb_instance_forall. - apply/wf_universe_instanceP. + apply/wf_instanceP. eapply consistent_instance_ext_wf; eauto. } pose proof (declared_inductive_inv wf_universes_weaken wf X isdecl). - cbn in X1. eapply onArity in X1. cbn in X1. - move: X1 => [s /andP[Hind ?]]. + move/onArity : X0 => /andP[] /andP[] /= _ Hind _. unshelve eapply declared_inductive_to_gen in isdecl; eauto. eapply wf_universes_inst; eauto. exact (weaken_lookup_on_global_env' Σ.1 _ _ wf (proj1 isdecl)). @@ -1132,30 +1118,28 @@ Qed. - apply/andP; split. { rewrite wf_universeb_instance_forall. - apply/wf_universe_instanceP. + apply/wf_instanceP. eapply consistent_instance_ext_wf; eauto. } pose proof (declared_constructor_inv wf_universes_weaken wf X isdecl) as [sc [nthe onc]]. unfold type_of_constructor. rewrite wf_universes_subst. { apply wf_universes_inds. now eapply consistent_instance_ext_wf. } - eapply on_ctype in onc. cbn in onc. - move: onc=> [_ /andP[onc _]]. + move/on_ctype : onc => /andP[] /andP[] /= _ onc _. clear nthe. unshelve eapply declared_constructor_to_gen in isdecl; eauto. eapply wf_universes_inst; eauto. exact (weaken_lookup_on_global_env' Σ.1 _ _ wf (proj1 (proj1 isdecl))). now eapply consistent_instance_ext_wf. - - rewrite wf_universes_mkApps in H5. - move/andP: H5 => /= [] wfu; rewrite forallb_app. + - rewrite wf_universes_mkApps in H7. + move/andP: H7 => /= [] wfu; rewrite forallb_app. move/andP=> [] wfpars wfinds. cbn in wfu. rewrite wfu /= wfpars wf_universes_mkApps /= - forallb_app wfinds /= H /= !andb_true_r. + forallb_app wfinds /= H0 /= !andb_true_r. pose proof (declared_inductive_inv wf_universes_weaken wf X isdecl). - destruct X5. destruct onArity as [s Hs]. - move/andP: Hs => [] /= hty hs. - rewrite ind_arity_eq in hty. + move:(onArity X3) => /andP[] /andP[] /= _ hty _. + rewrite (ind_arity_eq X3) in hty. rewrite !wf_universes_it_mkProd_or_LetIn in hty. move/and3P: hty => [] wfp wfindis wfisort. have ond : on_udecl_prop Σ (ind_universes mdecl). @@ -1164,7 +1148,7 @@ Qed. } eapply wf_ctx_universes_closed in wfp => //. eapply wf_ctx_universes_closed in wfindis => //. - rewrite (consistent_instance_length H1). + rewrite (consistent_instance_length H2). erewrite closedu_compare_decls; tea. rewrite closed_ind_predicate_context // /=. unfold test_branch. @@ -1172,10 +1156,10 @@ Qed. * have wfbrctx : All (fun cdecl => closedu_ctx #|polymorphic_instance (ind_universes mdecl)| (cstr_args cdecl)) (ind_ctors idecl). - { clear -wf ond onConstructors. + { pose proof (onConstructors := onConstructors X3). + clear -wf ond onConstructors. red in onConstructors. solve_all. destruct X. - do 2 red in on_ctype. destruct on_ctype as [s Hs]. - move/andP: Hs => [] wfty _. + move : on_ctype => /andP[] /andP[] /= _ wfty _. rewrite cstr_eq in wfty. rewrite !wf_universes_it_mkProd_or_LetIn in wfty. move/and3P: wfty => [] _ clargs _. @@ -1185,30 +1169,32 @@ Qed. rewrite /cstr_branch_context. eapply closedu_expand_lets_ctx. rewrite wfp. eapply closedu_subst_context. - rewrite a1. + rewrite a. now rewrite closedu_inds. - * rewrite /ptm. - rewrite wf_universes_it_mkLambda_or_LetIn H4 andb_true_r. - rewrite /predctx. - destruct X3 as [_ hctx]. move: hctx. - now rewrite test_context_app => /andP[]. + * rewrite wf_universes_it_mkLambda_or_LetIn H6 andb_true_r. + move: H4. + rewrite /wf_ctx_universes forallb_app => /andP[hctx _]. + apply (MCReflect.introT onctxP). + solve_all. unfold ondecl, wf_decl_universes, on_decl_universes in *. + move/andP:H3 => [Hb Ht] //; split; tas. + now destruct decl_body. - rewrite /subst1. rewrite wf_universes_subst. constructor => //. eapply All_rev. - rewrite wf_universes_mkApps in H1. - move/andP: H1 => []. + rewrite wf_universes_mkApps in H2. + move/andP: H2 => []. now intros _ hargs%forallb_All. pose proof (declared_projection_inv wf_universes_weaken wf X isdecl). destruct (declared_inductive_inv); simpl in *. destruct ind_ctors as [|cs []] => //. destruct ind_cunivs as [|cunivs []] => //; - destruct X1 as [[[? ?] ?] ?] => //. + destruct X0 as [[[? ?] ?] ?] => //. red in o0. destruct nth_error eqn:heq => //. destruct o0 as [_ ->]. rewrite wf_universes_mkApps {1}/wf_universes /= -!/(wf_universes _ _) - wf_universeb_instance_forall in H1. - move/andP: H1 => [/wf_universe_instanceP wfu wfargs]. + wf_universeb_instance_forall in H2. + move/andP: H2 => [/wf_instanceP wfu wfargs]. unshelve eapply declared_projection_to_gen in isdecl; eauto. eapply (wf_universes_inst (ind_universes mdecl)); eauto. exact (weaken_lookup_on_global_env' Σ.1 _ _ wf (proj1 (proj1 (proj1 isdecl)))). @@ -1223,38 +1209,36 @@ Qed. rewrite nth_error_subst_context in heq. autorewrite with len in heq. simpl in heq. epose proof (nth_error_lift_context_eq _ (smash_context [] (ind_params mdecl)) _ _). - autorewrite with len in H. simpl in H. rewrite -> H in heq. clear H. + autorewrite with len in H0. simpl in H0. rewrite -> H0 in heq. clear H0. autorewrite with len in heq. simpl in heq. destruct nth_error eqn:hnth; simpl in * => //. noconf heq. simpl. rewrite wf_universes_subst. apply wf_extended_subst. - rewrite ind_arity_eq in onArity. destruct onArity as [s' Hs]. - rewrite wf_universes_it_mkProd_or_LetIn in Hs. - now move/andP: Hs => /andP /andP [] /andP []. + rewrite ind_arity_eq in onArity. move:onArity => /andP[] /andP[] /= _ Hs _. + rewrite !wf_universes_it_mkProd_or_LetIn in Hs. + now move/andP: Hs => /andP /andP []. rewrite wf_universes_lift. - eapply wf_sorts_local_ctx_smash in s. - eapply wf_sorts_local_ctx_nth_error in s as [? [? H]]; eauto. - red in H. destruct x0. now move/andP: H => []. - now destruct H as [s [Hs _]%andb_and]. + eapply wf_universes_local_ctx_smash in s. + eapply wf_universes_local_ctx_nth_error in s; eauto. - - apply/andP; split; auto. - solve_all; destruct a0 as (? & _ & ?), b0; rtoProp; tas. - eapply nth_error_all in X0; eauto. - simpl in X0. now move: X0 => [s [Hty /andP[wfty _]]]. + - clear X X1. unfold on_def_type, on_def_body in *; cbn in *. + apply/andP; split; auto. + solve_all. + eapply nth_error_all in X0 as (? & _)%andb_and; eauto. - - apply/andP; split; auto. - solve_all; destruct a0 as (? & _ & ?), b0; rtoProp; tas. - eapply nth_error_all in X0; eauto. - simpl in X0. now move: X0 => [s [Hty /andP[wfty _]]]. + - clear X X1. unfold on_def_type, on_def_body in *; cbn in *. + apply/andP; split; auto. + solve_all. + eapply nth_error_all in X0 as (? & _)%andb_and; eauto. - apply/andP; split; eauto. - destruct X1; cbn => //. + destruct X0; cbn => //. rtoProp; intuition eauto. solve_all. - destruct X1; cbn => //. + destruct X0; cbn => //. simp prim_type. cbn. rtoProp; intuition. - Qed. +Qed. Lemma typing_wf_universes {Σ : global_env_ext} {Γ t T} : wf Σ -> @@ -1264,9 +1248,9 @@ Qed. exact (env_prop_typing wf_types _ wfΣ _ _ _ Hty). Qed. - Lemma typing_wf_universe {Σ : global_env_ext} {Γ t s} : + Lemma typing_wf_sort {Σ : global_env_ext} {Γ t s} : wf Σ -> - Σ ;;; Γ |- t : tSort s -> wf_universe Σ s. + Σ ;;; Γ |- t : tSort s -> wf_sort Σ s. Proof using Type. intros wfΣ Hty. apply typing_wf_universes in Hty as [_ wfs]%andb_and; auto. @@ -1275,16 +1259,16 @@ Qed. Lemma isType_wf_universes {Σ Γ T} : wf Σ.1 -> isType Σ Γ T -> wf_universes Σ T. Proof using Type. - intros wfΣ [s Hs]. now eapply typing_wf_universes in Hs as [HT _]%andb_and. + intros wfΣ (_ & s & Hs & _). now eapply typing_wf_universes in Hs as [HT _]%andb_and. Qed. End CheckerFlags. Arguments wf_universe_reflect {Σ u}. -#[global] Hint Resolve wf_universe_type1 wf_universe_super wf_universe_sup wf_universe_product : pcuic. +#[global] Hint Resolve wf_sort_type1 wf_sort_super wf_sort_sup wf_sort_product : pcuic. #[global] -Hint Extern 4 (wf_universe _ ?u) => +Hint Extern 4 (wf_sort _ ?u) => match goal with - [ H : typing _ _ _ (tSort u) |- _ ] => apply (typing_wf_universe _ H) + [ H : typing _ _ _ (tSort u) |- _ ] => apply (typing_wf_sort _ H) end : pcuic. diff --git a/pcuic/theories/Syntax/PCUICClosed.v b/pcuic/theories/Syntax/PCUICClosed.v index 8cddd17e5..0b58fee7a 100644 --- a/pcuic/theories/Syntax/PCUICClosed.v +++ b/pcuic/theories/Syntax/PCUICClosed.v @@ -43,7 +43,7 @@ Proof. rewrite /=. rewrite Nat.add_comm. bool_congr. Qed. -Lemma test_context_k_app p n Γ Γ' : +Lemma test_context_k_app (p : nat -> term -> bool) n Γ Γ' : test_context_k p n (Γ ,,, Γ') = test_context_k p n Γ && test_context_k p (n + #|Γ|) Γ'. Proof. @@ -370,11 +370,6 @@ Proof. now rewrite app_length // plus_n_Sm. Qed. -Definition Pclosed := - (fun (_ : global_env_ext) (Γ : context) (t T : term) => - closedn #|Γ| t && closedn #|Γ| T). - - Lemma closed_subst_context n (Δ Δ' : context) t : closedn (n + #|Δ|) t -> diff --git a/pcuic/theories/Syntax/PCUICDepth.v b/pcuic/theories/Syntax/PCUICDepth.v index 71d941a11..71a94c8b1 100644 --- a/pcuic/theories/Syntax/PCUICDepth.v +++ b/pcuic/theories/Syntax/PCUICDepth.v @@ -384,12 +384,11 @@ Proof. - case: a => [na [b|] ty] /=; rewrite {1}/decl_depth_gen /context_depth_gen /= => Hlt; constructor; auto. + eapply IHΔ => //. unfold context_depth. lia. - + simpl. apply aux => //. red. lia. - + simpl. split. + + split. * apply aux => //. red. lia. - * apply aux=> //; red; lia. + * apply aux => //. cbn. lia. + apply IHΔ => //; unfold context_depth; lia. - + apply aux => //. red. lia. } + + split => //. apply aux => //. cbn. lia. } assert (forall m, list_depth_gen (fun x : def term => depth (dtype x)) m < S (mfixpoint_depth_gen depth m)). { clear. unfold mfixpoint_depth_gen, def_depth_gen. induction m. simpl. auto. simpl. lia. } assert (forall m, list_depth_gen (fun x : def term => depth (dbody x)) m < S (mfixpoint_depth_gen depth m)). @@ -508,14 +507,13 @@ Proof. { move=> Γ h. induction Γ; cbn. - constructor. - - case: a h => [na [b|] ty] /=; + - move: a h => [na bo ty] /=. rewrite {1}/decl_depth_gen /context_depth_gen /= => Hlt; constructor; auto. - + simpl. split. - * apply aux => //. red. cbn. lia. - * apply aux=> //; red; lia. - + apply IHΓ => //; unfold context_depth; lia. - + red; cbn. split. apply aux => //. red. lia. exact tt. - + apply IHΓ => //. unfold context_depth; lia. } + 2: apply IHΓ => //; unfold context_depth; lia. + split. + + destruct bo => //. cbn in Hlt. + apply aux => //. cbn. lia. + + apply aux => //. cbn. lia. } assert (forall m, list_depth_gen (fun x : def term => depth (dtype x)) m < S (mfixpoint_depth_gen depth m)). { clear. unfold mfixpoint_depth_gen, def_depth_gen. induction m. simpl. auto. simpl. lia. } assert (forall m, list_depth_gen (fun x : def term => depth (dbody x)) m < S (mfixpoint_depth_gen depth m)). diff --git a/pcuic/theories/Syntax/PCUICInduction.v b/pcuic/theories/Syntax/PCUICInduction.v index c2bb01af3..98e5231cf 100644 --- a/pcuic/theories/Syntax/PCUICInduction.v +++ b/pcuic/theories/Syntax/PCUICInduction.v @@ -62,7 +62,7 @@ Proof. + generalize (pcontext p). fix auxc 1. destruct l; constructor; [|apply auxc]. - destruct c. split. apply auxt. + destruct c. split. 2: apply auxt. simpl. destruct decl_body; simpl. apply auxt. constructor. + apply auxt. @@ -73,7 +73,7 @@ Proof. + generalize (bcontext b). fix auxc 1. destruct l; constructor; [|apply auxc]. - destruct c. split. apply auxt. + destruct c. split. 2: apply auxt. simpl. destruct decl_body; simpl. apply auxt. constructor. + apply auxt. @@ -238,9 +238,9 @@ Lemma liftP_ctx_ind (P : term -> Type) (ctx : context) : Proof. induction ctx; simpl; constructor; auto. * split. - + apply X; cbn. unfold decl_size. simpl. lia. + destruct decl_body eqn:db; cbn. apply X; unfold decl_size. rewrite db; simpl; lia. exact tt. + + apply X; cbn. unfold decl_size. simpl. lia. * apply IHctx; intros; apply X. lia. Qed. @@ -366,9 +366,9 @@ Lemma liftP_ctx (P : term -> Type) : Proof. induction ctx; simpl; constructor; auto. split. - + apply X; cbn. + destruct decl_body eqn:db; cbn. apply X; unfold decl_size. exact tt. + + apply X; cbn. Qed. Lemma ctx_length_ind (P : context -> Type) (p0 : P []) @@ -450,12 +450,7 @@ Proof. rewrite !size_lift. rewrite list_size_map_hom; auto. Qed. -Definition on_local_decl (P : context -> term -> Type) - (Γ : context) (t : term) (T : typ_or_sort) := - match T with - | Typ T => (P Γ t * P Γ T)%type - | Sort => P Γ t - end. +Definition on_local_decl (P : context -> term -> Type) := lift_wf_term1 P. (* TODO: remove List.rev *) Lemma list_size_rev {A} size (l : list A) @@ -534,12 +529,11 @@ Proof. - case: a => [na [b|] ty] /=; rewrite {1}/decl_size /context_size /= => Hlt; constructor; auto. + eapply IHΔ => //. unfold context_size. lia. - + simpl. apply aux => //. red. lia. - + simpl. split. + + split. * apply aux => //. red. lia. - * apply aux=> //; red; lia. + * apply aux => //. cbn. lia. + apply IHΔ => //; unfold context_size; lia. - + apply aux => //. red. lia. } + + split => //. apply aux => //. cbn. lia. } assert (forall m, list_size (fun x : def term => size (dtype x)) m < S (mfixpoint_size size m)). { clear. unfold mfixpoint_size, def_size. induction m. simpl. auto. simpl. lia. } assert (forall m, list_size (fun x : def term => size (dbody x)) m < S (mfixpoint_size size m)). diff --git a/pcuic/theories/Syntax/PCUICNamelessDef.v b/pcuic/theories/Syntax/PCUICNamelessDef.v index 6f7482eff..4f3376eed 100644 --- a/pcuic/theories/Syntax/PCUICNamelessDef.v +++ b/pcuic/theories/Syntax/PCUICNamelessDef.v @@ -30,8 +30,7 @@ Definition anon (na : name) : bool := Definition banon (na : binder_annot name) : bool := anon na.(binder_name). Definition nameless_decl nameless (d : context_decl) := - banon (decl_name d) && nameless d.(decl_type) && - option_default nameless d.(decl_body) true. + banon (decl_name d) && option_default nameless d.(decl_body) true && nameless d.(decl_type). Fixpoint nameless (t : term) : bool := match t with diff --git a/pcuic/theories/Syntax/PCUICOnFreeVars.v b/pcuic/theories/Syntax/PCUICOnFreeVars.v index ac499f54f..011c13580 100644 --- a/pcuic/theories/Syntax/PCUICOnFreeVars.v +++ b/pcuic/theories/Syntax/PCUICOnFreeVars.v @@ -1049,13 +1049,13 @@ Lemma on_free_vars_case_predicate_context {cf} {Σ} {wfΣ : wf Σ} {P ci mdecl i rewrite alli_app; len; rewrite andb_true_r. apply andb_true_iff. split. - rewrite -/(on_free_vars_ctx P _). - rewrite (on_free_vars_ctx_all_term _ _ Universe.type0). + rewrite (on_free_vars_ctx_all_term _ _ Sort.type0). rewrite -(subst_it_mkProd_or_LetIn _ _ _ (tSort _)). apply on_free_vars_subst. { rewrite forallb_rev => //. } rewrite -on_free_vars_ctx_all_term. rewrite on_free_vars_ctx_subst_instance. - rewrite (on_free_vars_ctx_all_term _ _ (Universe.type0)). + rewrite (on_free_vars_ctx_all_term _ _ Sort.type0). rewrite -(expand_lets_it_mkProd_or_LetIn _ _ 0 (tSort _)). eapply on_free_vars_expand_lets_k; len. * rewrite (wf_predicate_length_pars wfp). @@ -1090,7 +1090,7 @@ Lemma on_free_vars_case_predicate_context {cf} {Σ} {wfΣ : wf Σ} {P ci mdecl i Proof. intros brctx decli wfp wfb havp. rewrite /brctx /case_branch_context /case_branch_context_gen. - rewrite (on_free_vars_ctx_all_term _ _ Universe.type0). + rewrite (on_free_vars_ctx_all_term _ _ Sort.type0). rewrite -(subst_it_mkProd_or_LetIn _ _ _ (tSort _)). apply on_free_vars_subst => //. { rewrite forallb_rev //. } @@ -1127,7 +1127,7 @@ Proof. intros hpars hn. rewrite /inst_case_context. rewrite test_context_k_eq. - rewrite (on_free_vars_ctx_all_term _ _ Universe.type0). + rewrite (on_free_vars_ctx_all_term _ _ Sort.type0). rewrite -(subst_it_mkProd_or_LetIn _ _ _ (tSort _)). intros a. apply on_free_vars_subst => //. @@ -1234,7 +1234,7 @@ Lemma on_free_vars_ctx_subst_context P s k ctx : on_free_vars_ctx (shiftnP k P) (subst_context s k ctx). Proof. intros onctx ons. - rewrite (on_free_vars_ctx_all_term _ _ Universe.type0). + rewrite (on_free_vars_ctx_all_term _ _ Sort.type0). rewrite -(subst_it_mkProd_or_LetIn _ _ _ (tSort _)). eapply on_free_vars_impl; revgoals. - eapply on_free_vars_subst_gen => //; tea. @@ -1255,7 +1255,7 @@ Lemma on_free_vars_ctx_lift_context p k n ctx : on_free_vars_ctx p ctx = on_free_vars_ctx (strengthenP k n p) (lift_context n k ctx). Proof. - rewrite !(on_free_vars_ctx_all_term _ _ Universe.type0). + rewrite !(on_free_vars_ctx_all_term _ _ Sort.type0). rewrite -(lift_it_mkProd_or_LetIn _ _ _ (tSort _)). rewrite on_free_vars_lift => //. Qed. @@ -1492,7 +1492,7 @@ Proof. Defined. Lemma alpha_eq_on_free_vars P (Γ Δ : context) : - All2 (PCUICEquality.compare_decls eq eq) Γ Δ -> + PCUICEquality.eq_context_upto_names Γ Δ -> on_free_vars_ctx P Γ -> on_free_vars_ctx P Δ. Proof. induction 1; cbn; auto. diff --git a/pcuic/theories/Syntax/PCUICPosition.v b/pcuic/theories/Syntax/PCUICPosition.v index ee75e7d03..589b0fb60 100644 --- a/pcuic/theories/Syntax/PCUICPosition.v +++ b/pcuic/theories/Syntax/PCUICPosition.v @@ -180,11 +180,11 @@ Definition dlet_ty na b B t (p : pos B) : pos (tLetIn na b B t) := Definition dlet_in na b B t (p : pos t) : pos (tLetIn na b B t) := exist (let_in :: proj1_sig p) (proj2_sig p). -Lemma eq_context_upto_context_choice_term Σ Re Rle Γ Γ' c : - eq_context_upto Σ Re Rle Γ Γ' -> - rel_option (eq_term_upto_univ Σ Re (match c.2 with - | choose_decl_body => Re - | choose_decl_type => Rle +Lemma eq_context_upto_context_choice_term Σ cmp_universe cmp_sort pb Γ Γ' c : + eq_context_upto Σ cmp_universe cmp_sort pb Γ Γ' -> + rel_option (eq_term_upto_univ Σ cmp_universe cmp_sort (match c.2 with + | choose_decl_body => Conv + | choose_decl_type => pb end) ) (context_choice_term Γ c) (context_choice_term Γ' c). @@ -203,13 +203,13 @@ Proof. Qed. Lemma eq_term_upto_valid_pos : - forall {Σ u v p Re Rle napp}, + forall {Σ u v p cmp_universe cmp_sort pb napp}, validpos u p -> - eq_term_upto_univ_napp Σ Re Rle napp u v -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp u v -> validpos v p. Proof. - intros Σ u v p Re Rle napp vp e. - induction p as [| c p ih ] in u, v, Re, Rle, napp, vp, e |- *. + intros Σ u v p cmp_universe cmp_sort pb napp vp e. + induction p as [| c p ih ] in u, v, pb, napp, vp, e |- *. - reflexivity. - destruct c, u. all: try discriminate. all: try solve [ @@ -219,50 +219,38 @@ Proof. + dependent destruction e. simpl in *. destruct (nth_error (pparams p0) n) as [par|] eqn:enth. 2: discriminate. destruct e. - induction a0 in n, par, enth, ih, vp |- *. 1: rewrite enth. 1: assumption. + induction a in n, par, enth, ih, vp |- *. 1: rewrite enth. 1: assumption. destruct n. * simpl in *. apply some_inj in enth. subst. intuition eauto. - * simpl in *. eapply IHa0. all: eauto. + * simpl in *. eapply IHa. all: eauto. + dependent destruction e. simpl in *. eapply ih; eauto. apply e. + dependent destruction e. simpl in *. destruct nth_error eqn:nth; [|congruence]. - eapply All2_nth_error_Some in a; eauto. - destruct a as (?&->&_&eq). + eapply All2_nth_error_Some in e1; eauto. + destruct e1 as (?&->&_&eq). eauto. + dependent destruction e. simpl in *. - destruct (nth_error mfix n) as [[na ty bo ra]|] eqn:e. 2: discriminate. - induction a in n, na, ty, bo, ra, e, ih, vp |- *. - 1:{ rewrite e. assumption. } - destruct n. - * simpl in *. apply some_inj in e. subst. - destruct y as [na' ty' bo' ra']. simpl in *. intuition eauto. - * simpl in *. eapply IHa. all: eauto. + destruct (nth_error mfix n) as [[na ty bo ra]|] eqn:e1. 2: discriminate. + eapply All2_nth_error_Some in e; eauto. + destruct e as (?&->&H&_). + now eapply ih. + dependent destruction e. simpl in *. - destruct (nth_error mfix n) as [[na ty bo ra]|] eqn:e. 2: discriminate. - induction a in n, na, ty, bo, ra, e, ih, vp |- *. - 1:{ rewrite e. assumption. } - destruct n. - * simpl in *. apply some_inj in e. subst. - destruct y as [na' ty' bo' ra']. simpl in *. intuition eauto. - * simpl in *. eapply IHa. all: eauto. + destruct (nth_error mfix n) as [[na ty bo ra]|] eqn:e1. 2: discriminate. + eapply All2_nth_error_Some in e; eauto. + destruct e as (?&->&_&H&_). + now eapply ih. + dependent destruction e. simpl in *. - destruct (nth_error mfix n) as [[na ty bo ra]|] eqn:e. 2: discriminate. - induction a in n, na, ty, bo, ra, e, ih, vp |- *. - 1:{ rewrite e. assumption. } - destruct n. - * simpl in *. apply some_inj in e. subst. - destruct y as [na' ty' bo' ra']. simpl in *. intuition eauto. - * simpl in *. eapply IHa. all: eauto. + destruct (nth_error mfix n) as [[na ty bo ra]|] eqn:e1. 2: discriminate. + eapply All2_nth_error_Some in e; eauto. + destruct e as (?&->&H&_). + now eapply ih. + dependent destruction e. simpl in *. - destruct (nth_error mfix n) as [[na ty bo ra]|] eqn:e. 2: discriminate. - induction a in n, na, ty, bo, ra, e, ih, vp |- *. - 1:{ rewrite e. assumption. } - destruct n. - * simpl in *. apply some_inj in e. subst. - destruct y as [na' ty' bo' ra']. simpl in *. intuition eauto. - * simpl in *. eapply IHa. all: eauto. + destruct (nth_error mfix n) as [[na ty bo ra]|] eqn:e1. 2: discriminate. + eapply All2_nth_error_Some in e; eauto. + destruct e as (?&->&_&H&_). + now eapply ih. + depelim e. depelim o; eauto. + depelim e; depelim o; eauto. + depelim e. depelim o; eauto. diff --git a/pcuic/theories/Syntax/PCUICReflect.v b/pcuic/theories/Syntax/PCUICReflect.v index 47fafc878..15a6b788e 100644 --- a/pcuic/theories/Syntax/PCUICReflect.v +++ b/pcuic/theories/Syntax/PCUICReflect.v @@ -25,7 +25,7 @@ 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.t, u' : Universe.t |- _ => fcase (eq_dec u u') + | u : sort, u' : sort |- _ => fcase (eq_dec u u') | x : Instance.t, y : Instance.t |- _ => fcase (eq_dec x y) | x : list Level.t, y : Instance.t |- _ => diff --git a/pcuic/theories/Typing/PCUICClosedTyp.v b/pcuic/theories/Typing/PCUICClosedTyp.v index a3c619762..483122f4f 100644 --- a/pcuic/theories/Typing/PCUICClosedTyp.v +++ b/pcuic/theories/Typing/PCUICClosedTyp.v @@ -12,9 +12,7 @@ From Equations Require Import Equations. Lemma declared_projection_closed_ind {cf:checker_flags} {Σ : global_env} {wfΣ : wf Σ}{mdecl idecl cdecl p pdecl} : declared_projection Σ p mdecl idecl cdecl pdecl -> - Forall_decls_typing - (fun _ (Γ : context) (t T : term) => - closedn #|Γ| t && closedn #|Γ| T) Σ -> + on_global_env cumulSpec0 (fun _ => lift_wfb_term1 (fun Γ t => closedn #|Γ| t)) Σ -> closedn (S (ind_npars mdecl)) pdecl.(proj_type). Proof. intros isdecl X0. @@ -75,26 +73,21 @@ Qed. Lemma declared_decl_closed_ind {cf : checker_flags} {Σ : global_env} {wfΣ : wf Σ} {cst decl} : lookup_env Σ cst = Some decl -> - Forall_decls_typing (fun (_ : global_env_ext) (Γ : context) (t T : term) => closedn #|Γ| t && closedn #|Γ| T) Σ -> - on_global_decl cumulSpec0 (fun Σ Γ b t => closedn #|Γ| b && typ_or_sort_default (closedn #|Γ|) t true) + on_global_env cumulSpec0 (fun _ => lift_wfb_term1 (fun Γ t => closedn #|Γ| t)) Σ -> + on_global_decl cumulSpec0 (fun _ => lift_wfb_term1 (fun Γ t => closedn #|Γ| t)) (Σ, universes_decl_of_decl decl) cst decl. Proof. intros. eapply weaken_lookup_on_global_env; eauto. do 2 red; eauto. - eapply @on_global_env_impl with (Σ := (empty_ext Σ)); cycle 1. tea. - red; intros. destruct T; intuition auto with wf. - destruct X2 as [s0 Hs0]. simpl. rtoProp; intuition. Qed. Lemma declared_minductive_closed_ind {cf:checker_flags} {Σ : global_env} {wfΣ : wf Σ}{mdecl mind} : - Forall_decls_typing - (fun (_ : global_env_ext) (Γ : context) (t T : term) => - closedn #|Γ| t && closedn #|Γ| T) Σ -> + on_global_env cumulSpec0 (fun _ => lift_wfb_term1 (fun Γ t => closedn #|Γ| t)) Σ -> declared_minductive Σ mind mdecl -> closed_inductive_decl mdecl. Proof. intros HΣ decl. pose proof (decl_ := decl). - eapply declared_minductive_to_gen in decl. + unshelve eapply declared_minductive_to_gen in decl; tea. pose proof (declared_decl_closed_ind decl) as decl'. specialize (decl' HΣ). red in decl'. @@ -115,19 +108,19 @@ Proof. rewrite /closed_inductive_body. apply andb_and; split. apply andb_and. split. - rewrite andb_and. split. - * apply onArity in oib. hnf in oib. - now move/andP: oib => [] /= ->. - * pose proof (onArity oib). + * apply onArity in oib. + now move/andP: oib => [] /=. + * pose proof (onArity oib) as X. unfold on_type in X. rewrite oib.(ind_arity_eq) in X. - red in X. simpl in X. + cbn in X. rewrite !closedn_it_mkProd_or_LetIn /= !andb_true_r in X. now move/andP: X. - pose proof (onConstructors oib). red in X. eapply All_forallb. eapply All2_All_left; eauto. intros cdecl cs X0; - move/andP: (on_ctype X0) => []. + move: (on_ctype X0) => /=. unfold on_type. cbn. simpl. unfold closed_constructor_body. - intros Hty _. + intros Hty. rewrite arities_context_length in Hty. rewrite Hty. rewrite X0.(cstr_eq) closedn_it_mkProd_or_LetIn in Hty. @@ -150,7 +143,6 @@ Proof. intros. split; auto. split; auto. cbn. now rewrite hcdecl. } eapply (Alli_All X0). intros. now eapply declared_projection_closed_ind in H. - Unshelve. all:eauto. Qed. @@ -158,14 +150,19 @@ Qed. Lemma typecheck_closed `{cf : checker_flags} : env_prop (fun Σ Γ t T => closedn #|Γ| t && closedn #|Γ| T) + (fun _ => lift_wfb_term1 (fun Γ t => closedn #|Γ| t)) (fun Σ Γ => closed_ctx Γ). Proof. assert (X := weaken_env_prop_closed). - apply typing_ind_env; intros * wfΣ Γ wfΓ *; intros; cbn in *; + apply typing_ind_env; intros * wfΣ Γ wfΓ *; intros; unfold lift_wfb_term in *; cbn in *; rewrite -> ?andb_and in *; try solve [intuition auto]. - - induction X0; auto; rewrite closedn_ctx_cons /= IHX0 /= //. - now move/andP: Hs => [] /=. + - destruct X0 as (Htm & s & (_ & Hty) & _). + rewrite /= andb_true_r in Hty. split; tas. + destruct j_term => //=. + move:Htm => [] _ /andP [] //. + + - now apply closedn_All_local_closed in X0. - pose proof (nth_error_Some_length H). elim (Nat.ltb_spec n #|Γ|); intuition auto. all: try lia. clear H1. @@ -191,22 +188,17 @@ Proof. move=> Hs. apply: Hs => /=. simpl. rewrite H1 => //. rewrite Nat.add_1_r. auto. - - eapply declared_constant_to_gen in H0. + - unshelve eapply declared_constant_to_gen in H0; tea. rewrite closedn_subst_instance. eapply lookup_on_global_env in X0; eauto. - destruct X0 as [Σ' [hext [onu HΣ'] IH]]. - repeat red in IH. destruct decl, cst_body0. simpl in *. - rewrite -> andb_and in IH. intuition auto. - eauto using closed_upwards with arith. + destruct X0 as (Σ' & hext & [onu HΣ'] & [IHtm IHty]%andb_and). simpl in *. - repeat red in IH. destruct IH as [s Hs]. - rewrite -> andb_and in Hs. intuition auto. eauto using closed_upwards with arith. - rewrite closedn_subst_instance. eapply declared_inductive_inv in X0; eauto with extends. - apply onArity in X0. repeat red in X0. - destruct X0 as [s Hs]. rewrite -> andb_and in Hs. + apply onArity in X0. + rewrite /on_type /lift_wfb_term /= in X0. intuition eauto using closed_upwards with arith. - destruct isdecl as [Hidecl Hcdecl]. @@ -219,9 +211,9 @@ Proof. pose proof X0.(onConstructors) as XX. eapply All2_nth_error_Some in Hcdecl; eauto. destruct Hcdecl as [? [? ?]]. cbn in *. - destruct o as [? ? [s Hs] _]. rewrite -> andb_and in Hs. - apply proj1 in Hs. - rewrite arities_context_length in Hs. + destruct o as [? ? HT _]. + rewrite /on_type /lift_wfb_term /= in HT. + rewrite arities_context_length in HT. eauto using closed_upwards with arith. - destruct H3 as [clret _]. @@ -272,29 +264,21 @@ Proof. eapply closed_upwards; eauto. lia. - clear H. - split. solve_all. destruct b. - destruct x; simpl in *. - unfold test_def. simpl. rtoProp. - split. - rewrite -> app_context_length in *. rewrite -> Nat.add_comm in *. - eapply closedn_lift_inv in H3; eauto. lia. + split. solve_all. + rewrite /on_def_type /on_def_body ?andb_and in a0, b |- *. cbn in *. + simpl in *. unfold test_def. + rewrite -> app_context_length in *. subst types. - now rewrite app_context_length fix_context_length in H. - eapply nth_error_all in X0; eauto. simpl in X0. intuition auto. rtoProp. - destruct X0 as [s [Hs cl]]. now rewrite andb_true_r in cl. + rewrite fix_context_length in b. intuition auto. + eapply nth_error_all in X1 as (_ & X1)%andb_and; eauto. - - split. solve_all. destruct b. - destruct x; simpl in *. - unfold test_def. simpl. rtoProp. - split. - destruct a as [s [Hs cl]]. - now rewrite andb_true_r in cl. - rewrite -> app_context_length in *. rewrite -> Nat.add_comm in *. - subst types. now rewrite fix_context_length in H3. - eapply nth_error_all in X0; eauto. - destruct X0 as [s [Hs cl]]. - now rewrite andb_true_r in cl. - Unshelve. all:eauto. + - split. solve_all. + rewrite /on_def_type /on_def_body ?andb_and in a0, b |- *. cbn in *. + simpl in *. unfold test_def. + rewrite -> app_context_length in *. + subst types. + rewrite fix_context_length in b. intuition auto. + eapply nth_error_all in X1 as (_ & X1)%andb_and; eauto. - destruct p as [[] pv]; cbn in X0 |- *; simp prim_type => //. depelim pv. simp prim_type. cbn. depelim X1. @@ -391,7 +375,7 @@ Qed. #[global] Hint Extern 10 => progress unfold PCUICTypingDef.typing in * : fvs. Lemma isType_closed {cf:checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ.1} {Γ T} : isType Σ Γ T -> closedn #|Γ| T. -Proof. intros [s Hs]; fvs. Qed. +Proof. intros (_ & s & Hs & _); cbn in Hs; fvs. Qed. #[global] Hint Extern 4 (closedn #|?Γ| ?A = true) => match goal with @@ -431,6 +415,22 @@ Proof. move/type_closed. now rewrite is_open_term_closed. Qed. +Lemma lift_typing_is_open_term {cf:checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ.1} {Γ j} : + lift_typing typing Σ Γ j -> + lift_wf_term (is_open_term Γ) j. +Proof. + intros (Htm & s & Hty & _). + split. 1: destruct j_term; cbn in *; auto. + all: rewrite -is_open_term_closed; now eapply subject_closed. +Qed. + +Lemma isType_is_open_term {cf:checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ.1} {Γ T} : + isType Σ Γ T -> + is_open_term Γ T. +Proof. + move/isType_closed. now rewrite is_open_term_closed. +Qed. + Lemma closed_wf_local `{checker_flags} {Σ Γ} : wf Σ.1 -> wf_local Σ Γ -> @@ -450,6 +450,11 @@ Qed. | [ H : _ ;;; Γ |- _ : A |- _ ] => exact (type_is_open_term H) end : fvs. +#[global] Hint Extern 4 (is_open_term ?Γ ?A = true) => + match goal with + | [ H : isType _ Γ A |- _ ] => exact (isType_is_open_term H) + end : fvs. + Lemma closed_ctx_on_ctx_free_vars Γ : closed_ctx Γ = on_ctx_free_vars (closedP #|Γ| xpredT) Γ. Proof. rewrite closedn_ctx_on_free_vars. @@ -475,7 +480,7 @@ Qed. end : fvs. Lemma ctx_inst_closed {cf:checker_flags} (Σ : global_env_ext) Γ i Δ : - wf Σ.1 -> ctx_inst typing Σ Γ i Δ -> All (closedn #|Γ|) i. + wf Σ.1 -> ctx_inst (typing Σ) Γ i Δ -> All (closedn #|Γ|) i. Proof. intros wfΣ; induction 1; auto; constructor; auto; fvs. Qed. @@ -485,8 +490,7 @@ Qed. Lemma declared_decl_closed `{checker_flags} {Σ : global_env} {cst decl} : wf Σ -> lookup_env Σ cst = Some decl -> - on_global_decl cumulSpec0 (fun Σ Γ b t => closedn #|Γ| b && typ_or_sort_default (closedn #|Γ|) t true) - (Σ, universes_decl_of_decl decl) cst decl. + on_global_decl cumulSpec0 (fun _ => lift_wfb_term1 (fun Γ t => closedn #|Γ| t)) (Σ, universes_decl_of_decl decl) cst decl. Proof. intros. apply declared_decl_closed_ind; eauto. @@ -499,15 +503,11 @@ Lemma declared_constant_closed_type {cf:checker_flags} {Σ : global_env} {wfΣ : declared_constant Σ cst decl -> closed decl.(cst_type). Proof. - intros h. eapply declared_constant_to_gen in h. + apply (env_prop_sigma typecheck_closed) in wfΣ. + intros h. unshelve eapply declared_constant_to_gen in h; tea. eapply lookup_on_global_env in h. 2: eauto. - destruct h as [Σ' [ext wfΣ' decl']]. - red in decl'. red in decl'. - destruct decl as [ty bo un]. simpl in *. - destruct bo as [t|]. - - now eapply type_closed in decl'. - - cbn in decl'. destruct decl' as [s h]. - now eapply subject_closed in h. Unshelve. all:tea. + destruct h as (Σ' & ext & wfΣ' & [Htm Hty]%andb_and). + easy. Qed. @@ -519,14 +519,13 @@ Lemma declared_constant_closed_body {cf : checker_flags} : closed body. Proof. intros Σ cst decl body hΣ h e. - eapply declared_constant_to_gen in h. + apply (env_prop_sigma typecheck_closed) in hΣ. + unshelve eapply declared_constant_to_gen in h; tea. eapply lookup_on_global_env in h. 2: eauto. - destruct h as [Σ' [ext wfΣ' decl']]. - red in decl'. red in decl'. - destruct decl as [ty bo un]. simpl in *. - rewrite e in decl'. - now eapply subject_closed in decl'. - Unshelve. all:tea. + destruct h as (Σ' & ext & wfΣ' & [Htm _]%andb_and). + simpl in Htm. + rewrite e in Htm. + easy. Qed. @@ -537,17 +536,15 @@ Lemma declared_inductive_closed_type {cf:checker_flags} : closed idecl.(ind_type). Proof. intros Σ mdecl ind idecl hΣ h. + apply (env_prop_sigma typecheck_closed) in hΣ. unfold declared_inductive in h. destruct h as [h1 h2]. - eapply declared_minductive_to_gen in h1. + unshelve eapply declared_minductive_to_gen in h1; tea. eapply lookup_on_global_env in h1. 2: eauto. - destruct h1 as [Σ' [ext wfΣ' decl']]. + destruct h1 as (Σ' & ext & wfΣ' & decl'). red in decl'. destruct decl' as [h ? ? ?]. eapply Alli_nth_error in h. 2: eassumption. - simpl in h. destruct h as [? [? h] ? ? ?]. - eapply typecheck_closed in h as [? e]. 2: auto. - now move: e => [_ /andP []]. - Unshelve. all:eauto. + simpl in h. now destruct h as [? [? h]%andb_and ? ? ?]. Qed. @@ -557,7 +554,7 @@ Lemma declared_inductive_closed_params {cf:checker_flags} {Σ mdecl ind idecl} { Proof. intros h. pose proof (on_declared_inductive h) as [onmind _]. - eapply onParams in onmind. + eapply onParams in onmind. unfold on_context in onmind. eapply closed_wf_local. 2:tea. eauto. Qed. @@ -1001,8 +998,8 @@ Proof. generalize (pcontext p). fix auxl' 1; destruct l; [constructor|]; simpl; rewrite ?Nat.sub_0_r. move/andP => [] tl /andP [tdef tty]. constructor. - + rewrite Nat.sub_0_r. simpl. split; [apply auxt|]; tas. - destruct (decl_body c); simpl in *; auto. exact tt. + + rewrite Nat.sub_0_r. simpl. split; [|apply auxt]; tas. + destruct (decl_body c); simpl in *; auto. + eapply Alli_shift, Alli_impl; eauto. simpl. intros n' x. now replace (Nat.pred #|l| - n' + #|pparams p|) with (#|l| - S n' + #|pparams p|). @@ -1018,8 +1015,8 @@ Proof. generalize (bcontext b). fix auxl'' 1; destruct l; [constructor|]; simpl; rewrite ?Nat.sub_0_r. move/andP => [] tl /andP [tdef tty]. constructor. - + rewrite Nat.sub_0_r. simpl. split; [apply auxt|]; tas. - destruct (decl_body c); simpl in *; auto. exact tt. + + rewrite Nat.sub_0_r. simpl. split; [|apply auxt]; tas. + destruct (decl_body c); simpl in *; auto. + eapply Alli_shift, Alli_impl; eauto. simpl. intros n' x. now replace (Nat.pred #|l| - n' + #|pparams p|) with (#|l| - S n' + #|pparams p|). diff --git a/pcuic/theories/Typing/PCUICContextConversionTyp.v b/pcuic/theories/Typing/PCUICContextConversionTyp.v index fa65ae745..c1ac54485 100644 --- a/pcuic/theories/Typing/PCUICContextConversionTyp.v +++ b/pcuic/theories/Typing/PCUICContextConversionTyp.v @@ -146,99 +146,93 @@ Proof. eapply All2_fold_refl. intros ? ?; reflexivity. Qed. -Lemma cumul_context_Algo_Spec {cf:checker_flags} Σ Γ' Γ : - Σ ⊢ Γ' ≤ Γ -> PCUICCumulativitySpec.cumul_context cumulSpec0 Σ Γ' Γ. +Lemma eq_context_cumul_Spec {cf:checker_flags} Σ pb Γ Δ : + compare_context Σ pb Γ Δ -> PCUICCumulativitySpec.cumul_pb_context cumulSpec0 pb Σ Γ Δ. Proof. intros e. eapply All2_fold_impl. 1: tea. cbn; intros. destruct X. - - econstructor 1; eauto. eapply (@cumulAlgo_cumulSpec _ _ Cumul); eauto. + - econstructor 1; eauto. eapply eq_term_upto_univ_cumulSpec; eauto. - econstructor 2; eauto. - + eapply (@cumulAlgo_cumulSpec _ _ Conv); eauto. - + eapply (@cumulAlgo_cumulSpec _ _ Cumul); eauto. + all: eapply eq_term_upto_univ_cumulSpec; eauto. +Defined. + +Lemma eq_context_cumul_Spec_rel {cf:checker_flags} Σ Ξ0 Γ Δ : + compare_context Σ Cumul Γ Δ -> cumul_ctx_rel cumulSpec0 Σ Ξ0 Γ Δ. +Proof. + intros e. + eapply All2_fold_impl. 1: tea. cbn; intros. + destruct X. + - econstructor 1; eauto. eapply eq_term_upto_univ_cumulSpec; eauto. + - econstructor 2; eauto. + all: eapply eq_term_upto_univ_cumulSpec; eauto. +Defined. + +Lemma cumul_context_Algo_Spec {cf:checker_flags} Σ pb Γ' Γ : + Σ ⊢ Γ' ≤[pb] Γ -> PCUICCumulativitySpec.cumul_pb_context cumulSpec0 pb Σ Γ' Γ. +Proof. + intros e. + eapply All2_fold_impl. 1: tea. cbn; intros. + destruct X. + - econstructor 1; eauto. eapply cumulAlgo_cumulSpec; eauto. + - econstructor 2; eauto. + all: eapply cumulAlgo_cumulSpec; eauto. Defined. Lemma context_cumulativity_prop {cf:checker_flags} : env_prop (fun Σ Γ t T => forall Γ', cumul_context cumulAlgo_gen Σ Γ' Γ -> wf_local Σ Γ' -> Σ ;;; Γ' |- t : T) + (fun Σ Γ j => + forall Γ', cumul_context cumulAlgo_gen Σ Γ' Γ -> wf_local Σ Γ' -> lift_typing0 (fun t T => Σ ;;; Γ' |- t : T) j) (fun Σ Γ => - All_local_env - (lift_typing (fun Σ (Γ : context) (t T : term) => - forall Γ' : context, cumul_context cumulAlgo_gen Σ Γ' Γ -> wf_local Σ Γ' -> Σ;;; Γ' |- t : T) Σ) Γ). + All_local_env (fun Γ j => + forall Γ' : context, cumul_context cumulAlgo_gen Σ Γ' Γ -> wf_local Σ Γ' -> + (lift_typing0 (fun t T => Σ;;; Γ' |- t : T) j)) Γ). Proof. apply typing_ind_env; intros Σ wfΣ Γ wfΓ; intros **; rename_all_hyps; try solve [econstructor; eauto]. - - induction X; constructor; auto. - all: now apply infer_typing_sort_impl with id tu. + - apply lift_typing_impl with (1 := X) => ?? [_ IH]. + now apply IH. + + - assumption. - pose proof heq_nth_error. eapply (All2_fold_nth_r X0) in H as [d' [Hnth [Hrel Hconv]]]. - unshelve eapply nth_error_All_local_env in X; tea. 2:eapply nth_error_Some_length in heq_nth_error; lia. - rewrite heq_nth_error /= in X. - destruct decl as [na [b|] ty] => /=. - + red in X. cbn in X. destruct X as [Hb Hty]. - destruct Hty as [s Hty]. specialize (Hty _ Hrel). - forward Hty by now eapply All_local_env_skipn. - eapply type_Cumul with _ s. - * econstructor. auto. eauto. - * rewrite -(firstn_skipn (S n) Γ'). - change (tSort s) with (lift0 (S n) (tSort s)). - eapply weakening_length. auto. - rewrite firstn_length_le. eapply nth_error_Some_length in Hnth. lia. auto. - now rewrite /app_context firstn_skipn. - assumption. - * depelim Hconv; simpl in *. - destruct (split_closed_context (S n) (wf_local_closed_context X1)) as [Δ [Δ' [eqΔ eqΔ' -> hn]]]. - eapply nth_error_Some_length in Hnth. lia. - rewrite -eqΔ in Hty, Hrel. - rewrite -eqΔ in eqb, eqt. - assert (is_open_term Δ t). - { eapply nth_error_closed_context in Hnth. 2:eauto with fvs. - rewrite -eqΔ in Hnth. now move/andP: Hnth => []. } - eapply PCUICClosedTyp.subject_closed in Hty. - eapply (@closedn_on_free_vars xpred0) in Hty. - eapply (weakening_cumulSpec0 (Γ := Δ) (Γ'' := Δ') (M := exist t H) (N := exist ty Hty)); cbn. lia. - unshelve eapply (@cumulAlgo_cumulSpec _ _ Cumul). apply into_ws_cumul_pb; eauto. - intuition. - + cbn in X. destruct X as [s ondecl]. - specialize (ondecl _ Hrel). - depelim Hconv. - forward ondecl by now eapply All_local_env_skipn. - eapply type_Cumul with _ s. - * econstructor. auto. eauto. - * rewrite -(firstn_skipn (S n) Γ'). - change (tSort s) with (lift0 (S n) (tSort s)). - eapply weakening_length. auto. - rewrite firstn_length_le. eapply nth_error_Some_length in Hnth. lia. auto. - now rewrite /app_context firstn_skipn. - assumption. - * destruct (split_closed_context (S n) (wf_local_closed_context X1)) as [Δ [Δ' [eqΔ eqΔ' -> hn]]]. - eapply nth_error_Some_length in Hnth. lia. - rewrite -eqΔ in ondecl, Hrel. - rewrite -eqΔ in eqt. - assert (is_open_term Δ t). - { rewrite nth_error_app_lt in Hnth. rewrite -hn. lia. - destruct Δ' as [Δ' hΔ']. cbn in *. - move: hΔ'. - rewrite -on_free_vars_ctx_on_ctx_free_vars -[shiftnP _ _]addnP0 => hΔ'. - eapply nth_error_on_free_vars_ctx in hΔ'; tea. - 2:{ rewrite shiftnP_add /shiftnP /= orb_false_r. apply Nat.ltb_lt. lia. } - rewrite /test_decl /= in hΔ'. move: hΔ'. - now rewrite hn addnP_shiftnP. } - eapply PCUICClosedTyp.subject_closed in ondecl. - eapply (@closedn_on_free_vars xpred0) in ondecl. - eapply (weakening_cumulSpec0 (Γ := Δ) (Γ'' := Δ') (M := exist t H) (N := exist ty ondecl)); cbn. lia. - unshelve eapply (@cumulAlgo_cumulSpec _ _ Cumul). apply into_ws_cumul_pb; eauto. - intuition. + eapply All_local_env_nth_error in X; tea. + red in X. specialize (X _ Hrel). + forward X by now eapply All_local_env_skipn. + destruct X as (_ & s & Hty & _). + eapply type_Cumul with _ s. + * econstructor. auto. eauto. + * rewrite -(firstn_skipn (S n) Γ'). + change (tSort s) with (lift0 (S n) (tSort s)). + eapply weakening_length. auto. + rewrite firstn_length_le. eapply nth_error_Some_length in Hnth. lia. auto. + now rewrite /app_context firstn_skipn. + assumption. + * destruct (split_closed_context (S n) (wf_local_closed_context X1)) as [Δ [Δ' [eqΔ eqΔ' -> hn]]]. + eapply nth_error_Some_length in Hnth. lia. + rewrite -eqΔ in Hty, Hrel. + eapply PCUICClosedTyp.subject_closed in Hty. + eapply (@closedn_on_free_vars xpred0) in Hty. + assert (is_open_term Δ (decl_type d') × cumulAlgo Σ Δ (decl_type d') (decl_type decl)) as [H H']. { + eapply nth_error_closed_context in Hnth. 2: eauto with fvs. + rewrite -eqΔ in Hnth. + depelim Hconv; simpl in *. + all: rewrite -eqΔ in eqt. + all: now move/andP: Hnth => []. + } + eapply (weakening_cumulSpec0 (Γ := Δ) (Γ'' := Δ') (M := exist (decl_type d') H) (N := exist (decl_type decl) Hty)); cbn. lia. + unshelve eapply (@cumulAlgo_cumulSpec _ _ Cumul). apply into_ws_cumul_pb; eauto. + now intuition. - constructor; pcuic. - eapply forall_Γ'0. repeat (constructor; pcuic). - constructor; auto. red. eexists; eapply forall_Γ'; auto. + eapply forall_Γ'0; constructor; pcuic. eapply lift_sorting_forget_univ. now eapply forall_Γ'. - econstructor; pcuic. eapply forall_Γ'0; repeat (constructor; pcuic). - econstructor; pcuic. - eapply forall_Γ'1; repeat (constructor; pcuic). + eapply forall_Γ'0; repeat (constructor; pcuic). - econstructor; eauto. 2,3: constructor; eauto. * eapply IHp0. rewrite /predctx. eapply All2_fold_app => //. @@ -263,34 +257,28 @@ Proof. eapply cumul_context_Algo_Spec; eauto. eapply into_ws_cumul_ctx_pb; eauto. + apply wf_local_closed_context; eauto. + apply wf_local_closed_context; eauto. - * eapply (All_impl X0). - intros d Ht. - apply infer_typing_sort_impl with id Ht; now intros [_ IH]. * eapply (All_impl X1). - intros d [Hs IH]. - eapply IH. - now apply cumul_context_app_same. + intros d Ht. now apply Ht. + * eapply (All_impl X3). + intros d Hb. apply Hb. + 1: now apply cumul_context_app_same. eapply (All_mfix_wf); auto. - apply (All_impl X0); simpl. - intros d' Ht. - apply infer_typing_sort_impl with id Ht; now intros [_ IH']. + apply (All_impl X1); simpl. + intros d' Ht. now apply Ht. - econstructor. all:pcuic. * eapply cofix_guard_context_cumulativity; eauto. eapply cumul_context_Algo_Spec; eauto. eapply into_ws_cumul_ctx_pb; eauto. + apply wf_local_closed_context; eauto. + apply wf_local_closed_context; eauto. - * eapply (All_impl X0). - intros d Ht. - apply infer_typing_sort_impl with id Ht; now intros [_ IH]. - * eapply (All_impl X1). - intros d [Hs IH]. - eapply IH. - now apply cumul_context_app_same. + * eapply (All_impl X1). + intros d Ht. now apply Ht. + * eapply (All_impl X3). + intros d Hb. apply Hb. + 1: now apply cumul_context_app_same. eapply (All_mfix_wf); auto. - apply (All_impl X0); simpl. - intros d' Ht. - apply infer_typing_sort_impl with id Ht; now intros [_ IH']. + apply (All_impl X1); simpl. + intros d' Ht. now apply Ht. - econstructor; tea. depelim X1; constructor; eauto. solve_all. - econstructor; eauto. pose proof (wf_local_closed_context wfΓ). @@ -330,7 +318,8 @@ Lemma closed_context_cumulativity {cf:checker_flags} {Σ} {wfΣ : wf Σ.1} Γ {p Proof. intros h hΓ' e. pose proof (ws_cumul_ctx_pb_forget e). - destruct pb; eapply context_cumulativity_prop; eauto. + eapply context_cumulativity_prop; eauto. + destruct pb; tas. eapply conv_cumul_context in e; tea. eapply (ws_cumul_ctx_pb_forget e). Qed. diff --git a/pcuic/theories/Typing/PCUICInstTyp.v b/pcuic/theories/Typing/PCUICInstTyp.v index a5ba61b4d..bc6dc5704 100644 --- a/pcuic/theories/Typing/PCUICInstTyp.v +++ b/pcuic/theories/Typing/PCUICInstTyp.v @@ -187,9 +187,9 @@ Proof. + eapply All2_map. repeat toAll. eapply All2_impl. 1: tea. cbn; intros. destruct_head'_prod; eauto. + unfold preturn. cbn. - unshelve erewrite (All2_fold_length _ : #|pcontext _| = #|pcontext _|); shelve_unifiable; tea. + unshelve erewrite (All2_length _ : #|pcontext _| = #|pcontext _|); shelve_unifiable; tea. exactly_once (idtac; multimatch goal with H : _ |- _ => eapply H end); eauto. - ++ unshelve erewrite <- (All2_fold_length _ : #|pcontext _| = #|pcontext _|); shelve_unifiable; tea. + ++ unshelve erewrite <- (All2_length _ : #|pcontext _| = #|pcontext _|); shelve_unifiable; tea. rewrite <- inst_case_predicate_context_length. rewrite inst_case_predicate_context_inst; eauto. eapply closed_subst_ext. 2: symmetry; apply up_Upn. @@ -209,17 +209,17 @@ Proof. unfold is_open_term. rewrite app_length. rewrite <- shiftnP_add. rewrite inst_case_predicate_context_length. - unshelve erewrite (All2_fold_length _ : #|pcontext _| = #|pcontext _|); shelve_unifiable; tea. + unshelve erewrite (All2_length _ : #|pcontext _| = #|pcontext _|); shelve_unifiable; tea. * eauto. - * rename Hbody into Hbrsbrs'. + * rename Hbody into Hbrsbrs'. unfold cumul_branches, cumul_branch in *. repeat toAll. eapply All2_impl. 1: tea. cbn; intros; destruct_head'_prod. - unshelve erewrite (All2_fold_length _ : #|bcontext _| = #|bcontext _|); shelve_unifiable; tea. + unshelve erewrite (All2_length _ : #|bcontext _| = #|bcontext _|); shelve_unifiable; tea. split; eauto. repeat match goal with H : is_true (andb _ _) |- _ => apply andb_and in H; destruct H end. rewrite -> test_context_k_closed_on_free_vars_ctx in *. exactly_once (idtac; multimatch goal with H : _ |- _ => eapply H end); eauto. - + unshelve erewrite <- (All2_fold_length _ : #|bcontext _| = #|bcontext _|); shelve_unifiable; tea. + + unshelve erewrite <- (All2_length _ : #|bcontext _| = #|bcontext _|); shelve_unifiable; tea. rewrite <- (inst_case_branch_context_length p). rewrite inst_case_branch_context_inst; eauto. eapply closed_subst_ext. 2: symmetry; apply up_Upn. @@ -242,9 +242,9 @@ Proof. unfold is_open_term. rewrite app_length. rewrite <- shiftnP_add. rewrite inst_case_branch_context_length. - unshelve erewrite (All2_fold_length _ : #|bcontext _| = #|bcontext _|); shelve_unifiable; tea. + unshelve erewrite (All2_length _ : #|bcontext _| = #|bcontext _|); shelve_unifiable; tea. - eapply cumul_Proj; try apply X0; eauto. - - cbn. eapply cumul_Fix. cbn in HfreeA, HfreeB. + - cbn. eapply cumul_Fix. cbn in HfreeA, HfreeB. unfold cumul_mfixpoint in *. repeat toAll. eapply All2_impl. 1: tea. cbn; intros. destruct_head'_prod. repeat split; eauto. @@ -285,7 +285,7 @@ Proof. rewrite fix_context_length. rewrite (All2_length X). eauto. - cbn. rewrite (All2_length X). - eapply cumul_CoFix. cbn in HfreeA, HfreeB. + eapply cumul_CoFix. cbn in HfreeA, HfreeB. unfold cumul_mfixpoint in *. repeat toAll. eapply All2_impl. 1: tea. cbn; intros. destruct_head'_prod. @@ -365,57 +365,51 @@ Lemma type_inst {cf : checker_flags} : env_prop wf_local Σ Δ -> Σ ;;; Δ ⊢ σ : Γ -> Σ ;;; Δ |- t.[σ] : A.[σ]) + (fun Σ Γ j => + forall Δ σ, + wf_local Σ Δ -> + Σ ;;; Δ ⊢ σ : Γ -> + lift_typing0 (fun t T => Σ ;;; Δ |- t.[σ] : T.[σ]) j) (fun Σ Γ => - All_local_env - (lift_typing (fun (Σ : global_env_ext) (Γ : context) (t T : term) - => + All_local_env (fun Γ j => forall Δ σ, wf_local Σ Δ -> Σ ;;; Δ ⊢ σ : Γ -> - Σ ;;; Δ |- t.[σ] : T.[σ]) Σ) Γ). + (lift_typing0 (fun (t T : term) => Σ ;;; Δ |- t.[σ] : T.[σ])) j) Γ). Proof. apply typing_ind_env. - - intros Σ wfΣ Γ wfΓ. auto. - induction 1; constructor; tas. - all: apply infer_typing_sort_impl with id tu; auto. + - intros Σ wfΣ Γ j H Δ σ wfΔ Hσ. + apply lift_typing_impl with (1 := H) => t T [_ IH]. + now apply IH. + + - intros Σ wfΣ Γ wfΓ _ X. assumption. - intros Σ wfΣ Γ wfΓ n decl e X Δ σ hΔ hσ. simpl. eapply hσ. assumption. - intros Σ wfΣ Γ wfΓ l X H0 Δ σ hΔ hσ. simpl. econstructor. all: assumption. - intros Σ wfΣ Γ wfΓ na A B s1 s2 X hA ihA hB ihB Δ σ hΔ hσ. autorewrite with sigma. simpl. - econstructor. - + eapply ihA ; auto. - + eapply ihB. - * econstructor ; auto. - eexists. eapply ihA ; auto. - * eapply well_subst_Up. 2: assumption. - econstructor ; auto. - eexists. eapply ihA. all: auto. - - intros Σ wfΣ Γ wfΓ na A t s1 bty X hA ihA ht iht Δ σ hΔ hσ. + assert (ihA' : lift_typing0 (typing Σ Δ) (j_vass_s na A.[σ] s1)) by now eapply ihA. + econstructor; tas. + apply lift_sorting_forget_univ in ihA'. + eassert (wf_local Σ (Δ ,, _)) by (constructor; eassumption). + eapply ihB; tea. + eapply well_subst_Up ; assumption. + - intros Σ wfΣ Γ wfΓ na A t bty X hA ihA ht iht Δ σ hΔ hσ. autorewrite with sigma. - econstructor. - + eapply ihA ; auto. - + eapply iht. - * econstructor ; auto. - eexists. eapply ihA ; auto. - * eapply well_subst_Up. 2: assumption. - constructor. 1: assumption. - eexists. eapply ihA. all: auto. - - intros Σ wfΣ Γ wfΓ na b B t s1 A X hB ihB hb ihb ht iht Δ σ hΔ hσ. + assert (ihA' : lift_typing0 (typing Σ Δ) (j_vass na A.[σ])) by now eapply ihA. + econstructor; tas. + eassert (wf_local Σ (Δ ,, _)) by (constructor; eassumption). + eapply iht; tea. + eapply well_subst_Up ; assumption. + - intros Σ wfΣ Γ wfΓ na b B t A X hbB ihbB ht iht Δ σ hΔ hσ. autorewrite with sigma. - econstructor. - + eapply ihB. all: auto. - + eapply ihb. all: auto. - + eapply iht. - * econstructor. all: auto. - -- eexists. eapply ihB. all: auto. - -- simpl. eapply ihb. all: auto. - * eapply well_subst_Up'; try assumption. - constructor; auto. - ** exists s1. apply ihB; auto. - ** apply ihb; auto. + assert (ihbB' : lift_typing0 (typing Σ Δ) (j_vdef na b.[σ] B.[σ])) by now eapply ihbB. + econstructor; tas. + eassert (wf_local Σ (Δ ,, _)) by (constructor; eassumption). + eapply iht; tea. + eapply well_subst_Up'; try assumption. - intros Σ wfΣ Γ wfΓ t na A B s u X hty ihty ht iht hu ihu Δ σ hΔ hσ. autorewrite with sigma. econstructor. @@ -460,10 +454,9 @@ Proof. apply All_local_env_app_inv in IHpredctx as []. eapply IHpret. ++ eapply wf_local_app_inst; eauto. - apply a2. ++ relativize #|pcontext p|; [eapply well_subst_app_up|] => //; rewrite /predctx; len. 2:{ rewrite case_predicate_context_length //. } - eapply wf_local_app_inst; eauto. apply a2. + eapply wf_local_app_inst; eauto. + simpl. unfold id. specialize (IHc _ _ HΔ Hf). now rewrite inst_mkApps map_app in IHc. @@ -506,7 +499,7 @@ Proof. assert (wf_local Σ (Δ,,, brctx'.1)). { rewrite /brctx'. cbn. apply All_local_env_app_inv in Hbrctx as []. - eapply wf_local_app_inst; tea. apply a0. } + eapply wf_local_app_inst; tea. } split => //. split. ++ eapply IHbr => //. rewrite /brctx' /brctxty; cbn. @@ -538,45 +531,47 @@ Proof. rewrite Upn_comp; cbn; try now repeat len. rewrite subst_consn_lt /=; cbn; len; try lia. now rewrite map_rev. - - intros Σ wfΣ Γ wfΓ mfix n decl types hguard hnth htypes hmfix ihmfix wffix Δ σ hΔ hσ. + - intros Σ wfΣ Γ wfΓ mfix n decl types hguard hnth htypes hmfixt ihmfixt hmfixb ihmfixb wffix Δ σ hΔ hσ. simpl. eapply meta_conv; [econstructor;eauto|]. * now eapply fix_guard_inst. * now rewrite nth_error_map hnth. - * solve_all. - apply infer_typing_sort_impl with id a; intros [_ IH]. - now apply IH. - * solve_all. destruct b as [? b0]. - len. rewrite /types in b0. len in b0. - pose proof (inst_fix_context mfix σ). + * apply All_map, (All_impl ihmfixt). + intros x t. eapply lift_typing_map with (j := TermoptTyp None _) => //. eapply t; eauto. + * pose proof (inst_fix_context mfix σ). setoid_rewrite <-up_Upn at 1 in H. rewrite H. - eapply All_local_env_app_inv in htypes as []. - eapply meta_conv; [eapply b0; eauto|]. - + eapply wf_local_app_inst; eauto. eapply a1. + apply All_map, (All_impl ihmfixb). + unfold on_def_body. + intros x t. relativize (lift0 _ _). + 1: eenough (wf_local Σ (Δ ,,, _)). + 1: eapply lift_typing_map with (j := TermoptTyp (Some _) _) => //; eapply t; eauto. + rewrite -(fix_context_length mfix). eapply well_subst_app_up => //. - eapply wf_local_app_inst; eauto. apply a1. - + rewrite lift0_inst. now sigma. + + eapply wf_local_app_inst; eauto. + now eapply All_local_env_app_inv in htypes as []. + + rewrite lift0_inst /types inst_context_length fix_context_length. + now sigma. * now apply inst_wf_fixpoint. * reflexivity. - - intros Σ wfΣ Γ wfΓ mfix n decl types hguard hnth htypes hmfix ihmfix wffix Δ σ hΔ hσ. + - intros Σ wfΣ Γ wfΓ mfix n decl types hguard hnth htypes hmfixt ihmfixt hmfixb ihmfixb wffix Δ σ hΔ hσ. simpl. eapply meta_conv; [econstructor;eauto|]. * now eapply cofix_guard_inst. * now rewrite nth_error_map hnth. - * solve_all. - apply infer_typing_sort_impl with id a; intros [_ IH]. - now apply IH. - * solve_all. destruct b as [? b0]. - len. rewrite /types in b0. len in b0. - pose proof (inst_fix_context mfix σ). + * apply All_map, (All_impl ihmfixt). + intros x t. eapply lift_typing_map with (j := TermoptTyp None _) => //. eapply t; eauto. + * pose proof (inst_fix_context mfix σ). setoid_rewrite <-up_Upn at 1 in H. rewrite H. - eapply All_local_env_app_inv in htypes as []. - eapply meta_conv; [eapply b0; eauto|]. - + eapply wf_local_app_inst; eauto. eapply a1. + apply All_map, (All_impl ihmfixb). + unfold on_def_body. + intros x t. relativize (lift0 _ _). + 1: eenough (wf_local Σ (Δ ,,, _)). + 1: eapply lift_typing_map with (j := TermoptTyp (Some _) _) => //; eapply t; eauto. + rewrite -(fix_context_length mfix). eapply well_subst_app_up => //. - eapply wf_local_app_inst; eauto. apply a1. - + rewrite lift0_inst. now sigma. + + eapply wf_local_app_inst; eauto. + now eapply All_local_env_app_inv in htypes as []. + + rewrite lift0_inst /types inst_context_length fix_context_length. + now sigma. * now apply inst_wf_cofixpoint. * reflexivity. diff --git a/pcuic/theories/Typing/PCUICNamelessTyp.v b/pcuic/theories/Typing/PCUICNamelessTyp.v index e4974c121..9b60d2709 100644 --- a/pcuic/theories/Typing/PCUICNamelessTyp.v +++ b/pcuic/theories/Typing/PCUICNamelessTyp.v @@ -102,7 +102,7 @@ Proof. now setoid_rewrite <- nl_check_one_cofix. Qed. -Lemma wf_universe_nl Σ u : wf_universe Σ u -> wf_universe (nlg Σ) u. +Lemma wf_sort_nl Σ u : wf_sort Σ u -> wf_sort (nlg Σ) u. Proof. destruct u; simpl; auto. intros. @@ -150,7 +150,7 @@ Proof. constructor. 1: eauto using nlg_wf_local. unfold nlctx. rewrite nth_error_map. now rewrite H. - constructor; eauto using nlg_wf_local. - now apply wf_universe_nl. + now apply wf_sort_nl. - replace (nl (cst_type decl)) with (cst_type (map_constant_body nl decl)); [|destruct decl; reflexivity]. constructor; eauto using nlg_wf_local. @@ -258,7 +258,7 @@ Proof. constructor. + now eapply fix_guard_nl. + now rewrite nth_error_map H0. - + rewrite nlctx_app_context in X. now eapply wf_local_app_inv in X. + + rewrite nlctx_app_context in X. now eapply All_local_env_app_inv in X. + clear -X0. apply All_map. eapply All_impl; tea. simpl. intros x [s Hs]. now exists s. @@ -277,7 +277,7 @@ Proof. constructor; auto. + now apply cofix_guard_nl. + now rewrite nth_error_map H0. - + now rewrite nlctx_app_context in X; apply wf_local_app_inv in X. + + now rewrite nlctx_app_context in X; apply All_local_env_app_inv in X. + clear -X0. apply All_map. eapply All_impl; tea. simpl. intros x [s Hs]. now exists s. diff --git a/pcuic/theories/Typing/PCUICRenameTyp.v b/pcuic/theories/Typing/PCUICRenameTyp.v index ce8cd6246..e94c5f32e 100644 --- a/pcuic/theories/Typing/PCUICRenameTyp.v +++ b/pcuic/theories/Typing/PCUICRenameTyp.v @@ -260,11 +260,11 @@ Proof. repeat split; eauto. + eapply All2_map. apply forallb_All in Hp, Hp'. repeat toAll. eapply All2_impl. 1: tea. cbn; intros. destruct X as [[X [X''' X']] X'']. apply X'; eauto. - + unfold preturn. cbn. unshelve erewrite (All2_fold_length _ : #|pcontext _| = #|pcontext _|); shelve_unifiable; tea. + + unfold preturn. cbn. unshelve erewrite (All2_length _ : #|pcontext _| = #|pcontext _|); shelve_unifiable; tea. exactly_once (idtac; multimatch goal with H : _ |- _ => eapply H end); eauto. ++ rewrite app_context_length. eapply urenaming_ext; try apply shiftnP_add; try reflexivity. - unshelve erewrite <- (All2_fold_length _ : #|pcontext _| = #|pcontext _|); shelve_unifiable; tea. + unshelve erewrite <- (All2_length _ : #|pcontext _| = #|pcontext _|); shelve_unifiable; tea. rewrite <- inst_case_predicate_context_length. rewrite test_context_k_closed_on_free_vars_ctx in Hcontext. rewrite inst_case_predicate_context_rename; eauto. @@ -281,25 +281,26 @@ Proof. unfold is_open_term. rewrite app_length. rewrite <- shiftnP_add. rewrite inst_case_predicate_context_length. - unshelve erewrite (All2_fold_length _ : #|pcontext _| = #|pcontext _|); shelve_unifiable; tea. + unshelve erewrite (All2_length _ : #|pcontext _| = #|pcontext _|); shelve_unifiable; tea. ++ rewrite test_context_k_closed_on_free_vars_ctx in Hcontext. unfold inst_case_predicate_context. apply on_free_vars_ctx_inst_case_context; eauto. +++ eapply All_forallb. apply All_map. apply forallb_All in Hp; eapply All_impl. 1: tea. cbn; intros. eapply urename_is_open_term; eauto. +++ unfold pparams. cbn. rewrite map_length. exact Hcontext. * eauto. - * let X2 := match goal with H : All2 _ brs brs' |- _ => H end in + * unfold cumul_branches, cumul_branch in *. + let X2 := match goal with H : All2 _ brs brs' |- _ => H end in rename X2 into Hbrsbrs'. apply forallb_All in Hbrs, Hbrs'. apply (All2_All_mix_left Hbrs) in Hbrsbrs'. clear Hbrs. apply (All2_All_mix_right Hbrs') in Hbrsbrs'. clear Hbrs'. apply All2_map. repeat toAll. eapply All2_impl. 1: tea. cbn; intros; destruct_head'_prod. - unshelve erewrite (All2_fold_length _ : #|bcontext _| = #|bcontext _|); shelve_unifiable; tea. + unshelve erewrite (All2_length _ : #|bcontext _| = #|bcontext _|); shelve_unifiable; tea. split; eauto. repeat match goal with H : is_true (andb _ _) |- _ => apply andb_and in H; destruct H end. exactly_once (idtac; multimatch goal with H : _ |- _ => apply H end); eauto. + rewrite app_context_length. eapply urenaming_ext; try apply shiftnP_add; try reflexivity. - unshelve erewrite <- (All2_fold_length _ : #|bcontext _| = #|bcontext _|); shelve_unifiable; tea. + unshelve erewrite <- (All2_length _ : #|bcontext _| = #|bcontext _|); shelve_unifiable; tea. rewrite <- (inst_case_branch_context_length p). rewrite -> test_context_k_closed_on_free_vars_ctx in *. rewrite inst_case_branch_context_rename; eauto. @@ -317,7 +318,7 @@ Proof. unfold is_open_term. rewrite app_length. rewrite <- shiftnP_add. rewrite inst_case_branch_context_length. - unshelve erewrite (All2_fold_length _ : #|bcontext _| = #|bcontext _|); shelve_unifiable; tea. + unshelve erewrite (All2_length _ : #|bcontext _| = #|bcontext _|); shelve_unifiable; tea. + rewrite test_context_k_closed_on_free_vars_ctx in Hcontext. unfold inst_case_predicate_context. apply on_free_vars_ctx_inst_case_context; eauto. ++ eapply All_forallb. apply All_map. repeat toAll. eapply All_impl. 1: tea. @@ -610,35 +611,25 @@ Lemma rename_predicate_preturn f p : Proof. reflexivity. Qed. Lemma wf_local_app_renaming P Σ Γ Δ : - All_local_env (lift_typing (fun (Σ : global_env_ext) (Γ' : context) (t T : term) => - forall P (Δ : PCUICEnvironment.context) (f : nat -> nat), - renaming (shiftnP #|Γ ,,, Γ'| P) Σ (Γ ,,, Γ') Δ f -> Σ ;;; Δ |- rename f t : rename f T) Σ) - Δ -> + All_local_rel (fun (Γ : context) j => + forall P Δ (f : nat -> nat), + renaming (shiftnP #|Γ| P) Σ Γ Δ f -> + lift_typing0 (fun t T => Σ ;;; Δ |- rename f t : rename f T) j) + Γ Δ -> forall Δ' f, renaming (shiftnP #|Γ| P) Σ Γ Δ' f -> wf_local Σ (Δ' ,,, rename_context f Δ). Proof. intros. destruct X0. - induction X. - - apply a. - - rewrite rename_context_snoc /=. constructor; auto. - apply infer_typing_sort_impl with id t0; intros Hs. - eapply (Hs P (Δ' ,,, rename_context f Γ0) (shiftn #|Γ0| f)). - split => //. - eapply urenaming_ext. - { now rewrite app_length -shiftnP_add. } - { reflexivity. } now eapply urenaming_context. - - rewrite rename_context_snoc /=. constructor; auto. - * apply infer_typing_sort_impl with id t0; intros Hs. - apply (Hs P (Δ' ,,, rename_context f Γ0) (shiftn #|Γ0| f)). - split => //. - eapply urenaming_ext. - { now rewrite app_length -shiftnP_add. } - { reflexivity. } now eapply urenaming_context. - * red. apply (t1 P). split => //. - eapply urenaming_ext. - { now rewrite app_length -shiftnP_add. } - { reflexivity. } now eapply urenaming_context. + induction X using All_local_rel_ind1. + 1: now apply a. + rewrite rename_context_snoc /=. apply All_local_env_snoc; auto. + eapply lift_typing_map with (j := TermoptTyp _ _) => //. + eapply X. + split; auto. + eapply urenaming_ext. + { now rewrite app_length -shiftnP_add. } + { reflexivity. } now eapply urenaming_context. Qed. Lemma rename_decompose_prod_assum f Γ t : @@ -796,26 +787,27 @@ Proof. Qed. (** For an unconditional renaming defined on all variables in the source context *) -Lemma typing_rename_prop : env_prop +Lemma typing_rename_prop : + let Pj := (fun Σ Γ j => + forall P Δ f, + renaming (shiftnP #|Γ| P) Σ Γ Δ f -> + lift_typing (fun Σ Γ t T => Σ ;;; Δ |- rename f t : rename f T) Σ Γ j) + in + env_prop (fun Σ Γ t A => forall P Δ f, renaming (shiftnP #|Γ| P) Σ Γ Δ f -> Σ ;;; Δ |- rename f t : rename f A) - (fun Σ Γ => - wf_local Σ Γ × - All_local_env - (lift_typing (fun (Σ : global_env_ext) (Γ : context) (t T : term) - => - forall P (Δ : PCUICEnvironment.context) (f : nat -> nat), - renaming (shiftnP #|Γ| P) Σ Γ Δ f -> - Σ;;; Δ |- rename f t : rename f T) Σ) Γ). + Pj + (fun Σ Γ => All_local_env (Pj Σ) Γ). Proof. apply typing_ind_env. - - intros Σ wfΣ Γ wfΓ HΓ. split; auto. - induction HΓ; constructor; tas. - all: apply infer_typing_sort_impl with id tu; intros Hty. - all: eauto. + - intros Σ wfΣ Γ j H P Δ f Hr. + apply lift_typing_impl with (1 := H) => t T [_ IH]. + now eapply IH. + + - intros Σ wfΣ Γ _ _ HΓ. assumption. - intros Σ wfΣ Γ wfΓ n decl isdecl ihΓ P Δ f hf. simpl in *. @@ -829,37 +821,36 @@ Proof. simpl. constructor. all: auto. - intros Σ wfΣ Γ wfΓ na A B s1 s2 X hA ihA hB ihB P Δ f hf. - rewrite /=. econstructor. - + eapply ihA; eauto. - + eapply ihB; eauto. + assert (lift_typing0 (typing Σ Δ) (j_vass_s na (rename f A) s1)). + + eapply lift_typing_map with (j := TypUniv _ _) => //. + eapply ihA; eauto. + + rewrite /=. econstructor; tas. + eapply ihB; eauto. simpl. eapply renaming_extP. { now rewrite -(shiftnP_add 1). } eapply renaming_vass. 2: eauto. constructor. - * destruct hf as [hΔ hf]. auto. - * simpl. exists s1. eapply ihA; eauto. - - intros Σ wfΣ Γ wfΓ na A t s1 B X hA ihA ht iht P Δ f hf. - simpl. - (* /andP [_ havB]. *) - simpl. econstructor. - + eapply ihA; eauto. - + eapply iht; eauto; simpl. + * destruct hf as [hΔ hf]. assumption. + * eapply lift_sorting_forget_univ, X0. + - intros Σ wfΣ Γ wfΓ na A t B X hA ihA ht iht P Δ f hf. + assert (lift_typing0 (typing Σ Δ) (j_vass na (rename f A))). + + eapply lift_typing_map with (j := Typ _) => //. + eapply ihA; eauto. + + simpl. econstructor; tas. + eapply iht; eauto; simpl. eapply renaming_extP. { now rewrite -(shiftnP_add 1). } eapply renaming_vass. 2: eauto. - constructor. - * destruct hf as [hΔ hf]. auto. - * simpl. exists s1. eapply ihA; eauto. - - intros Σ wfΣ Γ wfΓ na b B t s1 A X hB ihB hb ihb ht iht P Δ f hf. - simpl. econstructor. - + eapply ihB; tea. - + eapply ihb; tea. - + eapply iht; tea. + constructor; tas. + destruct hf as [hΔ hf]. assumption. + - intros Σ wfΣ Γ wfΓ na b B t A X hbB ihbB ht iht P Δ f hf. + assert (lift_typing0 (typing Σ Δ) (j_vdef na (rename f b) (rename f B))). + + eapply ihbB; tea. + + simpl. econstructor; tas. + eapply iht; tea. eapply renaming_extP. { now rewrite -(shiftnP_add 1). } eapply renaming_vdef. 2: eauto. - constructor. - * destruct hf. assumption. - * simpl. eexists. eapply ihB; tea. - * simpl. eapply ihb; tea. + constructor; tas. + destruct hf as [hΔ hf]. assumption. - intros Σ wfΣ Γ wfΓ t na A B s u X hty ihty ht iht hu ihu P Δ f hf. simpl. eapply meta_conv. + eapply type_App. @@ -891,7 +882,7 @@ Proof. + rewrite rename_closed. 2: reflexivity. eapply declared_constructor_closed_type. all: eauto. - intros Σ wfΣ Γ wfΓ ci p c brs indices ps mdecl idecl isdecl HΣ. - intros [_ IHΔ] ci_npar eqpctx predctx wfp cup wfpctx Hpret IHpret [_ IHpredctx] isallowed. + intros IHΔ ci_npar eqpctx predctx wfp cup wfpctx Hpret IHpret IHpredctx isallowed. intros IHctxi Hc IHc iscof ptm wfbrs Hbrs P Δ f Hf. simpl. rewrite rename_mkApps. @@ -906,7 +897,7 @@ Proof. + simpl. eapply IHpret. split. ++ apply All_local_env_app_inv in IHpredctx as []. - eapply wf_local_app_renaming; eauto. apply a0. + eapply wf_local_app_renaming; eauto. ++ rewrite /predctx app_length. eapply urenaming_ext. { now rewrite -shiftnP_add. } @@ -920,7 +911,7 @@ Proof. + now eapply rename_wf_predicate. + simpl. eauto. apply All_local_env_app_inv in IHpredctx as []. - eapply wf_local_app_renaming; eauto. apply a0. + eapply wf_local_app_renaming; eauto. + revert IHctxi. rewrite /= /id -map_app. rewrite -{2}[subst_instance _ _](rename_closedn_ctx f 0). @@ -943,7 +934,7 @@ Proof. eapply All2i_nth_hyp in Hbrs. eapply All2i_map_right, (All2i_impl Hbrs) => i cdecl br. set (brctxty := case_branch_type _ _ _ _ _ _ _ _). - intros (Hnth & wfbr & eqbctx & (wfbctx & IHbctx) & (Hbod & IHbod) & Hbty & IHbty). + intros (Hnth & wfbr & eqbctx & IHbctx & (Hbod & IHbod) & Hbty & IHbty). rewrite -(rename_closed_constructor_body mdecl cdecl f). { eapply (declared_constructor_closed (c:=(ci.(ci_ind),i))); eauto. split; eauto. } @@ -963,7 +954,7 @@ Proof. set (brctx' := rename_context f _). assert (wf_local Σ (Δ ,,, brctx')). { apply All_local_env_app_inv in IHbctx as []. - eapply wf_local_app_renaming; tea. simpl. apply a0. } + eapply wf_local_app_renaming; tea. } split => //. rewrite rename_case_predicate_context // rename_case_branch_type //. split. @@ -1000,7 +991,7 @@ Proof. eapply declared_projection_closed_type in isdecl. rewrite List.rev_length. rewrite e. assumption. - - intros Σ wfΣ Γ wfΓ mfix n decl types H1 hdecl [_ X] ihmfixt ihmfixb wffix P Δ f hf. + - intros Σ wfΣ Γ wfΓ mfix n decl types H1 hdecl X hmfixt ihmfixt hmfixb ihmfixb wffix P Δ f hf. apply All_local_env_app_inv in X as [_ X]. eapply wf_local_app_renaming in X; tea. simpl. eapply meta_conv. @@ -1009,24 +1000,24 @@ Proof. * destruct hf; eapply fix_guard_rename; eauto. * rewrite nth_error_map. rewrite hdecl. simpl. reflexivity. * apply All_map, (All_impl ihmfixt). - intros x t. apply infer_typing_sort_impl with id t. - intros [_ IHs]; now eapply IHs. + intros x t. eapply lift_typing_map with (j := TermoptTyp None _) => //. eapply t. eauto. * apply All_map, (All_impl ihmfixb). - intros x [Hb IHb]. - destruct x as [na ty bo rarg]. simpl in *. + unfold on_def_body. rewrite fix_context_length map_length {2}/map_def /=. + intros x t. + relativize (lift0 _ _). + 1: eapply lift_typing_map with (j := TermoptTyp (Some _) _) => //; eapply t; eauto. + 2: len; now sigma. rewrite rename_fix_context. - eapply meta_conv. - ++ apply (IHb P (Δ ,,, rename_context f types) (shiftn #|mfix| f)). - split; auto. subst types. rewrite -(fix_context_length mfix). - eapply urenaming_ext. - { now rewrite app_context_length -shiftnP_add. } - { reflexivity. } - apply urenaming_context; auto. apply hf. - ++ len. now sigma. + split; auto. + subst types. rewrite -(fix_context_length mfix). + eapply urenaming_ext. + { now rewrite app_context_length -shiftnP_add. } + { reflexivity. } + apply urenaming_context; auto. apply hf. * now eapply rename_wf_fixpoint. + reflexivity. - - intros Σ wfΣ Γ wfΓ mfix n decl types guard hdecl [_ X] ihmfixt ihmfixb wfcofix P Δ f hf. + - intros Σ wfΣ Γ wfΓ mfix n decl types guard hdecl X hmfixt ihmfixt hmfixb ihmfixb wfcofix P Δ f hf. apply All_local_env_app_inv in X as [_ X]. eapply wf_local_app_renaming in X; eauto. simpl. eapply meta_conv. @@ -1035,20 +1026,20 @@ Proof. * destruct hf; eapply cofix_guard_rename; eauto. * rewrite nth_error_map. rewrite hdecl. simpl. reflexivity. * apply All_map, (All_impl ihmfixt). - intros x t. apply infer_typing_sort_impl with id t. - intros [_ IHs]; now eapply IHs. + intros x t. eapply lift_typing_map with (j := TermoptTyp None _) => //. eapply t. eauto. * apply All_map, (All_impl ihmfixb). - intros x [Hb IHb]. - destruct x as [na ty bo rarg]. simpl in *. + unfold on_def_body. rewrite fix_context_length map_length {2}/map_def /=. + intros x t. + relativize (lift0 _ _). + 1: eapply lift_typing_map with (j := TermoptTyp (Some _) _) => //; eapply t; eauto. + 2: len; now sigma. rewrite rename_fix_context. - eapply meta_conv. - ++ apply (IHb P (Δ ,,, rename_context f types) (shiftn #|mfix| f)). - split; auto. subst types. rewrite -(fix_context_length mfix). - eapply urenaming_ext. - { now rewrite app_context_length -shiftnP_add. } - { reflexivity. } - apply urenaming_context; auto. apply hf. - ++ len. now sigma. + split; auto. + subst types. rewrite -(fix_context_length mfix). + eapply urenaming_ext. + { now rewrite app_context_length -shiftnP_add. } + { reflexivity. } + apply urenaming_context; auto. apply hf. * now eapply rename_wf_cofixpoint. + reflexivity. diff --git a/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v b/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v index 75594d5d3..f5a3d55c4 100644 --- a/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v +++ b/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v @@ -21,17 +21,17 @@ Local Ltac aa := rdest; eauto with univ_subst. Section SubstIdentity. Context `{cf:checker_flags}. -Lemma compare_universe_subst_instance pb {Σ : global_env_ext} univs u : +Lemma compare_sort_subst_instance pb {Σ : global_env_ext} univs u : valid_constraints (global_ext_constraints (Σ.1, univs)) (subst_instance_cstrs u Σ) -> - RelationClasses.subrelation (compare_universe pb Σ) - (fun x y : Universe.t => - compare_universe pb (global_ext_constraints (Σ.1, univs)) (subst_instance_univ u x) - (subst_instance_univ u y)). + RelationClasses.subrelation (compare_sort Σ pb) + (fun x y : sort => + compare_sort (global_ext_constraints (Σ.1, univs)) pb (subst_instance_sort u x) + (subst_instance_sort u y)). Proof using Type. intros v. destruct pb; cbn. - - now apply eq_universe_subst_instance. - - now apply leq_universe_subst_instance. + - now apply eq_sort_subst_instance. + - now apply leq_sort_subst_instance. Qed. Lemma cumulSpec_subst_instance (Σ : global_env_ext) Γ u A B pb univs : @@ -105,12 +105,13 @@ Proof. - eapply cumul_Case; try solve [intuition eauto; eauto]. * cbv [cumul_predicate] in *; destruct_head'_prod. repeat split; eauto; cbn. + apply All2_map. repeat toAll. eapply All2_impl. 1: tea. cbn; intros. destruct_head'_prod; eauto. - + apply precompose_subst_instance. eapply R_universe_instance_impl. + + apply precompose_subst_instance. eapply cmp_universe_instance_impl. 1:now apply eq_universe_subst_instance. eauto. + rewrite -> subst_instance_app, inst_case_predicate_context_subst_instance in *. eauto. - * eapply All2_map. repeat toAll. eapply All2_impl. 1: tea. cbn; intros. + * unfold cumul_branches, cumul_branch in *. + eapply All2_map. repeat toAll. eapply All2_impl. 1: tea. cbn; intros. repeat split; intuition eauto. rewrite -> subst_instance_app, inst_case_branch_context_subst_instance in *; eauto. - eapply cumul_Fix. apply All2_map. repeat toAll. eapply All2_impl. 1: tea. @@ -120,26 +121,26 @@ Proof. cbn; intros; intuition eauto. rewrite -> subst_instance_app, fix_context_subst_instance in *; eauto. - eapply cumul_Prim. depelim e0; depelim X; cbn in H; cbn; noconf H; cbn in H; constructor; cbn -[Universe.make]; eauto. - + rewrite -!subst_instance_univ_make. + + rewrite -!subst_instance_universe_make. eapply eq_universe_subst_instance; tea. + solve_all. - repeat rewrite subst_instance_mkApps. eapply cumul_Ind. * apply precompose_subst_instance_global. - rewrite map_length. eapply R_global_instance_impl_same_napp; try eapply H; eauto. + rewrite map_length. eapply cmp_global_instance_impl_same_napp; try eapply H; eauto. { now apply eq_universe_subst_instance. } { now apply compare_universe_subst_instance. } * eapply All2_map. repeat toAll. eapply All2_impl. 1: tea. cbn; intros. eapply X0.2; eauto. - repeat rewrite subst_instance_mkApps. eapply cumul_Construct. * apply precompose_subst_instance_global. cbn. - rewrite map_length. eapply R_global_instance_impl_same_napp; try eapply H; eauto. + rewrite map_length. eapply cmp_global_instance_impl_same_napp; try eapply H; eauto. { now apply eq_universe_subst_instance. } { now apply compare_universe_subst_instance. } * eapply All2_map. repeat toAll. eapply All2_impl. 1: tea. cbn; intros. eapply X0.2; eauto. - - eapply cumul_Sort. now apply compare_universe_subst_instance. + - eapply cumul_Sort. now apply compare_sort_subst_instance. - eapply cumul_Const. apply precompose_subst_instance. - eapply R_universe_instance_impl; eauto. + eapply cmp_universe_instance_impl; eauto. now apply compare_universe_subst_instance. Defined. @@ -235,22 +236,27 @@ Lemma typing_subst_instance : consistent_instance_ext (Σ.1, univs) Σ.2 u -> (Σ.1,univs) ;;; subst_instance u Γ |- subst_instance u t : subst_instance u T) + (fun Σ Γ j => forall u univs, + wf_ext_wk Σ -> + consistent_instance_ext (Σ.1, univs) Σ.2 u -> + lift_typing (fun Σ Γ t T => + (Σ.1,univs) ;;; Γ@[u] |- t@[u] : T@[u]) Σ Γ j) (fun Σ Γ => forall u univs, wf_ext_wk Σ -> consistent_instance_ext (Σ.1, univs) Σ.2 u -> wf_local(Σ.1,univs) (subst_instance u Γ)). Proof using Type. apply typing_ind_env; intros Σ wfΣ Γ wfΓ; cbn -[Universe.make] in *. + - intros X ????. + apply lift_typing_impl with (1 := X) => t T [_ IH]. + now eapply IH. - rewrite /subst_instance /=. - induction 1. - + constructor. - + simpl. constructor; auto. - eapply infer_typing_sort_impl; tea. - intros Hty. eapply Hs; auto. - + simpl. constructor; auto. - ++ eapply infer_typing_sort_impl; tea. - intros Hty. eapply Hs; auto. - ++ apply Hc; auto. + intros _; clear wfΓ. + induction 1 using All_local_env_ind1. 1: constructor. + intros. simpl. + apply All_local_env_snoc. 1: now apply IHX. + eapply lift_typing_mapu with (u := None) => //. + now apply X. - intros n decl eq X u univs wfΣ' H. rewrite subst_instance_lift. rewrite map_decl_type. econstructor; aa. @@ -260,17 +266,19 @@ Proof using Type. rewrite subst_instance_univ_super. + econstructor. * aa. - * now apply wf_universe_subst_instance. + * now apply wf_sort_subst_instance. - intros n t0 b s1 s2 X X0 X1 X2 X3 u univs wfΣ' H. rewrite product_subst_instance; aa. econstructor. - + eapply X1; eauto. + + eapply lift_typing_mapu with (tm := None) (u := Some _) => //. apply X1; eauto. + eapply X3; eauto. - - intros n t0 b s1 bty X X0 X1 X2 X3 u univs wfΣ' H. + - intros n t0 b bty X X0 X1 X2 X3 u univs wfΣ' H. econstructor. - + eapply X1; aa. + + eapply lift_typing_mapu with (tm := None) (u := None) => //. eapply X1; aa. + + eapply X3; aa. + - intros n b b_ty b' b'_ty X X0 X1 X2 X3 u univs wfΣ' H. + econstructor; eauto. + + eapply lift_typing_mapu with (tm := Some _) (u := None) => //. eapply X1; aa. + eapply X3; aa. - - intros n b b_ty b' s1 b'_ty X X0 X1 X2 X3 X4 X5 u univs wfΣ' H. - econstructor; eauto. eapply X5; aa. - intros t0 na A B s u X X0 X1 X2 X3 X4 X5 u0 univs wfΣ' H. rewrite subst_instance_subst. cbn. econstructor. + eapply X1; eauto. @@ -300,7 +308,7 @@ Proof using Type. change (map_predicate _ _ _ _ _) with (subst_instance i p). rewrite subst_instance_case_predicate_context. eapply type_Case with (p:=subst_instance i p) - (ps:=subst_instance_univ i u); eauto with pcuic. + (ps:=subst_instance_sort i u); eauto with pcuic. 3,4: constructor; eauto with pcuic. + rewrite -subst_instance_case_predicate_context - !subst_instance_app_ctx. eapply Hpty; eauto. @@ -337,26 +345,20 @@ Proof using Type. eapply X2 in H0; tas. rewrite subst_instance_mkApps in H0. eassumption. - - intros mfix n decl H H0 H1 X X0 wffix u univs wfΣ'. + - intros mfix n decl H H0 H1 X IHX X0 IHX0 wffix u univs wfΣ'. rewrite (map_dtype _ (subst_instance u)). econstructor. + specialize (H1 u univs wfΣ' H2). rewrite subst_instance_app in H1. - now eapply wf_local_app_inv in H1 as []. + now eapply All_local_env_app_inv in H1 as []. + now eapply fix_guard_subst_instance. + rewrite nth_error_map H0. reflexivity. - + apply All_map, (All_impl X); simpl; intuition auto. - eapply infer_typing_sort_impl with (tu := X1). - intros [_ Hs]; now apply Hs. - + eapply All_map, All_impl; tea. - intros x [X1 X3]. - specialize (X3 u univs wfΣ' H2). - rewrite (map_dbody (subst_instance u)) in X3. - rewrite subst_instance_lift in X3. - rewrite fix_context_length ?map_length in X0, X1, X3. - rewrite (map_dtype _ (subst_instance u) x) in X3. - rewrite subst_instance_app in X3. - rewrite <- (fix_context_subst_instance u mfix). - now len. + + apply All_map, (All_impl IHX); simpl. intros d X1. + eapply lift_typing_mapu with (tm := None) (u := None); cbn; eauto. + + eapply All_map, (All_impl IHX0); simpl. intros d X1. + unfold map_def at 2, on_def_body. cbn. + rewrite -fix_context_subst_instance /app_context -(subst_instance_app u (fix_context mfix) Γ) -/(app_context Γ _). + rewrite -subst_instance_lift map_length. + eapply lift_typing_mapu with (tm := Some _) (u := None); cbn; eauto. + red; rewrite <- wffix. unfold wf_fixpoint, wf_fixpoint_gen. f_equal. @@ -365,30 +367,24 @@ Proof using Type. rewrite map_map_compose. now rewrite subst_instance_check_one_fix. - - intros mfix n decl H H0 H1 X X0 wffix u univs wfΣ'. - rewrite (map_dtype _ (subst_instance u)). econstructor. - + specialize (H1 u univs wfΣ' H2). - rewrite subst_instance_app in H1. - now eapply wf_local_app_inv in H1 as []. - + now eapply cofix_guard_subst_instance. - + rewrite nth_error_map H0. reflexivity. - + apply All_map, (All_impl X); simpl; intuition auto. - eapply infer_typing_sort_impl with (tu := X1). - intros [_ Hs]; now apply Hs. - + eapply All_map, All_impl; tea. - intros x [X1 X3]. - specialize (X3 u univs wfΣ' H2). - rewrite (map_dbody (subst_instance u)) in X3. - rewrite subst_instance_lift in X3. - rewrite fix_context_length ?map_length in X0, X1, X3. - rewrite (map_dtype _ (subst_instance u) x) in X3. - rewrite subst_instance_app in X3. - rewrite <- (fix_context_subst_instance u mfix). - now len. - + red; rewrite <- wffix. - unfold wf_cofixpoint, wf_cofixpoint_gen. - rewrite map_map_compose. - now rewrite subst_instance_check_one_cofix. + - intros mfix n decl H H0 H1 X IHX X0 IHX0 wffix u univs wfΣ'. + rewrite (map_dtype _ (subst_instance u)). econstructor. + + specialize (H1 u univs wfΣ' H2). + rewrite subst_instance_app in H1. + now eapply All_local_env_app_inv in H1 as []. + + now eapply cofix_guard_subst_instance. + + rewrite nth_error_map H0. reflexivity. + + apply All_map, (All_impl IHX); simpl. intros d X1. + eapply lift_typing_mapu with (tm := None) (u := None); cbn; eauto. + + eapply All_map, (All_impl IHX0); simpl. intros d X1. + unfold map_def at 2, on_def_body. cbn. + rewrite -fix_context_subst_instance /app_context -(subst_instance_app u (fix_context mfix) Γ) -/(app_context Γ _). + rewrite -subst_instance_lift map_length. + eapply lift_typing_mapu with (tm := Some _) (u := None); cbn; eauto. + + red; rewrite <- wffix. + unfold wf_cofixpoint, wf_cofixpoint_gen. + rewrite map_map_compose. + now rewrite subst_instance_check_one_cofix. - intros. rewrite subst_instance_prim_type. @@ -398,9 +394,10 @@ Proof using Type. + exact H0. + now rewrite subst_instance_prim_val_tag. + destruct p as [? []]; depelim X1; constructor; eauto. - * rewrite -subst_instance_univ_make. eapply wf_universe_subst_instance => //. - * cbn -[Universe.make] in hty. - specialize (hty u univs). rewrite subst_instance_univ_make in hty. now eapply hty. + * rewrite -subst_instance_universe_make. eapply wf_universe_subst_instance => //. + * cbn -[Universe.make'] in hty. + specialize (hty u univs). + rewrite /subst_instance subst_instance_universe_make in hty. now eapply hty. * cbn. solve_all. - intros t0 A B X X0 X1 X2 X3 X4 cum u univs wfΣ' H. @@ -479,7 +476,7 @@ Lemma wf_local_instantiate_poly {Σ ctx Γ u} : wf_local Σ (subst_instance u Γ). Proof using Type. intros wfΣ Huniv wf. - epose proof (type_Sort _ _ Universes.Universe.lProp wf) as ty. forward ty. + epose proof (type_Sort _ _ sProp wf) as ty. forward ty. - now simpl. - eapply typing_subst_instance_ctx in ty; cbn; eauto using typing_wf_local. @@ -495,7 +492,7 @@ Lemma wf_local_instantiate {Σ} {decl : global_decl} {Γ u c} : wf_local Σ (subst_instance u Γ). Proof using Type. intros wfΣ Hdecl Huniv wf. - epose proof (type_Sort _ _ Universes.Universe.lProp wf) as ty. forward ty. + epose proof (type_Sort _ _ sProp wf) as ty. forward ty. - now simpl. - eapply typing_subst_instance_decl in ty; cbn; eauto using typing_wf_local. @@ -509,8 +506,8 @@ Lemma isType_subst_instance_decl Σ Γ T c decl u : isType Σ (subst_instance u Γ) (subst_instance u T). Proof using Type. intros wfΣ look isty cu. - eapply infer_typing_sort_impl with (tu := isty). - intros Hs; now eapply (typing_subst_instance_decl _ _ _ (tSort _)). + eapply lift_typing_fu_impl with (1 := isty) => //. + intros ?? Hs; now eapply typing_subst_instance_decl. Qed. Lemma isArity_subst_instance u T : @@ -528,8 +525,7 @@ Lemma wf_local_subst_instance Σ Γ ext u : Proof using Type. destruct Σ as [Σ φ]. intros X X0 X1. simpl in *. induction X1; cbn; constructor; auto. - 1,2: eapply infer_typing_sort_impl with (tu := t0); intros Hs. - 3: rename t1 into Hs. + all: eapply lift_typing_fu_impl with (1 := t0) => //= ?? Hs. all: eapply typing_subst_instance'' in Hs; eauto; apply X. Qed. @@ -542,27 +538,35 @@ Lemma wf_local_subst_instance_decl Σ Γ c decl u : Proof using Type. destruct Σ as [Σ φ]. intros X X0 X1 X2. induction X1; cbn; constructor; auto. - 1,2: eapply infer_typing_sort_impl with (tu := t0); intros Hs. - 3: rename t1 into Hs. + all: eapply lift_typing_fu_impl with (1 := t0) => // ?? Hs. all: eapply typing_subst_instance_decl in Hs; eauto; apply X. Qed. + Lemma isType_subst_instance_id Σ Γ T : + wf_ext_wk Σ -> + let u := abstract_instance Σ.2 in + isType Σ Γ T -> subst_instance u T = T. + Proof using Type. + intros wf_ext u isT. + destruct isT as (_ & s & Hs & _). + eapply typed_subst_abstract_instance in Hs; auto. + Qed. + Lemma subst_instance_ind_sort_id Σ mdecl ind idecl : wf Σ -> declared_inductive Σ ind mdecl idecl -> let u := abstract_instance (ind_universes mdecl) in - subst_instance_univ u (ind_sort idecl) = ind_sort idecl. + subst_instance_sort u (ind_sort idecl) = ind_sort idecl. Proof using Type. intros wfΣ decli u. pose proof (on_declared_inductive decli) as [onmind oib]. pose proof (onArity oib) as ona. rewrite (oib.(ind_arity_eq)) in ona. - red in ona. destruct ona. - eapply typed_subst_abstract_instance in t. + apply isType_subst_instance_id in ona. 2:split; simpl; auto. - - rewrite !subst_instance_it_mkProd_or_LetIn in t. - eapply (f_equal (destArity [])) in t. - rewrite !destArity_it_mkProd_or_LetIn in t. simpl in t. noconf t. + - rewrite !subst_instance_it_mkProd_or_LetIn in ona. + eapply (f_equal (destArity [])) in ona. + rewrite !destArity_it_mkProd_or_LetIn in ona. simpl in ona. noconf ona. simpl in H; noconf H. apply H0. - destruct decli as [declm _]. eapply declared_inductive_wf_global_ext in declm; auto. @@ -579,19 +583,9 @@ Qed. pose proof (on_declared_inductive decli) as [_ oib]. pose proof (onArity oib) as ona. rewrite (oib.(ind_arity_eq)) in ona |- *. - red in ona. destruct ona. - eapply typed_subst_abstract_instance in t; eauto. + apply isType_subst_instance_id in ona; eauto. destruct decli as [declm _]. eapply declared_inductive_wf_global_ext in declm; auto. Qed. - Lemma isType_subst_instance_id Σ Γ T : - wf_ext_wk Σ -> - let u := abstract_instance Σ.2 in - isType Σ Γ T -> subst_instance u T = T. - Proof using Type. - intros wf_ext u isT. - destruct isT. eapply typed_subst_abstract_instance in t; auto. - Qed. - End SubstIdentity. diff --git a/pcuic/theories/Typing/PCUICWeakeningConfigTyp.v b/pcuic/theories/Typing/PCUICWeakeningConfigTyp.v index 68d149d01..841911e2f 100644 --- a/pcuic/theories/Typing/PCUICWeakeningConfigTyp.v +++ b/pcuic/theories/Typing/PCUICWeakeningConfigTyp.v @@ -22,16 +22,13 @@ Local Ltac constructor_per_goal _ := Lemma weakening_config_wf_local_sized {cf1 cf2 : checker_flags} Σ Γ (Hwf : All_local_env (lift_typing (@typing cf1) Σ) Γ) (IH : forall Γ0 t0 T0 (H0 : @typing cf1 Σ Γ0 t0 T0), - typing_size H0 < S (All_local_env_size (fun _ _ _ _ H => typing_size H) Σ Γ Hwf) + typing_size H0 < S (All_local_env_size (fun _ _ _ H => typing_size H) Γ Hwf) -> @typing cf2 Σ Γ0 t0 T0) : wf_local Σ Γ. Proof. simpl in *. induction Hwf; [ constructor 1 | constructor 2 | constructor 3 ]. - all: try assumption. - all: unfold lift_typing, lift_judgment, infer_sort in *. - all: rdest. - all: simpl in *. + 2,4: eapply @lift_typing_size_impl with (Psize := @typing_size _ _ _) (tu := t0); unfold lift_sorting_size; cbn; intros. all: try (unshelve eapply IH; [ eassumption | ]; try solve [ constructor; simpl in *; cbn in *; lia ]). all: repeat (exactly_once (idtac; multimatch goal with H : _ |- _ => unshelve eapply H; [ try eassumption .. | intros; simpl in * ]; clear H end)). @@ -73,10 +70,14 @@ Proof. all: simpl in IH. all: try (set (k := fix_context _) in *; clearbody k). all: match goal with - | [ H : All (fun d => ∑ s : ?S, _) ?l |- All (fun d => ∑ s' : ?S, _) ?l ] - => is_var l; clear -H IH; induction H as [|???? IH']; constructor - | [ H : All (fun d => _ ;;; _ |- _ : _) ?l |- All (fun d => _ ;;; _ |- _ : _) ?l ] - => is_var l; clear -H IH; induction H as [|???? IH']; constructor + | [ H : lift_typing0 _ ?j |- lift_typing0 _ ?j ] + => eapply @lift_typing_size_impl with (Psize := @typing_size _ _ _) (tu := H); unfold lift_sorting_size, on_def_type_sorting_size in *; cbn; intros + | [ H : All (on_def_type _ _) ?l |- All (on_def_type _ _) ?l ] + => is_var l; clear -H IH; induction H as [|???? IH']; constructor; + [eapply @lift_typing_size_impl with (Psize := @typing_size _ _ _) (tu := p); unfold lift_sorting_size, on_def_type_sorting_size in *; cbn; intros|] + | [ H : All (on_def_body _ _ _) ?l |- All (on_def_body _ _ _) ?l ] + => is_var l; clear -H IH; induction H as [|???? IH']; constructor; + [eapply @lift_typing_size_impl with (Psize := @typing_size _ _ _) (tu := p); unfold lift_sorting_size, on_def_body_sorting_size in *; cbn; intros|] | [ H : case_side_conditions _ _ _ _ _ _ _ _ _ _ _ |- case_side_conditions _ _ _ _ _ _ _ _ _ _ _ ] => destruct H; constructor | [ H : case_branch_typing _ _ _ _ _ _ _ _ _ _ _ |- case_branch_typing _ _ _ _ _ _ _ _ _ _ _ ] => destruct H; constructor | _ => idtac @@ -93,7 +94,7 @@ Proof. all: try now repeat destruct ?; subst; simpl in *; assumption. all: repeat destruct ?; subst. all: lazymatch goal with - | [ H : ctx_inst _ _ _ _ _ |- ctx_inst _ _ _ _ _ ] + | [ H : ctx_inst _ _ _ _ |- ctx_inst _ _ _ _ ] => revert dependent H; repeat match goal with | [ |- context[typing_size ?x] ] @@ -120,9 +121,9 @@ Proof. all: cbn in *; try lia. all: try assumption. all: repeat match goal with - | [ |- context[All_local_env_size ?x ?y ?z ?w] ] + | [ |- context[All_local_env_size ?x ?y ?w] ] => let v := fresh in - set (v := All_local_env_size _ y z w) in * + set (v := All_local_env_size _ y w) in * end. all: try lia. Qed. @@ -132,7 +133,7 @@ Lemma weakening_config_wf_local {cf1 cf2 : checker_flags} Σ Γ : -> match cf1 with _ => wf_local Σ Γ end -> match cf2 with _ => wf_local Σ Γ end. Proof. - intros Hcf H; eapply All_local_env_impl; [ eassumption | ]. + intros Hcf H; eapply All_local_env_impl; [ eassumption | cbn ]. intros * H'; eapply lift_typing_impl; [ eassumption | ]. intros *; eapply (@weakening_config cf1 cf2); assumption. Qed. @@ -142,8 +143,8 @@ Lemma weakening_config_wf {cf1 cf2 : checker_flags} Σ : -> @wf cf1 Σ -> @wf cf2 Σ. Proof. - rewrite /wf/Forall_decls_typing. - intros; eapply (@on_global_env_impl_config cf1 cf2); try eassumption. + rewrite /wf. + intros; eapply (@on_global_env_impl_config cf1 cf2); try eassumption; cbn. { intros; eapply @lift_typing_impl; [ eassumption | ]. intros; eapply (@weakening_config cf1 cf2); eassumption. } { intros; eapply (@weakening_config_cumulSpec0 cf1 cf2); eassumption. } diff --git a/pcuic/theories/Typing/PCUICWeakeningEnvTyp.v b/pcuic/theories/Typing/PCUICWeakeningEnvTyp.v index af047afdd..f4cb79137 100644 --- a/pcuic/theories/Typing/PCUICWeakeningEnvTyp.v +++ b/pcuic/theories/Typing/PCUICWeakeningEnvTyp.v @@ -24,19 +24,19 @@ Ltac my_rename_hyp h th := Ltac rename_hyp h ht ::= my_rename_hyp h ht. Lemma extends_wf_local_gen `{cf : checker_flags} R Σ Γ (wfΓ : wf_local Σ Γ) : - All_local_env_over typing - (fun Σ0 Γ0 wfΓ (t T : term) ty => + All_local_env_over (typing Σ) + (fun Γ0 wfΓ (t T : term) ty => forall Σ' : global_env, wf Σ' -> - R Σ0 Σ' -> - (Σ', Σ0.2);;; Γ0 |- t : T - ) Σ Γ wfΓ -> + R Σ Σ' -> + (Σ', Σ.2);;; Γ0 |- t : T + ) Γ wfΓ -> forall Σ' : global_env, wf Σ' -> R Σ Σ' -> wf_local (Σ', Σ.2) Γ. Proof. intros X0 Σ' H0. induction X0 in H0 |- *; try econstructor; simpl in *; intuition auto. - - apply infer_typing_sort_impl with id tu. intro; eauto. - - apply infer_typing_sort_impl with id tu. intro; eauto. + - apply lift_sorting_it_impl with tu => //. intro; eauto. + - apply lift_sorting_it_impl with tu => /=. all: intro; eauto. Qed. Definition extends_wf_local `{cf : checker_flags} := extends_wf_local_gen extends. @@ -97,38 +97,37 @@ Local Hint Resolve extends_primitive_constant : extends. Lemma weakening_env `{checker_flags} : env_prop (fun Σ Γ t T => forall Σ', wf Σ -> wf Σ' -> extends Σ.1 Σ' -> (Σ', Σ.2) ;;; Γ |- t : T) + (fun Σ Γ j => + forall Σ', wf Σ -> wf Σ' -> extends Σ.1 Σ' -> lift_typing typing (Σ', Σ.2) Γ j) (fun Σ Γ => forall Σ', wf Σ -> wf Σ' -> extends Σ.1 Σ' -> wf_local (Σ', Σ.2) Γ). Proof. apply typing_ind_env; intros; rename_all_hyps; try solve [econstructor; eauto 2 with extends]. - - induction X; constructor; eauto 2 with extends. - + apply infer_typing_sort_impl with id tu; intro; auto. - + apply infer_typing_sort_impl with id tu; intro; auto. - + eapply Hc; eauto. + - apply lift_typing_impl with (1 := X) => t T [] _ IH //. + now apply IH. + + - apply All_local_env_impl with (1 := X0) => Γ' j Hj. + now apply Hj. - econstructor; eauto 2 with extends. - now apply extends_wf_universe. + now apply extends_wf_sort. - econstructor; eauto 2 with extends. all: econstructor; eauto 2 with extends. * revert X5. clear -Σ' wf0 wfΣ' extΣ. induction 1; constructor; try destruct t0; eauto with extends. * close_Forall. intros; intuition eauto with extends. - econstructor; eauto with extends. + specialize (forall_Σ' _ wf0 wfΣ' extΣ). - now apply wf_local_app_inv in forall_Σ'. + now apply All_local_env_app_inv in forall_Σ'. + eapply fix_guard_extends; eauto. - + eapply (All_impl X0); intros d X. - apply (lift_typing_impl X); now intros ? []. - + eapply (All_impl X1); intros d X. - apply (lift_typing_impl X); now intros ? []. + + eapply (All_impl X1); intros d X. now apply X. + + eapply (All_impl X3); intros d X. now apply X. - econstructor; eauto with extends. + specialize (forall_Σ' _ wf0 wfΣ' extΣ). - now apply wf_local_app_inv in forall_Σ'. + now apply All_local_env_app_inv in forall_Σ'. + eapply cofix_guard_extends; eauto. - + eapply (All_impl X0); intros d X. - apply (lift_typing_impl X); now intros ? []. - + eapply (All_impl X1); intros d X. - apply (lift_typing_impl X); now intros ? []. + + eapply (All_impl X1); intros d X. now apply X. + + eapply (All_impl X3); intros d X. now apply X. - econstructor; eauto with extends. destruct X1; constructor; eauto with extends. * now eapply extends_wf_universe. @@ -146,71 +145,18 @@ Lemma weakening_on_global_decl_gen `{checker_flags} R P Σ Σ' φ kn decl : on_global_decl cumulSpec0 P (Σ', φ) kn decl. Proof. unfold weaken_env_prop_gen. - intros HRext HPΣ wfΣ wfΣ' Hext Hdecl. - destruct decl. - 1:{ - destruct c. destruct cst_body0. - - simpl in *. - red in Hdecl |- *. simpl in *. - eapply (HPΣ Σ Σ'); eauto. - - eapply (HPΣ Σ Σ'); eauto. - } - simpl in *. - destruct Hdecl as [onI onP onnP]; constructor; eauto. - - eapply Alli_impl; eauto. intros. - destruct X. unshelve econstructor; eauto. - + unfold on_type in *; intuition eauto. - + unfold on_constructors in *. eapply All2_impl; eauto. - intros. - destruct X as [? ? ? ?]. unshelve econstructor; eauto. - * unfold on_type in *; eauto. - * clear on_cindices cstr_eq cstr_args_length. - revert on_cargs. - induction (cstr_args x0) in y |- *; destruct y; simpl in *; eauto. - ** destruct a as [na [b|] ty]; simpl in *; intuition eauto. - ** destruct a as [na [b|] ty]; simpl in *; intuition eauto. - * clear on_ctype on_cargs. - revert on_cindices. - generalize (List.rev (lift_context #|cstr_args x0| 0 (ind_indices x))). - generalize (cstr_indices x0). - induction 1; constructor; eauto. - * simpl. - intros v indv. specialize (on_ctype_variance v indv). - simpl in *. move: on_ctype_variance. - unfold cstr_respects_variance. destruct variance_universes as [[[univs u] u']|]; auto. - intros [args idxs]. split. - ** eapply (All2_fold_impl args); intros. - inversion X; constructor; auto. - ++ eapply weakening_env_cumulSpec; eauto. - ++ eapply weakening_env_convSpec; eauto. - ++ eapply weakening_env_cumulSpec; eauto. - ** eapply (All2_impl idxs); intros. - eapply weakening_env_convSpec; eauto. - + unfold check_ind_sorts in *. - destruct Universe.is_prop; auto. - destruct Universe.is_sprop; auto. - split; [apply fst in ind_sorts|apply snd in ind_sorts]. - * eapply Forall_impl; tea; cbn. - intros. eapply Forall_impl; tea; simpl; intros. - eapply leq_universe_subset; tea. - apply weakening_env_global_ext_constraints; tea; eauto. - * destruct indices_matter; [|trivial]. clear -ind_sorts HPΣ wfΣ wfΣ' HRext Hext. - induction ind_indices; simpl in *; auto. - -- eapply (extends_wf_universe (Σ:=(Σ,φ)) Σ'); auto. - -- destruct a as [na [b|] ty]; simpl in *; intuition eauto. - + destruct ind_variance => //. - move: onIndices. unfold ind_respects_variance. - destruct variance_universes as [[[univs u] u']|] => //. - intros idx; eapply (All2_fold_impl idx); simpl. - intros par par' t t' d. - inv d; constructor; auto. - ++ eapply weakening_env_cumulSpec; eauto. - ++ eapply weakening_env_convSpec; eauto. - ++ eapply weakening_env_cumulSpec; eauto. - - red in onP |- *. eapply All_local_env_impl; eauto. - - move: onVariance. - rewrite /on_variance. destruct ind_universes => //. - destruct ind_variance => //. + intros HRext HPΣ wfΣ wfΣ' Hext. + apply on_global_decl_impl_full; eauto. + - reflexivity. + - intros ?????. eapply weakening_env_cumulSpec0; eauto. + - intro. eapply (extends_wf_sort (Σ:=(Σ,φ)) Σ'); auto. + - unfold check_constructors_smaller. intros ???. + eapply Forall_impl; tea; cbn. intros. + eapply Forall_impl; tea; simpl; intros. + eapply leq_sort_subset; tea. + apply weakening_env_global_ext_constraints; tea; eauto. + - rewrite /on_variance. intros u l. destruct u => //. + destruct l => //. intros [univs' [i [i' []]]]. exists univs', i, i'. split => //. all:eapply weakening_env_consistent_instance; tea; eauto. Qed. @@ -303,34 +249,32 @@ Proof. Qed. Lemma declared_constant_inv `{checker_flags} Σ P cst decl : - weaken_env_strictly_decls_prop cumulSpec0 (lift_typing typing) (lift_typing P) -> - wf Σ -> Forall_decls_typing P Σ -> + weaken_env_strictly_decls_prop cumulSpec0 (lift_typing typing) P -> + wf Σ -> on_global_env cumulSpec0 P Σ -> declared_constant Σ cst decl -> - on_constant_decl (lift_typing P) (Σ, cst_universes decl) decl. + on_constant_decl P (Σ, cst_universes decl) decl. Proof. intros. - eapply declared_constant_to_gen in H0. + unshelve eapply declared_constant_to_gen in H0; tea. eapply weaken_lookup_on_global_env in X1; eauto. apply X1. - Unshelve. all:eauto. Qed. Lemma declared_minductive_inv `{checker_flags} {Σ P ind mdecl} : - weaken_env_strictly_decls_prop cumulSpec0 (lift_typing typing) (lift_typing P) -> - wf Σ -> Forall_decls_typing P Σ -> + weaken_env_strictly_decls_prop cumulSpec0 (lift_typing typing) P -> + wf Σ -> on_global_env cumulSpec0 P Σ -> declared_minductive Σ ind mdecl -> - on_inductive cumulSpec0 (lift_typing P) (Σ, ind_universes mdecl) ind mdecl. + on_inductive cumulSpec0 P (Σ, ind_universes mdecl) ind mdecl. Proof. intros. - eapply declared_minductive_to_gen in H0. + unshelve eapply declared_minductive_to_gen in H0; tea. eapply weaken_lookup_on_global_env in X1; eauto. apply X1. - Unshelve. all:eauto. Qed. Lemma declared_inductive_inv `{checker_flags} {Σ P ind mdecl idecl} : - weaken_env_strictly_decls_prop cumulSpec0 (lift_typing typing) (lift_typing P) -> - wf Σ -> Forall_decls_typing P Σ -> + weaken_env_strictly_decls_prop cumulSpec0 (lift_typing typing) P -> + wf Σ -> on_global_env cumulSpec0 P Σ -> declared_inductive Σ ind mdecl idecl -> - on_ind_body cumulSpec0 (lift_typing P) (Σ, ind_universes mdecl) (inductive_mind ind) mdecl (inductive_ind ind) idecl. + on_ind_body cumulSpec0 P (Σ, ind_universes mdecl) (inductive_mind ind) mdecl (inductive_ind ind) idecl. Proof. intros. destruct H0 as [Hmdecl Hidecl]. @@ -352,14 +296,14 @@ Proof. Qed. Lemma declared_constructor_inv `{checker_flags} {Σ P mdecl idecl ref cdecl} - (HP : weaken_env_strictly_decls_prop cumulSpec0 (lift_typing typing) (lift_typing P)) + (HP : weaken_env_strictly_decls_prop cumulSpec0 (lift_typing typing) P) (wfΣ : wf Σ) - (HΣ : Forall_decls_typing P Σ) + (HΣ : on_global_env cumulSpec0 P Σ) (Hdecl : declared_constructor Σ ref mdecl idecl cdecl) : ∑ cs, let onib := declared_inductive_inv HP wfΣ HΣ (let (x, _) := Hdecl in x) in nth_error onib.(ind_cunivs) ref.2 = Some cs - × on_constructor cumulSpec0 (lift_typing P) (Σ, ind_universes mdecl) mdecl + × on_constructor cumulSpec0 P (Σ, ind_universes mdecl) mdecl (inductive_ind ref.1) idecl idecl.(ind_indices) cdecl cs. Proof. intros. @@ -370,22 +314,21 @@ Proof. Defined. Lemma declared_minductive_inv_decls `{checker_flags} {Σ P ind mdecl} : - weaken_env_decls_prop cumulSpec0 (lift_typing typing) (lift_typing P) -> - wf Σ -> Forall_decls_typing P Σ -> + weaken_env_decls_prop cumulSpec0 (lift_typing typing) P -> + wf Σ -> on_global_env cumulSpec0 P Σ -> declared_minductive Σ ind mdecl -> - on_inductive cumulSpec0 (lift_typing P) (Σ, ind_universes mdecl) ind mdecl. + on_inductive cumulSpec0 P (Σ, ind_universes mdecl) ind mdecl. Proof. intros. - eapply declared_minductive_to_gen in H0. + unshelve eapply declared_minductive_to_gen in H0; tea. eapply weaken_decls_lookup_on_global_env in X1; eauto. apply X1. - Unshelve. all:eauto. Qed. Lemma declared_inductive_inv_decls `{checker_flags} {Σ P ind mdecl idecl} : - weaken_env_decls_prop cumulSpec0 (lift_typing typing) (lift_typing P) -> - wf Σ -> Forall_decls_typing P Σ -> + weaken_env_decls_prop cumulSpec0 (lift_typing typing) P -> + wf Σ -> on_global_env cumulSpec0 P Σ -> declared_inductive Σ ind mdecl idecl -> - on_ind_body cumulSpec0 (lift_typing P) (Σ, ind_universes mdecl) (inductive_mind ind) mdecl (inductive_ind ind) idecl. + on_ind_body cumulSpec0 P (Σ, ind_universes mdecl) (inductive_mind ind) mdecl (inductive_ind ind) idecl. Proof. intros. destruct H0 as [Hmdecl Hidecl]. @@ -396,14 +339,14 @@ Proof. Qed. Lemma declared_constructor_inv_decls `{checker_flags} {Σ P mdecl idecl ref cdecl} - (HP : weaken_env_decls_prop cumulSpec0 (lift_typing typing) (lift_typing P)) + (HP : weaken_env_decls_prop cumulSpec0 (lift_typing typing) P) (wfΣ : wf Σ) - (HΣ : Forall_decls_typing P Σ) + (HΣ : on_global_env cumulSpec0 P Σ) (Hdecl : declared_constructor Σ ref mdecl idecl cdecl) : ∑ cs, let onib := declared_inductive_inv_decls HP wfΣ HΣ (let (x, _) := Hdecl in x) in nth_error onib.(ind_cunivs) ref.2 = Some cs - × on_constructor cumulSpec0 (lift_typing P) (Σ, ind_universes mdecl) mdecl + × on_constructor cumulSpec0 P (Σ, ind_universes mdecl) mdecl (inductive_ind ref.1) idecl idecl.(ind_indices) cdecl cs. Proof. intros. @@ -414,16 +357,16 @@ Proof. Defined. Lemma declared_projection_inv `{checker_flags} {Σ P mdecl idecl cdecl ref pdecl} : - forall (HP : weaken_env_strictly_decls_prop cumulSpec0 (lift_typing typing) (lift_typing P)) + forall (HP : weaken_env_strictly_decls_prop cumulSpec0 (lift_typing typing) P) (wfΣ : wf Σ) - (HΣ : Forall_decls_typing P Σ) + (HΣ : on_global_env cumulSpec0 P Σ) (Hdecl : declared_projection Σ ref mdecl idecl cdecl pdecl), match idecl.(ind_ctors) return Type with | [c] => let oib := declared_inductive_inv HP wfΣ HΣ (let (x, _) := Hdecl in let (x, _) := x in x) in (match oib.(ind_cunivs) with - | [cs] => sorts_local_ctx (lift_typing P) (Σ, ind_universes mdecl) (arities_context (ind_bodies mdecl) ,,, ind_params mdecl) (cstr_args c) cs + | [cs] => sorts_local_ctx P (Σ, ind_universes mdecl) (arities_context (ind_bodies mdecl) ,,, ind_params mdecl) (cstr_args c) cs | _ => False end) * on_projections mdecl (inductive_mind ref.(proj_ind)) (inductive_ind ref.(proj_ind)) idecl (idecl.(ind_indices)) c * @@ -454,8 +397,8 @@ Qed. Lemma weaken_env_prop_typing `{checker_flags} : weaken_env_prop cumulSpec0 (lift_typing typing) (lift_typing typing). Proof. - do 2 red. intros * wfΣ wfΣ' Hext Γ t T HT. - apply lift_typing_impl with (1 := HT); intros ? Hty. + do 2 red. intros * wfΣ wfΣ' Hext Γ j Hj. + apply lift_typing_impl with (1 := Hj); intros ?? HT. eapply (weakening_env (_, _)). 2: eauto. all: auto. diff --git a/pcuic/theories/Typing/PCUICWeakeningTyp.v b/pcuic/theories/Typing/PCUICWeakeningTyp.v index 56202a081..71b644118 100644 --- a/pcuic/theories/Typing/PCUICWeakeningTyp.v +++ b/pcuic/theories/Typing/PCUICWeakeningTyp.v @@ -27,28 +27,21 @@ Lemma weakening_wf_local {cf: checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ wf_local Σ (Γ ,,, Γ'' ,,, lift_context #|Γ''| 0 Γ'). Proof. intros wfΓ' wfΓ''. - pose proof (env_prop_wf_local typing_rename_prop _ wfΣ _ wfΓ') as [_ X]. simpl in X. + pose proof (env_prop_wf_local typing_rename_prop _ wfΣ _ wfΓ') as X. simpl in X. eapply All_local_env_app_inv in X as [XΓ XΓ']. - apply wf_local_app => //. + apply All_local_env_app => //. rewrite /lift_context. apply All_local_env_fold. eapply (All_local_env_impl_ind XΓ'). - intros Δ t [T|] IH; simpl. - - intros Hf. rewrite -/(lift_context #|Γ''| 0 Δ). - rewrite Nat.add_0_r. rewrite !lift_rename. - eapply (Hf xpredT). - split. - + apply wf_local_app; auto. - apply (All_local_env_fold (fun Δ => lift_typing typing Σ (Γ ,,, Γ'' ,,, Δ))) in IH. apply IH. - + apply weakening_renaming. - - intros Hty. simple apply (infer_typing_sort_impl (P := fun Σ Γ T s => forall P Δ f, renaming _ Σ Γ Δ _ -> Σ;;; Δ |- rename f T : rename f s)) with id Hty; intros Hs. - rewrite -/(lift_context #|Γ''| 0 Δ). - rewrite Nat.add_0_r !lift_rename. - eapply (Hs xpredT). - split. - + apply wf_local_app; auto. - apply (All_local_env_fold (fun Δ => lift_typing typing Σ (Γ ,,, Γ'' ,,, Δ))) in IH. apply IH. - + apply weakening_renaming. + intros Δ j IH H; simpl. + eapply lift_typing_f_impl => //. + 2: intros ?? Hf; rewrite -/(lift_context #|Γ''| 0 Δ). + 2: rewrite Nat.add_0_r; rewrite !lift_rename; apply Hf. + eapply (H xpredT). + split. + + apply All_local_env_app; auto. + apply All_local_env_fold, IH. + + apply weakening_renaming. Qed. Lemma weakening_wf_local_eq {cf: checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} {Γ Γ' Γ'' n} : @@ -184,20 +177,20 @@ Proof. generalize (@nil context_decl) as Δ. rewrite /fix_context_gen. intros Δ wfΔ. - eapply All_local_env_app. split; auto. + eapply All_local_env_app; auto. induction a in Δ, wfΔ |- *; simpl; auto. + constructor. + simpl. - eapply All_local_env_app. split; auto. - * repeat constructor. - apply infer_typing_sort_impl with id p; intros Hs. - eapply (weakening Σ Γ Δ _ (tSort _)); auto. + eapply All_local_env_app; auto. + * constructor. 1: constructor. + apply lift_typing_f_impl with (1 := p) => // ?? Hs. + eapply (weakening Σ Γ Δ); auto. * specialize (IHa (Δ ,,, [vass (dname x) (lift0 #|Δ| (dtype x))])). rewrite app_length in IHa. simpl in IHa. forward IHa. ** simpl; constructor; auto. - apply infer_typing_sort_impl with id p; intros Hs. - eapply (weakening Σ Γ Δ _ (tSort _)); auto. + apply lift_typing_f_impl with (1 := p) => // ?? Hs. + eapply (weakening Σ Γ Δ); auto. ** eapply All_local_env_impl; eauto. simpl; intros. rewrite app_context_assoc. apply X. @@ -212,8 +205,8 @@ Proof. intros wfΣ wfΓ wfty. rewrite <- (firstn_skipn n Γ) in wfΓ |- *. assert (n = #|firstn n Γ|). { rewrite firstn_length_le; auto with arith. } - apply infer_typing_sort_impl with id wfty; intros Hs. - rewrite {3}H. - eapply (weakening_typing (Γ := skipn n Γ) (Γ' := []) (Γ'' := firstn n Γ) (T := tSort _)); + apply lift_typing_f_impl with (1 := wfty) => // ?? Hs. + rewrite {3 4}H. + eapply (weakening_typing (Γ := skipn n Γ) (Γ' := []) (Γ'' := firstn n Γ)); eauto with wf. Qed. diff --git a/pcuic/theories/utils/PCUICOnOne.v b/pcuic/theories/utils/PCUICOnOne.v index e77d79454..6c6ac1c9d 100644 --- a/pcuic/theories/utils/PCUICOnOne.v +++ b/pcuic/theories/utils/PCUICOnOne.v @@ -77,32 +77,34 @@ Definition map_decl_na (f : aname -> aname) (g : term -> term) d := (** We do not allow alpha-conversion and P applies to only one of the fields in the context declaration. Used to define one-step context reduction. *) -Definition on_one_decl (P : context -> term -> term -> Type) - Γ (d : context_decl) (d' : context_decl) : Type := +Definition on_one_decl (P : term -> term -> Type) + (d : context_decl) (d' : context_decl) : Type := match d, d' with | {| decl_name := na; decl_body := None; decl_type := ty |}, {| decl_name := na'; decl_body := None; decl_type := ty' |} => - na = na' × P Γ ty ty' + na = na' × P ty ty' | {| decl_name := na; decl_body := Some b; decl_type := ty |}, {| decl_name := na'; decl_body := Some b'; decl_type := ty' |} => na = na' × - ((P Γ ty ty' × b = b') + - (P Γ b b' × ty = ty')) + ((P ty ty' × b = b') + + (P b b' × ty = ty')) | _, _ => False end. -Lemma on_one_decl_impl (P Q : context -> term -> term -> Type) : - (forall Γ, inclusion (P Γ) (Q Γ)) -> - forall Γ, inclusion (on_one_decl P Γ) (on_one_decl Q Γ). +Notation on_one_decl1 P Γ := (on_one_decl (P Γ)). + +Lemma on_one_decl_impl (P Q : term -> term -> Type) : + (inclusion P Q) -> + inclusion (on_one_decl P) (on_one_decl Q). Proof. - intros HP Γ x y. + intros HP x y. destruct x as [na [b|] ty], y as [na' [b'|] ty']; simpl; firstorder auto. Qed. Lemma on_one_decl_map_na (P : context -> term -> term -> Type) f g : forall Γ, - inclusion (on_one_decl (fun Γ => on_Trel (P (map (map_decl_na f g) Γ)) g) Γ) - (on_Trel (on_one_decl P (map (map_decl_na f g) Γ)) (map_decl_na f g)). + inclusion (on_one_decl (on_Trel (P (map (map_decl_na f g) Γ)) g)) + (on_Trel (on_one_decl (P (map (map_decl_na f g) Γ))) (map_decl_na f g)). Proof. intros Γ x y. destruct x as [na [b|] ty], y as [na' [b'|] ty']; simpl in *; firstorder auto; subst; simpl; @@ -111,8 +113,8 @@ Qed. Lemma on_one_decl_map (P : context -> term -> term -> Type) f : forall Γ, - inclusion (on_one_decl (fun Γ => on_Trel (P (map (map_decl f) Γ)) f) Γ) - (on_Trel (on_one_decl P (map (map_decl f) Γ)) (map_decl f)). + inclusion (on_one_decl (on_Trel (P (map (map_decl f) Γ)) f)) + (on_Trel (on_one_decl (P (map (map_decl f) Γ))) (map_decl f)). Proof. intros Γ x y. destruct x as [na [b|] ty], y as [na' [b'|] ty']; simpl in *; firstorder auto; subst; simpl; @@ -121,22 +123,22 @@ Qed. Lemma on_one_decl_mapi_context (P : context -> term -> term -> Type) f : forall Γ, - inclusion (on_one_decl (fun Γ => on_Trel (P (mapi_context f Γ)) (f #|Γ|)) Γ) - (on_Trel (on_one_decl P (mapi_context f Γ)) (map_decl (f #|Γ|))). + inclusion (on_one_decl (on_Trel (P (mapi_context f Γ)) (f #|Γ|))) + (on_Trel (on_one_decl (P (mapi_context f Γ))) (map_decl (f #|Γ|))). Proof. intros Γ x y. destruct x as [na [b|] ty], y as [na' [b'|] ty']; simpl in *; firstorder auto; subst; simpl; auto. Qed. -Lemma on_one_decl_test_impl (P Q : context -> term -> term -> Type) (p : term -> bool) : - forall Γ d d', - on_one_decl P Γ d d' -> +Lemma on_one_decl_test_impl (P Q : term -> term -> Type) (p : term -> bool) : + forall d d', + on_one_decl P d d' -> test_decl p d -> - (forall x y, p x -> P Γ x y -> Q Γ x y) -> - on_one_decl Q Γ d d'. + (forall x y, p x -> P x y -> Q x y) -> + on_one_decl Q d d'. Proof. - intros Γ [na [b|] ty] [na' [b'|] ty'] ond []%andb_and; simpl; firstorder auto. + intros [na [b|] ty] [na' [b'|] ty'] ond []%andb_and; simpl; firstorder auto. Qed. Section OnOne_local_2. @@ -172,9 +174,9 @@ Qed. Lemma OnOne2_local_env_ondecl_impl P Q : (forall Γ, inclusion (P Γ) (Q Γ)) -> - inclusion (OnOne2_local_env (on_one_decl P)) (OnOne2_local_env (on_one_decl P)). + inclusion (OnOne2_local_env (fun Γ => on_one_decl1 P Γ)) (OnOne2_local_env (fun Γ => on_one_decl1 Q Γ)). Proof. - intros HP. now apply OnOne2_local_env_impl, on_one_decl_impl. + intros HP. apply OnOne2_local_env_impl => Γ. apply on_one_decl_impl. easy. Qed. Lemma OnOne2_local_env_map R Γ Δ (f : aname -> aname) (g : term -> term) : @@ -236,11 +238,11 @@ Proof. now rewrite -(length_of onenv). Qed. -Lemma on_one_decl_test_decl (P : context -> term -> term -> Type) Γ +Lemma on_one_decl_test_decl (P : term -> term -> Type) (p q : term -> bool) d d' : (forall t, p t -> q t) -> - (forall t t', P Γ t t' -> p t -> q t') -> - on_one_decl P Γ d d' -> + (forall t t', P t t' -> p t -> q t') -> + on_one_decl P d d' -> test_decl p d -> test_decl q d'. Proof. diff --git a/pcuic/theories/utils/PCUICPrimitive.v b/pcuic/theories/utils/PCUICPrimitive.v index b9bcae6fa..9807da587 100644 --- a/pcuic/theories/utils/PCUICPrimitive.v +++ b/pcuic/theories/utils/PCUICPrimitive.v @@ -173,7 +173,7 @@ Inductive onPrims {term} (eq_term : term -> term -> Type) Re : prim_val term -> | onPrimsInt i : onPrims eq_term Re (primInt; primIntModel i) (primInt; primIntModel i) | onPrimsFloat f : onPrims eq_term Re (primFloat; primFloatModel f) (primFloat; primFloatModel f) | onPrimsArray a a' : - Re (Universe.make a.(array_level)) (Universe.make a'.(array_level)) -> + Re (Universe.make' a.(array_level)) (Universe.make' a'.(array_level)) -> eq_term a.(array_default) a'.(array_default) -> eq_term a.(array_type) a'.(array_type) -> All2 eq_term a.(array_value) a'.(array_value) -> @@ -192,7 +192,7 @@ Inductive onPrims_dep {term} (eq_term : term -> term -> Type) (Re : Universe.t - | onPrimsInt_dep i : onPrims_dep eq_term Re eq_term_dep Re' (primInt; primIntModel i) (primInt; primIntModel i) (onPrimsInt eq_term Re i) | onPrimsFloat_dep f : onPrims_dep eq_term Re eq_term_dep Re' (primFloat; primFloatModel f) (primFloat; primFloatModel f) (onPrimsFloat _ _ f) | onPrimsArray_dep a a' : - forall (hre : Re (Universe.make a.(array_level)) (Universe.make a'.(array_level))) + forall (hre : Re (Universe.make' a.(array_level)) (Universe.make' a'.(array_level))) (eqdef : eq_term a.(array_default) a'.(array_default)) (eqty : eq_term a.(array_type) a'.(array_type)) (eqt : All2 eq_term a.(array_value) a'.(array_value)), diff --git a/quotation/theories/CommonUtils.v b/quotation/theories/CommonUtils.v index 27f703c07..1364ef661 100644 --- a/quotation/theories/CommonUtils.v +++ b/quotation/theories/CommonUtils.v @@ -178,7 +178,7 @@ Module WithTemplate. end. Section with_monad. - Context {M} {M_monad : Monad M} (in_domain : bool) (U : Universe.t -> M term). + Context {M} {M_monad : Monad M} (in_domain : bool) (U : sort -> M term). #[local] Fixpoint tmRelaxSortsM (t : term) {struct t} : M term @@ -235,20 +235,20 @@ Module WithTemplate. end. End with_monad. - #[local] Definition is_set (s : Universe.t) : bool - := match option_map Level.is_set (Universe.get_is_level s) with + #[local] Definition is_set (s : sort) : bool + := match option_map Level.is_set (Sort.get_is_level s) with | Some true => true | _ => false end. - #[local] Definition is_type (s : Universe.t) : bool - := match Universe.get_is_level s with + #[local] Definition is_type (s : sort) : bool + := match Sort.get_is_level s with | Some _ => true | _ => false end. - #[local] Definition is_only_type (s : Universe.t) : bool - := match option_map Level.is_set (Universe.get_is_level s) with + #[local] Definition is_only_type (s : sort) : bool + := match option_map Level.is_set (Sort.get_is_level s) with | Some false => true | _ => false end. @@ -256,7 +256,7 @@ Module WithTemplate. Definition tmRelaxSet (in_domain : bool) (prefix : string) (t : term) : term := tmRelaxSortsM (M:=fun T => T) in_domain - (fun u => tSort (if is_set u then Universe.of_levels (inr (Level.level (prefix ++ "._Set.0")%bs)) else u)) + (fun u => tSort (if is_set u then Sort.of_levels (inr (Level.level (prefix ++ "._Set.0")%bs)) else u)) t. Module Import PrefixUniverse. @@ -302,21 +302,21 @@ Module WithTemplate. ; t_ne := eq_trans LevelExprSet.is_empty_prefix_with l.(t_ne) |}. End nonEmptyLevelExprSet. - Module LevelAlgExpr := nonEmptyLevelExprSet. + Module Universe := nonEmptyLevelExprSet. - Module Universe. - Definition prefix_with (prefix : string) (u : Universe.t) : Universe.t + Module Sort. + Definition prefix_with (prefix : string) (u : sort) : sort := match u with - | Universe.lProp | Universe.lSProp => u - | Universe.lType u => Universe.lType (LevelAlgExpr.prefix_with prefix u) + | sProp | sSProp => u + | sType u => sType (Universe.prefix_with prefix u) end. - End Universe. + End Sort. End PrefixUniverse. Definition tmRelaxOnlyType (in_domain : bool) (prefix : string) (t : term) : term := tmRelaxSortsM (M:=fun T => T) in_domain - (fun u => tSort (PrefixUniverse.Universe.prefix_with prefix u)) + (fun u => tSort (PrefixUniverse.Sort.prefix_with prefix u)) t. Polymorphic Definition tmRetypeMagicRelaxSetInCodomain@{a b t u} (prefix : string) {A : Type@{a}} (B : Type@{b}) (x : A) : TemplateMonad@{t u} B diff --git a/quotation/theories/ToPCUIC/Common/BasicAst.v b/quotation/theories/ToPCUIC/Common/BasicAst.v index 574bba57f..f0bc398f9 100644 --- a/quotation/theories/ToPCUIC/Common/BasicAst.v +++ b/quotation/theories/ToPCUIC/Common/BasicAst.v @@ -16,7 +16,7 @@ From MetaCoq.Template Require Import AstUtils (* for tFixType *). #[export] Hint Unfold aname : quotation. #[export] Typeclasses Transparent aname. #[export] Instance quote_def {term} {qterm : quotation_of term} {quote_term : ground_quotable term} : ground_quotable (def term) := ltac:(destruct 1; exact _). -#[export] Instance quote_typ_or_sort_ {term} {qterm : quotation_of term} {quote_term : ground_quotable term} : ground_quotable (typ_or_sort_ term) := ltac:(destruct 1; exact _). +#[export] Instance quote_judgment_ {term univ} {qterm : quotation_of term} {uterm : quotation_of univ} {quote_term : ground_quotable term} {quote_univ : ground_quotable univ} : ground_quotable (judgment_ term univ) := ltac:(destruct 1; exact _). #[export] Instance quote_context_decl {term} {qterm : quotation_of term} {quote_term : ground_quotable term} : ground_quotable (context_decl term) := ltac:(destruct 1; exact _). #[export] Hint Unfold mfixpoint : quotation. #[export] Typeclasses Transparent mfixpoint. diff --git a/quotation/theories/ToPCUIC/Common/Environment.v b/quotation/theories/ToPCUIC/Common/Environment.v index 52dda940c..5381daf3e 100644 --- a/quotation/theories/ToPCUIC/Common/Environment.v +++ b/quotation/theories/ToPCUIC/Common/Environment.v @@ -81,13 +81,13 @@ Module QuoteEnvironment (T : Term) (Import E : EnvironmentSig T) (Import QEH : Q context global_declarations global_env_ext - typ_or_sort + judgment : quotation. #[export] Typeclasses Transparent context global_declarations global_env_ext - typ_or_sort + judgment . Import PolymorphicInstances. diff --git a/quotation/theories/ToPCUIC/Common/EnvironmentTyping.v b/quotation/theories/ToPCUIC/Common/EnvironmentTyping.v index 3627d5476..e138b1418 100644 --- a/quotation/theories/ToPCUIC/Common/EnvironmentTyping.v +++ b/quotation/theories/ToPCUIC/Common/EnvironmentTyping.v @@ -21,8 +21,10 @@ Module QuoteLookup (Import T : Term) (Import E : EnvironmentSig T) (Import L : L End with_refl. #[export] Instance quote_consistent_instance {cf lvs ϕ uctx u} : ground_quotable (@consistent_instance cf lvs ϕ uctx u) := ltac:(cbv [consistent_instance]; exact _). - #[export] Instance quote_wf_universe {Σ s} : ground_quotable (@wf_universe Σ s) - := ground_quotable_of_dec (@wf_universe_dec Σ s). + #[export] Instance quote_wf_universe {Σ u} : ground_quotable (@wf_universe Σ u) + := ground_quotable_of_dec (@wf_universe_dec Σ u). + #[export] Instance quote_wf_sort {Σ s} : ground_quotable (@wf_sort Σ s) + := ground_quotable_of_dec (@wf_sort_dec Σ s). #[export] Instance quote_declared_constant {Σ id decl} : ground_quotable (@declared_constant Σ id decl) := ltac:(cbv [declared_constant]; exact _). #[export] Instance quote_declared_minductive {Σ mind decl} : ground_quotable (@declared_minductive Σ mind decl) := ltac:(cbv [declared_minductive]; exact _). @@ -34,42 +36,36 @@ End QuoteLookup. Module QuoteEnvTyping (Import T : Term) (Import E : EnvironmentSig T) (Import TU : TermUtils T E) (Import ET : EnvTypingSig T E TU) (Import qT : QuotationOfTerm T) (Import qE : QuotationOfEnvironment T E) (Import qET : QuotationOfEnvTyping T E TU ET) (Import QuoteT : QuoteTerm T) (Import QuoteE : QuoteEnvironmentSig T E) <: QuoteEnvTypingSig T E TU ET. #[export] Hint Unfold - infer_sort - lift_typing + on_def_type + on_def_body + lift_typing0 : quotation. #[export] Typeclasses Transparent - infer_sort - lift_typing + on_def_type + on_def_body + lift_typing0 . - #[export] Instance quote_All_local_env {typing Γ} {qtyping : quotation_of typing} {quote_typing : forall Γ t T, ground_quotable (typing Γ t T)} : ground_quotable (@All_local_env typing Γ) := ltac:(induction 1; exact _). + #[export] Instance quote_All_local_env {typing Γ} {qtyping : quotation_of typing} {quote_typing : forall Γ j, ground_quotable (typing Γ j)} : ground_quotable (@All_local_env typing Γ) := ltac:(induction 1; exact _). Import StrongerInstances. #[local] Hint Extern 2 (_ = _) => reflexivity : typeclass_instances. - #[export] Instance quote_on_local_decl {P Γ d} {quoteP1 : forall b, d.(decl_body) = Some b -> ground_quotable (P Γ b (Typ d.(decl_type)))} {quoteP2 : d.(decl_body) = None -> ground_quotable (P Γ d.(decl_type) Sort)} : ground_quotable (@on_local_decl P Γ d) := ltac:(cbv [on_local_decl]; exact _). - #[export] Instance quote_lift_judgment {check infer_sort Σ Γ t T} {quote_check : forall T', T = Typ T' -> ground_quotable (check Σ Γ t T')} {quote_infer_sort : T = Sort -> ground_quotable (infer_sort Σ Γ t)} : ground_quotable (@lift_judgment check infer_sort Σ Γ t T) := ltac:(cbv [lift_judgment]; exact _). - #[local] Typeclasses Transparent lift_judgment. - #[export] Instance quote_All_local_env_over_gen - {checking sorting cproperty sproperty Σ Γ H} + #[export] Instance quote_lift_sorting {checking sorting j} {qcheck : quotation_of checking} {qsorting : quotation_of sorting} {quote_check : forall tm, j_term j = Some tm -> ground_quotable (checking tm (j_typ j))} {quote_sorting : forall u, ground_quotable (sorting (j_typ j) u)} : ground_quotable (@lift_sorting checking sorting j) := ltac:(cbv [lift_sorting]; exact _). + #[local] Typeclasses Transparent lift_sorting. + #[export] Instance quote_All_local_env_over_sorting + {checking sorting cproperty sproperty Γ H} {qchecking : quotation_of checking} {qsorting : quotation_of sorting} {qcproperty : quotation_of cproperty} {qsproperty : quotation_of sproperty} - {quote_checking : forall Γ t T, ground_quotable (checking Σ Γ t T)} {quote_sorting : forall Γ T, ground_quotable (sorting Σ Γ T)} {quote_sproperty : forall Γ all t tu, ground_quotable (sproperty Σ Γ all t tu)} {quote_cproperty : forall Γ all b t tb, ground_quotable (cproperty Σ Γ all b t tb)} - : ground_quotable (@All_local_env_over_gen checking sorting cproperty sproperty Σ Γ H) - := ltac:(induction 1; exact _). - #[export] Instance quote_All_local_env_over {typing property Σ Γ H} + {quote_checking : forall Γ t T, ground_quotable (checking Γ t T)} {quote_sorting : forall Γ T u, ground_quotable (sorting Γ T u)} {quote_sproperty : forall Γ all t u tu, ground_quotable (sproperty Γ all t u tu)} {quote_cproperty : forall Γ all b t tb, ground_quotable (cproperty Γ all b t tb)} + : ground_quotable (@All_local_env_over_sorting checking sorting cproperty sproperty Γ H) := ltac:(induction 1; cbv [lift_sorting j_term MCOption.option_default] in *; exact _). + #[export] Instance quote_All_local_env_over {typing property Γ H} {qtyping : quotation_of typing} {qproperty : quotation_of property} - {quote_typing : forall Γ t T, ground_quotable (typing Σ Γ t T)} {quote_property : forall Γ all b t tb, ground_quotable (property Σ Γ all b t tb)} - : ground_quotable (@All_local_env_over typing property Σ Γ H) + {quote_typing : forall Γ t T, ground_quotable (typing Γ t T)} {quote_property : forall Γ all b t tb, ground_quotable (property Γ all b t tb)} + : ground_quotable (@All_local_env_over typing property Γ H) := ltac:(cbv [All_local_env_over]; exact _). - #[export] Instance quote_All_local_env_over_sorting - {checking sorting cproperty sproperty Σ Γ H} - {qchecking : quotation_of checking} {qsorting : quotation_of sorting} {qcproperty : quotation_of cproperty} {qsproperty : quotation_of sproperty} - {quote_checking : forall Γ t T, ground_quotable (checking Σ Γ t T)} {quote_sorting : forall Γ T U, ground_quotable (sorting Σ Γ T U)} {quote_sproperty : forall Γ all t tu U, ground_quotable (sproperty Σ Γ all t tu U)} {quote_cproperty : forall Γ all b t tb, ground_quotable (cproperty Σ Γ all b t tb)} - : ground_quotable (@All_local_env_over_sorting checking sorting cproperty sproperty Σ Γ H) - := ltac:(cbv [All_local_env_over_sorting]; exact _). - #[export] Instance quote_ctx_inst {typing Σ Γ ctx inst} + #[export] Instance quote_ctx_inst {typing Γ ctx inst} {qtyping : quotation_of typing} - {quote_typing : forall i t, ground_quotable (typing Σ Γ i t)} - : ground_quotable (@ctx_inst typing Σ Γ ctx inst) + {quote_typing : forall i t, ground_quotable (typing Γ i t)} + : ground_quotable (@ctx_inst typing Γ ctx inst) := ltac:(induction 1; exact _). End QuoteEnvTyping. @@ -115,10 +111,10 @@ Module QuoteGlobalMaps (Import T : Term) (Import E : EnvironmentSig T) (Import T Section GlobalMaps. Context {cf : config.checker_flags} {Pcmp: global_env_ext -> context -> conv_pb -> term -> term -> Type} - {P : global_env_ext -> context -> term -> typ_or_sort -> Type} + {P : global_env_ext -> context -> judgment -> Type} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} - {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)}. + {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)}. #[export] Instance quote_on_context {Σ ctx} : ground_quotable (@on_context P Σ ctx) := ltac:(cbv [on_context]; exact _). diff --git a/quotation/theories/ToPCUIC/Common/Universes.v b/quotation/theories/ToPCUIC/Common/Universes.v index b5d367728..3077ec65a 100644 --- a/quotation/theories/ToPCUIC/Common/Universes.v +++ b/quotation/theories/ToPCUIC/Common/Universes.v @@ -58,7 +58,7 @@ End QuoteUniverses1. Export (hints) QuoteUniverses1. #[export] Hint Unfold - LevelAlgExpr.t + Universe.t Instance.t UContext.t AUContext.t @@ -68,34 +68,20 @@ Export (hints) QuoteUniverses1. : quotation. #[export] Typeclasses Transparent - LevelAlgExpr.t + Universe.t Instance.t UContext.t AUContext.t ContextSet.t . -#[export] Instance quote_nonEmptyLevelExprSet : ground_quotable nonEmptyLevelExprSet := ltac:(destruct 1; exact _). - -#[export] Instance quote_concreteUniverses : ground_quotable concreteUniverses := ltac:(destruct 1; exact _). -Import StrongerInstances. -#[export] Instance quote_leq_cuniverse_n {cf n u u'} : ground_quotable (@leq_cuniverse_n cf n u u') := ltac:(cbv [leq_cuniverse_n]; exact _). -#[export] Instance quote_is_uprop {u} : ground_quotable (@is_uprop u) := ltac:(cbv [is_uprop]; exact _). -#[export] Instance quote_is_usprop {u} : ground_quotable (@is_usprop u) := ltac:(cbv [is_usprop]; exact _). -#[export] Instance quote_is_uproplevel {u} : ground_quotable (@is_uproplevel u) := ltac:(cbv [is_uproplevel]; exact _). -#[export] Instance quote_is_uproplevel_or_set {u} : ground_quotable (@is_uproplevel_or_set u) := ltac:(cbv [is_uproplevel_or_set]; exact _). -#[export] Instance quote_is_utype {u} : ground_quotable (@is_utype u) := ltac:(cbv [is_utype]; exact _). - -#[export] Instance quote_allowed_eliminations : ground_quotable allowed_eliminations := ltac:(destruct 1; exact _). -#[export] Instance quote_is_allowed_elimination_cuniv {allowed u} : ground_quotable (is_allowed_elimination_cuniv allowed u) := ltac:(destruct allowed; cbv [is_allowed_elimination_cuniv]; exact _). - Module QuoteUniverses2. Module Import Universe. - #[export] Instance quote_t_ : ground_quotable Universe.t_ := ltac:(destruct 1; exact _). - #[export] Hint Unfold Universe.t : quotation. - #[export] Typeclasses Transparent Universe.t. + #[export] Instance quote_t_ {univ} {quniv : quotation_of univ} {quote_univ : ground_quotable univ} : ground_quotable (Sort.t_ univ) := ltac:(destruct 1; exact _). + #[export] Hint Unfold sort : quotation. + #[export] Typeclasses Transparent sort. #[local] Hint Constructors or eq : typeclass_instances. - #[export] Instance quote_on_sort {P def s} {quoteP : forall l, s = Universe.lType l -> ground_quotable (P l:Prop)} {quote_def : s = Universe.lProp \/ s = Universe.lSProp -> ground_quotable (def:Prop)} : ground_quotable (@Universe.on_sort P def s) := ltac:(cbv [Universe.on_sort]; exact _). + #[export] Instance quote_on_sort {univ P def s} {quniv : quotation_of univ} {quote_univ : ground_quotable univ} {quoteP : forall l, s = sType l -> ground_quotable (P l : Prop)} {quote_def : s = sProp \/ s = sSProp -> ground_quotable (def : Prop)} : ground_quotable (@Sort.on_sort univ Prop P def s) := ltac:(cbv [Sort.on_sort]; exact _). End Universe. Export (hints) Universe. @@ -129,6 +115,14 @@ Module QuoteUniverses2. End QuoteUniverses2. Export (hints) QuoteUniverses2. +#[export] Instance quote_nonEmptyLevelExprSet : ground_quotable nonEmptyLevelExprSet := ltac:(destruct 1; exact _). +#[export] Instance quote_Universe : ground_quotable Universe.t := ltac:(destruct 1; exact _). + +#[export] Instance quote_concrete_sort : ground_quotable concrete_sort := ltac:(destruct 1; exact _). +Import StrongerInstances. + +#[export] Instance quote_allowed_eliminations : ground_quotable allowed_eliminations := ltac:(destruct 1; exact _). + #[export] Instance quote_declared_cstr_levels {levels cstr} : ground_quotable (declared_cstr_levels levels cstr) := ltac:(cbv [declared_cstr_levels]; exact _). #[export] Instance quote_universes_decl : ground_quotable universes_decl := ltac:(destruct 1; exact _). #[export] Instance quote_satisfies0 {v s} {qv : quotation_of v} : ground_quotable (@satisfies0 v s) @@ -139,16 +133,19 @@ Export (hints) QuoteUniverses2. := ground_quotable_of_dec (@consistent_dec ctrs). #[export] Instance quote_consistent_extension_on {cs cstr} : ground_quotable (@consistent_extension_on cs cstr) := ground_quotable_of_dec (@consistent_extension_on_dec cs cstr). -#[export] Instance quote_leq_levelalg_n {cf n ϕ u u'} : ground_quotable (@leq_levelalg_n cf (uGraph.Z_of_bool n) ϕ u u') - := ground_quotable_of_dec (@leq_levelalg_n_dec cf _ ϕ u u'). -#[export] Instance quote_leq_universe_n_ {cf CS leq_levelalg_n n ϕ s s'} {quote_leq_levelalg_n : forall u u', ground_quotable (leq_levelalg_n n ϕ u u':Prop)} : ground_quotable (@leq_universe_n_ cf CS leq_levelalg_n n ϕ s s') := ltac:(cbv [leq_universe_n_]; exact _). -#[export] Instance quote_leq_universe_n {cf n ϕ s s'} : ground_quotable (@leq_universe_n cf (uGraph.Z_of_bool n) ϕ s s') := ltac:(cbv [leq_universe_n]; exact _). +#[export] Instance quote_leq_universe_n {cf n ϕ u u'} : ground_quotable (@leq_universe_n cf (uGraph.Z_of_bool n) ϕ u u') + := ground_quotable_of_dec (@leq_universe_n_dec cf _ ϕ u u'). #[export] Instance quote_leq_universe {cf ϕ s s'} : ground_quotable (@leq_universe cf ϕ s s') := @quote_leq_universe_n cf false ϕ s s'. -#[export] Instance quote_eq_levelalg {cf ϕ u u'} : ground_quotable (@eq_levelalg cf ϕ u u') - := ground_quotable_of_dec (@eq_levelalg_dec cf ϕ u u'). -#[export] Instance quote_eq_universe_ {CS eq_levelalg ϕ s s'} {quote_eq_levelalg : forall u u', ground_quotable (eq_levelalg ϕ u u':Prop)} : ground_quotable (@eq_universe_ CS eq_levelalg ϕ s s') := ltac:(cbv [eq_universe_]; exact _). -#[export] Instance quote_eq_universe {cf ϕ s s'} : ground_quotable (@eq_universe cf ϕ s s') := ltac:(cbv [eq_universe]; exact _). -#[export] Instance quote_compare_universe {cf pb ϕ u u'} : ground_quotable (@compare_universe cf pb ϕ u u') := ltac:(destruct pb; cbv [compare_universe]; exact _). +#[export] Instance quote_leq_sort_n_ {cf univ leq_universe_n n s s'} {quote_leq_universe_n : forall u u', ground_quotable (leq_universe_n n u u':Prop)} : ground_quotable (@leq_sort_n_ cf univ leq_universe_n n s s') := ltac:(cbv [leq_sort_n_]; exact _). +#[export] Instance quote_leq_csort_n {cf n s s'} : ground_quotable (@leq_csort_n cf n s s') := @quote_leq_sort_n_ cf nat _ n s s' _. +#[export] Instance quote_leq_sort_n {cf n ϕ s s'} : ground_quotable (@leq_sort_n cf (uGraph.Z_of_bool n) ϕ s s') := ltac:(cbv [leq_sort_n]; exact _). +#[export] Instance quote_leq_sort {cf ϕ s s'} : ground_quotable (@leq_sort cf ϕ s s') := @quote_leq_sort_n cf false ϕ s s'. +#[export] Instance quote_eq_universe {cf ϕ u u'} : ground_quotable (@eq_universe cf ϕ u u') + := ground_quotable_of_dec (@eq_universe_dec cf ϕ u u'). +#[export] Instance quote_eq_sort_ {univ eq_universe s s'} {quote_eq_universe : forall u u', ground_quotable (eq_universe u u':Prop)} : ground_quotable (@eq_sort_ univ eq_universe s s') := ltac:(cbv [eq_sort_]; exact _). +#[export] Instance quote_eq_sort {cf ϕ s s'} : ground_quotable (@eq_sort cf ϕ s s') := ltac:(cbv [eq_sort]; exact _). +#[export] Instance quote_compare_universe {cf univ pb u u'} : ground_quotable (@compare_universe cf univ pb u u') := ltac:(destruct pb; cbv [compare_universe]; exact _). +#[export] Instance quote_compare_sort {cf univ pb u u'} : ground_quotable (@compare_sort cf univ pb u u') := ltac:(destruct pb; cbv [compare_sort]; exact _). #[export] Instance quote_valid_constraints0 {ϕ ctrs} : ground_quotable (@valid_constraints0 ϕ ctrs) := ground_quotable_of_dec (@valid_constraints0_dec ϕ ctrs). #[export] Instance quote_valid_constraints {cf ϕ ctrs} : ground_quotable (@valid_constraints cf ϕ ctrs) diff --git a/quotation/theories/ToPCUIC/Coq/Init.v b/quotation/theories/ToPCUIC/Coq/Init.v index 02d39f555..b0a969fa0 100644 --- a/quotation/theories/ToPCUIC/Coq/Init.v +++ b/quotation/theories/ToPCUIC/Coq/Init.v @@ -75,6 +75,12 @@ End Init. #[export] Instance quote_sumbool {A B : Prop} {qA : quotation_of A} {qB : quotation_of B} {quoteA : ground_quotable A} {quoteB : ground_quotable B} : ground_quotable (sumbool A B) := ltac:(destruct 1; exact _). #[export] Instance quote_sumor {A} {B : Prop} {qA : quotation_of A} {qB : quotation_of B} {quoteA : ground_quotable A} {quoteB : ground_quotable B} : ground_quotable (sumor A B) := ltac:(destruct 1; exact _). +Definition quote_or_dec {P Q : Prop} (dec : P \/ Q -> {P} + {Q}) {qP : quotation_of P} {qQ : quotation_of Q} {quoteP : ground_quotable P} {quoteQ : ground_quotable Q} : ground_quotable (or P Q). +Proof. + intro pf. + destruct (dec pf). + all: adjust_quotation_of_by_econstructor_then ltac:(fun _ => destruct pf; first [ eassumption | tauto ]) ltac:(fun _ => exact _). +Defined. Definition quote_or_dec_l {P Q : Prop} (decP : {P} + {~P}) {qP : quotation_of P} {qQ : quotation_of Q} {quoteP : ground_quotable P} {quoteQ : ground_quotable Q} : ground_quotable (or P Q). Proof. destruct decP. diff --git a/quotation/theories/ToPCUIC/PCUIC/PCUICCumulativitySpec.v b/quotation/theories/ToPCUIC/PCUIC/PCUICCumulativitySpec.v index 284810bb3..2522a9241 100644 --- a/quotation/theories/ToPCUIC/PCUIC/PCUICCumulativitySpec.v +++ b/quotation/theories/ToPCUIC/PCUIC/PCUICCumulativitySpec.v @@ -6,10 +6,16 @@ From MetaCoq.Quotation.ToPCUIC.Common Require Import (hints) config BasicAst Uni From MetaCoq.Quotation.ToPCUIC.PCUIC Require Import (hints) PCUICAst PCUICEquality utils.PCUICPrimitive. From MetaCoq.Quotation.ToPCUIC.QuotationOf.Common Require Import EnvironmentTyping.Sig. -#[export] Instance quote_cumul_predicate {cumul Γ Re p p'} {qcumul : quotation_of cumul} {qRe : quotation_of Re} {quote_cumul : forall x y Γ, ground_quotable (cumul Γ x y)} {quote_Re : forall x y, ground_quotable (Re x y:Prop)} : ground_quotable (@cumul_predicate cumul Γ Re p p') +#[export] Instance quote_cumul_predicate {cumul Γ Re p p'} {qcumul : quotation_of cumul} {qRe : quotation_of Re} {quote_cumul : forall x y Γ, ground_quotable (cumul Γ x y)} {quote_Re : forall x y, ground_quotable (Re x y:Prop)} : ground_quotable (@cumul_predicate cumul Re Γ p p') := ltac:(cbv [cumul_predicate]; exact _). -Definition quote_cumul_predicate_via_dep {cumul Γ Re Re' p p'} {c : @cumul_predicate cumul Γ Re p p'} (qc : cumul_predicate_dep c (fun _ _ _ c => quotation_of c) Re') {qcumul : quotation_of cumul} {qRe : quotation_of Re} {quote_Re : forall x y, ground_quotable (Re x y:Prop)} : quotation_of c. +#[export] Instance quote_cumul_branch {cumul Γ p br br'} {qcumul : quotation_of cumul} {quote_cumul : forall x y Γ, ground_quotable (cumul Γ x y)} : ground_quotable (@cumul_branch cumul Γ p br br') +:= ltac:(cbv [cumul_branch]; exact _). + +#[export] Instance quote_cumul_mfixpoint {cumul Γ mfix mfix'} {qcumul : quotation_of cumul} {quote_cumul : forall x y Γ, ground_quotable (cumul Γ x y)} : ground_quotable (@cumul_mfixpoint cumul Γ mfix mfix') +:= ltac:(cbv [cumul_mfixpoint]; exact _). + +Definition quote_cumul_predicate_via_dep {cumul Γ Re Re' p p'} {c : @cumul_predicate cumul Re Γ p p'} (qc : cumul_predicate_dep c (fun _ _ _ c => quotation_of c) Re') {qcumul : quotation_of cumul} {qRe : quotation_of Re} {quote_Re : forall x y, ground_quotable (Re x y:Prop)} : quotation_of c. Proof. cbv [cumul_predicate cumul_predicate_dep] in *. repeat match goal with H : _ * _ |- _ => destruct H end. @@ -17,7 +23,7 @@ Proof. Defined. #[export] Hint Extern 0 (cumul_predicate_dep _ (fun _ _ _ r => quotation_of r) _) => eassumption : typeclass_instances. -#[export] Hint Extern 0 (@quotation_of (@cumul_predicate ?cumul ?Γ ?Re ?p ?p') ?x) +#[export] Hint Extern 0 (@quotation_of (@cumul_predicate ?cumul ?Re ?Γ ?p ?p') ?x) => lazymatch goal with | [ H : cumul_predicate_dep x (fun _ _ _ r => quotation_of r) ?Re' |- _ ] => simple eapply (@quote_cumul_predicate_via_dep cumul Γ Re Re' p p' x H) @@ -42,7 +48,9 @@ Defined. #[export] Instance quote_cumulSpec0 {cf Σ Γ pb t u} : ground_quotable (@cumulSpec0 cf Σ Γ pb t u). Proof. + pose proof (compare_universe_subrel Σ : forall pb, RelationClasses.subrelation (compare_universe Σ Conv) (compare_universe Σ pb)). induction 1. + all: unfold cumul_branches, cumul_branch, cumul_mfixpoint in *. all: lazymatch goal with | [ H : All_Forall.All2_dep ?T ?x |- _ ] => lazymatch T with diff --git a/quotation/theories/ToPCUIC/PCUIC/PCUICEquality.v b/quotation/theories/ToPCUIC/PCUIC/PCUICEquality.v index 904baa88f..ba56223fd 100644 --- a/quotation/theories/ToPCUIC/PCUIC/PCUICEquality.v +++ b/quotation/theories/ToPCUIC/PCUIC/PCUICEquality.v @@ -5,51 +5,61 @@ From MetaCoq.Quotation.ToPCUIC.Utils Require Import (hints) utils All_Forall. From MetaCoq.Quotation.ToPCUIC.Common Require Import (hints) config Reflect Environment Universes BasicAst Kernames. From MetaCoq.Quotation.ToPCUIC.PCUIC Require Import (hints) PCUICAst utils.PCUICPrimitive (*PCUICAstUtils*) (*Induction*). -#[export] Instance quote_R_universe_instance {R u u'} {qR : quotation_of R} {quoteR : forall x y, ground_quotable (R x y:Prop)} : ground_quotable (@R_universe_instance R u u') := ltac:(cbv [R_universe_instance]; exact _). +#[export] Instance quote_on_rel {T T'} {R} {f: T -> T'} {x y : T} {qR : quotation_of R} {quoteR : forall x y, ground_quotable (R x y:Prop)}: ground_quotable (MCRelations.on_rel R f x y) := ltac:(cbv [MCRelations.on_rel]; exact _). + +#[export] Instance quote_cmp_universe_instance {R u u'} {qR : quotation_of R} {quoteR : forall x y, ground_quotable (R x y:Prop)} : ground_quotable (@cmp_universe_instance R u u') := ltac:(cbv [cmp_universe_instance]; exact _). Section with_R. - Context {Re Rle : Universe.t -> Universe.t -> Prop} - {qRe : quotation_of Re} {qRle : quotation_of Rle} - {quoteRe : forall x y, ground_quotable (Re x y)} {quoteRle : forall x y, ground_quotable (Rle x y)}. + Context {cmp_universe : conv_pb -> Universe.t -> Universe.t -> Prop} + {qRe : quotation_of cmp_universe} + {quoteRe : forall pb x y, ground_quotable (cmp_universe pb x y)}. - #[export] Instance quote_R_universe_variance {v u u'} : ground_quotable (@R_universe_variance Re Rle v u u') := ltac:(cbv [R_universe_variance]; exact _). + #[export] Instance quote_cmp_universe_variance {pb v u u'} : ground_quotable (@cmp_universe_variance cmp_universe pb v u u') := ltac:(cbv [cmp_universe_variance]; exact _). - #[export] Instance quote_R_universe_instance_variance {v u u'} : ground_quotable (@R_universe_instance_variance Re Rle v u u') := ltac:(revert v u'; induction u, u'; cbn; exact _). + #[export] Instance quote_cmp_universe_instance_variance {pb v u u'} : ground_quotable (@cmp_universe_instance_variance cmp_universe pb v u u') := ltac:(cbv [cmp_universe_instance_variance]; exact _). - #[export] Instance quote_R_opt_variance {v u u'} : ground_quotable (@R_opt_variance Re Rle v u u') := ltac:(destruct v; cbv [R_opt_variance]; exact _). + #[export] Instance quote_cmp_opt_variance {pb v u u'} (subr : RelationClasses.subrelation (cmp_universe Conv) (cmp_universe pb)) : ground_quotable (@cmp_opt_variance cmp_universe pb v u u'). + Proof using cmp_universe qRe quoteRe. + destruct v; cbv [cmp_opt_variance]; try exact _. + eapply Coq.Init.quote_or_dec; try exact _. + now apply cmp_opt_variance_var_dec. + Defined. - #[export] Instance quote_R_global_instance {Σ gr napp u u'} : ground_quotable (@R_global_instance Σ Re Rle gr napp u u') := ltac:(cbv [R_global_instance]; exact _). + #[export] Instance quote_cmp_global_instance {Σ pb gr napp u u'} (subr : RelationClasses.subrelation (cmp_universe Conv) (cmp_universe pb)) : ground_quotable (@cmp_global_instance Σ cmp_universe pb gr napp u u') := ltac:(cbv [cmp_global_instance]; exact _). End with_R. #[export] Existing Instances - quote_R_universe_variance - quote_R_universe_instance_variance - quote_R_opt_variance - quote_R_global_instance + quote_cmp_universe_variance + quote_cmp_universe_instance_variance + quote_cmp_opt_variance + quote_cmp_global_instance . -#[export] Instance quote_compare_decls {eq_term leq_term u u'} {qeq_term : quotation_of eq_term} {qleq_term : quotation_of leq_term} {quote_eq_term : forall x y, ground_quotable (eq_term x y)} {quote_leq_term : forall x y, ground_quotable (leq_term x y)} : ground_quotable (@compare_decls eq_term leq_term u u') +#[export] Instance quote_compare_decls {d d'} : ground_quotable (@eq_decl_upto_names d d') := ltac:(destruct 1; exact _). #[export] Hint Unfold - eq_predicate + eq_predicate eq_branches eq_branch eq_mfixpoint : quotation. #[export] Typeclasses Transparent - eq_predicate + eq_predicate eq_branches eq_branch eq_mfixpoint . #[export] Instance quote_eq_term_upto_univ_napp - {Re Rle : Universe.t -> Universe.t -> Prop} - {qRe : quotation_of Re} {qRle : quotation_of Rle} - {quoteRe : forall x y, ground_quotable (Re x y)} {quoteRle : forall x y, ground_quotable (Rle x y)} - {Σ napp x y} - : ground_quotable (@eq_term_upto_univ_napp Σ Re Rle napp x y). + {cmp_universe : conv_pb -> Universe.t -> Universe.t -> Prop} + {cmp_sort : conv_pb -> sort -> sort -> Prop} + {qRe : quotation_of cmp_universe} {qRle : quotation_of cmp_sort} + {quoteRe : forall pb x y, ground_quotable (cmp_universe pb x y)} {quoteRle : forall pb x y, ground_quotable (cmp_sort pb x y)} + {Σ pb napp x y} + {subr : RelationClasses.subrelation (cmp_universe Conv) (cmp_universe pb)} + : ground_quotable (@eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp x y). Proof. - unfold ground_quotable; revert Σ Re Rle napp x y qRe qRle quoteRe quoteRle. - fix quote_eq_term_upto_univ_napp 11; intros. + unfold ground_quotable; revert Σ cmp_universe cmp_sort pb napp x y qRe qRle quoteRe quoteRle subr. + fix quote_eq_term_upto_univ_napp 13; intros. lazymatch type of quote_eq_term_upto_univ_napp with - | forall (x1 : ?X1) (x2 : ?X2) (x3 : ?X3) (x4 : ?X4) (x5 : ?X5) (x6 : ?X6) (x7 : ?X7) (x8 : ?X8) (x9 : ?X9) (x10 : ?X10) (t : ?X11), quotation_of t - => change (forall (x1 : X1) (x2 : X2) (x3 : X3) (x4 : X4) (x5 : X5) (x6 : X6) (x7 : X7) (x8 : X8) (x9 : X9) (x10 : X10), ground_quotable X11) in quote_eq_term_upto_univ_napp + | forall (x1 : ?X1) (x2 : ?X2) (x3 : ?X3) (x4 : ?X4) (x5 : ?X5) (x6 : ?X6) (x7 : ?X7) (x8 : ?X8) (x9 : ?X9) (x10 : ?X10) (x11 : ?X11) (x12 : ?X12) (t : ?X13), quotation_of t + => change (forall (x1 : X1) (x2 : X2) (x3 : X3) (x4 : X4) (x5 : X5) (x6 : X6) (x7 : X7) (x8 : X8) (x9 : X9) (x10 : X10) (x11 : X11) (x12 : X12), ground_quotable X13) in quote_eq_term_upto_univ_napp end. destruct t; replace_quotation_of_goal (). Defined. -#[export] Instance quote_compare_term {cf pb Σ ϕ x y} : ground_quotable (@compare_term cf pb Σ ϕ x y) := ltac:(cbv [compare_term]; exact _). +#[export] Instance quote_compare_term {cf pb Σ ϕ x y} : ground_quotable (@compare_term cf Σ ϕ pb x y). +Proof. unshelve eapply quote_eq_term_upto_univ_napp. apply compare_universe_subrel. Defined. diff --git a/quotation/theories/ToPCUIC/QuotationOf/Common/Environment/Sig.v b/quotation/theories/ToPCUIC/QuotationOf/Common/Environment/Sig.v index 3b5d13ea9..df429decb 100644 --- a/quotation/theories/ToPCUIC/QuotationOf/Common/Environment/Sig.v +++ b/quotation/theories/ToPCUIC/QuotationOf/Common/Environment/Sig.v @@ -23,13 +23,13 @@ Module Type QuoteEnvironmentSig (T : Term) (Import E : EnvironmentSig T). context global_declarations global_env_ext - typ_or_sort + judgment : quotation. #[export] Typeclasses Transparent context global_declarations global_env_ext - typ_or_sort + judgment . #[export] Declare Instance quote_constructor_body : ground_quotable constructor_body. diff --git a/quotation/theories/ToPCUIC/QuotationOf/Common/EnvironmentTyping/Sig.v b/quotation/theories/ToPCUIC/QuotationOf/Common/EnvironmentTyping/Sig.v index 7b9cf9b9e..9aed1cc53 100644 --- a/quotation/theories/ToPCUIC/QuotationOf/Common/EnvironmentTyping/Sig.v +++ b/quotation/theories/ToPCUIC/QuotationOf/Common/EnvironmentTyping/Sig.v @@ -14,7 +14,7 @@ Module Type QuoteLookupSig (Import T : Term) (Import E : EnvironmentSig T) (Impo . #[export] Declare Instance quote_consistent_instance {cf lvs ϕ uctx u} : ground_quotable (@consistent_instance cf lvs ϕ uctx u). - #[export] Declare Instance quote_wf_universe {Σ s} : ground_quotable (@wf_universe Σ s). + #[export] Declare Instance quote_wf_sort {Σ s} : ground_quotable (@wf_sort Σ s). #[export] Declare Instance quote_declared_constant {Σ id decl} : ground_quotable (@declared_constant Σ id decl). #[export] Declare Instance quote_declared_minductive {Σ mind decl} : ground_quotable (@declared_minductive Σ mind decl). #[export] Declare Instance quote_declared_inductive {Σ ind mdecl decl} : ground_quotable (@declared_inductive Σ ind mdecl decl). @@ -29,41 +29,32 @@ End QuotationOfEnvTyping. Module Type QuoteEnvTypingSig (Import T : Term) (Import E : EnvironmentSig T) (Import TU : TermUtils T E) (Import ET : EnvTypingSig T E TU). #[export] Hint Unfold - infer_sort - lift_typing + lift_typing0 All_local_rel : quotation. #[export] Typeclasses Transparent - infer_sort - lift_typing + lift_typing0 All_local_rel . - #[export] Declare Instance quote_All_local_env {typing Γ} {qtyping : quotation_of typing} {quote_typing : forall Γ t T, ground_quotable (typing Γ t T)} : ground_quotable (@All_local_env typing Γ). - #[export] Declare Instance quote_on_local_decl {P Γ d} {quoteP1 : forall b, d.(decl_body) = Some b -> ground_quotable (P Γ b (Typ d.(decl_type)))} {quoteP2 : d.(decl_body) = None -> ground_quotable (P Γ d.(decl_type) Sort)} : ground_quotable (@on_local_decl P Γ d). - #[export] Declare Instance quote_lift_judgment {check infer_sort Σ Γ t T} {quote_check : forall T', T = Typ T' -> ground_quotable (check Σ Γ t T')} {quote_infer_sort : T = Sort -> ground_quotable (infer_sort Σ Γ t)} : ground_quotable (@lift_judgment check infer_sort Σ Γ t T). + #[export] Declare Instance quote_All_local_env {typing Γ} {qtyping : quotation_of typing} {quote_typing : forall Γ j, ground_quotable (typing Γ j)} : ground_quotable (@All_local_env typing Γ). + #[export] Declare Instance quote_lift_sorting {checking sorting j} {qcheck : quotation_of checking} {qsorting : quotation_of sorting} {quote_check : forall tm, j_term j = Some tm -> ground_quotable (checking tm (j_typ j))} {quote_sorting : forall u, ground_quotable (sorting (j_typ j) u)} : ground_quotable (@lift_sorting checking sorting j). - #[export] Declare Instance quote_All_local_env_over_gen - {checking sorting cproperty sproperty Σ Γ H} + #[export] Declare Instance quote_All_local_env_over_sorting + {checking sorting cproperty sproperty Γ H} {qchecking : quotation_of checking} {qsorting : quotation_of sorting} {qcproperty : quotation_of cproperty} {qsproperty : quotation_of sproperty} - {quote_checking : forall Γ t T, ground_quotable (checking Σ Γ t T)} {quote_sorting : forall Γ T, ground_quotable (sorting Σ Γ T)} {quote_sproperty : forall Γ all t tu, ground_quotable (sproperty Σ Γ all t tu)} {quote_cproperty : forall Γ all b t tb, ground_quotable (cproperty Σ Γ all b t tb)} - : ground_quotable (@All_local_env_over_gen checking sorting cproperty sproperty Σ Γ H). + {quote_checking : forall Γ t T, ground_quotable (checking Γ t T)} {quote_sorting : forall Γ T u, ground_quotable (sorting Γ T u)} {quote_sproperty : forall Γ all t u tu, ground_quotable (sproperty Γ all t u tu)} {quote_cproperty : forall Γ all b t tb, ground_quotable (cproperty Γ all b t tb)} + : ground_quotable (@All_local_env_over_sorting checking sorting cproperty sproperty Γ H). - #[export] Declare Instance quote_All_local_env_over {typing property Σ Γ H} + #[export] Declare Instance quote_All_local_env_over {typing property Γ H} {qtyping : quotation_of typing} {qproperty : quotation_of property} - {quote_typing : forall Γ t T, ground_quotable (typing Σ Γ t T)} {quote_property : forall Γ all b t tb, ground_quotable (property Σ Γ all b t tb)} - : ground_quotable (@All_local_env_over typing property Σ Γ H). - - #[export] Declare Instance quote_All_local_env_over_sorting - {checking sorting cproperty sproperty Σ Γ H} - {qchecking : quotation_of checking} {qsorting : quotation_of sorting} {qcproperty : quotation_of cproperty} {qsproperty : quotation_of sproperty} - {quote_checking : forall Γ t T, ground_quotable (checking Σ Γ t T)} {quote_sorting : forall Γ T U, ground_quotable (sorting Σ Γ T U)} {quote_sproperty : forall Γ all t tu U, ground_quotable (sproperty Σ Γ all t tu U)} {quote_cproperty : forall Γ all b t tb, ground_quotable (cproperty Σ Γ all b t tb)} - : ground_quotable (@All_local_env_over_sorting checking sorting cproperty sproperty Σ Γ H). + {quote_typing : forall Γ t T, ground_quotable (typing Γ t T)} {quote_property : forall Γ all b t tb, ground_quotable (property Γ all b t tb)} + : ground_quotable (@All_local_env_over typing property Γ H). - #[export] Declare Instance quote_ctx_inst {typing Σ Γ ctx inst} + #[export] Declare Instance quote_ctx_inst {typing Γ ctx inst} {qtyping : quotation_of typing} - {quote_typing : forall i t, ground_quotable (typing Σ Γ i t)} - : ground_quotable (@ctx_inst typing Σ Γ ctx inst). + {quote_typing : forall i t, ground_quotable (typing Γ i t)} + : ground_quotable (@ctx_inst typing Γ ctx inst). End QuoteEnvTypingSig. Module Type QuotationOfConversion (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E) (ET : EnvTypingSig T E TU) (C : ConversionSig T E TU ET). @@ -110,13 +101,13 @@ Module Type QuoteGlobalMapsSig (Import T: Term) (Import E: EnvironmentSig T) (Im fresh_global . - #[export] Declare Instance quote_on_context {P} {qP : quotation_of P} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {Σ ctx} : ground_quotable (@on_context P Σ ctx). + #[export] Declare Instance quote_on_context {P} {qP : quotation_of P} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {Σ ctx} : ground_quotable (@on_context P Σ ctx). - #[export] Declare Instance quote_type_local_ctx {P} {qP : quotation_of P} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {Σ Γ Δ u} : ground_quotable (@type_local_ctx P Σ Γ Δ u). + #[export] Declare Instance quote_type_local_ctx {P} {qP : quotation_of P} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {Σ Γ Δ u} : ground_quotable (@type_local_ctx P Σ Γ Δ u). - #[export] Declare Instance quote_sorts_local_ctx {P} {qP : quotation_of P} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {Σ Γ Δ us} : ground_quotable (@sorts_local_ctx P Σ Γ Δ us). + #[export] Declare Instance quote_sorts_local_ctx {P} {qP : quotation_of P} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {Σ Γ Δ us} : ground_quotable (@sorts_local_ctx P Σ Γ Δ us). - #[export] Declare Instance quote_on_type {P} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {Σ Γ T} : ground_quotable (@on_type P Σ Γ T). + #[export] Declare Instance quote_on_type {P} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {Σ Γ T} : ground_quotable (@on_type P Σ Γ T). #[export] Declare Instance quote_on_udecl {univs udecl} : ground_quotable (@on_udecl univs udecl). #[export] Declare Instance quote_satisfiable_udecl {univs ϕ} : ground_quotable (@satisfiable_udecl univs ϕ). @@ -127,21 +118,21 @@ Module Type QuoteGlobalMapsSig (Import T: Term) (Import E: EnvironmentSig T) (Im #[export] Declare Instance quote_ind_respects_variance {Pcmp} {qPcmp : quotation_of Pcmp} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {Σ mdecl v indices} : ground_quotable (@ind_respects_variance Pcmp Σ mdecl v indices). #[export] Declare Instance quote_cstr_respects_variance {Pcmp} {qPcmp : quotation_of Pcmp} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {Σ mdecl v cs} : ground_quotable (@cstr_respects_variance Pcmp Σ mdecl v cs). - #[export] Declare Instance quote_on_constructor {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {Σ mdecl i idecl ind_indices cdecl cunivs} : ground_quotable (@on_constructor cf Pcmp P Σ mdecl i idecl ind_indices cdecl cunivs). + #[export] Declare Instance quote_on_constructor {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {Σ mdecl i idecl ind_indices cdecl cunivs} : ground_quotable (@on_constructor cf Pcmp P Σ mdecl i idecl ind_indices cdecl cunivs). #[export] Declare Instance quote_on_proj {mdecl mind i k p decl} : ground_quotable (@on_proj mdecl mind i k p decl). #[export] Declare Instance quote_on_projection {mdecl mind i cdecl k p} : ground_quotable (@on_projection mdecl mind i cdecl k p). #[export] Declare Instance quote_on_projections {mdecl mind i idecl ind_indices cdecl} : ground_quotable (@on_projections mdecl mind i idecl ind_indices cdecl). - #[export] Declare Instance quote_check_ind_sorts {cf P} {qP : quotation_of P} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {Σ params kelim ind_indices cdecls ind_sort} : ground_quotable (@check_ind_sorts cf P Σ params kelim ind_indices cdecls ind_sort). - #[export] Declare Instance quote_on_ind_body {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {Σ mind mdecl i idecl} : ground_quotable (@on_ind_body cf Pcmp P Σ mind mdecl i idecl). + #[export] Declare Instance quote_check_ind_sorts {cf P} {qP : quotation_of P} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {Σ params kelim ind_indices cdecls ind_sort} : ground_quotable (@check_ind_sorts cf P Σ params kelim ind_indices cdecls ind_sort). + #[export] Declare Instance quote_on_ind_body {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {Σ mind mdecl i idecl} : ground_quotable (@on_ind_body cf Pcmp P Σ mind mdecl i idecl). #[export] Declare Instance quote_on_variance {cf Σ univs variances} : ground_quotable (@on_variance cf Σ univs variances). - #[export] Declare Instance quote_on_inductive {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {Σ mind mdecl} : ground_quotable (@on_inductive cf Pcmp P Σ mind mdecl). - #[export] Declare Instance quote_on_constant_decl {P} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {Σ d} : ground_quotable (@on_constant_decl P Σ d). - #[export] Declare Instance quote_on_global_decl {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {Σ kn d} : ground_quotable (@on_global_decl cf Pcmp P Σ kn d). - #[export] Declare Instance quote_on_global_decls_data {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {univs retro Σ kn d} : ground_quotable (@on_global_decls_data cf Pcmp P univs retro Σ kn d). - #[export] Declare Instance quote_on_global_decls {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {univs retro Σ} : ground_quotable (@on_global_decls cf Pcmp P univs retro Σ). + #[export] Declare Instance quote_on_inductive {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {Σ mind mdecl} : ground_quotable (@on_inductive cf Pcmp P Σ mind mdecl). + #[export] Declare Instance quote_on_constant_decl {P} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {Σ d} : ground_quotable (@on_constant_decl P Σ d). + #[export] Declare Instance quote_on_global_decl {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {Σ kn d} : ground_quotable (@on_global_decl cf Pcmp P Σ kn d). + #[export] Declare Instance quote_on_global_decls_data {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {univs retro Σ kn d} : ground_quotable (@on_global_decls_data cf Pcmp P univs retro Σ kn d). + #[export] Declare Instance quote_on_global_decls {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {univs retro Σ} : ground_quotable (@on_global_decls cf Pcmp P univs retro Σ). #[export] Declare Instance quote_on_global_univs {univs} : ground_quotable (@on_global_univs univs). - #[export] Declare Instance quote_on_global_env {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {g} : ground_quotable (@on_global_env cf Pcmp P g). - #[export] Declare Instance quote_on_global_env_ext {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {Σ} : ground_quotable (@on_global_env_ext cf Pcmp P Σ). + #[export] Declare Instance quote_on_global_env {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {g} : ground_quotable (@on_global_env cf Pcmp P g). + #[export] Declare Instance quote_on_global_env_ext {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {Σ} : ground_quotable (@on_global_env_ext cf Pcmp P Σ). End QuoteGlobalMapsSig. Module Type QuotationOfConversionPar (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E) (ET : EnvTypingSig T E TU) (CS : ConversionParSig T E TU ET). diff --git a/quotation/theories/ToPCUIC/Utils/All_Forall.v b/quotation/theories/ToPCUIC/Utils/All_Forall.v index c46a9e524..0031d0442 100644 --- a/quotation/theories/ToPCUIC/Utils/All_Forall.v +++ b/quotation/theories/ToPCUIC/Utils/All_Forall.v @@ -31,3 +31,9 @@ Definition quote_All2_via_dep {A B R lsA lsB} {a : @All2 A B R lsA lsB} (qa : Al #[export] Instance quote_All2i_len {A B R lsA lsB} {qA : quotation_of A} {qB : quotation_of B} {qR : quotation_of R} {quoteA : ground_quotable A} {quoteB : ground_quotable B} {quoteR : forall n x y, ground_quotable (R n x y)} : ground_quotable (@All2i_len A B R lsA lsB) := ltac:(induction 1; exact _). #[export] Instance quote_All_fold {A P ls} {qA : quotation_of A} {qP : quotation_of P} {quoteA : ground_quotable A} {quoteP : forall x y, ground_quotable (P x y)} : ground_quotable (@All_fold A P ls) := ltac:(induction 1; exact _). #[export] Instance quote_All2_fold {A P ls1 ls2} {qA : quotation_of A} {qP : quotation_of P} {quoteA : ground_quotable A} {quoteP : forall x y z w, ground_quotable (P x y z w)} : ground_quotable (@All2_fold A P ls1 ls2) := ltac:(induction 1; exact _). +#[export] Instance quote_Forall3 {A B C R lsA lsB lsC} {qA : quotation_of A} {qB : quotation_of B} {qC : quotation_of C} {qR : quotation_of R} {quoteA : ground_quotable A} {quoteB : ground_quotable B} {quoteC : ground_quotable C} {quoteR : forall x y z, ground_quotable (R x y z:Prop)} : ground_quotable (@Forall3 A B C R lsA lsB lsC). +Proof. + revert lsB lsC; induction lsA as [|a lsA IH], lsB as [|b lsB], lsC as [|c lsC]. + all: try solve [ intro pf; exfalso; inversion pf ]. + all: adjust_ground_quotable_by_econstructor_inversion (). +Defined. diff --git a/quotation/theories/ToTemplate/Common/BasicAst.v b/quotation/theories/ToTemplate/Common/BasicAst.v index 62c1e5a21..81bf30b19 100644 --- a/quotation/theories/ToTemplate/Common/BasicAst.v +++ b/quotation/theories/ToTemplate/Common/BasicAst.v @@ -16,7 +16,7 @@ From MetaCoq.Template Require Import AstUtils (* for tFixType *). #[export] Hint Unfold aname : quotation. #[export] Typeclasses Transparent aname. #[export] Instance quote_def {term} {qterm : quotation_of term} {quote_term : ground_quotable term} : ground_quotable (def term) := ltac:(destruct 1; exact _). -#[export] Instance quote_typ_or_sort_ {term} {qterm : quotation_of term} {quote_term : ground_quotable term} : ground_quotable (typ_or_sort_ term) := ltac:(destruct 1; exact _). +#[export] Instance quote_judgment_ {term univ} {qterm : quotation_of term} {uterm : quotation_of univ} {quote_term : ground_quotable term} {quote_univ : ground_quotable univ} : ground_quotable (judgment_ term univ) := ltac:(destruct 1; exact _). #[export] Instance quote_context_decl {term} {qterm : quotation_of term} {quote_term : ground_quotable term} : ground_quotable (context_decl term) := ltac:(destruct 1; exact _). #[export] Hint Unfold mfixpoint : quotation. #[export] Typeclasses Transparent mfixpoint. diff --git a/quotation/theories/ToTemplate/Common/Environment.v b/quotation/theories/ToTemplate/Common/Environment.v index 5b7544a12..adfaebe72 100644 --- a/quotation/theories/ToTemplate/Common/Environment.v +++ b/quotation/theories/ToTemplate/Common/Environment.v @@ -81,13 +81,13 @@ Module QuoteEnvironment (T : Term) (Import E : EnvironmentSig T) (Import QEH : Q context global_declarations global_env_ext - typ_or_sort + judgment : quotation. #[export] Typeclasses Transparent context global_declarations global_env_ext - typ_or_sort + judgment . Import PolymorphicInstances. diff --git a/quotation/theories/ToTemplate/Common/EnvironmentTyping.v b/quotation/theories/ToTemplate/Common/EnvironmentTyping.v index 7601387cf..bb7dc2708 100644 --- a/quotation/theories/ToTemplate/Common/EnvironmentTyping.v +++ b/quotation/theories/ToTemplate/Common/EnvironmentTyping.v @@ -21,8 +21,10 @@ Module QuoteLookup (Import T : Term) (Import E : EnvironmentSig T) (Import L : L End with_refl. #[export] Instance quote_consistent_instance {cf lvs ϕ uctx u} : ground_quotable (@consistent_instance cf lvs ϕ uctx u) := ltac:(cbv [consistent_instance]; exact _). - #[export] Instance quote_wf_universe {Σ s} : ground_quotable (@wf_universe Σ s) - := ground_quotable_of_dec (@wf_universe_dec Σ s). + #[export] Instance quote_wf_universe {Σ u} : ground_quotable (@wf_universe Σ u) + := ground_quotable_of_dec (@wf_universe_dec Σ u). + #[export] Instance quote_wf_sort {Σ s} : ground_quotable (@wf_sort Σ s) + := ground_quotable_of_dec (@wf_sort_dec Σ s). #[export] Instance quote_declared_constant {Σ id decl} : ground_quotable (@declared_constant Σ id decl) := ltac:(cbv [declared_constant]; exact _). #[export] Instance quote_declared_minductive {Σ mind decl} : ground_quotable (@declared_minductive Σ mind decl) := ltac:(cbv [declared_minductive]; exact _). @@ -34,42 +36,36 @@ End QuoteLookup. Module QuoteEnvTyping (Import T : Term) (Import E : EnvironmentSig T) (Import TU : TermUtils T E) (Import ET : EnvTypingSig T E TU) (Import qT : QuotationOfTerm T) (Import qE : QuotationOfEnvironment T E) (Import qET : QuotationOfEnvTyping T E TU ET) (Import QuoteT : QuoteTerm T) (Import QuoteE : QuoteEnvironmentSig T E) <: QuoteEnvTypingSig T E TU ET. #[export] Hint Unfold - infer_sort - lift_typing + on_def_type + on_def_body + lift_typing0 : quotation. #[export] Typeclasses Transparent - infer_sort - lift_typing + on_def_type + on_def_body + lift_typing0 . - #[export] Instance quote_All_local_env {typing Γ} {qtyping : quotation_of typing} {quote_typing : forall Γ t T, ground_quotable (typing Γ t T)} : ground_quotable (@All_local_env typing Γ) := ltac:(induction 1; exact _). + #[export] Instance quote_All_local_env {typing Γ} {qtyping : quotation_of typing} {quote_typing : forall Γ j, ground_quotable (typing Γ j)} : ground_quotable (@All_local_env typing Γ) := ltac:(induction 1; exact _). Import StrongerInstances. #[local] Hint Extern 2 (_ = _) => reflexivity : typeclass_instances. - #[export] Instance quote_on_local_decl {P Γ d} {quoteP1 : forall b, d.(decl_body) = Some b -> ground_quotable (P Γ b (Typ d.(decl_type)))} {quoteP2 : d.(decl_body) = None -> ground_quotable (P Γ d.(decl_type) Sort)} : ground_quotable (@on_local_decl P Γ d) := ltac:(cbv [on_local_decl]; exact _). - #[export] Instance quote_lift_judgment {check infer_sort Σ Γ t T} {quote_check : forall T', T = Typ T' -> ground_quotable (check Σ Γ t T')} {quote_infer_sort : T = Sort -> ground_quotable (infer_sort Σ Γ t)} : ground_quotable (@lift_judgment check infer_sort Σ Γ t T) := ltac:(cbv [lift_judgment]; exact _). - #[local] Typeclasses Transparent lift_judgment. - #[export] Instance quote_All_local_env_over_gen - {checking sorting cproperty sproperty Σ Γ H} + #[export] Instance quote_lift_sorting {checking sorting j} {qcheck : quotation_of checking} {qsorting : quotation_of sorting} {quote_check : forall tm, j_term j = Some tm -> ground_quotable (checking tm (j_typ j))} {quote_sorting : forall u, ground_quotable (sorting (j_typ j) u)} : ground_quotable (@lift_sorting checking sorting j) := ltac:(cbv [lift_sorting]; exact _). + #[local] Typeclasses Transparent lift_sorting. + #[export] Instance quote_All_local_env_over_sorting + {checking sorting cproperty sproperty Γ H} {qchecking : quotation_of checking} {qsorting : quotation_of sorting} {qcproperty : quotation_of cproperty} {qsproperty : quotation_of sproperty} - {quote_checking : forall Γ t T, ground_quotable (checking Σ Γ t T)} {quote_sorting : forall Γ T, ground_quotable (sorting Σ Γ T)} {quote_sproperty : forall Γ all t tu, ground_quotable (sproperty Σ Γ all t tu)} {quote_cproperty : forall Γ all b t tb, ground_quotable (cproperty Σ Γ all b t tb)} - : ground_quotable (@All_local_env_over_gen checking sorting cproperty sproperty Σ Γ H) - := ltac:(induction 1; exact _). - #[export] Instance quote_All_local_env_over {typing property Σ Γ H} + {quote_checking : forall Γ t T, ground_quotable (checking Γ t T)} {quote_sorting : forall Γ T u, ground_quotable (sorting Γ T u)} {quote_sproperty : forall Γ all t u tu, ground_quotable (sproperty Γ all t u tu)} {quote_cproperty : forall Γ all b t tb, ground_quotable (cproperty Γ all b t tb)} + : ground_quotable (@All_local_env_over_sorting checking sorting cproperty sproperty Γ H) := ltac:(induction 1; cbv [lift_sorting j_term MCOption.option_default] in *; exact _). + #[export] Instance quote_All_local_env_over {typing property Γ H} {qtyping : quotation_of typing} {qproperty : quotation_of property} - {quote_typing : forall Γ t T, ground_quotable (typing Σ Γ t T)} {quote_property : forall Γ all b t tb, ground_quotable (property Σ Γ all b t tb)} - : ground_quotable (@All_local_env_over typing property Σ Γ H) + {quote_typing : forall Γ t T, ground_quotable (typing Γ t T)} {quote_property : forall Γ all b t tb, ground_quotable (property Γ all b t tb)} + : ground_quotable (@All_local_env_over typing property Γ H) := ltac:(cbv [All_local_env_over]; exact _). - #[export] Instance quote_All_local_env_over_sorting - {checking sorting cproperty sproperty Σ Γ H} - {qchecking : quotation_of checking} {qsorting : quotation_of sorting} {qcproperty : quotation_of cproperty} {qsproperty : quotation_of sproperty} - {quote_checking : forall Γ t T, ground_quotable (checking Σ Γ t T)} {quote_sorting : forall Γ T U, ground_quotable (sorting Σ Γ T U)} {quote_sproperty : forall Γ all t tu U, ground_quotable (sproperty Σ Γ all t tu U)} {quote_cproperty : forall Γ all b t tb, ground_quotable (cproperty Σ Γ all b t tb)} - : ground_quotable (@All_local_env_over_sorting checking sorting cproperty sproperty Σ Γ H) - := ltac:(cbv [All_local_env_over_sorting]; exact _). - #[export] Instance quote_ctx_inst {typing Σ Γ ctx inst} + #[export] Instance quote_ctx_inst {typing Γ ctx inst} {qtyping : quotation_of typing} - {quote_typing : forall i t, ground_quotable (typing Σ Γ i t)} - : ground_quotable (@ctx_inst typing Σ Γ ctx inst) + {quote_typing : forall i t, ground_quotable (typing Γ i t)} + : ground_quotable (@ctx_inst typing Γ ctx inst) := ltac:(induction 1; exact _). End QuoteEnvTyping. @@ -115,10 +111,10 @@ Module QuoteGlobalMaps (Import T : Term) (Import E : EnvironmentSig T) (Import T Section GlobalMaps. Context {cf : config.checker_flags} {Pcmp: global_env_ext -> context -> conv_pb -> term -> term -> Type} - {P : global_env_ext -> context -> term -> typ_or_sort -> Type} + {P : global_env_ext -> context -> judgment -> Type} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} - {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)}. + {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)}. #[export] Instance quote_on_context {Σ ctx} : ground_quotable (@on_context P Σ ctx) := ltac:(cbv [on_context]; exact _). diff --git a/quotation/theories/ToTemplate/Common/Universes.v b/quotation/theories/ToTemplate/Common/Universes.v index f89016683..bf1ad2912 100644 --- a/quotation/theories/ToTemplate/Common/Universes.v +++ b/quotation/theories/ToTemplate/Common/Universes.v @@ -58,7 +58,7 @@ End QuoteUniverses1. Export (hints) QuoteUniverses1. #[export] Hint Unfold - LevelAlgExpr.t + Universe.t Instance.t UContext.t AUContext.t @@ -68,34 +68,20 @@ Export (hints) QuoteUniverses1. : quotation. #[export] Typeclasses Transparent - LevelAlgExpr.t + Universe.t Instance.t UContext.t AUContext.t ContextSet.t . -#[export] Instance quote_nonEmptyLevelExprSet : ground_quotable nonEmptyLevelExprSet := ltac:(destruct 1; exact _). - -#[export] Instance quote_concreteUniverses : ground_quotable concreteUniverses := ltac:(destruct 1; exact _). -Import StrongerInstances. -#[export] Instance quote_leq_cuniverse_n {cf n u u'} : ground_quotable (@leq_cuniverse_n cf n u u') := ltac:(cbv [leq_cuniverse_n]; exact _). -#[export] Instance quote_is_uprop {u} : ground_quotable (@is_uprop u) := ltac:(cbv [is_uprop]; exact _). -#[export] Instance quote_is_usprop {u} : ground_quotable (@is_usprop u) := ltac:(cbv [is_usprop]; exact _). -#[export] Instance quote_is_uproplevel {u} : ground_quotable (@is_uproplevel u) := ltac:(cbv [is_uproplevel]; exact _). -#[export] Instance quote_is_uproplevel_or_set {u} : ground_quotable (@is_uproplevel_or_set u) := ltac:(cbv [is_uproplevel_or_set]; exact _). -#[export] Instance quote_is_utype {u} : ground_quotable (@is_utype u) := ltac:(cbv [is_utype]; exact _). - -#[export] Instance quote_allowed_eliminations : ground_quotable allowed_eliminations := ltac:(destruct 1; exact _). -#[export] Instance quote_is_allowed_elimination_cuniv {allowed u} : ground_quotable (is_allowed_elimination_cuniv allowed u) := ltac:(destruct allowed; cbv [is_allowed_elimination_cuniv]; exact _). - Module QuoteUniverses2. Module Import Universe. - #[export] Instance quote_t_ : ground_quotable Universe.t_ := ltac:(destruct 1; exact _). - #[export] Hint Unfold Universe.t : quotation. - #[export] Typeclasses Transparent Universe.t. + #[export] Instance quote_t_ {univ} {quniv : quotation_of univ} {quote_univ : ground_quotable univ} : ground_quotable (Sort.t_ univ) := ltac:(destruct 1; exact _). + #[export] Hint Unfold sort : quotation. + #[export] Typeclasses Transparent sort. #[local] Hint Constructors or eq : typeclass_instances. - #[export] Instance quote_on_sort {P def s} {quoteP : forall l, s = Universe.lType l -> ground_quotable (P l:Prop)} {quote_def : s = Universe.lProp \/ s = Universe.lSProp -> ground_quotable (def:Prop)} : ground_quotable (@Universe.on_sort P def s) := ltac:(cbv [Universe.on_sort]; exact _). + #[export] Instance quote_on_sort {univ P def s} {quniv : quotation_of univ} {quote_univ : ground_quotable univ} {quoteP : forall l, s = sType l -> ground_quotable (P l : Prop)} {quote_def : s = sProp \/ s = sSProp -> ground_quotable (def : Prop)} : ground_quotable (@Sort.on_sort univ Prop P def s) := ltac:(cbv [Sort.on_sort]; exact _). End Universe. Export (hints) Universe. @@ -129,6 +115,14 @@ Module QuoteUniverses2. End QuoteUniverses2. Export (hints) QuoteUniverses2. +#[export] Instance quote_nonEmptyLevelExprSet : ground_quotable nonEmptyLevelExprSet := ltac:(destruct 1; exact _). +#[export] Instance quote_Universe : ground_quotable Universe.t := ltac:(destruct 1; exact _). + +#[export] Instance quote_concrete_sort : ground_quotable concrete_sort := ltac:(destruct 1; exact _). +Import StrongerInstances. + +#[export] Instance quote_allowed_eliminations : ground_quotable allowed_eliminations := ltac:(destruct 1; exact _). + #[export] Instance quote_declared_cstr_levels {levels cstr} : ground_quotable (declared_cstr_levels levels cstr) := ltac:(cbv [declared_cstr_levels]; exact _). #[export] Instance quote_universes_decl : ground_quotable universes_decl := ltac:(destruct 1; exact _). #[export] Instance quote_satisfies0 {v s} {qv : quotation_of v} : ground_quotable (@satisfies0 v s) @@ -139,16 +133,19 @@ Export (hints) QuoteUniverses2. := ground_quotable_of_dec (@consistent_dec ctrs). #[export] Instance quote_consistent_extension_on {cs cstr} : ground_quotable (@consistent_extension_on cs cstr) := ground_quotable_of_dec (@consistent_extension_on_dec cs cstr). -#[export] Instance quote_leq_levelalg_n {cf n ϕ u u'} : ground_quotable (@leq_levelalg_n cf (uGraph.Z_of_bool n) ϕ u u') - := ground_quotable_of_dec (@leq_levelalg_n_dec cf _ ϕ u u'). -#[export] Instance quote_leq_universe_n_ {cf CS leq_levelalg_n n ϕ s s'} {quote_leq_levelalg_n : forall u u', ground_quotable (leq_levelalg_n n ϕ u u':Prop)} : ground_quotable (@leq_universe_n_ cf CS leq_levelalg_n n ϕ s s') := ltac:(cbv [leq_universe_n_]; exact _). -#[export] Instance quote_leq_universe_n {cf n ϕ s s'} : ground_quotable (@leq_universe_n cf (uGraph.Z_of_bool n) ϕ s s') := ltac:(cbv [leq_universe_n]; exact _). +#[export] Instance quote_leq_universe_n {cf n ϕ u u'} : ground_quotable (@leq_universe_n cf (uGraph.Z_of_bool n) ϕ u u') + := ground_quotable_of_dec (@leq_universe_n_dec cf _ ϕ u u'). #[export] Instance quote_leq_universe {cf ϕ s s'} : ground_quotable (@leq_universe cf ϕ s s') := @quote_leq_universe_n cf false ϕ s s'. -#[export] Instance quote_eq_levelalg {cf ϕ u u'} : ground_quotable (@eq_levelalg cf ϕ u u') - := ground_quotable_of_dec (@eq_levelalg_dec cf ϕ u u'). -#[export] Instance quote_eq_universe_ {CS eq_levelalg ϕ s s'} {quote_eq_levelalg : forall u u', ground_quotable (eq_levelalg ϕ u u':Prop)} : ground_quotable (@eq_universe_ CS eq_levelalg ϕ s s') := ltac:(cbv [eq_universe_]; exact _). -#[export] Instance quote_eq_universe {cf ϕ s s'} : ground_quotable (@eq_universe cf ϕ s s') := ltac:(cbv [eq_universe]; exact _). -#[export] Instance quote_compare_universe {cf pb ϕ u u'} : ground_quotable (@compare_universe cf pb ϕ u u') := ltac:(destruct pb; cbv [compare_universe]; exact _). +#[export] Instance quote_leq_sort_n_ {cf univ leq_universe_n n s s'} {quote_leq_universe_n : forall u u', ground_quotable (leq_universe_n n u u':Prop)} : ground_quotable (@leq_sort_n_ cf univ leq_universe_n n s s') := ltac:(cbv [leq_sort_n_]; exact _). +#[export] Instance quote_leq_csort_n {cf n s s'} : ground_quotable (@leq_csort_n cf n s s') := @quote_leq_sort_n_ cf nat _ n s s' _. +#[export] Instance quote_leq_sort_n {cf n ϕ s s'} : ground_quotable (@leq_sort_n cf (uGraph.Z_of_bool n) ϕ s s') := ltac:(cbv [leq_sort_n]; exact _). +#[export] Instance quote_leq_sort {cf ϕ s s'} : ground_quotable (@leq_sort cf ϕ s s') := @quote_leq_sort_n cf false ϕ s s'. +#[export] Instance quote_eq_universe {cf ϕ u u'} : ground_quotable (@eq_universe cf ϕ u u') + := ground_quotable_of_dec (@eq_universe_dec cf ϕ u u'). +#[export] Instance quote_eq_sort_ {univ eq_universe s s'} {quote_eq_universe : forall u u', ground_quotable (eq_universe u u':Prop)} : ground_quotable (@eq_sort_ univ eq_universe s s') := ltac:(cbv [eq_sort_]; exact _). +#[export] Instance quote_eq_sort {cf ϕ s s'} : ground_quotable (@eq_sort cf ϕ s s') := ltac:(cbv [eq_sort]; exact _). +#[export] Instance quote_compare_universe {cf univ pb u u'} : ground_quotable (@compare_universe cf univ pb u u') := ltac:(destruct pb; cbv [compare_universe]; exact _). +#[export] Instance quote_compare_sort {cf univ pb u u'} : ground_quotable (@compare_sort cf univ pb u u') := ltac:(destruct pb; cbv [compare_sort]; exact _). #[export] Instance quote_valid_constraints0 {ϕ ctrs} : ground_quotable (@valid_constraints0 ϕ ctrs) := ground_quotable_of_dec (@valid_constraints0_dec ϕ ctrs). #[export] Instance quote_valid_constraints {cf ϕ ctrs} : ground_quotable (@valid_constraints cf ϕ ctrs) diff --git a/quotation/theories/ToTemplate/Coq/Init.v b/quotation/theories/ToTemplate/Coq/Init.v index 826a29845..0151e2ad0 100644 --- a/quotation/theories/ToTemplate/Coq/Init.v +++ b/quotation/theories/ToTemplate/Coq/Init.v @@ -75,6 +75,12 @@ End Init. #[export] Instance quote_sumbool {A B : Prop} {qA : quotation_of A} {qB : quotation_of B} {quoteA : ground_quotable A} {quoteB : ground_quotable B} : ground_quotable (sumbool A B) := ltac:(destruct 1; exact _). #[export] Instance quote_sumor {A} {B : Prop} {qA : quotation_of A} {qB : quotation_of B} {quoteA : ground_quotable A} {quoteB : ground_quotable B} : ground_quotable (sumor A B) := ltac:(destruct 1; exact _). +Definition quote_or_dec {P Q : Prop} (dec : P \/ Q -> {P} + {Q}) {qP : quotation_of P} {qQ : quotation_of Q} {quoteP : ground_quotable P} {quoteQ : ground_quotable Q} : ground_quotable (or P Q). +Proof. + intro pf. + destruct (dec pf). + all: adjust_quotation_of_by_econstructor_then ltac:(fun _ => destruct pf; first [ eassumption | tauto ]) ltac:(fun _ => exact _). +Defined. Definition quote_or_dec_l {P Q : Prop} (decP : {P} + {~P}) {qP : quotation_of P} {qQ : quotation_of Q} {quoteP : ground_quotable P} {quoteQ : ground_quotable Q} : ground_quotable (or P Q). Proof. destruct decP. diff --git a/quotation/theories/ToTemplate/QuotationOf/Common/Environment/Sig.v b/quotation/theories/ToTemplate/QuotationOf/Common/Environment/Sig.v index 86e7a9309..6d61c490e 100644 --- a/quotation/theories/ToTemplate/QuotationOf/Common/Environment/Sig.v +++ b/quotation/theories/ToTemplate/QuotationOf/Common/Environment/Sig.v @@ -23,13 +23,13 @@ Module Type QuoteEnvironmentSig (T : Term) (Import E : EnvironmentSig T). context global_declarations global_env_ext - typ_or_sort + judgment : quotation. #[export] Typeclasses Transparent context global_declarations global_env_ext - typ_or_sort + judgment . #[export] Declare Instance quote_constructor_body : ground_quotable constructor_body. diff --git a/quotation/theories/ToTemplate/QuotationOf/Common/EnvironmentTyping/Sig.v b/quotation/theories/ToTemplate/QuotationOf/Common/EnvironmentTyping/Sig.v index 8ed05be49..b6742d89f 100644 --- a/quotation/theories/ToTemplate/QuotationOf/Common/EnvironmentTyping/Sig.v +++ b/quotation/theories/ToTemplate/QuotationOf/Common/EnvironmentTyping/Sig.v @@ -14,7 +14,7 @@ Module Type QuoteLookupSig (Import T : Term) (Import E : EnvironmentSig T) (Impo . #[export] Declare Instance quote_consistent_instance {cf lvs ϕ uctx u} : ground_quotable (@consistent_instance cf lvs ϕ uctx u). - #[export] Declare Instance quote_wf_universe {Σ s} : ground_quotable (@wf_universe Σ s). + #[export] Declare Instance quote_wf_sort {Σ s} : ground_quotable (@wf_sort Σ s). #[export] Declare Instance quote_declared_constant {Σ id decl} : ground_quotable (@declared_constant Σ id decl). #[export] Declare Instance quote_declared_minductive {Σ mind decl} : ground_quotable (@declared_minductive Σ mind decl). #[export] Declare Instance quote_declared_inductive {Σ ind mdecl decl} : ground_quotable (@declared_inductive Σ ind mdecl decl). @@ -29,41 +29,32 @@ End QuotationOfEnvTyping. Module Type QuoteEnvTypingSig (Import T : Term) (Import E : EnvironmentSig T) (Import TU : TermUtils T E) (Import ET : EnvTypingSig T E TU). #[export] Hint Unfold - infer_sort - lift_typing + lift_typing0 All_local_rel : quotation. #[export] Typeclasses Transparent - infer_sort - lift_typing + lift_typing0 All_local_rel . - #[export] Declare Instance quote_All_local_env {typing Γ} {qtyping : quotation_of typing} {quote_typing : forall Γ t T, ground_quotable (typing Γ t T)} : ground_quotable (@All_local_env typing Γ). - #[export] Declare Instance quote_on_local_decl {P Γ d} {quoteP1 : forall b, d.(decl_body) = Some b -> ground_quotable (P Γ b (Typ d.(decl_type)))} {quoteP2 : d.(decl_body) = None -> ground_quotable (P Γ d.(decl_type) Sort)} : ground_quotable (@on_local_decl P Γ d). - #[export] Declare Instance quote_lift_judgment {check infer_sort Σ Γ t T} {quote_check : forall T', T = Typ T' -> ground_quotable (check Σ Γ t T')} {quote_infer_sort : T = Sort -> ground_quotable (infer_sort Σ Γ t)} : ground_quotable (@lift_judgment check infer_sort Σ Γ t T). + #[export] Declare Instance quote_All_local_env {typing Γ} {qtyping : quotation_of typing} {quote_typing : forall Γ j, ground_quotable (typing Γ j)} : ground_quotable (@All_local_env typing Γ). + #[export] Declare Instance quote_lift_sorting {checking sorting j} {qcheck : quotation_of checking} {qsorting : quotation_of sorting} {quote_check : forall tm, j_term j = Some tm -> ground_quotable (checking tm (j_typ j))} {quote_sorting : forall u, ground_quotable (sorting (j_typ j) u)} : ground_quotable (@lift_sorting checking sorting j). - #[export] Declare Instance quote_All_local_env_over_gen - {checking sorting cproperty sproperty Σ Γ H} + #[export] Declare Instance quote_All_local_env_over_sorting + {checking sorting cproperty sproperty Γ H} {qchecking : quotation_of checking} {qsorting : quotation_of sorting} {qcproperty : quotation_of cproperty} {qsproperty : quotation_of sproperty} - {quote_checking : forall Γ t T, ground_quotable (checking Σ Γ t T)} {quote_sorting : forall Γ T, ground_quotable (sorting Σ Γ T)} {quote_sproperty : forall Γ all t tu, ground_quotable (sproperty Σ Γ all t tu)} {quote_cproperty : forall Γ all b t tb, ground_quotable (cproperty Σ Γ all b t tb)} - : ground_quotable (@All_local_env_over_gen checking sorting cproperty sproperty Σ Γ H). + {quote_checking : forall Γ t T, ground_quotable (checking Γ t T)} {quote_sorting : forall Γ T u, ground_quotable (sorting Γ T u)} {quote_sproperty : forall Γ all t u tu, ground_quotable (sproperty Γ all t u tu)} {quote_cproperty : forall Γ all t T tu, ground_quotable (cproperty Γ all t T tu)} + : ground_quotable (@All_local_env_over_sorting checking sorting cproperty sproperty Γ H). - #[export] Declare Instance quote_All_local_env_over {typing property Σ Γ H} + #[export] Declare Instance quote_All_local_env_over {typing property Γ H} {qtyping : quotation_of typing} {qproperty : quotation_of property} - {quote_typing : forall Γ t T, ground_quotable (typing Σ Γ t T)} {quote_property : forall Γ all b t tb, ground_quotable (property Σ Γ all b t tb)} - : ground_quotable (@All_local_env_over typing property Σ Γ H). - - #[export] Declare Instance quote_All_local_env_over_sorting - {checking sorting cproperty sproperty Σ Γ H} - {qchecking : quotation_of checking} {qsorting : quotation_of sorting} {qcproperty : quotation_of cproperty} {qsproperty : quotation_of sproperty} - {quote_checking : forall Γ t T, ground_quotable (checking Σ Γ t T)} {quote_sorting : forall Γ T U, ground_quotable (sorting Σ Γ T U)} {quote_sproperty : forall Γ all t tu U, ground_quotable (sproperty Σ Γ all t tu U)} {quote_cproperty : forall Γ all b t tb, ground_quotable (cproperty Σ Γ all b t tb)} - : ground_quotable (@All_local_env_over_sorting checking sorting cproperty sproperty Σ Γ H). + {quote_typing : forall Γ t T, ground_quotable (typing Γ t T)} {quote_property : forall Γ all b t tb, ground_quotable (property Γ all b t tb)} + : ground_quotable (@All_local_env_over typing property Γ H). - #[export] Declare Instance quote_ctx_inst {typing Σ Γ ctx inst} + #[export] Declare Instance quote_ctx_inst {typing Γ ctx inst} {qtyping : quotation_of typing} - {quote_typing : forall i t, ground_quotable (typing Σ Γ i t)} - : ground_quotable (@ctx_inst typing Σ Γ ctx inst). + {quote_typing : forall i t, ground_quotable (typing Γ i t)} + : ground_quotable (@ctx_inst typing Γ ctx inst). End QuoteEnvTypingSig. Module Type QuotationOfConversion (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E) (ET : EnvTypingSig T E TU) (C : ConversionSig T E TU ET). @@ -110,13 +101,13 @@ Module Type QuoteGlobalMapsSig (Import T: Term) (Import E: EnvironmentSig T) (Im fresh_global . - #[export] Declare Instance quote_on_context {P} {qP : quotation_of P} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {Σ ctx} : ground_quotable (@on_context P Σ ctx). + #[export] Declare Instance quote_on_context {P} {qP : quotation_of P} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {Σ ctx} : ground_quotable (@on_context P Σ ctx). - #[export] Declare Instance quote_type_local_ctx {P} {qP : quotation_of P} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {Σ Γ Δ u} : ground_quotable (@type_local_ctx P Σ Γ Δ u). + #[export] Declare Instance quote_type_local_ctx {P} {qP : quotation_of P} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {Σ Γ Δ u} : ground_quotable (@type_local_ctx P Σ Γ Δ u). - #[export] Declare Instance quote_sorts_local_ctx {P} {qP : quotation_of P} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {Σ Γ Δ us} : ground_quotable (@sorts_local_ctx P Σ Γ Δ us). + #[export] Declare Instance quote_sorts_local_ctx {P} {qP : quotation_of P} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {Σ Γ Δ us} : ground_quotable (@sorts_local_ctx P Σ Γ Δ us). - #[export] Declare Instance quote_on_type {P} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {Σ Γ T} : ground_quotable (@on_type P Σ Γ T). + #[export] Declare Instance quote_on_type {P} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {Σ Γ T} : ground_quotable (@on_type P Σ Γ T). #[export] Declare Instance quote_on_udecl {univs udecl} : ground_quotable (@on_udecl univs udecl). #[export] Declare Instance quote_satisfiable_udecl {univs ϕ} : ground_quotable (@satisfiable_udecl univs ϕ). @@ -127,21 +118,21 @@ Module Type QuoteGlobalMapsSig (Import T: Term) (Import E: EnvironmentSig T) (Im #[export] Declare Instance quote_ind_respects_variance {Pcmp} {qPcmp : quotation_of Pcmp} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {Σ mdecl v indices} : ground_quotable (@ind_respects_variance Pcmp Σ mdecl v indices). #[export] Declare Instance quote_cstr_respects_variance {Pcmp} {qPcmp : quotation_of Pcmp} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {Σ mdecl v cs} : ground_quotable (@cstr_respects_variance Pcmp Σ mdecl v cs). - #[export] Declare Instance quote_on_constructor {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {Σ mdecl i idecl ind_indices cdecl cunivs} : ground_quotable (@on_constructor cf Pcmp P Σ mdecl i idecl ind_indices cdecl cunivs). + #[export] Declare Instance quote_on_constructor {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {Σ mdecl i idecl ind_indices cdecl cunivs} : ground_quotable (@on_constructor cf Pcmp P Σ mdecl i idecl ind_indices cdecl cunivs). #[export] Declare Instance quote_on_proj {mdecl mind i k p decl} : ground_quotable (@on_proj mdecl mind i k p decl). #[export] Declare Instance quote_on_projection {mdecl mind i cdecl k p} : ground_quotable (@on_projection mdecl mind i cdecl k p). #[export] Declare Instance quote_on_projections {mdecl mind i idecl ind_indices cdecl} : ground_quotable (@on_projections mdecl mind i idecl ind_indices cdecl). - #[export] Declare Instance quote_check_ind_sorts {cf P} {qP : quotation_of P} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {Σ params kelim ind_indices cdecls ind_sort} : ground_quotable (@check_ind_sorts cf P Σ params kelim ind_indices cdecls ind_sort). - #[export] Declare Instance quote_on_ind_body {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {Σ mind mdecl i idecl} : ground_quotable (@on_ind_body cf Pcmp P Σ mind mdecl i idecl). + #[export] Declare Instance quote_check_ind_sorts {cf P} {qP : quotation_of P} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {Σ params kelim ind_indices cdecls ind_sort} : ground_quotable (@check_ind_sorts cf P Σ params kelim ind_indices cdecls ind_sort). + #[export] Declare Instance quote_on_ind_body {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {Σ mind mdecl i idecl} : ground_quotable (@on_ind_body cf Pcmp P Σ mind mdecl i idecl). #[export] Declare Instance quote_on_variance {cf Σ univs variances} : ground_quotable (@on_variance cf Σ univs variances). - #[export] Declare Instance quote_on_inductive {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {Σ mind mdecl} : ground_quotable (@on_inductive cf Pcmp P Σ mind mdecl). - #[export] Declare Instance quote_on_constant_decl {P} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {Σ d} : ground_quotable (@on_constant_decl P Σ d). - #[export] Declare Instance quote_on_global_decl {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {Σ kn d} : ground_quotable (@on_global_decl cf Pcmp P Σ kn d). - #[export] Declare Instance quote_on_global_decls_data {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {univs retro Σ kn d} : ground_quotable (@on_global_decls_data cf Pcmp P univs retro Σ kn d). - #[export] Declare Instance quote_on_global_decls {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {univs retro Σ} : ground_quotable (@on_global_decls cf Pcmp P univs retro Σ). + #[export] Declare Instance quote_on_inductive {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {Σ mind mdecl} : ground_quotable (@on_inductive cf Pcmp P Σ mind mdecl). + #[export] Declare Instance quote_on_constant_decl {P} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {Σ d} : ground_quotable (@on_constant_decl P Σ d). + #[export] Declare Instance quote_on_global_decl {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {Σ kn d} : ground_quotable (@on_global_decl cf Pcmp P Σ kn d). + #[export] Declare Instance quote_on_global_decls_data {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {univs retro Σ kn d} : ground_quotable (@on_global_decls_data cf Pcmp P univs retro Σ kn d). + #[export] Declare Instance quote_on_global_decls {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {univs retro Σ} : ground_quotable (@on_global_decls cf Pcmp P univs retro Σ). #[export] Declare Instance quote_on_global_univs {univs} : ground_quotable (@on_global_univs univs). - #[export] Declare Instance quote_on_global_env {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {g} : ground_quotable (@on_global_env cf Pcmp P g). - #[export] Declare Instance quote_on_global_env_ext {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ t T, ground_quotable (P Σ Γ t T)} {Σ} : ground_quotable (@on_global_env_ext cf Pcmp P Σ). + #[export] Declare Instance quote_on_global_env {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {g} : ground_quotable (@on_global_env cf Pcmp P g). + #[export] Declare Instance quote_on_global_env_ext {cf Pcmp P} {qPcmp : quotation_of Pcmp} {qP : quotation_of P} {quotePcmp : forall Σ Γ pb t T, ground_quotable (Pcmp Σ Γ pb t T)} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {Σ} : ground_quotable (@on_global_env_ext cf Pcmp P Σ). End QuoteGlobalMapsSig. Module Type QuotationOfConversionPar (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E) (ET : EnvTypingSig T E TU) (CS : ConversionParSig T E TU ET). diff --git a/quotation/theories/ToTemplate/Template/TermEquality.v b/quotation/theories/ToTemplate/Template/TermEquality.v index eac3907cc..fdbc90d07 100644 --- a/quotation/theories/ToTemplate/Template/TermEquality.v +++ b/quotation/theories/ToTemplate/Template/TermEquality.v @@ -5,44 +5,54 @@ From MetaCoq.Quotation.ToTemplate.Utils Require Import (hints) utils All_Forall. From MetaCoq.Quotation.ToTemplate.Common Require Import (hints) config Reflect Environment Universes BasicAst Kernames. From MetaCoq.Quotation.ToTemplate.Template Require Import (hints) Ast AstUtils Induction. -#[export] Instance quote_R_universe_instance {R u u'} {qR : quotation_of R} {quoteR : forall x y, ground_quotable (R x y:Prop)} : ground_quotable (@R_universe_instance R u u') := ltac:(cbv [R_universe_instance]; exact _). +#[export] Instance quote_on_rel {T T'} {R} {f: T -> T'} {x y : T} {qR : quotation_of R} {quoteR : forall x y, ground_quotable (R x y:Prop)}: ground_quotable (MCRelations.on_rel R f x y) := ltac:(cbv [MCRelations.on_rel]; exact _). + +#[export] Instance quote_cmp_universe_instance {R u u'} {qR : quotation_of R} {quoteR : forall x y, ground_quotable (R x y:Prop)} : ground_quotable (@cmp_universe_instance R u u') := ltac:(cbv [cmp_universe_instance]; exact _). Section with_R. - Context {Re Rle : Universe.t -> Universe.t -> Prop} - {qRe : quotation_of Re} {qRle : quotation_of Rle} - {quoteRe : forall x y, ground_quotable (Re x y)} {quoteRle : forall x y, ground_quotable (Rle x y)}. + Context {cmp_universe : conv_pb -> Universe.t -> Universe.t -> Prop} + {qRe : quotation_of cmp_universe} + {quoteRe : forall pb x y, ground_quotable (cmp_universe pb x y)}. - #[export] Instance quote_R_universe_variance {v u u'} : ground_quotable (@R_universe_variance Re Rle v u u') := ltac:(cbv [R_universe_variance]; exact _). + #[export] Instance quote_cmp_universe_variance {pb v u u'} : ground_quotable (@cmp_universe_variance cmp_universe pb v u u') := ltac:(cbv [cmp_universe_variance]; exact _). - #[export] Instance quote_R_universe_instance_variance {v u u'} : ground_quotable (@R_universe_instance_variance Re Rle v u u') := ltac:(revert v u'; induction u, u'; cbn; exact _). + #[export] Instance quote_cmp_universe_instance_variance {pb v u u'} : ground_quotable (@cmp_universe_instance_variance cmp_universe pb v u u') := ltac:(cbv [cmp_universe_instance_variance]; exact _). - #[export] Instance quote_R_opt_variance {v u u'} : ground_quotable (@R_opt_variance Re Rle v u u') := ltac:(destruct v; cbv [R_opt_variance]; exact _). + #[export] Instance quote_cmp_opt_variance {pb v u u'} (subr : RelationClasses.subrelation (cmp_universe Conv) (cmp_universe pb)) : ground_quotable (@cmp_opt_variance cmp_universe pb v u u'). + Proof using cmp_universe qRe quoteRe. + destruct v; cbv [cmp_opt_variance]; try exact _. + eapply Coq.Init.quote_or_dec; try exact _. + now apply cmp_opt_variance_var_dec. + Defined. - #[export] Instance quote_R_global_instance {Σ gr napp u u'} : ground_quotable (@R_global_instance Σ Re Rle gr napp u u') := ltac:(cbv [R_global_instance]; exact _). + #[export] Instance quote_cmp_global_instance {Σ pb gr napp u u'} (subr : RelationClasses.subrelation (cmp_universe Conv) (cmp_universe pb)) : ground_quotable (@cmp_global_instance Σ cmp_universe pb gr napp u u') := ltac:(cbv [cmp_global_instance]; exact _). End with_R. #[export] Existing Instances - quote_R_universe_variance - quote_R_universe_instance_variance - quote_R_opt_variance - quote_R_global_instance + quote_cmp_universe_variance + quote_cmp_universe_instance_variance + quote_cmp_opt_variance + quote_cmp_global_instance . -#[export] Instance quote_compare_decls {eq_term leq_term u u'} {qeq_term : quotation_of eq_term} {qleq_term : quotation_of leq_term} {quote_eq_term : forall x y, ground_quotable (eq_term x y)} {quote_leq_term : forall x y, ground_quotable (leq_term x y)} : ground_quotable (@compare_decls eq_term leq_term u u') +#[export] Instance quote_compare_decls {d d'} : ground_quotable (@eq_decl_upto_names d d') := ltac:(destruct 1; exact _). #[export] Instance quote_eq_term_upto_univ_napp - {Re Rle : Universe.t -> Universe.t -> Prop} - {qRe : quotation_of Re} {qRle : quotation_of Rle} - {quoteRe : forall x y, ground_quotable (Re x y)} {quoteRle : forall x y, ground_quotable (Rle x y)} - {Σ napp x y} - : ground_quotable (@eq_term_upto_univ_napp Σ Re Rle napp x y). + {cmp_universe : conv_pb -> Universe.t -> Universe.t -> Prop} + {cmp_sort : conv_pb -> sort -> sort -> Prop} + {qRe : quotation_of cmp_universe} {qRle : quotation_of cmp_sort} + {quoteRe : forall pb x y, ground_quotable (cmp_universe pb x y)} {quoteRle : forall pb x y, ground_quotable (cmp_sort pb x y)} + {Σ pb napp x y} + {subr : RelationClasses.subrelation (cmp_universe Conv) (cmp_universe pb)} + : ground_quotable (@eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp x y). Proof. - unfold ground_quotable; revert Σ Re Rle napp x y qRe qRle quoteRe quoteRle. - fix quote_eq_term_upto_univ_napp 11; intros. + unfold ground_quotable; revert Σ cmp_universe cmp_sort pb napp x y qRe qRle quoteRe quoteRle subr. + fix quote_eq_term_upto_univ_napp 13; intros. lazymatch type of quote_eq_term_upto_univ_napp with - | forall (x1 : ?X1) (x2 : ?X2) (x3 : ?X3) (x4 : ?X4) (x5 : ?X5) (x6 : ?X6) (x7 : ?X7) (x8 : ?X8) (x9 : ?X9) (x10 : ?X10) (t : ?X11), quotation_of t - => change (forall (x1 : X1) (x2 : X2) (x3 : X3) (x4 : X4) (x5 : X5) (x6 : X6) (x7 : X7) (x8 : X8) (x9 : X9) (x10 : X10), ground_quotable X11) in quote_eq_term_upto_univ_napp + | forall (x1 : ?X1) (x2 : ?X2) (x3 : ?X3) (x4 : ?X4) (x5 : ?X5) (x6 : ?X6) (x7 : ?X7) (x8 : ?X8) (x9 : ?X9) (x10 : ?X10) (x11 : ?X11) (x12 : ?X12) (t : ?X13), quotation_of t + => change (forall (x1 : X1) (x2 : X2) (x3 : X3) (x4 : X4) (x5 : X5) (x6 : X6) (x7 : X7) (x8 : X8) (x9 : X9) (x10 : X10) (x11 : X11) (x12 : X12), ground_quotable X13) in quote_eq_term_upto_univ_napp end. destruct t; replace_quotation_of_goal (). Defined. -#[export] Instance quote_compare_term {cf pb Σ ϕ x y} : ground_quotable (@compare_term cf pb Σ ϕ x y) := ltac:(cbv [compare_term]; exact _). +#[export] Instance quote_compare_term {cf pb Σ ϕ x y} : ground_quotable (@compare_term cf Σ ϕ pb x y). +Proof. unshelve eapply quote_eq_term_upto_univ_napp. apply compare_universe_subrel. Defined. diff --git a/quotation/theories/ToTemplate/Template/Typing.v b/quotation/theories/ToTemplate/Template/Typing.v index 656b1ac90..34283aa43 100644 --- a/quotation/theories/ToTemplate/Template/Typing.v +++ b/quotation/theories/ToTemplate/Template/Typing.v @@ -72,7 +72,6 @@ End quote_typing. Definition quote_wf_local {cf : config.checker_flags} {Σ Γ} : ground_quotable (wf_local Σ Γ) := _. #[export] Instance quote_has_nparams {npars ty} : ground_quotable (@has_nparams npars ty) := ltac:(cbv [has_nparams]; exact _). -#[export] Instance quote_infer_sorting {cf Σ Γ T} : ground_quotable (@infer_sorting cf Σ Γ T) := ltac:(cbv [infer_sorting]; exact _). Module QuoteTemplateTyping <: QuoteTyping TemplateTerm Env TemplateTermUtils TemplateEnvTyping TemplateConversion TemplateConversionPar TemplateTyping. diff --git a/quotation/theories/ToTemplate/Template/WfAst.v b/quotation/theories/ToTemplate/Template/WfAst.v index 9ba8c3d34..41c72d916 100644 --- a/quotation/theories/ToTemplate/Template/WfAst.v +++ b/quotation/theories/ToTemplate/Template/WfAst.v @@ -1,7 +1,7 @@ From MetaCoq.Template Require Import Ast WfAst. From MetaCoq.Quotation.ToTemplate Require Import Init. From MetaCoq.Quotation.ToTemplate Require Import (hints) Coq.Init Coq.Lists Coq.Numbers Coq.Floats. -From MetaCoq.Quotation.ToTemplate.Utils Require Import (hints) utils All_Forall MCProd. +From MetaCoq.Quotation.ToTemplate.Utils Require Import (hints) utils All_Forall MCProd MCOption. From MetaCoq.Quotation.ToTemplate.Common Require Import (hints) config BasicAst Universes Kernames. From MetaCoq.Quotation.ToTemplate.Template Require Import (hints) Ast AstUtils Induction UnivSubst. @@ -13,6 +13,5 @@ Proof. Defined. #[export] Instance quote_wf_Inv {Σ t} : ground_quotable (@wf_Inv Σ t) := ltac:(cbv [wf_Inv]; exact _). -Import StrongerInstances. -#[export] Instance quote_wf_decl {Σ d} : ground_quotable (@wf_decl Σ d) := ltac:(cbv [wf_decl]; destruct decl_body; exact _). -#[export] Instance quote_wf_decl_pred {Σ Γ t T} : ground_quotable (@wf_decl_pred Σ Γ t T) := ltac:(cbv [wf_decl_pred]; exact _). +#[export] Instance quote_wf_decl {Σ d} : ground_quotable (@wf_decl Σ d) := ltac:(cbv [wf_decl]; exact _). +#[export] Instance quote_wf_decl_pred {Σ Γ j} : ground_quotable (@wf_decl_pred Σ Γ j) := ltac:(cbv [wf_decl_pred]; exact _). diff --git a/quotation/theories/ToTemplate/Utils/All_Forall.v b/quotation/theories/ToTemplate/Utils/All_Forall.v index 78f35a215..9fbe15c9f 100644 --- a/quotation/theories/ToTemplate/Utils/All_Forall.v +++ b/quotation/theories/ToTemplate/Utils/All_Forall.v @@ -31,3 +31,9 @@ Definition quote_All2_via_dep {A B R lsA lsB} {a : @All2 A B R lsA lsB} (qa : Al #[export] Instance quote_All2i_len {A B R lsA lsB} {qA : quotation_of A} {qB : quotation_of B} {qR : quotation_of R} {quoteA : ground_quotable A} {quoteB : ground_quotable B} {quoteR : forall n x y, ground_quotable (R n x y)} : ground_quotable (@All2i_len A B R lsA lsB) := ltac:(induction 1; exact _). #[export] Instance quote_All_fold {A P ls} {qA : quotation_of A} {qP : quotation_of P} {quoteA : ground_quotable A} {quoteP : forall x y, ground_quotable (P x y)} : ground_quotable (@All_fold A P ls) := ltac:(induction 1; exact _). #[export] Instance quote_All2_fold {A P ls1 ls2} {qA : quotation_of A} {qP : quotation_of P} {quoteA : ground_quotable A} {quoteP : forall x y z w, ground_quotable (P x y z w)} : ground_quotable (@All2_fold A P ls1 ls2) := ltac:(induction 1; exact _). +#[export] Instance quote_Forall3 {A B C R lsA lsB lsC} {qA : quotation_of A} {qB : quotation_of B} {qC : quotation_of C} {qR : quotation_of R} {quoteA : ground_quotable A} {quoteB : ground_quotable B} {quoteC : ground_quotable C} {quoteR : forall x y z, ground_quotable (R x y z:Prop)} : ground_quotable (@Forall3 A B C R lsA lsB lsC). +Proof. + revert lsB lsC; induction lsA as [|a lsA IH], lsB as [|b lsB], lsC as [|c lsC]. + all: try solve [ intro pf; exfalso; inversion pf ]. + all: adjust_ground_quotable_by_econstructor_inversion (). +Defined. diff --git a/safechecker/theories/PCUICEqualityDec.v b/safechecker/theories/PCUICEqualityDec.v index 1d6c7ccc6..80e0e4976 100644 --- a/safechecker/theories/PCUICEqualityDec.v +++ b/safechecker/theories/PCUICEqualityDec.v @@ -16,11 +16,11 @@ Set Default Goal Selector "!". (* TODO move*) -Lemma consistent_instance_wf_universe `{checker_flags} Σ uctx u : +Lemma consistent_instance_wf_sort `{checker_flags} Σ uctx u : consistent_instance_ext Σ uctx u -> - Forall (wf_universe Σ) (map Universe.make u). + Forall (wf_universe Σ) (map Universe.make' u). Proof. - move => /consistent_instance_ext_wf /wf_universe_instanceP. + move => /consistent_instance_ext_wf /wf_instanceP. rewrite -wf_universeb_instance_forall. move => /forallb_Forall ?. eapply Forall_impl ; tea. @@ -28,7 +28,7 @@ Proof. Qed. Lemma ctx_inst_on_universes Σ Γ ts Ts : - ctx_inst (fun Σ' _ t' _ => wf_universes Σ' t') Σ Γ ts Ts -> + ctx_inst (fun _ t _ => wf_universes Σ t) Γ ts Ts -> Forall (wf_universes Σ) ts. Proof. induction 1. @@ -39,112 +39,117 @@ Qed. (** ** Boolean of equality ** *) -Definition compare_universe_variance (equ lequ : Universe.t -> Universe.t -> bool) v u u' := +Definition compare_universe_variance (cmpu : conv_pb -> Universe.t -> Universe.t -> bool) pb v u u' := match v with | Variance.Irrelevant => true - | Variance.Covariant => lequ (Universe.make u) (Universe.make u') - | Variance.Invariant => equ (Universe.make u) (Universe.make u') + | Variance.Covariant => cmpu pb (Universe.make' u) (Universe.make' u') + | Variance.Invariant => cmpu Conv (Universe.make' u) (Universe.make' u') end. Definition compare_universe_instance equ u u' := - forallb2 equ (map Universe.make u) (map Universe.make u'). - -Fixpoint compare_universe_instance_variance equ lequ v u u' := - match u, u' with - | u :: us, u' :: us' => - match v with - | [] => compare_universe_instance_variance equ lequ v us us' - (* Missing variance stands for irrelevance *) - | v :: vs => compare_universe_variance equ lequ v u u' && - compare_universe_instance_variance equ lequ vs us us' - end - | [], [] => true - | _, _ => false - end. + forallb2 (fun u u' => equ (Universe.make' u) (Universe.make' u')) u u'. + +Definition compare_universe_instance_variance cmpu pb v u u' := + forallb3 (compare_universe_variance cmpu pb) v u u'. -Definition compare_global_instance lookup equ lequ gr napp := +Definition compare_global_instance lookup cmpu pb gr napp u u' := match global_variance_gen lookup gr napp with - | Some v => compare_universe_instance_variance equ lequ v - | None => compare_universe_instance equ + | AllEqual => compare_universe_instance (cmpu Conv) u u' + | Variance v => + compare_universe_instance_variance cmpu pb v u u' || compare_universe_instance (cmpu Conv) u u' + | AllIrrelevant => #|u| == #|u'| end. Definition eqb_binder_annots (x y : list aname) : bool := forallb2 eqb_binder_annot x y. +Definition eqb_decl_upto_names (d d' : context_decl) : bool := + match d, d' with + | {| decl_name := na; decl_body := None; decl_type := T |}, + {| decl_name := na'; decl_body := None; decl_type := T' |} => + eqb_binder_annot na na' && eqb T T' + | {| decl_name := na; decl_body := Some b; decl_type := T |}, + {| decl_name := na'; decl_body := Some b'; decl_type := T' |} => + eqb_binder_annot na na' && eqb b b' && eqb T T' + | _, _ => false + end. + +Notation eqb_context_upto_names := (forallb2 eqb_decl_upto_names). + Fixpoint eqb_term_upto_univ_napp - (equ lequ : Universe.t -> Universe.t -> bool) - (gen_compare_global_instance : (Universe.t -> Universe.t -> bool) -> global_reference -> nat -> list Level.t -> list Level.t -> bool) - napp (u v : term) : bool := + (cmpu : conv_pb -> Universe.t -> Universe.t -> bool) + (cmps : conv_pb -> sort -> sort -> bool) + (gen_compare_global_instance : conv_pb -> global_reference -> nat -> list Level.t -> list Level.t -> bool) + pb napp (u v : term) : bool := match u, v with | tRel n, tRel m => eqb n m | tEvar e args, tEvar e' args' => eqb e e' && - forallb2 (eqb_term_upto_univ_napp equ equ gen_compare_global_instance 0) args args' + forallb2 (eqb_term_upto_univ_napp cmpu cmps gen_compare_global_instance Conv 0) args args' | tVar id, tVar id' => eqb id id' | tSort u, tSort u' => - lequ u u' + cmps pb u u' | tApp u v, tApp u' v' => - eqb_term_upto_univ_napp equ lequ gen_compare_global_instance (S napp) u u' && - eqb_term_upto_univ_napp equ equ gen_compare_global_instance 0 v v' + eqb_term_upto_univ_napp cmpu cmps gen_compare_global_instance pb (S napp) u u' && + eqb_term_upto_univ_napp cmpu cmps gen_compare_global_instance Conv 0 v v' | tConst c u, tConst c' u' => eqb c c' && - forallb2 equ (map Universe.make u) (map Universe.make u') + compare_universe_instance (cmpu Conv) u u' | tInd i u, tInd i' u' => eqb i i' && - gen_compare_global_instance lequ (IndRef i) napp u u' + gen_compare_global_instance pb (IndRef i) napp u u' | tConstruct i k u, tConstruct i' k' u' => eqb i i' && eqb k k' && - gen_compare_global_instance lequ (ConstructRef i k) napp u u' + gen_compare_global_instance pb (ConstructRef i k) napp u u' | tLambda na A t, tLambda na' A' t' => eqb_binder_annot na na' && - eqb_term_upto_univ_napp equ equ gen_compare_global_instance 0 A A' && - eqb_term_upto_univ_napp equ lequ gen_compare_global_instance 0 t t' + eqb_term_upto_univ_napp cmpu cmps gen_compare_global_instance Conv 0 A A' && + eqb_term_upto_univ_napp cmpu cmps gen_compare_global_instance pb 0 t t' | tProd na A B, tProd na' A' B' => eqb_binder_annot na na' && - eqb_term_upto_univ_napp equ equ gen_compare_global_instance 0 A A' && - eqb_term_upto_univ_napp equ lequ gen_compare_global_instance 0 B B' + eqb_term_upto_univ_napp cmpu cmps gen_compare_global_instance Conv 0 A A' && + eqb_term_upto_univ_napp cmpu cmps gen_compare_global_instance pb 0 B B' | tLetIn na B b u, tLetIn na' B' b' u' => eqb_binder_annot na na' && - eqb_term_upto_univ_napp equ equ gen_compare_global_instance 0 B B' && - eqb_term_upto_univ_napp equ equ gen_compare_global_instance 0 b b' && - eqb_term_upto_univ_napp equ lequ gen_compare_global_instance 0 u u' + eqb_term_upto_univ_napp cmpu cmps gen_compare_global_instance Conv 0 B B' && + eqb_term_upto_univ_napp cmpu cmps gen_compare_global_instance Conv 0 b b' && + eqb_term_upto_univ_napp cmpu cmps gen_compare_global_instance pb 0 u u' | tCase indp p c brs, tCase indp' p' c' brs' => eqb indp indp' && eqb_predicate_gen - (fun u u' => forallb2 equ (map Universe.make u) (map Universe.make u')) - (bcompare_decls eqb eqb) - (eqb_term_upto_univ_napp equ equ gen_compare_global_instance 0) p p' && - eqb_term_upto_univ_napp equ equ gen_compare_global_instance 0 c c' && + (compare_universe_instance (cmpu Conv)) + eqb_decl_upto_names + (eqb_term_upto_univ_napp cmpu cmps gen_compare_global_instance Conv 0) p p' && + eqb_term_upto_univ_napp cmpu cmps gen_compare_global_instance Conv 0 c c' && forallb2 (fun x y => forallb2 - (bcompare_decls eqb eqb) - x.(bcontext) y.(bcontext) && - eqb_term_upto_univ_napp equ equ gen_compare_global_instance 0 (bbody x) (bbody y) + eqb_decl_upto_names x.(bcontext) y.(bcontext) && + eqb_term_upto_univ_napp cmpu cmps gen_compare_global_instance Conv 0 (bbody x) (bbody y) ) brs brs' | tProj p c, tProj p' c' => eqb p p' && - eqb_term_upto_univ_napp equ equ gen_compare_global_instance 0 c c' + eqb_term_upto_univ_napp cmpu cmps gen_compare_global_instance Conv 0 c c' | tFix mfix idx, tFix mfix' idx' => eqb idx idx' && forallb2 (fun x y => - eqb_term_upto_univ_napp equ equ gen_compare_global_instance 0 x.(dtype) y.(dtype) && - eqb_term_upto_univ_napp equ equ gen_compare_global_instance 0 x.(dbody) y.(dbody) && + eqb_term_upto_univ_napp cmpu cmps gen_compare_global_instance Conv 0 x.(dtype) y.(dtype) && + eqb_term_upto_univ_napp cmpu cmps gen_compare_global_instance Conv 0 x.(dbody) y.(dbody) && eqb x.(rarg) y.(rarg) && eqb_binder_annot x.(dname) y.(dname) ) mfix mfix' @@ -152,72 +157,131 @@ Fixpoint eqb_term_upto_univ_napp | tCoFix mfix idx, tCoFix mfix' idx' => eqb idx idx' && forallb2 (fun x y => - eqb_term_upto_univ_napp equ equ gen_compare_global_instance 0 x.(dtype) y.(dtype) && - eqb_term_upto_univ_napp equ equ gen_compare_global_instance 0 x.(dbody) y.(dbody) && + eqb_term_upto_univ_napp cmpu cmps gen_compare_global_instance Conv 0 x.(dtype) y.(dtype) && + eqb_term_upto_univ_napp cmpu cmps gen_compare_global_instance Conv 0 x.(dbody) y.(dbody) && eqb x.(rarg) y.(rarg) && eqb_binder_annot x.(dname) y.(dname) ) mfix mfix' - | tPrim p, tPrim p' => eqb_prim_val (equ := fun l l' => equ (Universe.make l) (Universe.make l')) - (req := eqb_term_upto_univ_napp equ equ gen_compare_global_instance 0) p p' + | tPrim p, tPrim p' => eqb_prim_val (equ := fun l l' => compare_universe_instance (cmpu Conv) [l] [l']) + (req := eqb_term_upto_univ_napp cmpu cmps gen_compare_global_instance Conv 0) p p' | _, _ => false end. -Notation eqb_term_upto_univ eq leq gen_compare_global_instance := - (eqb_term_upto_univ_napp eq leq gen_compare_global_instance 0). +Notation eqb_term_upto_univ cmpu cmps gen_compare_global_instance pb := + (eqb_term_upto_univ_napp cmpu cmps gen_compare_global_instance pb 0). -Definition conv_pb_relb_gen pb (eq leq : Universe.t -> Universe.t -> bool) := +Definition conv_pb_relb_gen {T} (eq leq : T -> T -> bool) pb := match pb with | Conv => eq | Cumul => leq end. -Definition eqb_termp_napp_gen pb eq leq compare_global_instance_gen napp := - eqb_term_upto_univ_napp eq (conv_pb_relb_gen pb eq leq) - compare_global_instance_gen napp. - +Lemma equiv_reflect P p : + ssrbool.reflect P p <~> + (P <-> p). +Proof. + split. + - destruct 1; split; intro; eauto. + now exfalso. + - destruct p; intros []; constructor; auto. +Qed. -Ltac eqspec := - lazymatch goal with - | |- context [ eqb ?u ?v ] => - destruct (eqb_specT u v) ; nodec ; subst - end. +Lemma reflect_cmp_universe_instance (p : Universe.t -> bool) cmpu cmp_universe ui ui' : + (forall u u', p u -> p u' -> reflect (cmp_universe u u') (cmpu u u')) -> + forallb p (map Universe.make' ui) -> + forallb p (map Universe.make' ui') -> + reflect (cmp_universe_instance cmp_universe ui ui') (compare_universe_instance cmpu ui ui'). +Proof. + intros he hui hui'. + apply equiv_reflect; split. + all: unfold cmp_universe_instance, compare_universe_instance in *. + all: solve_all. + - now apply/he. + - now apply/he. +Qed. -Ltac eqspecs := - repeat eqspec. +Lemma reflect_cmp_universe_instance_variance (p : Universe.t -> bool) cmpu cmp_universe pb v ui ui' : + (forall u u', p u -> p u' -> reflect (cmp_universe Conv u u') (cmpu Conv u u')) -> + (forall u u', p u -> p u' -> reflect (cmp_universe pb u u') (cmpu pb u u')) -> + forallb p (map Universe.make' ui) -> + forallb p (map Universe.make' ui') -> + reflect (cmp_universe_instance_variance cmp_universe pb v ui ui') (compare_universe_instance_variance cmpu pb v ui ui'). +Proof. + intros he hle hui hui'. + apply equiv_reflect; split. + all: unfold cmp_universe_instance_variance, compare_universe_instance_variance in *. + - induction 1 in hui, hui' |- *; cbnr. + cbn in hui, hui'; rtoProp. + split; auto. + destruct x; cbnr. + + now apply/hle. + + now apply/he. + - intro X. apply forallb3_Forall3 in X. + induction X in hui, hui' |- *; cbnr; try solve [ constructor ]. + cbn in hui, hui'; rtoProp. + constructor; auto. + destruct x; cbnr. + + now apply/hle. + + now apply/he. +Qed. -Local Ltac equspec equ h := - repeat lazymatch goal with - | |- context [ equ ?x ?y ] => - destruct (h x y) ; nodec ; subst - end. +Lemma reflect_cmp_global_instance' lookup (p : Universe.t -> bool) cmpu cmp_universe pb gr napp ui ui' : + (forall u u', p u -> p u' -> reflect (cmp_universe Conv u u') (cmpu Conv u u')) -> + (forall u u', p u -> p u' -> reflect (cmp_universe pb u u') (cmpu pb u u')) -> + forallb p (map Universe.make' ui) -> + forallb p (map Universe.make' ui') -> + reflect (cmp_global_instance_gen lookup cmp_universe pb gr napp ui ui') + (compare_global_instance lookup cmpu pb gr napp ui ui'). +Proof. + intros he hle hui hui'. + rewrite /compare_global_instance /cmp_global_instance_gen /cmp_opt_variance. + destruct (global_variance_gen _ gr napp) as [| |v]. + - eapply reflect_cmp_universe_instance; tea. + - apply eqb_specT. + - apply equiv_reflect; split. + + intros [X|X]; apply orb_true_iff; [right|left]. + * now apply/reflect_cmp_universe_instance. + * now apply/reflect_cmp_universe_instance_variance. + + intros [X|X]%orb_true_iff; [right|left]. + * now apply/reflect_cmp_universe_instance_variance. + * now apply/reflect_cmp_universe_instance. +Qed. -Local Ltac ih := - repeat lazymatch goal with - | ih : forall lequ Rle napp hle t' ht ht', reflectT (eq_term_upto_univ_napp _ _ _ napp ?t _) _, - hle : forall u u' hu hu', reflect (?Rle u u') (?lequ u u') , - hcompare : forall R leq H ref n l l' _ _ , _ <-> _ - |- context [ eqb_term_upto_univ _ ?lequ _ ?t ?t' ] => - destruct (ih lequ Rle 0 hle t') - ; nodec ; subst - end. +Lemma reflect_cmp_global_instance Σ lookup (p : Universe.t -> bool) cmpu cmp_universe pb gr napp ui ui' : + (forall u u', p u -> p u' -> reflect (cmp_universe Conv u u') (cmpu Conv u u')) -> + (forall u u', p u -> p u' -> reflect (cmp_universe pb u u') (cmpu pb u u')) -> + (forall kn, lookup_env Σ kn = lookup kn) -> + forallb p (map Universe.make' ui) -> + forallb p (map Universe.make' ui') -> + reflect (cmp_global_instance Σ cmp_universe pb gr napp ui ui') + (compare_global_instance lookup cmpu pb gr napp ui ui'). +Proof. + intros he hleh hlookup hui hui'. + pose proof (Hglobal := reflect_cmp_global_instance' lookup p + cmpu cmp_universe pb gr napp ui ui' he hleh hui hui'). + rewrite /cmp_global_instance_gen /compare_global_instance /cmp_opt_variance. + rewrite /global_variance_gen /lookup_constructor_gen /lookup_inductive_gen /lookup_minductive_gen. + destruct gr; auto; now repeat rewrite hlookup. +Qed. -(* Lemma compare_global_instance_impl (equ lequ : _ -> _ -> bool) Σ Re Rle gr napp : - RelationClasses.subrelation equ Re -> - RelationClasses.subrelation lequ Rle -> - subrelation (compare_global_instance Σ equ lequ gr napp) (R_global_instance Σ Re Rle gr napp). +Lemma forallb_true {A : Type} (l : list A) : forallb xpredT l. Proof. - intros hre hrle x y. - unfold compare_global_instance, R_global_instance, R_opt_variance. - destruct global_variance as [v|]. - - induction x in v, y |- *; destruct v, y; simpl; auto. - rtoProp. intros [Hat Hxy]. split; auto. - destruct t; simpl in *; auto. - - intro. eapply forallb2_Forall2 in H. - eapply Forall2_impl; tea; eauto. -Qed. *) + now induction l. +Qed. +Lemma compare_global_instance_impl Σ lookup cmpu cmp_universe pb gr napp : + (forall u u', reflect (cmp_universe Conv u u') (cmpu Conv u u')) -> + (forall u u', reflect (cmp_universe pb u u') (cmpu pb u u')) -> + (forall kn, lookup_env Σ kn = lookup kn) -> + subrelation (compare_global_instance lookup cmpu pb gr napp) (cmp_global_instance Σ cmp_universe pb gr napp). +Proof. + intros hre hrle hlookup x y. + move/reflect_cmp_global_instance => H. eapply H. + all: intros; eauto. + all: apply forallb_true. +Qed. Lemma Forall2_forallb2: forall (A : Type) (p : A -> A -> bool) (l l' : list A), @@ -261,65 +325,27 @@ Proof. red in H'. unfold eqb_binder_annots in H'. congruence. Qed. -Lemma forallb2_bcompare_decl_All2_fold - (P : term -> term -> bool) Γ Δ : - forallb2 (bcompare_decls P P) Γ Δ -> - All2_fold (fun _ _ => bcompare_decls P P) Γ Δ. +Lemma reflect_eq_decl_upto_names d d' : + reflectT (eq_decl_upto_names d d') (eqb_decl_upto_names d d'). Proof. - induction Γ as [|[na [b|] ty] Γ] in Δ |- *; destruct Δ as [|[na' [b'|] ty'] Δ]; simpl => //; constructor; auto; - now move/andb_and: H => []. + destruct d as [na [b|] ty], d' as [na' [b'|] ty'] => //=. + 2,3: constructor; intro X; inv X. + all: eapply equiv_reflectT; intro X; [inv X|]; rtoProp. + - repeat split; tas. 1: now apply/eqb_annot_reflect. all: apply eqb_refl. + - apply eqb_eq in H0 as <-, H1 as <-. apply eqb_annot_spec in H. + now constructor. + - repeat split; tas. 1: now apply/eqb_annot_reflect. all: apply eqb_refl. + - apply eqb_eq in H0 as <-. apply eqb_annot_spec in H. + now constructor. Qed. -Lemma reflect_eq_context_IH {Σ equ lequ} - {Re Rle : Universe.t -> Universe.t -> Prop} - {gen_compare_global_instance : (Universe.t -> Universe.t -> Prop) -> global_reference -> nat -> list Level.t -> list Level.t -> bool }: - (forall u u', reflect (Re u u') (equ u u')) -> - (forall u u', reflect (Rle u u') (lequ u u')) -> - forall ctx ctx', - onctx - (fun t : term => - forall (lequ : Universe.t -> Universe.t -> bool) - (Rle : Universe.t -> Universe.t -> Prop) - (napp : nat), - (forall u u' : Universe.t, reflect (Rle u u') (lequ u u')) -> - forall t' : term, - reflectT (eq_term_upto_univ_napp Σ Re Rle napp t t') - (eqb_term_upto_univ_napp equ lequ gen_compare_global_instance napp t t')) - ctx -> - reflectT - (eq_context_gen (eq_term_upto_univ Σ Re Re) (eq_term_upto_univ Σ Re Re) ctx ctx') - (forallb2 (bcompare_decls (eqb_term_upto_univ equ equ gen_compare_global_instance) - (eqb_term_upto_univ equ equ gen_compare_global_instance)) ctx ctx'). +Lemma reflect_eq_context_upto_names Γ Δ : + reflectT (eq_context_upto_names Γ Δ) (eqb_context_upto_names Γ Δ). Proof. - intros hRe hRle ctx ctx' onc. - eapply equiv_reflectT. - - intros hcc'. - eapply All2_fold_forallb2, All2_fold_impl_onctx; tea. - unfold ondecl; intuition auto. - depelim X0; cbn in * => //; - intuition auto. - + destruct (eqb_annot_reflect na na') => //. - destruct (a equ Re 0 hRe T') => //. - + destruct (eqb_annot_reflect na na') => //. - destruct (b0 equ Re 0 hRe b') => //. - destruct (a equ Re 0 hRe T') => //. - - intros hcc'. - eapply forallb2_bcompare_decl_All2_fold in hcc'; tea. - eapply All2_fold_impl_onctx in onc; tea; simpl; intuition eauto. - destruct X0. - move: H. - destruct d as [na [bod|] ty], d' as [na' [bod'|] ty']; cbn in * => //. - + destruct (eqb_annot_reflect na na') => //. - destruct (r equ Re 0 hRe ty') => //. - 2: now rewrite andb_false_r. - destruct (o equ Re 0 hRe bod') => //. - now constructor. - + destruct (eqb_annot_reflect na na') => //. - destruct (r equ Re 0 hRe ty') => //. - now constructor. + eapply All2P, reflect_eq_decl_upto_names. Qed. -Definition eqb_univ_reflect : forall u u' : Universe.t, reflectProp (u = u') (eqb u u'). +Definition eqb_univ_reflect : forall u u' : sort, reflectProp (u = u') (eqb u u'). Proof. intros u u'. destruct (eqb_spec u u'); constructor; auto. @@ -332,40 +358,11 @@ Proof. now destruct (Classes.eq_dec x x). Qed. -Lemma reflect_eq_ctx ctx ctx' : - reflectT - (eq_context_gen eq eq ctx ctx') - (forallb2 (bcompare_decls eqb eqb) ctx ctx'). -Proof. - eapply equiv_reflectT. - - intros hcc'. - eapply All2_fold_forallb2; tea. - unfold ondecl; intuition auto. - eapply All2_fold_impl; tea. intros. - destruct X; cbn ; subst. - all: destruct (eqb_annot_reflect na na') => /= //. - 2: apply/andb_and; split. - all: eapply eqb_refl. - - intros hcc'. - eapply forallb2_bcompare_decl_All2_fold in hcc'; tea. - eapply All2_fold_impl in hcc'; tea; simpl; intuition eauto. - move: H. - destruct d as [na [bod|] ty], d' as [na' [bod'|] ty']; cbn in * => //. - + destruct (eqb_annot_reflect na na') => // /=. - repeat case: eqb_specT => //; intros; subst; cbn; auto; constructor; auto. - + destruct (eqb_annot_reflect na na') => //. - repeat case: eqb_specT => //; intros; subst; cbn; auto; constructor; auto. -Qed. - -Lemma forallb_true {A : Type} (l : list A) : forallb xpredT l. -Proof. - now induction l. -Qed. - Lemma on_universes_true (t : term) : on_universes xpredT (fun _ => xpredT) t. Proof. induction t using term_forall_list_ind => //=. - now apply All_forallb. + - now destruct s. - now rewrite IHt1 IHt2. - now rewrite IHt1 IHt2. - now rewrite IHt1 IHt2 IHt3. @@ -401,512 +398,216 @@ Proof. solve_all. Qed. -Definition reflect_eq_predicate {Σ equ lequ} - {gen_compare_global_instance : (Universe.t -> Universe.t -> bool) -> global_reference -> nat -> list Level.t -> list Level.t -> bool} - {Re Rle : Universe.t -> Universe.t -> Prop} - {p : Universe.t -> bool} - {q : nat -> term -> bool} : - (forall u u', p u -> p u' -> reflect (Re u u') (equ u u')) -> - (forall u u', p u -> p u' -> reflect (Rle u u') (lequ u u')) -> - (forall R leq, - (forall u u', p u -> p u' -> reflect (R u u') (leq u u')) -> - forall ref n l l', - forallb p (map Universe.make l) -> - forallb p (map Universe.make l') -> - R_global_instance Σ Re R ref n l l' <-> - gen_compare_global_instance leq ref n l l') -> - forall pr pr', - Forall p (map Universe.make pr.(puinst)) -> - Forall (on_universes p q) pr.(pparams) -> - on_universes p q pr.(preturn) -> - Forall p (map Universe.make pr'.(puinst)) -> - Forall (on_universes p q) pr'.(pparams) -> - on_universes p q pr'.(preturn) -> - tCasePredProp - (fun t : term => - forall (lequ : Universe.t -> Universe.t -> bool) (Rle : Universe.t -> Universe.t -> Prop) (napp : nat), - (forall u u' : Universe.t, p u -> p u' -> reflect (Rle u u') (lequ u u')) -> - forall t' : term, - on_universes p q t -> - on_universes p q t' -> - reflectT (eq_term_upto_univ_napp Σ Re Rle napp t t') (eqb_term_upto_univ_napp equ lequ gen_compare_global_instance napp t t')) - (fun t : term => - forall (lequ : Universe.t -> Universe.t -> bool) (Rle : Universe.t -> Universe.t -> Prop) (napp : nat), - (forall u u' : Universe.t, p u -> p u' -> reflect (Rle u u') (lequ u u')) -> - forall t' : term, - on_universes p q t -> - on_universes p q t' -> - reflectT (eq_term_upto_univ_napp Σ Re Rle napp t t') (eqb_term_upto_univ_napp equ lequ gen_compare_global_instance napp t t')) pr -> - reflectT (eq_predicate (eq_term_upto_univ_napp Σ Re Re 0) Re pr pr') - (eqb_predicate_gen (fun u u' => forallb2 equ (map Universe.make u) (map Universe.make u')) - (bcompare_decls eqb eqb) - (eqb_term_upto_univ_napp equ equ gen_compare_global_instance 0) pr pr'). -Proof. - intros. - solve_all. unfold eq_predicate, eqb_predicate, eqb_predicate_gen. - cbn -[eqb]; apply equiv_reflectT. - - intros Hp; rtoProp. - destruct Hp as [onpars [onuinst [pctx pret]]]. - intuition auto; rtoProp; intuition auto. - * solve_all. destruct (b1 _ Re 0 X y); auto; try contradiction. - * red in onuinst. - solve_all. - eapply All2_impl. - 1: eapply (All2_All_mix_left H0), (All2_All_mix_right H3), All2_map_inv ; - eassumption. - now move=> x y [? []] /X. - * case: (reflect_eq_ctx (pcontext pr) (pcontext pr')) => //. - * now destruct (r equ Re 0 X (preturn pr')) ; nodec ; subst. - - move/andb_and => [/andb_and [/andb_and [ppars pinst] pctx] pret]. - intuition auto. - * solve_all. - now destruct (b1 _ _ 0 X y). - * red. - solve_all. - eapply All2_impl. - 1: eapply (All2_All_mix_left H0), (All2_All_mix_right H3), All2_map_inv; - eassumption. - now move=> x y [? []] /X. - * move: pctx; case: (reflect_eq_ctx (pcontext pr) (pcontext pr')) => //. - * now destruct (r _ _ 0 X (preturn pr')). -Qed. +Ltac eqspec := + lazymatch goal with + | |- context [ eqb ?u ?v ] => + destruct (eqb_specT u v) ; nodec ; subst + end. + +Ltac eqspecs := + repeat eqspec. + +Local Ltac equspec equ pb h := + repeat lazymatch goal with + | |- context [ equ pb ?x ?y ] => + destruct (h x y); tas ; nodec ; subst + end. + +Local Ltac ih := + repeat lazymatch goal with + | ih : forall pb napp hle hspb hglpb t' ht ht', reflectT (eq_term_upto_univ_napp _ _ _ pb napp ?t _) _ + |- context [ eqb_term_upto_univ_napp _ _ _ ?pb ?napp ?t ?t' ] => + destruct (ih pb napp ltac:(assumption) ltac:(assumption) ltac:(assumption) t'); eauto; + nodec ; subst + end. Arguments eqb : simpl never. -Lemma reflect_R_global_instance' lookup equ lequ (p : Universe.t -> bool) - (Re Rle : Universe.t -> Universe.t -> Prop) gr napp ui ui' : - (forall u u', p u -> p u' -> reflect (Re u u') (equ u u')) -> - (forall u u', p u -> p u' -> reflect (Rle u u') (lequ u u')) -> - forallb p (map Universe.make ui) -> - forallb p (map Universe.make ui') -> - reflect (R_global_instance_gen lookup Re Rle gr napp ui ui') - (compare_global_instance lookup equ lequ gr napp ui ui'). +Lemma reflectT_change_left2 P1 P2 R p1 p2 : + CRelationClasses.iffT (P1 × P2) R -> reflectT P1 p1 -> reflectT P2 p2 -> reflectT R (p1 && p2). Proof. - intros he hle hui hui'. - rewrite /compare_global_instance /R_global_instance_gen. - destruct (global_variance_gen _ gr napp) as [v|]. - - induction ui as [|u ui IHui] in ui', v, hui, hui' |- * ; cbn in *. - all: destruct ui' as [|u' ui']. - 1-3: by constructor. - move: hui => /andP [? ?]. - move: hui' => /andP [? ?]. - destruct v as [|hv v]; cbn in *. - 1: now apply IHui. - apply: (iffP andP) => [] [?] /IHui H. - all: split. - 2,4: now apply H. - all: destruct hv => //=. - + apply /hle => //. - + apply /he => //. - + edestruct (hle (Universe.make u) (Universe.make u')) => //. - + edestruct (he (Universe.make u) (Universe.make u')) => //. - - cbn. - unfold R_universe_instance, compare_universe_instance. - apply (iffP idP). - + solve_all. - eapply All2_impl. - 1: eapply (All2_All_mix_left hui), (All2_All_mix_right hui'), All2_map_inv ; - eassumption. - now move => u u' [? []] /he. - + solve_all. - eapply All2_impl. - 1: eapply (All2_All_mix_left hui), (All2_All_mix_right hui'), All2_map_inv ; - eassumption. - now move => u u' [? []] /he. + intros [H H'] [] []; constructor. + 1: apply H; now split. + all: intro HR; destruct (H' HR); now apply f. Qed. -Lemma reflect_R_global_instance Σ lookup equ lequ (p : Universe.t -> bool) - (Re Rle : Universe.t -> Universe.t -> Prop) gr napp ui ui' : - (forall u u', p u -> p u' -> reflect (Re u u') (equ u u')) -> - (forall u u', p u -> p u' -> reflect (Rle u u') (lequ u u')) -> - (forall kn, lookup_env Σ kn = lookup kn) -> - forallb p (map Universe.make ui) -> - forallb p (map Universe.make ui') -> - reflect (R_global_instance Σ Re Rle gr napp ui ui') - (compare_global_instance lookup equ lequ gr napp ui ui'). +Lemma reflectT_change_left3 P1 P2 P3 R p1 p2 p3 : + CRelationClasses.iffT [× P1, P2 & P3] R -> reflectT P1 p1 -> reflectT P2 p2 -> reflectT P3 p3 -> reflectT R (p1 && p2 && p3). Proof. - intros he hleh hlookup hui hui'. - pose proof (Hglobal := reflect_R_global_instance' lookup equ lequ p - Re Rle gr napp ui ui' he hleh hui hui'). - rewrite /R_global_instance_gen /compare_global_instance /R_opt_variance. - rewrite /global_variance_gen /lookup_constructor_gen /lookup_inductive_gen /lookup_minductive_gen. - destruct gr; auto; now repeat rewrite hlookup. + intros [H H'] [] [] []; constructor. + 1: apply H; now split. + all: intro HR; destruct (H' HR); now apply f. +Qed. + +Lemma reflectT_change_left4 P1 P2 P3 P4 R p1 p2 p3 p4 : + CRelationClasses.iffT [× P1, P2, P3 & P4] R -> reflectT P1 p1 -> reflectT P2 p2 -> reflectT P3 p3 -> reflectT P4 p4 -> reflectT R (p1 && p2 && p3 && p4). +Proof. + intros [H H'] [] [] [] []; constructor. + 1: apply H; now split. + all: intro HR; destruct (H' HR); now apply f. Qed. Transparent eqb_prim_val eqb_prim_model. -Lemma reflect_eq_term_upto_univ Σ equ lequ - (p : Universe.t -> bool) (q : nat -> term -> bool) - (Re Rle : Universe.t -> Universe.t -> Prop) - (gen_compare_global_instance : (Universe.t -> Universe.t -> bool) -> global_reference -> nat -> list Level.t -> list Level.t -> bool) - napp : - (forall u u', p u -> p u' -> reflect (Re u u') (equ u u')) -> - (forall u u', p u -> p u' -> reflect (Rle u u') (lequ u u')) -> - (forall R leq, - (forall u u', p u -> p u' -> reflect (R u u') (leq u u')) -> - forall ref n l l' , - forallb p (map Universe.make l) -> - forallb p (map Universe.make l') -> - R_global_instance Σ Re R ref n l l' <-> - gen_compare_global_instance leq ref n l l') -> +Lemma reflect_eq_term_upto_univ Σ (p : Universe.t -> bool) (q : nat -> term -> bool) cmpu cmps cmp_universe cmp_sort + (gen_compare_global_instance : conv_pb -> global_reference -> nat -> list Level.t -> list Level.t -> bool) + pb napp : + (forall u u', p u -> p u' -> reflect (cmp_universe Conv u u') (cmpu Conv u u')) -> + (forall u u', p u -> p u' -> reflect (cmp_universe pb u u') (cmpu pb u u')) -> + (forall s s', Sort.on_sort p true s -> Sort.on_sort p true s' -> reflect (cmp_sort Conv s s') (cmps Conv s s')) -> + (forall s s', Sort.on_sort p true s -> Sort.on_sort p true s' -> reflect (cmp_sort pb s s') (cmps pb s s')) -> + (forall gr napp ui ui', forallb p (map Universe.make' ui) -> forallb p (map Universe.make' ui') -> reflect (cmp_global_instance Σ cmp_universe Conv gr napp ui ui') (gen_compare_global_instance Conv gr napp ui ui')) -> + (forall gr napp ui ui', forallb p (map Universe.make' ui) -> forallb p (map Universe.make' ui') -> reflect (cmp_global_instance Σ cmp_universe pb gr napp ui ui') (gen_compare_global_instance pb gr napp ui ui')) -> forall t t', on_universes p q t -> on_universes p q t' -> - reflectT (eq_term_upto_univ_napp Σ Re Rle napp t t') - (eqb_term_upto_univ_napp equ lequ gen_compare_global_instance napp t t'). + reflectT (eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp t t') + (eqb_term_upto_univ_napp cmpu cmps gen_compare_global_instance pb napp t t'). Proof. - intros he hle hcompare t t' ht ht'. - induction t in t', napp, lequ, Rle, hle, ht, ht' |- * using term_forall_list_ind. + intros he hle hs hspb hgl hglpb t t' ht ht'. + induction t in t', pb, napp, hle, hspb, hglpb, ht, ht' |- * using term_forall_list_ind. all: destruct t' ; nodec. all: move: ht => /= ; (repeat move => /andP [?]) ; move => ht. all: move: ht' => /= ; (repeat move => /andP [?]) ; move => ht'. + all: cbn - [eqb]; eqspecs; try solve [ ih => //; constructor; constructor; assumption ]. + all: try solve [ match goal with |- context [eqb_binder_annot ?na ?na'] => + destruct (eqb_annot_reflect na na'); ih => //=; constructor; constructor; assumption end ]. - - cbn - [eqb]. eqspecs. equspec equ he. equspec lequ hle. ih. - constructor. constructor ; assumption. - - cbn - [eqb]. eqspecs. equspec equ he. equspec lequ hle. ih. - constructor. constructor ; assumption. - - cbn - [eqb]. eqspecs. equspec equ he. equspec lequ hle. ih. - cbn. - induction X in l0, ht, ht' |- *. - + destruct l0. - * constructor. constructor. constructor. - * constructor. intro bot. inversion bot. inversion X. - + destruct l0. - * constructor. intro bot. inversion bot. subst. inversion X0. - * move: ht => /= /andP [? ?]. - move: ht' => /= /andP [? ?]. - destruct (p0 _ _ 0 he t) => //. - -- cbn. destruct (IHX l0) => //. - ++ compute. constructor. constructor. constructor ; try assumption. - inversion e0. subst. assumption. - ++ constructor. intro bot. inversion bot. subst. - inversion X0. subst. - apply f. constructor. assumption. - -- constructor. intro bot. apply f. - inversion bot. subst. inversion X0. subst. assumption. - - cbn - [eqb]. equspec lequ hle. - all: auto. - ih. + - eapply reflectT_change_left. { split; intros XE. 1: constructor; now apply XE. now depelim XE. } + eapply All_reflect_reflect_All2. 2: now apply forallb_All. + solve_all. + - equspec cmps pb hspb. constructor. constructor. assumption. - - cbn - [eqb]. eqspecs. equspec equ he. equspec lequ hle. - destruct (eqb_annot_reflect n na); ih => //=. - constructor. constructor; assumption. - - cbn - [eqb]. eqspecs. equspec equ he. equspec lequ hle. - destruct (eqb_annot_reflect n na); ih => //. - constructor. constructor ; assumption. - - cbn - [eqb]. eqspecs. equspec equ he. equspec lequ hle. - destruct (eqb_annot_reflect n na); ih => //. - constructor. constructor ; assumption. - - cbn - [eqb]. eqspecs. equspec equ he. equspec lequ hle. ih => //. - + destruct (IHt1 lequ Rle (S napp) hle t'1) => //. - * constructor. constructor ; assumption. - * constructor. intros H. now inv H. - + destruct (IHt1 lequ Rle (S napp) hle t'1) => //. - * constructor; auto. - intros H; now inv H. - * constructor. intros H; now inv H. - - cbn - [eqb]. - destruct (eqb_specT s k) ; nodec. subst. - induction u in ui, ht, ht' |- *. - + destruct ui. - * constructor. constructor. constructor. - * constructor. intro bot. inversion bot. subst. inversion H0. - + destruct ui. - * constructor. intro bot. inversion bot. subst. inversion H0. - * move: ht => /= /andP [? ?]. - move: ht' => /= /andP [? ?]. - cbn in *. equspec equ he => //=. - -- destruct (IHu ui) => //=. - ++ constructor. constructor. - inversion e. subst. - constructor ; assumption. - ++ constructor. intro bot. apply f. - inversion bot. subst. constructor. inversion H0. - subst. assumption. - -- constructor. intro bot. apply n. - inversion bot. subst. inversion H0. subst. - assumption. - - cbn - [eqb]. eqspecs. - apply equiv_reflectT. - + inversion 1; subst. now eapply hcompare. - + intros H. - constructor. - now rewrite hcompare. - - cbn - [eqb]. eqspecs. - apply equiv_reflectT. - + inversion 1; subst. now eapply hcompare. - + intros H. - constructor. now eapply hcompare. - - cbn - [eqb]. eqspecs => /=. - unshelve epose proof (Hr := (reflect_eq_predicate he hle hcompare p0 p1 _ _ _ _ _ _ X)). - 7: ih. - 1-8: solve_all. - 2:{ rewrite andb_false_r /=. - constructor. intros bot; inv bot; contradiction. - } - case: Hr => e' /=. - 2:{ - constructor. intro bot. inversion bot. subst. - auto. - } - clear X. rename X0 into X. - induction l in brs, X, ht, ht' |- *. - + destruct brs. - * constructor. constructor ; try assumption. - constructor. - * constructor. intro bot. inversion bot. subst. inversion X2. - + destruct brs. - * constructor. intro bot. inversion bot. subst. inversion X2. - * move: ht; rewrite /= /test_branch /= => /andP [] /andP [? ?] ?. - move: ht'; rewrite /= /test_branch /= => /andP [] /andP [? ?] ?. - cbn - [eqb]. inversion X. subst. - destruct a, b. cbn - [eqb eqb_binder_annots]. - destruct X0 as [onc onbod]. - case: (reflect_eq_ctx bcontext bcontext0) => a // /=. - 2:{ constructor. intro bot. inversion bot. subst. - inversion X3. subst. destruct X4. cbn in *. subst. - contradiction. - } - cbn - [eqb]. - pose proof (onbod equ Re 0 he bbody0) as hh. cbn in hh. - case: hh => //= [e1|f]. - 2:{ constructor. intro bot. apply f. inversion bot. subst. - inversion X3. subst. destruct X4. assumption. } - cbn -[eqb eqb_binder_annots] in *. case: (IHl X1 brs) => // [e2|f]. - 2:{ constructor. intro bot. apply f. inversion bot. subst. - inversion X3. subst. constructor; auto. } - constructor. inversion e2. subst. - constructor ; auto. - - - cbn - [eqb]. eqspecs. equspec equ he. equspec lequ hle. ih => //. - constructor. constructor ; assumption. - - - cbn - [eqb]. eqspecs. equspec equ he. equspec lequ hle. ih. - cbn - [eqb]. - induction m in X, mfix, ht, ht' |- *. - + destruct mfix. - * constructor. constructor. constructor. - * constructor. intro bot. inversion bot. subst. inversion X0. - + destruct mfix. - * constructor. intro bot. inversion bot. subst. inversion X0. - * cbn - [eqb]. - move: ht => /= /andP [] /andP [? ?] ?. - move: ht' => /= /andP [] /andP [? ?] ?. - inversion X. subst. - destruct X0 as [h1 h2]. - destruct (h1 equ Re 0 he (dtype d)) => //. - 2:{ constructor. intro bot. apply f. - inversion bot. subst. inversion X0. subst. apply X2. } - destruct (h2 equ Re 0 he (dbody d)) => //. - 2:{ - constructor. intro bot. apply f. - inversion bot. subst. inversion X0. subst. - apply X2. - } - cbn - [eqb]. eqspecs. - 2:{ constructor. intro bot. inversion bot. subst. - apply n. inversion X0. subst. destruct X2 as [[? ?] ?]. - assumption. - } - destruct (eqb_annot_reflect (dname a) (dname d)). - 2:{ constructor. intro bot; inversion bot. subst. - apply n. inversion X0. subst. destruct X2 as [[? ?] ?]. - assumption. - } - cbn - [eqb]. destruct (IHm X1 mfix) => //. - 2:{ constructor. intro bot. apply f. - inversion bot. subst. constructor. - inversion X0. subst. assumption. } - constructor. constructor. constructor ; try easy. - now inversion e3. - - - cbn - [eqb]. eqspecs. equspec equ he. equspec lequ hle. ih. - cbn - [eqb]. - induction m in X, mfix, ht, ht' |- *. - + destruct mfix. - * constructor. constructor. constructor. - * constructor. intro bot. inversion bot. subst. inversion X0. - + destruct mfix. - * constructor. intro bot. inversion bot. subst. inversion X0. - * cbn - [eqb]. - move: ht => /= /andP [] /andP [? ?] ?. - move: ht' => /= /andP [] /andP [? ?] ?. - inversion X. subst. - destruct X0 as [h1 h2]. - destruct (h1 equ Re 0 he (dtype d)) => //. - 2:{ constructor. intro bot. apply f. - inversion bot. subst. inversion X0. subst. apply X2. } - destruct (h2 equ Re 0 he (dbody d)) => //. - 2:{ - constructor. intro bot. apply f. - inversion bot. subst. inversion X0. subst. - apply X2. - } - cbn - [eqb]. eqspecs. - 2:{ constructor. intro bot. inversion bot. subst. - apply n. inversion X0. subst. destruct X2 as [[? ?] ?]. - assumption. - } - destruct (eqb_annot_reflect (dname a) (dname d)). - 2:{ constructor. intro bot; inversion bot. subst. - apply n. inversion X0. subst. destruct X2 as [[? ?] ?]. - assumption. - } - cbn - [eqb]. destruct (IHm X1 mfix) => //. - 2:{ constructor. intro bot. apply f. - inversion bot. subst. constructor. - inversion X0. subst. assumption. } - constructor. constructor. constructor ; try easy. - now inversion e3. - - - destruct p0 as [? []], prim as [? []]; + - eapply reflectT_change_left. { split; intros XE. 1: constructor; now apply XE. now depelim XE. } + apply reflect_reflectT. eapply reflect_cmp_universe_instance; eauto. + - eapply reflectT_change_left. { split; intros XE. 1: constructor; now apply XE. now depelim XE. } + apply reflect_reflectT. now eapply hglpb. + - eapply reflectT_change_left. { split; intros XE. 1: constructor; now apply XE. now depelim XE. } + apply reflect_reflectT. now eapply hglpb. + + - eapply reflectT_change_left3. { split; intros XE. 1: destruct XE as [XE1 XE2 XE3]; constructor; [apply XE1|apply XE2|apply XE3]. now depelim XE. } + + eapply reflectT_change_left4. { split; intros XE. 1: destruct XE as [XE1 XE2 XE3 XE4]; repeat split; [apply XE1|apply XE2|apply XE3|apply XE4]. now depelim XE. } + * eapply All_reflect_reflect_All2. 2: now apply forallb_All. + solve_all. + * apply reflect_reflectT. eapply reflect_cmp_universe_instance; eauto. + * apply reflect_eq_context_upto_names. + * solve_all. + + ih => //=; constructor; assumption. + + eapply All_reflect_reflect_All2. 2: now apply forallb_All. + solve_all. + eapply reflectT_change_left2. { split; intros XE. 1: destruct XE as [XE1 XE2]; constructor; [apply XE1|apply XE2]. now depelim XE. } + 1: apply reflect_eq_context_upto_names. + unfold test_branch in *. rtoProp. + ih => //=. constructor; assumption. + + - eapply reflectT_change_left. { split; intros XE. 1: constructor; now apply XE. now depelim XE. } + eapply All_reflect_reflect_All2. 2: now apply forallb_All. + solve_all. + eapply reflectT_change_left4. { split; intros XE. 1: destruct XE as [XE1 XE2 XE3 XE4]; repeat split; [apply XE1|apply XE2|apply XE3|apply XE4]. now depelim XE. } + 1,2: ih => //; constructor; assumption. + + apply reflect_reflectT, eqb_specT. + + apply reflect_reflectT, eqb_annot_reflect. + + - eapply reflectT_change_left. { split; intros XE. 1: constructor; now apply XE. now depelim XE. } + eapply All_reflect_reflect_All2. 2: now apply forallb_All. + solve_all. + eapply reflectT_change_left4. { split; intros XE. 1: destruct XE as [XE1 XE2 XE3 XE4]; repeat split; [apply XE1|apply XE2|apply XE3|apply XE4]. now depelim XE. } + 1,2: ih => //; constructor; assumption. + + apply reflect_reflectT, eqb_specT. + + apply reflect_reflectT, eqb_annot_reflect. + + - destruct p0 as [? []], prim as [? []]; cbn in X, ht, ht'; rewrite /eqb_prim_val /eqb_prim_model; cbn -[eqb]; try eqspecs; try repeat constructor. 1-8:intros e; depelim e; try depelim o; intuition eauto. - unfold eqb_array. eqspecs; rewrite ?andb_true_r ?andb_false_r. - rtoProp; intuition auto. specialize (a1 equ Re 0 he _ H4 H0). - specialize (a2 equ Re 0 he _ H5 H1). - case: a1; rewrite ?andb_true_r ?andb_false_r; [intros eq|constructor; intros e; depelim e; depelim o]; eauto. - case: a2; rewrite ?andb_true_r ?andb_false_r; [intros eq'|constructor; intros e; depelim e; depelim o]; eauto. - equspec equ he; eauto; cbn. - 2:constructor; intros e; depelim e; depelim o; eauto. - repeat toAll. - revert b0 H2. - destruct a as [l d ty v], a0 as [l' d' ty' v']; cbn in *. - intros a. - induction a in v' |- *; intros. - * depelim H2; cbn; constructor; eauto; try repeat constructor; cbn; eauto. - intros heq; depelim heq; cbn in *; depelim o; eauto. depelim a0. - * intuition auto. depelim H2; cbn; try constructor; eauto. - ++ intros heq; depelim heq; depelim o; eauto. depelim a2. - ++ specialize (IHa _ H2). case: IHa; intros htl. - +++ rewrite andb_true_r. specialize (b equ Re 0 he _ a0 i). - case: b; repeat constructor; eauto; depelim htl; depelim o; eauto. - intros heq; depelim heq; depelim o; cbn in *; eauto. eapply f. now depelim a3. - +++ rewrite andb_false_r; constructor. - intros heq; depelim heq; depelim o; cbn in *; eauto. eapply htl. - repeat constructor; eauto. cbn. now depelim a2. + rtoProp. destruct X as (hty & hdef & harr). + eapply reflectT_change_left. { split; intros XE. 1: constructor; now apply XE. now depelim XE. } + eapply reflectT_change_left4. { split; intros XE. 1: destruct XE as [XE1 XE2 XE3 XE4]; constructor; [apply XE1|apply XE3|apply XE4|apply XE2]. now depelim XE. } + 3,4: ih => //; constructor; assumption. + + rewrite andb_true_r. eapply reflect_reflectT. eauto. + + eapply All_reflect_reflect_All2. 2: now apply forallb_All. + solve_all. Qed. - Lemma eqb_term_upto_univ_impl (equ lequ : _ -> _ -> bool) - (gen_compare_global_instance : (Universe.t -> Universe.t -> bool) -> global_reference -> nat -> list Level.t -> list Level.t -> bool) - (p : Universe.t -> bool) (q : nat -> term -> bool) Σ Re Rle napp: - (forall u u', p u -> p u' -> reflect (Re u u') (equ u u')) -> - (forall u u', p u -> p u' -> reflect (Rle u u') (lequ u u')) -> - (forall R leq, - (forall u u', p u -> p u' -> reflect (R u u') (leq u u')) -> - forall ref n l l' , - forallb p (map Universe.make l) -> - forallb p (map Universe.make l') -> - R_global_instance Σ Re R ref n l l' <-> - gen_compare_global_instance leq ref n l l') -> +Lemma eqb_term_upto_univ_impl Σ (p : Universe.t -> bool) (q : nat -> term -> bool) cmpu cmps cmp_universe cmp_sort + (gen_compare_global_instance : conv_pb -> global_reference -> nat -> list Level.t -> list Level.t -> bool) + pb napp : + (forall u u', p u -> p u' -> reflect (cmp_universe Conv u u') (cmpu Conv u u')) -> + (forall u u', p u -> p u' -> reflect (cmp_universe pb u u') (cmpu pb u u')) -> + (forall s s', Sort.on_sort p true s -> Sort.on_sort p true s' -> reflect (cmp_sort Conv s s') (cmps Conv s s')) -> + (forall s s', Sort.on_sort p true s -> Sort.on_sort p true s' -> reflect (cmp_sort pb s s') (cmps pb s s')) -> + (forall gr napp ui ui', forallb p (map Universe.make' ui) -> forallb p (map Universe.make' ui') -> reflect (cmp_global_instance Σ cmp_universe Conv gr napp ui ui') (gen_compare_global_instance Conv gr napp ui ui')) -> + (forall gr napp ui ui', forallb p (map Universe.make' ui) -> forallb p (map Universe.make' ui') -> reflect (cmp_global_instance Σ cmp_universe pb gr napp ui ui') (gen_compare_global_instance pb gr napp ui ui')) -> forall t t', on_universes p q t -> on_universes p q t' -> - eqb_term_upto_univ_napp equ lequ (gen_compare_global_instance) napp t t' -> - eq_term_upto_univ_napp Σ Re Rle napp t t'. + eqb_term_upto_univ_napp cmpu cmps gen_compare_global_instance pb napp t t' -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp t t'. Proof. - intros he hle Hcompare t t' ht ht'. - case: (reflect_eq_term_upto_univ Σ equ lequ p q Re Rle) => //; eauto. + intros he hle hs hspb hgl hglpb t t' ht ht'. + eapply elimT, reflect_eq_term_upto_univ. + all: eassumption. Qed. - -Definition compare_global_instance_proper Σ equ equ' eqlu eqlu' : - (forall u u', equ u u' = equ' u u') -> - (forall u u', eqlu u u' = eqlu' u u') -> - forall ref n l l', - compare_global_instance Σ equ eqlu ref n l l' = - compare_global_instance Σ equ' eqlu' ref n l l'. +Definition compare_global_instance_proper lookup cmpu cmpu' : + (forall pb u u', cmpu pb u u' = cmpu' pb u u') -> + forall ref pb n l l', + compare_global_instance lookup cmpu pb ref n l l' = + compare_global_instance lookup cmpu' pb ref n l l'. Proof. - intros Hequ Heqlu ref n l l'. + intros Hequ ref pb n l l'. apply eq_true_iff_eq. etransitivity. - symmetry. eapply reflect_iff. - eapply reflect_R_global_instance' with (p := xpredT); intros; eauto. + eapply reflect_cmp_global_instance' with (p := xpredT) (cmp_universe := cmpu); intros; eauto. 1-2: apply idP. 1-2: apply forallb_true. - eapply reflect_iff. - eapply reflect_R_global_instance' with (p := xpredT); intros; eauto. + eapply reflect_cmp_global_instance' with (p := xpredT); intros; eauto. 3-4: apply forallb_true. - + rewrite Hequ. destruct equ'; constructor; eauto. - + rewrite Heqlu. destruct eqlu'; constructor; eauto. + 1-2: rewrite Hequ; apply idP. Defined. -Definition eqb_term_upto_univ_proper Σ equ equ' eqlu eqlu' -(gen_compare_global_instance gen_compare_global_instance' : (Universe.t -> Universe.t -> bool) -> global_reference -> nat -> list Level.t -> list Level.t -> bool) -napp (t u : term) : -(forall u u', wf_universe Σ u -> wf_universe Σ u' -> equ u u' = equ' u u') -> -(forall u u', wf_universe Σ u -> wf_universe Σ u' -> eqlu u u' = eqlu' u u') -> -(forall leq ref n l l', compare_global_instance (lookup_env Σ) equ leq ref n l l' = - gen_compare_global_instance leq ref n l l') -> -(forall leq ref n l l', gen_compare_global_instance leq ref n l l' = - gen_compare_global_instance' leq ref n l l') -> -wf_universes Σ t -> wf_universes Σ u -> -eqb_term_upto_univ_napp equ eqlu gen_compare_global_instance napp t u = -eqb_term_upto_univ_napp equ' eqlu' gen_compare_global_instance' napp t u. +Definition eqb_term_upto_univ_proper Σ cmpu cmpu' cmps cmps' + (gen_compare_global_instance gen_compare_global_instance' : conv_pb -> global_reference -> nat -> list Level.t -> list Level.t -> bool) + pb napp (t u : term) : + (forall pb u u', wf_universe Σ u -> wf_universe Σ u' -> cmpu pb u u' = cmpu' pb u u') -> + (forall pb s s', wf_sort Σ s -> wf_sort Σ s' -> cmps pb s s' = cmps' pb s s') -> + (forall ref pb n l l', compare_global_instance (lookup_env Σ) cmpu pb ref n l l' = gen_compare_global_instance pb ref n l l') -> + (forall ref pb n l l', gen_compare_global_instance ref pb n l l' = + gen_compare_global_instance' ref pb n l l') -> + wf_universes Σ t -> wf_universes Σ u -> + eqb_term_upto_univ_napp cmpu cmps gen_compare_global_instance pb napp t u = + eqb_term_upto_univ_napp cmpu' cmps' gen_compare_global_instance' pb napp t u. Proof. -intros Hequ Heqlu Hcompare Hgen_compare Ht Hu. apply eq_true_iff_eq. split; intros. -- eapply introT. - * eapply reflect_eq_term_upto_univ; intros. - 1-2: apply idP. - + apply reflect_iff. rewrite <- Hgen_compare, <- Hcompare. - eapply reflect_R_global_instance with (p := wf_universeb Σ); eauto. - 1: intros; rewrite <- Hequ; eauto. all: try apply idP. - ++ revert H2. eapply reflect_iff; eapply wf_universe_reflect. - ++ revert H3. eapply reflect_iff; eapply wf_universe_reflect. - + eassumption. - + eassumption. - * eapply elimT. 2: exact H. - eapply reflect_eq_term_upto_univ; intros; try assumption. - 1: intros; rewrite <- Hequ; try apply idP. - 1: revert H0; eapply reflect_iff; eapply wf_universe_reflect. - 1: revert H1; eapply reflect_iff; eapply wf_universe_reflect. - 1: intros; rewrite <- Heqlu; try apply idP. - 1: revert H0; eapply reflect_iff; eapply wf_universe_reflect. - 1: revert H1; eapply reflect_iff; eapply wf_universe_reflect. - 2-3: eassumption. - + apply reflect_iff. rewrite <- Hcompare. - eapply reflect_R_global_instance; eauto. - 1: intros; rewrite <- Hequ; try apply idP. - 1: revert H2; eapply reflect_iff; eapply wf_universe_reflect. - 1: revert H3; eapply reflect_iff; eapply wf_universe_reflect. - - eapply introT. - * eapply reflect_eq_term_upto_univ; intros. - 1-2: apply idP. - 2-3: eassumption. - + apply reflect_iff. rewrite <- Hcompare. - eapply reflect_R_global_instance; eauto. - intros; rewrite Hequ; try apply idP. - 1: revert H2; eapply reflect_iff; eapply wf_universe_reflect. - 1: revert H3; eapply reflect_iff; eapply wf_universe_reflect. - * eapply elimT. 2: exact H. - eapply reflect_eq_term_upto_univ; intros; try eassumption. - 1: intros; rewrite <- Hequ; try apply idP. - 1: revert H0; eapply reflect_iff; eapply wf_universe_reflect. - 1: revert H1; eapply reflect_iff; eapply wf_universe_reflect. - 1: intros; rewrite <- Heqlu; try apply idP. - 1: revert H0; eapply reflect_iff; eapply wf_universe_reflect. - 1: revert H1; eapply reflect_iff; eapply wf_universe_reflect. - + apply reflect_iff. rewrite <- Hgen_compare, <- Hcompare. - eapply reflect_R_global_instance; eauto. - intros; rewrite Hequ; try apply idP. - 1: revert H2; eapply reflect_iff; eapply wf_universe_reflect. - 1: revert H3; eapply reflect_iff; eapply wf_universe_reflect. + intros Hequ Heqs Hcompare Hgen_compare Ht Hu. apply eq_true_iff_eq. + eassert _ as X. + 2: split; [apply introT|apply elimT]; apply X. + eapply reflectT_change_left. + 2: eapply reflect_eq_term_upto_univ with (cmp_universe := cmpu) (cmp_sort := cmps). + 1: eassert _ as X; [eapply reflect_eq_term_upto_univ with (cmp_universe := cmpu) (cmp_sort := cmps) | split; [apply introT|apply elimT]; eapply X]. + all: intros; eauto. + 1-4: apply idP. + 1-2: rewrite -Hcompare; eapply reflect_cmp_global_instance; intros; eauto using idP. + 1-2: rewrite Hequ; eauto using idP. + 1-4: now apply/wf_universe_reflect. + 1-2: rewrite Heqs; eauto using idP. + 1-4: now apply/wf_sort_reflect. + 1-2: rewrite -Hgen_compare -Hcompare; eapply reflect_cmp_global_instance; intros; eauto using idP. Defined. Lemma compare_global_instance_refl : - forall Σ (eqb leqb : Universe.t -> Universe.t -> bool) gr napp u, - (forall u, eqb u u) -> - (forall u, leqb u u) -> - compare_global_instance Σ eqb leqb gr napp u u. + forall Σ (cmpu : conv_pb -> Universe.t -> Universe.t -> bool) gr pb napp u, + (forall u, cmpu Conv u u) -> + (forall u, cmpu pb u u) -> + compare_global_instance Σ cmpu pb gr napp u u. Proof. - intros Σ eqb leqb gr napp u eqb_refl leqb_refl. + intros Σ cmpu gr pb napp u cmpu_refl cmpu_refl'. rewrite /compare_global_instance. - destruct global_variance_gen as [v|]. - - induction u in v |- *; destruct v; simpl; auto. - rtoProp. split; auto. - destruct t; simpl; auto. - - rewrite /compare_universe_instance. - rewrite forallb2_map; eapply forallb2_refl; intro; apply eqb_refl. + destruct global_variance_gen as [| |v]. + - apply All2_forallb2, All2_refl; auto. + - apply eqb_refl. + - apply orb_true_iff. right. + apply All2_forallb2, All2_refl; auto. Qed. Lemma All_All2_diag {A} {P : A -> A -> Prop} {l} : @@ -921,174 +622,169 @@ Proof. induction 1; constructor; auto. Qed. -Lemma R_global_instance_refl_wf Σ (Re Rle : Universe.t -> Universe.t -> Prop) gr napp l : -(forall u, wf_universe Σ u -> Re u u) -> -(forall u, wf_universe Σ u -> Rle u u) -> -forallb (wf_universeb Σ) (map Universe.make l) -> -R_global_instance Σ Re Rle gr napp l l. +Lemma cmp_universe_instance_refl_wf Σ (cmp_universe : Universe.t -> Universe.t -> Prop) l : + (forall u, wf_universe Σ u -> cmp_universe u u) -> + forallb (wf_universeb Σ) (map Universe.make' l) -> + cmp_universe_instance cmp_universe l l. Proof. - intros rRE rRle Hl. - rewrite /R_global_instance_gen. - destruct global_variance_gen as [v|] eqn:lookup. - - induction l in v , Hl |- *; simpl; auto. - apply andb_and in Hl as [? Hl]. revert a H. move => ? /wf_universe_reflect ?. - unfold R_opt_variance in IHl; destruct v; simpl; auto. - split; auto. - destruct t; simpl; eauto. - - eapply Forall_Forall2_diag. eapply forallb_Forall in Hl. - eapply Forall_impl; eauto. - move => ? /wf_universe_reflect ?; eauto. + intros rRE Hl. + unfold cmp_universe_instance. solve_all. + eapply All_All2; tea. intros. apply rRE. + now apply/wf_universe_reflect. Qed. -Definition eq_term_upto_univ_refl_wf Σ (Re Rle : Universe.t -> Universe.t -> Prop) napp : - (forall u, wf_universe Σ u -> Re u u) -> - (forall u, wf_universe Σ u -> Rle u u) -> - forall t, wf_universes Σ t -> eq_term_upto_univ_napp Σ Re Rle napp t t. +Lemma cmp_global_instance_refl_wf Σ (cmp_universe : conv_pb -> Universe.t -> Universe.t -> Prop) gr pb napp l : + (forall u, wf_universe Σ u -> cmp_universe Conv u u) -> + forallb (wf_universeb Σ) (map Universe.make' l) -> + cmp_global_instance Σ cmp_universe pb gr napp l l. Proof. - intros hRe hRle t wt. - induction t in napp, Rle, hRle, wt |- * using term_forall_list_ind. - all: repeat (apply andb_and in wt as [? wt]). + intros rRE Hl. + rewrite /cmp_global_instance_gen. + destruct global_variance_gen as [| |v] => //=. 2: left. all: now eapply cmp_universe_instance_refl_wf. +Qed. + +Definition eq_term_upto_univ_refl_wf Σ (cmp_universe : forall _ _ _, Prop) (cmp_sort : forall _ _ _, Prop) pb napp : + (forall u, wf_universe Σ u -> cmp_universe Conv u u) -> + (forall s, wf_sort Σ s -> cmp_sort Conv s s) -> + (forall s, wf_sort Σ s -> cmp_sort pb s s) -> + forall t, wf_universes Σ t -> eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp t t. +Proof. + intros hU hS hS' t wt. + induction t in pb, napp, hS', wt |- * using term_forall_list_ind. + all: repeat (cbn in wt; apply andb_and in wt as [? wt]). all: try constructor. all: eauto. - apply forallb_All in wt; eapply All_mix in wt; try exact X; eapply All_All2 ; try exact wt; intros ? [? ?]; eauto. - - revert s wt; move => ? /wf_universe_reflect ?; eauto. - - apply forallb_All in wt. apply All2_Forall2. induction u; cbn. - + eapply All2_nil. - + cbn in wt. inversion wt; subst. eapply All2_cons; eauto. - clear -a H0 hRe; revert a H0; move => ? /wf_universe_reflect ?; eauto. - - apply R_global_instance_refl_wf; auto. - - apply R_global_instance_refl_wf; auto. + - revert s wt; move => ? /wf_sort_reflect ?; eauto. + - eapply cmp_universe_instance_refl_wf; eauto. + - apply cmp_global_instance_refl_wf; auto. + - apply cmp_global_instance_refl_wf; auto. - destruct X as [? [? ?]]. unfold eq_predicate. repeat split; eauto. - + eapply forallb_All in H0. eapply All_mix in H0; try apply a. clear a. eapply All_All2; eauto. - clear H0. cbn; intros x [Hx Hclose]. apply Hx; eauto. - + unfold R_universe_instance. clear - H hRe. apply Forall_Forall2_diag. - apply forallb_Forall in H. eapply Forall_impl; eauto. - move => ? /wf_universe_reflect ?; eauto. - + eapply onctx_eq_ctx in a0; eauto. - - eapply forallb_All in wt; eapply All_mix in X0; try apply wt. - clear wt. eapply All_All2; eauto; simpl; intuition eauto. - + eapply onctx_eq_ctx in a0; eauto. - + eapply b0; eauto. apply andb_and in a as [? a]. exact a. + + solve_all. eapply All_All2; tea. solve_all. + + eapply cmp_universe_instance_refl_wf; eauto. + + reflexivity. + - solve_all. + eapply All_All2; tea. solve_all. + unfold test_branch in *. rtoProp. + split. 1: reflexivity. + eapply b0; tea. - eapply forallb_All in wt; eapply All_mix in X; try apply wt; clear wt. eapply All_All2; eauto; simpl; intuition eauto; apply andb_and in a as [? ?]; eauto. - eapply forallb_All in wt; eapply All_mix in X; try apply wt; clear wt. eapply All_All2; eauto; simpl; intuition eauto; apply andb_and in a as [? ?]; eauto. - - destruct p as [? []]; cbn -[Universe.make] in X, wt; rtoProp; intuition eauto; constructor; eauto. - + eapply hRe. now move/wf_universe_reflect: H. + - destruct p as [? []]; cbn -[Universe.make'] in X, wt; rtoProp; intuition eauto; constructor; eauto. + + eapply hU. now move/wf_universe_reflect: H. + solve_all. eapply All_All2; eauto; simpl; intuition eauto. Defined. -Lemma eqb_term_upto_univ_refl Σ (eqb leqb : Universe.t -> Universe.t -> bool) (Re : Universe.t -> Universe.t -> Prop) - (gen_compare_global_instance : (Universe.t -> Universe.t -> bool) -> global_reference -> nat -> list Level.t -> list Level.t -> bool) - napp t : - (forall u, wf_universe Σ u -> eqb u u) -> - (forall u, wf_universe Σ u -> leqb u u) -> - (forall u u', wf_universe Σ u -> wf_universe Σ u' -> reflect (Re u u') (eqb u u')) -> - (forall R leq , - (forall u u', wf_universe Σ u -> wf_universe Σ u' -> reflect (R u u') (leq u u')) -> - forall ref n l l' , - forallb (wf_universeb Σ) (map Universe.make l) -> - forallb (wf_universeb Σ) (map Universe.make l') -> - R_global_instance Σ Re R ref n l l' <-> - gen_compare_global_instance leq ref n l l') -> +Lemma eqb_term_upto_univ_refl Σ (cmpu : forall _ _ _, bool) (cmps : forall _ _ _, bool) gen_compare_global_instance pb napp t : + (forall u, wf_universe Σ u -> cmpu Conv u u) -> + (forall s, wf_sort Σ s -> cmps Conv s s) -> + (forall s, wf_sort Σ s -> cmps pb s s) -> + (forall gr napp ui ui', forallb (wf_universeb Σ) (map Universe.make' ui) -> forallb (wf_universeb Σ) (map Universe.make' ui') -> reflect (cmp_global_instance Σ cmpu Conv gr napp ui ui') (gen_compare_global_instance Conv gr napp ui ui')) -> + (forall gr napp ui ui', forallb (wf_universeb Σ) (map Universe.make' ui) -> forallb (wf_universeb Σ) (map Universe.make' ui') -> reflect (cmp_global_instance Σ cmpu pb gr napp ui ui') (gen_compare_global_instance pb gr napp ui ui')) -> wf_universes Σ t -> - eqb_term_upto_univ_napp eqb leqb gen_compare_global_instance napp t t. + eqb_term_upto_univ_napp cmpu cmps gen_compare_global_instance pb napp t t. Proof. - intros eqb_refl leqb_refl eqRe Hcompare Ht. - case: (reflect_eq_term_upto_univ Σ eqb leqb (wf_universeb Σ) closedu Re leqb gen_compare_global_instance napp _ _ _ t t) => //; eauto. - - move => ? ? /wf_universe_reflect ? - /wf_universe_reflect ?. now apply eqRe. - - intros; apply/idP. - - intros; apply Hcompare; eauto. move => ? ? /wf_universe_reflect ? - /wf_universe_reflect ?; now apply X. - - unshelve epose proof (eq_term_upto_univ_refl_wf Σ Re leqb napp _ _ t); eauto. - + intros u Hu; pose proof (eqb_refl u Hu) as H. revert H. apply reflect_iff; eauto. - + intro H. destruct (H (X Ht)). + intros eqb_refl leqb_refl eqRe Hglobal Hglobal' Ht. + eapply introT. + 2: eapply eq_term_upto_univ_refl_wf; eauto. + 1: eapply reflect_eq_term_upto_univ with (p := wf_universeb Σ) (cmp_universe := cmpu) (cmp_sort := cmps); eauto. + 1-4: intros; apply idP. + all: cbn; tas. Qed. -(* genereic equality for contexts *) - -Fixpoint eqb_ctx_gen equ gen_compare_global_instance - (Γ Δ : context) : bool := - match Γ, Δ with - | [], [] => true - | {| decl_name := na1 ; decl_body := None ; decl_type := t1 |} :: Γ, - {| decl_name := na2 ; decl_body := None ; decl_type := t2 |} :: Δ => - eqb_binder_annot na1 na2 && eqb_term_upto_univ equ equ gen_compare_global_instance t1 t2 && eqb_ctx_gen equ gen_compare_global_instance Γ Δ - | {| decl_name := na1 ; decl_body := Some b1 ; decl_type := t1 |} :: Γ, - {| decl_name := na2 ; decl_body := Some b2 ; decl_type := t2 |} :: Δ => - eqb_binder_annot na1 na2 && eqb_term_upto_univ equ equ gen_compare_global_instance b1 b2 && eqb_term_upto_univ equ equ gen_compare_global_instance t1 t2 && eqb_ctx_gen equ gen_compare_global_instance Γ Δ - | _, _ => false - end. +(* generic equality for contexts *) - Lemma eqb_binder_annot_spec {A} na na' : eqb_binder_annot (A:=A) na na' -> eq_binder_annot (A:=A) na na'. - Proof. case: eqb_annot_reflect => //. Qed. - - Lemma reflect_eqb_ctx_gen Σ equ - (p : Universe.t -> bool) (q : nat -> term -> bool) - (Re : Universe.t -> Universe.t -> Prop) - (gen_compare_global_instance : (Universe.t -> Universe.t -> bool) -> global_reference -> nat -> list Level.t -> list Level.t -> bool) : - (forall u u', p u -> p u' -> reflect (Re u u') (equ u u')) -> - (forall R leq, - (forall u u', p u -> p u' -> reflect (R u u') (leq u u')) -> - forall ref n l l' , - forallb p (map Universe.make l) -> - forallb p (map Universe.make l') -> - R_global_instance Σ Re R ref n l l' <-> - gen_compare_global_instance leq ref n l l') -> +Definition eqb_decl_gen eqb_term pb (d d' : context_decl) : bool := + match d, d' with + | {| decl_name := na; decl_body := None; decl_type := T |}, + {| decl_name := na'; decl_body := None; decl_type := T' |} => + eqb_binder_annot na na' && eqb_term pb T T' + | {| decl_name := na; decl_body := Some b; decl_type := T |}, + {| decl_name := na'; decl_body := Some b'; decl_type := T' |} => + eqb_binder_annot na na' && eqb_term Conv b b' && eqb_term pb T T' + | _, _ => false + end. + +Notation eqb_context_gen eqb_term pb := (forallb2 (eqb_decl_gen eqb_term pb)). + +Definition eqb_ctx_upto cmpu cmps gen_compare_global_instance pb : context -> context -> bool := + eqb_context_gen (fun pb => eqb_term_upto_univ cmpu cmps gen_compare_global_instance pb) pb. + +Lemma eqb_binder_annot_spec {A} na na' : eqb_binder_annot (A:=A) na na' -> eq_binder_annot (A:=A) na na'. +Proof. case: eqb_annot_reflect => //. Qed. + +Section reflectContext. + Context Σ (p : Universe.t -> bool) (q : nat -> term -> bool) cmpu cmps cmp_universe cmp_sort + (gen_compare_global_instance : conv_pb -> global_reference -> nat -> list Level.t -> list Level.t -> bool) + pb + (hu : forall u u', p u -> p u' -> reflect (cmp_universe Conv u u') (cmpu Conv u u')) + (hu' : forall u u', p u -> p u' -> reflect (cmp_universe pb u u') (cmpu pb u u')) + (hs : forall s s', Sort.on_sort p true s -> Sort.on_sort p true s' -> reflect (cmp_sort Conv s s') (cmps Conv s s')) + (hs' : forall s s', Sort.on_sort p true s -> Sort.on_sort p true s' -> reflect (cmp_sort pb s s') (cmps pb s s')) + (hglobal : forall gr napp ui ui', forallb p (map Universe.make' ui) -> forallb p (map Universe.make' ui') -> reflect (cmp_global_instance Σ cmp_universe Conv gr napp ui ui') (gen_compare_global_instance Conv gr napp ui ui')) + (hglobal' : forall gr napp ui ui', forallb p (map Universe.make' ui) -> forallb p (map Universe.make' ui') -> reflect (cmp_global_instance Σ cmp_universe pb gr napp ui ui') (gen_compare_global_instance pb gr napp ui ui')). + + Lemma reflect_eqb_decl_gen : + forall d d', + on_decl_universes p q d -> + on_decl_universes p q d' -> + reflectT (compare_decls (fun pb => eq_term_upto_univ Σ cmp_universe cmp_sort pb) pb d d') + (eqb_decl_gen (fun pb => eqb_term_upto_univ cmpu cmps gen_compare_global_instance pb) pb d d'). + Proof. + unfold on_decl_universes, compare_decls. + move => [na [b|] A] [na' [b'|] A'] /= ond ond'. + 2,3: constructor; intro X; depelim X. + all: rtoProp. + - eapply reflectT_change_left3. 1: { split; intros XE. 1: destruct XE as [XE1 XE2 XE3]; constructor; [apply XE1|apply XE2|apply XE3]. now depelim XE. } + + apply reflect_reflectT, eqb_annot_reflect. + + eapply reflect_eq_term_upto_univ; tea. + + eapply reflect_eq_term_upto_univ; tea. + - eapply reflectT_change_left2. 1: { split; intros XE. 1: destruct XE as [XE1 XE2]; constructor; [apply XE1|apply XE2]. now depelim XE. } + + apply reflect_reflectT, eqb_annot_reflect. + + eapply reflect_eq_term_upto_univ; tea. + Qed. + + Lemma reflect_eqb_ctx_gen : forall Γ Δ, on_ctx_universes p q Γ -> on_ctx_universes p q Δ -> - eqb_ctx_gen equ gen_compare_global_instance Γ Δ -> - eq_context_upto Σ Re Re Γ Δ. + reflectT (eq_context_upto Σ cmp_universe cmp_sort pb Γ Δ) + (eqb_ctx_upto cmpu cmps gen_compare_global_instance pb Γ Δ). Proof. - intros Hequ Hcompare Γ Δ hΓ hΔ h. - induction Γ as [| [na [b|] A] Γ ih ] in Δ, hΓ, hΔ, h |- *. - all: destruct Δ as [| [na' [b'|] A'] Δ]. - all: try discriminate. - - constructor. - - simpl in h. apply andb_andI in h as [[[h1 h2]%andb_and h3]%andb_and h4]. - simpl in hΓ. apply andb_andI in hΓ as [[hΓ1 hΓ2]%andb_and hΓ]. - simpl in hΔ. apply andb_andI in hΔ as [[hΔ1 hΔ2]%andb_and hΔ]. cbn in *. - constructor; auto. constructor; auto. - + now apply eqb_binder_annot_spec in h1. - + eapply eqb_term_upto_univ_impl; eauto. - + eapply eqb_term_upto_univ_impl; eauto. - - simpl in h. apply andb_and in h as [[h1 h2]%andb_and h3]. - simpl in hΓ. apply andb_andI in hΓ as [[hΓ1 hΓ2]%andb_and hΓ]. - simpl in hΔ. apply andb_andI in hΔ as [[hΔ1 hΔ2]%andb_and hΔ]. cbn in *. - constructor; auto; constructor. - + now apply eqb_binder_annot_spec. - + eapply eqb_term_upto_univ_impl; eauto. + intros. + eapply reflectT_change_left. 1: { split; apply All2_fold_All2. } + eapply All_reflect_reflect_All2. 2: apply forallb_All in H0; apply H0. + apply forallb_All in H. solve_all. + now apply reflect_eqb_decl_gen. Qed. - - Definition eqb_ctx_gen_proper (Σ:global_env_ext) equ equ' gen_compare_global_instance - gen_compare_global_instance' (Γ Δ : context) : - (forall u u', wf_universe Σ u -> wf_universe Σ u' -> equ u u' = equ' u u') -> - (forall leq ref n l l', compare_global_instance (lookup_env Σ) equ leq ref n l l' = - gen_compare_global_instance leq ref n l l') -> - (forall leq ref n l l', compare_global_instance (lookup_env Σ) equ' leq ref n l l' = - gen_compare_global_instance' leq ref n l l') -> - (forall leq ref n l l', gen_compare_global_instance' leq ref n l l' = - gen_compare_global_instance leq ref n l l') -> +End reflectContext. + +Definition eqb_ctx_gen_proper Σ cmpu cmpu' cmps cmps' + (gen_compare_global_instance gen_compare_global_instance' : conv_pb -> global_reference -> nat -> list Level.t -> list Level.t -> bool) + pb : + (forall pb u u', wf_universe Σ u -> wf_universe Σ u' -> cmpu pb u u' = cmpu' pb u u') -> + (forall pb s s', wf_sort Σ s -> wf_sort Σ s' -> cmps pb s s' = cmps' pb s s') -> + (forall ref pb n l l', compare_global_instance (lookup_env Σ) cmpu pb ref n l l' = gen_compare_global_instance pb ref n l l') -> + (forall ref pb n l l', gen_compare_global_instance ref pb n l l' = + gen_compare_global_instance' ref pb n l l') -> + forall Γ Δ, wf_ctx_universes Σ Γ -> wf_ctx_universes Σ Δ -> - eqb_ctx_gen equ gen_compare_global_instance Γ Δ = - eqb_ctx_gen equ' gen_compare_global_instance' Γ Δ. - Proof. - revert Δ; induction Γ; destruct Δ; simpl; eauto. - intros Hequ Hcompare Hcompare' Hgen_compare HΓ HΔ. - destruct a. destruct decl_body. - all: destruct c; destruct decl_body; eauto; cbn in *; - apply eq_true_iff_eq; split; intros. - all: repeat (let foo := fresh "H" in apply andb_and in H; destruct H as [H foo]); - repeat (let foo := fresh "HΓ" in apply andb_and in HΓ; destruct HΓ as [HΓ foo]); - repeat (let foo := fresh "HΔ" in apply andb_and in HΔ; destruct HΔ as [HΔ foo]); - repeat (apply andb_and; split; eauto); - try first[ rewrite <- IHΓ | rewrite IHΓ]; eauto. - all: erewrite <- eqb_term_upto_univ_proper; eauto. - all: intros; symmetry; eapply Hequ; eauto. - Defined. + eqb_ctx_upto cmpu cmps gen_compare_global_instance pb Γ Δ = + eqb_ctx_upto cmpu' cmps' gen_compare_global_instance' pb Γ Δ. +Proof. + intros hu hs hglob hglob'. + induction Γ; destruct Δ; simpl; eauto. + intros HΓ HΔ. rtoProp; f_equal; eauto. + destruct a as [na [b|] ty], c as [na' [b'|] ty']; cbnr. + all: unfold wf_decl_universes, on_decl_universes in *; rtoProp; cbn in *. + all: repeat apply (f_equal2 andb); cbnr. + all: eapply eqb_term_upto_univ_proper; tea. +Defined. (** Checking equality *) @@ -1200,10 +896,10 @@ Section EqualityDecGen. Qed. Lemma leq_universeP_gen leqb_level_n_gen - (leqb_correct : leqb_level_n_spec_gen uctx' leqb_level_n_gen) u u' : - wf_universe Σ u -> - wf_universe Σ u' -> - reflect (leq_universe Σ u u') (check_leqb_universe_gen leqb_level_n_gen u u'). + (leqb_correct : leqb_level_n_spec_gen uctx' leqb_level_n_gen) u u' : + wf_universe Σ u -> + wf_universe Σ u' -> + reflect (leq_universe Σ u u') (check_leqb_universe_gen leqb_level_n_gen u u'). Proof using hΣ. intros. apply (equivP idP) ; split. @@ -1220,48 +916,127 @@ Section EqualityDecGen. + now eapply global_ext_uctx_consistent. Qed. - Definition eqb_ctx leqb_level_n_gen := eqb_ctx_gen (check_eqb_universe_gen leqb_level_n_gen) (compare_global_instance (lookup_env Σ) (check_eqb_universe_gen leqb_level_n_gen)). + Definition check_cmpb_universe_gen leqb_level_n_gen := + (conv_pb_relb_gen (check_eqb_universe_gen leqb_level_n_gen) (check_leqb_universe_gen leqb_level_n_gen)). + + Lemma compare_universeP_gen leqb_level_n_gen + (leqb_correct : leqb_level_n_spec_gen uctx' leqb_level_n_gen) pb u u' : + wf_universe Σ u -> + wf_universe Σ u' -> + reflect (compare_universe Σ pb u u') (check_cmpb_universe_gen leqb_level_n_gen pb u u'). + Proof. + destruct pb. + - now apply eq_universeP_gen. + - now apply leq_universeP_gen. + Qed. + + Lemma eq_sortP_gen leqb_level_n_gen + (leqb_correct : leqb_level_n_spec_gen uctx' leqb_level_n_gen) + s s' : + wf_sort Σ s -> + wf_sort Σ s' -> + reflect (eq_sort Σ s s') (check_eqb_sort_gen leqb_level_n_gen s s'). + Proof using hΣ. + intros. destruct Σ as [Σ' φ]. + apply (equivP idP); split; sq. + all: pose proof hΣ as hΣ' ; sq. + - intros e. + eapply check_eqb_sort_spec_gen' + with (uctx := global_ext_uctx (Σ', φ)) in e ; eauto. + + now eapply wf_ext_global_uctx_invariants. + + now eapply global_ext_uctx_consistent. + - intros e. + eapply check_eqb_sort_complete_gen + with (uctx := global_ext_uctx (Σ', φ)); eauto. + + now eapply wf_ext_global_uctx_invariants. + + now eapply global_ext_uctx_consistent. + Qed. + + Lemma leq_sortP_gen leqb_level_n_gen + (leqb_correct : leqb_level_n_spec_gen uctx' leqb_level_n_gen) s s' : + wf_sort Σ s -> + wf_sort Σ s' -> + reflect (leq_sort Σ s s') (check_leqb_sort_gen leqb_level_n_gen s s'). + Proof using hΣ. + intros. + apply (equivP idP) ; split. + all: pose proof hΣ as hΣ' ; sq. + - intros e. + eapply check_leqb_sort_spec_gen' + with (uctx := global_ext_uctx Σ) in e ; eauto. + + now eapply wf_ext_global_uctx_invariants. + + now eapply global_ext_uctx_consistent. + - intros e. + eapply check_leqb_sort_complete_gen + with (uctx := global_ext_uctx Σ); eauto. + + now eapply wf_ext_global_uctx_invariants. + + now eapply global_ext_uctx_consistent. + Qed. + + Definition check_cmpb_sort_gen leqb_level_n_gen := + (conv_pb_relb_gen (check_eqb_sort_gen leqb_level_n_gen) (check_leqb_sort_gen leqb_level_n_gen)). + + Lemma compare_sortP_gen leqb_level_n_gen + (leqb_correct : leqb_level_n_spec_gen uctx' leqb_level_n_gen) pb s s' : + wf_sort Σ s -> + wf_sort Σ s' -> + reflect (compare_sort Σ pb s s') (check_cmpb_sort_gen leqb_level_n_gen pb s s'). + Proof. + destruct pb. + - now apply eq_sortP_gen. + - now apply leq_sortP_gen. + Qed. + + Definition eqb_ctx leqb_level_n_gen := + eqb_ctx_upto (check_cmpb_universe_gen leqb_level_n_gen) (check_cmpb_sort_gen leqb_level_n_gen) + (compare_global_instance (lookup_env Σ) (check_cmpb_universe_gen leqb_level_n_gen)). - Definition eqb_termp_napp pb leqb_level_n_gen := - eqb_termp_napp_gen pb (check_eqb_universe_gen leqb_level_n_gen) (check_leqb_universe_gen leqb_level_n_gen) - (compare_global_instance (lookup_env Σ) (check_eqb_universe_gen leqb_level_n_gen)). + Definition eqb_termp_napp leqb_level_n_gen := + eqb_term_upto_univ_napp (check_cmpb_universe_gen leqb_level_n_gen) (check_cmpb_sort_gen leqb_level_n_gen) + (compare_global_instance (lookup_env Σ) (check_cmpb_universe_gen leqb_level_n_gen)). Lemma reflect_eqb_termp_napp pb leqb_level_n_gen (leqb_correct : leqb_level_n_spec_gen uctx' leqb_level_n_gen) napp t u : wf_universes Σ t -> wf_universes Σ u -> - reflectT (eq_termp_napp pb Σ napp t u) (eqb_termp_napp pb leqb_level_n_gen napp t u). + reflectT (eq_termp_napp Σ pb napp t u) (eqb_termp_napp leqb_level_n_gen pb napp t u). Proof using hΣ. apply reflect_eq_term_upto_univ. - move => ? ? /wf_universe_reflect ? - /wf_universe_reflect ?. - now apply eq_universeP_gen. + now apply compare_universeP_gen. - move => ? ? /wf_universe_reflect ? - /wf_universe_reflect ?. - destruct pb. - + now apply eq_universeP_gen. - + now eapply leq_universeP_gen. - - intros. apply reflect_iff. - eapply reflect_R_global_instance with (p := wf_universeb Σ); eauto. - move => ? ? /wf_universe_reflect ? - /wf_universe_reflect ?; now apply eq_universeP_gen; eauto. + now apply compare_universeP_gen. + - move => ? ? /wf_sort_reflect ? - /wf_sort_reflect ?. + now apply compare_sortP_gen. + - move => ? ? /wf_sort_reflect ? - /wf_sort_reflect ?. + now apply compare_sortP_gen. + - intros. + eapply reflect_cmp_global_instance; eauto. + + move => ? ? /wf_universe_reflect ? - /wf_universe_reflect ?. + now apply compare_universeP_gen. + + move => ? ? /wf_universe_reflect ? - /wf_universe_reflect ?. + now apply compare_universeP_gen. + - intros. + eapply reflect_cmp_global_instance; eauto. + + move => ? ? /wf_universe_reflect ? - /wf_universe_reflect ?. + now apply compare_universeP_gen. + + move => ? ? /wf_universe_reflect ? - /wf_universe_reflect ?. + now apply compare_universeP_gen. Qed. Lemma eqb_termp_napp_spec pb leqb_level_n_gen (leqb_correct : leqb_level_n_spec_gen uctx' leqb_level_n_gen) napp t u : wf_universes Σ t -> wf_universes Σ u -> - eqb_termp_napp pb leqb_level_n_gen napp t u -> - eq_termp_napp pb Σ napp t u. + eqb_termp_napp leqb_level_n_gen pb napp t u -> + eq_termp_napp Σ pb napp t u. Proof using hΣ. - pose proof hΣ. - eapply eqb_term_upto_univ_impl. - - move => ? ? /wf_universe_reflect ? - /wf_universe_reflect ?; now apply eq_universeP_gen; eauto. - - destruct pb. - + move => ? ? /wf_universe_reflect ? - /wf_universe_reflect ?; now apply eq_universeP_gen; eauto. - + move => ? ? /wf_universe_reflect ? - /wf_universe_reflect ?; now apply leq_universeP_gen; eauto. - - intros. apply reflect_iff. eapply reflect_R_global_instance; eauto. - move => ? ? /wf_universe_reflect ? - /wf_universe_reflect ?; now apply eq_universeP_gen; eauto. + intros. + eapply elimT. 1: apply reflect_eqb_termp_napp. + all: eassumption. Qed. - Definition eqb_termp pb leq := (eqb_termp_napp pb leq 0). + Definition eqb_termp pb leq := (eqb_termp_napp leq pb 0). Definition eqb_term := (eqb_termp Conv). Definition leqb_term := (eqb_termp Cumul). @@ -1312,26 +1087,42 @@ Section EqualityDecGen. (leqb_correct : leqb_level_n_spec_gen uctx' leqb_level_n_gen) : forall t, wf_universes Σ t -> eqb_term leqb_level_n_gen t t. Proof using hΣ. - intro t. eapply eqb_term_upto_univ_refl with (Re := eq_universe Σ). - all: intros; try eapply check_eqb_universe_refl_gen; eauto. - - eapply eq_universeP_gen; eauto. - - apply reflect_iff. eapply reflect_R_global_instance; eauto. - + move => ? ? /wf_universe_reflect ? - /wf_universe_reflect ?; apply eq_universeP_gen; eauto. - + move => ? ? /wf_universe_reflect ? - /wf_universe_reflect ?; apply X; eauto. + intro t. eapply eqb_term_upto_univ_refl. + 4,5: intros; eapply reflect_cmp_global_instance; tea; intros; cbnr; try apply idP. + - intros. cbn. unfold check_eqb_universe_gen; toProp; left. toProp. right. apply eqb_refl. + - intros. eapply check_eqb_sort_refl_gen; eauto. + - intros. eapply check_eqb_sort_refl_gen; eauto. Qed. Lemma eqb_ctx_spec leqb_level_n_gen (leqb_correct : leqb_level_n_spec_gen uctx' leqb_level_n_gen) : - forall Γ Δ, + forall pb Γ Δ, wf_ctx_universes Σ Γ -> wf_ctx_universes Σ Δ -> - eqb_ctx leqb_level_n_gen Γ Δ -> - eq_context_upto Σ (eq_universe Σ) (eq_universe Σ) Γ Δ. + eqb_ctx leqb_level_n_gen pb Γ Δ -> + eq_context_upto Σ (compare_universe Σ) (compare_sort Σ) pb Γ Δ. Proof using hΣ. - intros Γ Δ hΓ hΔ h. eapply reflect_eqb_ctx_gen; eauto. - - move => ? ? /wf_universe_reflect ? - /wf_universe_reflect ?; now apply eq_universeP_gen; eauto. - - intros. apply reflect_iff. eapply reflect_R_global_instance; eauto. - move => ? ? /wf_universe_reflect ? - /wf_universe_reflect ?; now apply eq_universeP_gen; eauto. + intros pb Γ Δ hΓ hΔ h. eapply elimT. 1: eapply reflect_eqb_ctx_gen; eauto. 7: tea. + - move => ? ? /wf_universe_reflect ? - /wf_universe_reflect ?. + now apply compare_universeP_gen. + - move => ? ? /wf_universe_reflect ? - /wf_universe_reflect ?. + now apply compare_universeP_gen. + - move => ? ? /wf_sort_reflect ? - /wf_sort_reflect ?. + now apply compare_sortP_gen. + - move => ? ? /wf_sort_reflect ? - /wf_sort_reflect ?. + now apply compare_sortP_gen. + - intros. + eapply reflect_cmp_global_instance; eauto. + + move => ? ? /wf_universe_reflect ? - /wf_universe_reflect ?. + now apply compare_universeP_gen. + + move => ? ? /wf_universe_reflect ? - /wf_universe_reflect ?. + now apply compare_universeP_gen. + - intros. + eapply reflect_cmp_global_instance; eauto. + + move => ? ? /wf_universe_reflect ? - /wf_universe_reflect ?. + now apply compare_universeP_gen. + + move => ? ? /wf_universe_reflect ? - /wf_universe_reflect ?. + now apply compare_universeP_gen. Qed. (* diff --git a/safechecker/theories/PCUICSafeChecker.v b/safechecker/theories/PCUICSafeChecker.v index 8d2d670db..038c990c2 100644 --- a/safechecker/theories/PCUICSafeChecker.v +++ b/safechecker/theories/PCUICSafeChecker.v @@ -459,20 +459,24 @@ Section CheckEnv. Σ ;;; Γ |- it_mkProd_or_LetIn Δ (tSort s) : A -> isType Σ Γ (it_mkProd_or_LetIn Δ (tSort s)). Proof using Type. - unfold isType. unfold PCUICTypingDef.typing. intros Γ Δ s A h. revert Γ s A h. + intros Γ Δ s A h. + enough (∑ s', Σ ;;; Γ |- it_mkProd_or_LetIn Δ (tSort s) : tSort s'). + 1: destruct X as (? & ?); now eapply has_sort_isType. + revert Γ s A h. induction Δ using rev_ind; intros. - - simpl in h. eapply inversion_Sort in h as (?&?&?). + - simpl in h |- *. eapply inversion_Sort in h as (?&?&?). eexists; constructor; eauto. apply wfΣ. - destruct x as [na [b|] ty]; simpl in *; rewrite it_mkProd_or_LetIn_app /= /mkProd_or_LetIn /= in h *. - * eapply inversion_LetIn in h as [s' [? [? [? [? ?]]]]]; auto. - specialize (IHΔ _ _ _ t1) as [s'' vdefty]. + * eapply inversion_LetIn in h as [s' [? [? ?]]]; auto. + specialize (IHΔ _ _ _ t) as [s'' vdefty]. exists s''. eapply type_Cumul_alt. econstructor; eauto. pcuic. eapply red_cumul. repeat constructor. - * eapply inversion_Prod in h as [? [? [? [? ]]]]; auto. - specialize (IHΔ _ _ _ t0) as [s'' Hs'']. - eexists (Universe.sort_of_product x s''). + * eapply inversion_Prod in h as [? [? [h1 [? ?]]]]; auto. + specialize (IHΔ _ _ _ t) as [s'' Hs'']. + exists (Sort.sort_of_product x s''). + apply unlift_TypUniv in h1. eapply type_Cumul'; eauto. econstructor; pcuic. pcuic. reflexivity. Qed. @@ -481,7 +485,7 @@ Section CheckEnv. Γ Δ s (wfΓ : forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ wf_local Σ Γ ∥) : typing_result (forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ type_local_ctx (lift_typing typing) Σ Γ Δ s ∥) := match Δ with - | [] => match abstract_env_ext_wf_universeb X_ext s with true => ret _ | false => raise (Msg "Ill-formed universe") end + | [] => match abstract_env_ext_wf_sortb X_ext s with true => ret _ | false => raise (Msg "Ill-formed universe") end | {| decl_body := None; decl_type := ty |} :: Δ => checkΔ <- check_type_local_ctx X_ext Γ Δ s wfΓ ;; checkty <- check_type_wf_env X_ext (Γ ,,, Δ) _ ty (tSort s) ;; @@ -492,8 +496,8 @@ Section CheckEnv. ret _ end. Next Obligation. - erewrite <- abstract_env_ext_wf_universeb_correct in Heq_anonymous; eauto. - now sq; apply/PCUICWfUniverses.wf_universe_reflect. + erewrite <- abstract_env_ext_wf_sortb_correct in Heq_anonymous; eauto. + now sq; apply/PCUICWfUniverses.wf_sort_reflect. Qed. Next Obligation. specialize_Σ H. sq. now eapply PCUICContexts.type_local_ctx_wf_local in checkΔ. @@ -501,6 +505,7 @@ Section CheckEnv. Next Obligation. specialize_Σ H. sq. split; auto. + repeat (eexists; tea). Qed. Next Obligation. specialize_Σ H. sq. now eapply PCUICContexts.type_local_ctx_wf_local in checkΔ. @@ -508,7 +513,7 @@ Section CheckEnv. Next Obligation. pose proof (abstract_env_ext_wf _ H). specialize_Σ H. sq. split; auto. split; auto. - eapply PCUICValidity.validity in checkty; auto. + eapply PCUICValidity.validity in checkty as (_ & ?); auto. Qed. Program Fixpoint infer_sorts_local_ctx X_ext {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X_ext -> NormalizationIn Σ} Γ Δ (wfΓ : forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ wf_local Σ Γ ∥) : @@ -529,14 +534,14 @@ Section CheckEnv. specialize_Σ H. sq. now eapply PCUICContexts.sorts_local_ctx_wf_local in Δinfer. Qed. Next Obligation. - specialize_Σ H. sq. split; auto. + specialize_Σ H. sq. split; auto. repeat (eexists; tea). Qed. Next Obligation. specialize_Σ H. sq. now eapply PCUICContexts.sorts_local_ctx_wf_local in Δinfer. Qed. Next Obligation. pose proof (abstract_env_ext_wf _ H). specialize_Σ H. sq. split; auto. split; auto. - eapply PCUICValidity.validity in checkty; auto. + eapply PCUICValidity.validity in checkty as (_ & ?); auto. Qed. Definition cumul_decl Pcmp Σ Γ (d d' : context_decl) : Type := cumul_decls Pcmp Σ Γ Γ d d'. @@ -556,33 +561,23 @@ Section CheckEnv. typing_result (forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ ws_cumul_ctx_pb_rel le Σ Γ Δ Δ' ∥) := check_ws_cumul_ctx X_impl X_ext le Γ Δ Δ' wfΔ wfΔ'. - Notation eqb_term_conv X conv_pb := (eqb_term_upto_univ (abstract_env_eq X) (abstract_env_conv_pb_relb X conv_pb) (abstract_env_compare_global_instance _ X)). + Notation eqb_term_conv X conv_pb := (eqb_term_upto_univ (abstract_env_compare_universe X) (abstract_env_compare_sort X) (abstract_env_compare_global_instance _ X) conv_pb). Program Definition check_eq_term pb X_ext t u (wft : forall Σ, abstract_env_ext_rel X_ext Σ -> wf_universes Σ t) (wfu : forall Σ, abstract_env_ext_rel X_ext Σ -> wf_universes Σ u) : - typing_result (forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ compare_term pb Σ Σ t u ∥) := + typing_result (forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ compare_term Σ Σ pb t u ∥) := check <- check_eq_true (eqb_term_conv X_ext pb t u) (Msg "Terms are not equal") ;; ret _. Next Obligation. simpl in *; sq. - eapply eqb_term_upto_univ_impl in check; sq; eauto. - - intros u0 u'. repeat erewrite <- abstract_env_ext_wf_universeb_correct; eauto. - move => / wf_universe_reflect ? => /wf_universe_reflect ?. - apply iff_reflect. eapply (abstract_env_compare_universe_correct _ _ Conv); eauto. - - intros u0 u'. repeat erewrite <- abstract_env_ext_wf_universeb_correct; eauto. - move => /wf_universe_reflect ? => /wf_universe_reflect ?. - apply iff_reflect. eapply (abstract_env_compare_universe_correct _ _ pb); eauto. - - intros. apply compare_global_instance_correct; eauto. - + apply wf_universe_instance_iff. rewrite <- wf_universeb_instance_forall; eauto. - + apply wf_universe_instance_iff. rewrite <- wf_universeb_instance_forall; eauto. - Unshelve. all: eauto. + eapply cmpb_term_correct in check; sq; eauto. Qed. Program Definition check_eq_decl pb X_ext d d' (wfd : forall Σ, abstract_env_ext_rel X_ext Σ -> wf_decl_universes Σ d) (wfd' : forall Σ, abstract_env_ext_rel X_ext Σ -> wf_decl_universes Σ d') : - typing_result (forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ compare_decl pb Σ Σ d d' ∥) := + typing_result (forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ compare_decl Σ Σ pb d d' ∥) := match d, d' return (forall Σ, abstract_env_ext_rel X_ext Σ -> wf_decl_universes Σ d) -> (forall Σ, abstract_env_ext_rel X_ext Σ -> wf_decl_universes Σ d') -> typing_result _ with @@ -625,7 +620,7 @@ Section CheckEnv. Program Fixpoint check_compare_context (pb : conv_pb) X_ext Γ Δ (wfΓ : forall Σ, abstract_env_ext_rel X_ext Σ -> wf_ctx_universes Σ Γ) (wfΔ : forall Σ, abstract_env_ext_rel X_ext Σ -> wf_ctx_universes Σ Δ) : - typing_result (forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ PCUICEquality.compare_context pb Σ Σ Γ Δ ∥) := + typing_result (forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ PCUICEquality.compare_context Σ Σ pb Γ Δ ∥) := match Γ, Δ return (forall Σ, abstract_env_ext_rel X_ext Σ -> wf_ctx_universes Σ Γ) -> (forall Σ, abstract_env_ext_rel X_ext Σ -> wf_ctx_universes Σ Δ) -> typing_result _ with @@ -656,7 +651,7 @@ Section CheckEnv. Program Fixpoint check_leq_terms (pb : conv_pb) X_ext l l' (wfl : forall Σ, abstract_env_ext_rel X_ext Σ -> forallb (wf_universes Σ) l) (wfl' : forall Σ, abstract_env_ext_rel X_ext Σ -> forallb (wf_universes Σ) l') : - typing_result (forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ All2 (compare_term pb Σ Σ) l l' ∥) := + typing_result (forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ All2 (compare_term Σ Σ pb) l l' ∥) := match l, l' return (forall Σ, abstract_env_ext_rel X_ext Σ -> forallb (wf_universes Σ) l) -> (forall Σ, abstract_env_ext_rel X_ext Σ -> forallb (wf_universes Σ) l') -> _ with | [], [] => fun _ _ => ret (fun Σ wfΣ => sq _) @@ -794,14 +789,12 @@ Section CheckEnv. unfold arities_context. induction 1; simpl; auto. rewrite rev_map_cons /=. - eapply All_local_env_app; split. constructor; pcuic. + eapply All_local_env_app. constructor; pcuic. eapply All_local_env_impl; eauto. - move=> Γ t [] /=. - * move=> ty ht. eapply weaken_ctx; eauto. - constructor; pcuic. - * move=> [s Hs]; exists s. - eapply weaken_ctx; eauto. - constructor; pcuic. + intros Γ j Hj. + apply lift_typing_impl with (1 := Hj) => t T HT. + eapply weaken_ctx; eauto. + constructor; pcuic. Qed. Program Definition check_constructor X_ext {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X_ext -> NormalizationIn Σ} (ind : nat) (mdecl : mutual_inductive_body) @@ -858,7 +851,7 @@ Section CheckEnv. eapply eqb_eq in eqbindices. eapply eqb_eq in eqbpars. sq; red; cbn. intuition auto. - eexists; eauto. + now eapply has_sort_isType. symmetry in Heq_anonymous. eapply PCUICSubstitution.decompose_prod_assum_it_mkProd_or_LetIn in Heq_anonymous. simpl in Heq_anonymous. subst concl0. rewrite it_mkProd_or_LetIn_app. @@ -900,8 +893,8 @@ Section CheckEnv. isType Σ Γ (it_mkProd_or_LetIn Δ T) -> isType Σ (Γ ,,, Δ) T. Proof using Type. - move=> wfX [u Hu]. - exists u. unfold PCUICTypingDef.typing in *. + move=> wfX H. + apply lift_sorting_it_impl_gen with H => // Hu. now eapply PCUICSpine.inversion_it_mkProd_or_LetIn in Hu. Qed. @@ -910,7 +903,7 @@ Section CheckEnv. ∑ fty s, (Σ ;;; Γ |- f : fty) * typing_spine Σ Γ fty args (tSort s). Proof using Type. - intros [s Hs]. + intros (_ & s & Hs & _). eapply inversion_mkApps in Hs as [fty [Hf Hargs]]; auto. now exists fty, s. Qed. @@ -1028,8 +1021,8 @@ Section CheckEnv. apply on_free_vars_lift0. rewrite /app_context /snoc; len. replace (#|Δ| + S #|Γ|) with (S #|Δ| + #|Γ|). 2:lia. rewrite Nat.add_1_r. rewrite -shiftnP_add addnP_shiftnP. eapply on_free_vars_subst. - eapply isType_wf_local in isty. depelim isty. red in l0. cbn. - rewrite andb_true_r. red in l0. fvs. + eapply isType_wf_local in isty. depelim isty. apply unlift_TermTyp in l as l0. cbn. + rewrite andb_true_r. fvs. eapply isType_is_open_term in isty. cbn. now rewrite shiftnP_add. Qed. @@ -1065,7 +1058,7 @@ Section CheckEnv. This is a quite common operation in tactics. Making this work up-to unfolding of let-ins is the tricky part. *) - Lemma typing_spine_it_mkProd_or_LetIn_ext_list_inv_gen {Σ : global_env_ext} {wfX : wf Σ} {Γ Δ Δ' Δ'' s args s'} : + Lemma typing_spine_it_mkProd_or_LetIn_ext_list_inv_gen {Σ : global_env_ext} {wfX : wf Σ} {Γ Δ Δ' Δ'' : context} {s args s'} : wf_local Σ (Γ ,,, Δ ,,, Δ'') -> typing_spine Σ (Γ ,,, Δ ,,, Δ') (it_mkProd_or_LetIn (lift_context #|Δ ,,, Δ'| 0 (Δ ,,, Δ'')) (tSort s)) (to_extended_list_k Δ #|Δ'| ++ args) (tSort s') -> @@ -1078,10 +1071,10 @@ Section CheckEnv. * len in sp. now rewrite app_context_nil_l in sp. * set (decl := d) in *. - assert (wf_universe Σ s). + assert (wf_sort Σ s). { eapply typing_spine_isType_dom in sp. eapply isType_it_mkProd_or_LetIn_inv in sp; auto. - destruct sp as [? Hs]. now eapply inversion_Sort in Hs as [? []]. } + destruct sp as (_ & ? & Hs & _). now eapply inversion_Sort in Hs as [? []]. } destruct d as [na [b|] ty]. simpl in sp; unfold mkProd_or_LetIn in sp; simpl in *. - len in sp. rewrite lift_context_app /= it_mkProd_or_LetIn_app lift_context_app it_mkProd_or_LetIn_app /= @@ -1178,6 +1171,7 @@ Section CheckEnv. welltyped Σ Γ ty * welltyped Σ (Γ ,, vass na ty) ty'. Proof using Type. intros [s [s1 [s2 [hty [hty' hcum]]]]%inversion_Prod]; auto. + apply unlift_TypUniv in hty; cbn in hty. split; eexists; eauto. Qed. @@ -1187,7 +1181,8 @@ Section CheckEnv. welltyped Σ Γ b * welltyped Σ (Γ ,, vdef na b ty) t. Proof using Type. - intros [s [s1 [s2 [hty [hdef [hty' hcum]]]]]%inversion_LetIn]; auto. + intros [s [s1 [hty [hty' hcum]]]%inversion_LetIn]; auto. + destruct hty as (hdef & s2 & hty & _); cbn in *. split; [split|]; eexists; eauto. Qed. @@ -1195,8 +1190,8 @@ Section CheckEnv. welltyped Σ Γ (tLetIn na b ty t) -> welltyped Σ Γ (subst0 [b] t). Proof using Type. - intros [s [s1 [s2 [hty [hdef [hty' hcum]]]]]%inversion_LetIn]; auto. - exists (subst0 [b] s2). + intros [s [s1 [hdef [hty' hcum]]]%inversion_LetIn]; auto. + exists (subst0 [b] s1). now eapply substitution_let in hty'. Qed. @@ -1381,11 +1376,12 @@ Section CheckEnv. Next Obligation. pose proof (abstract_env_ext_wf _ H); specialize_Σ H. sq. constructor; eauto. - eapply validity in wfty. apply wfty. + split; auto. + eapply validity in wfty as []; auto. Qed. Next Obligation. pose proof (abstract_env_ext_wf _ H); specialize_Σ H. - sq. constructor; auto. now exists wfty. + sq. constructor; auto. repeat (eexists; tea). Qed. Definition wt_indices Σ mdecl indices cs := @@ -1402,7 +1398,7 @@ Section CheckEnv. Qed. (* Now in PCUIC *) - Lemma type_smash {Σ : global_env_ext} {wfΣ : wf Σ} {Γ Δ t T} : + Lemma type_smash {Σ : global_env_ext} {wfΣ : wf Σ} {Γ Δ : context} {t T} : Σ ;;; Γ ,,, Δ |- t : T -> Σ ;;; Γ ,,, smash_context [] Δ |- expand_lets Δ t : expand_lets Δ T. Proof using Type. @@ -1418,7 +1414,7 @@ Section CheckEnv. rewrite !subst_empty. eapply typing_wf_local in h. rewrite app_context_assoc in h. eapply All_local_env_app_inv in h as [h _]. - now depelim h. + depelim h. now eapply unlift_TermTyp. - intros Γ t T h. rewrite !smash_context_app_ass !expand_lets_vass. rewrite app_context_assoc. eapply X. lia. @@ -1441,65 +1437,6 @@ Section CheckEnv. eapply All2_fold_length in H. len in H. Qed. - Lemma eq_decl_eq_decl_upto (Σ : global_env_ext) x y : - compare_decl Cumul Σ Σ x y -> - eq_decl_upto_gen Σ (eq_universe Σ) (leq_universe Σ) x y. - Proof using Type. - intros []; constructor; intuition auto. cbn. constructor. - cbn. constructor; auto. - Qed. - - Lemma eq_decl_upto_refl (Σ : global_env_ext) x : eq_decl_upto_gen Σ (eq_universe Σ) (leq_universe Σ) x x. - Proof using Type. - destruct x as [na [b|] ty]; constructor; simpl; auto. - split; constructor; reflexivity. reflexivity. - split; constructor; reflexivity. reflexivity. - Qed. - - - Lemma eq_context_upto_cumul_context (Σ : global_env_ext) Re Rle : - RelationClasses.subrelation Re (eq_universe Σ) -> - RelationClasses.subrelation Rle (leq_universe Σ) -> - RelationClasses.subrelation Re Rle -> - CRelationClasses.subrelation (eq_context_upto Σ Re Rle) (fun Γ Γ' => cumul_context cumulSpec0 Σ Γ Γ'). - Proof using Type. - intros HRe HRle hR Γ Δ h. induction h. - - constructor. - - constructor; tas. - depelim p; constructor; auto. - eapply eq_term_upto_univ_cumulSpec. - eapply eq_term_upto_univ_impl. 5:eauto. all:tea. - now transitivity Rle. auto. - eapply eq_term_upto_univ_cumulSpec. - eapply eq_term_upto_univ_impl; eauto. - eapply eq_term_upto_univ_cumulSpec. - eapply eq_term_upto_univ_impl. 5:eauto. all:tea. - now transitivity Rle. auto. - Qed. - - Lemma eq_context_upto_univ_cumul_context {Σ : global_env_ext} Γ Δ : - eq_context_upto Σ.1 (eq_universe Σ) (leq_universe Σ) Γ Δ -> - cumul_context cumulSpec0 Σ Γ Δ. - Proof using Type. - intros h. eapply eq_context_upto_cumul_context; tea. - reflexivity. tc. tc. - Qed. - - Lemma leq_context_cumul_context (Σ : global_env_ext) Γ Δ Δ' : - PCUICEquality.compare_context Cumul Σ Σ Δ Δ' -> - cumul_ctx_rel cumulSpec0 Σ Γ Δ Δ'. - Proof using Type. - intros eqc. - apply cumul_ctx_rel_close'. - apply eq_context_upto_univ_cumul_context. - apply All2_eq_context_upto. - eapply All2_app. red in eqc. - eapply All2_fold_All2; eauto. - eapply All2_fold_impl; eauto. - intros; now eapply eq_decl_eq_decl_upto. - eapply All2_refl. intros. simpl. eapply (eq_decl_upto_refl Σ). - Qed. - Lemma wt_cstrs {n mdecl cstrs cs} X_ext : (forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ All2 (fun (cstr : constructor_body) (cs0 : constructor_univs) => @@ -1588,7 +1525,7 @@ Section CheckEnv. destruct (variance_universes_spec _ _ _ _ _ _ _ _ eqvu). exists univs', i, i'; split => //; now eapply eqb_eq in check_leq. - - sq. red. destruct univs; auto. exact tt. + - sq. red. destruct univs; auto. Qed. Definition Build_on_inductive_sq {X_ext ind mdecl} @@ -1647,7 +1584,7 @@ Section CheckEnv. Lemma wf_decl_universes_subst_instance Σ udecl udecl' d u : wf_ext (Σ, udecl) -> - wf_universe_instance (Σ, udecl') u -> + wf_instance (Σ, udecl') u -> wf_decl_universes (Σ, udecl) d -> wf_decl_universes (Σ, udecl') d@[u]. Proof using Type. @@ -1662,7 +1599,7 @@ Section CheckEnv. Lemma wf_ctx_universes_subst_instance Σ udecl udecl' Γ u : wf_ext (Σ, udecl) -> - wf_universe_instance (Σ, udecl') u -> + wf_instance (Σ, udecl') u -> wf_ctx_universes (Σ, udecl) Γ -> wf_ctx_universes (Σ, udecl') Γ@[u]. Proof using Type. @@ -1677,9 +1614,9 @@ Section CheckEnv. wf_local Σ Γ -> wf_ctx_universes Σ Γ. Proof using Type. intros wfΣ. unfold wf_ctx_universes. - induction 1 => //. cbn. rewrite IHX andb_true_r. - destruct t0 as [s Hs]. move/typing_wf_universes: Hs => /andP[] //. - cbn. move/typing_wf_universes: t1 => /=; rewrite /wf_decl_universes /on_decl_universes /= => -> /= //. + induction 1 => //=; rewrite IHX andb_true_r. + destruct t0 as (_ & s & Hs & _). move/typing_wf_universes: Hs => /andP[] //. + destruct t0 as (Hb & s & Hs & _). cbn in Hb. move/typing_wf_universes: Hb => /andP[] //. Qed. Lemma wf_ctx_universes_app {Σ Γ Δ} : @@ -1774,7 +1711,7 @@ Section CheckEnv. intros v0 [= <-]. red. rewrite -Heq_anonymous. split; auto. erewrite (abstract_env_irr _ _ wfΣ0); eauto. - now apply leq_context_cumul_context. + now eapply PCUICContextConversionTyp.eq_context_cumul_Spec_rel. clear check_args. eapply All2_impl. eauto. simpl; intros. erewrite (abstract_env_irr _ _ wfΣ0); eauto. now eapply eq_term_upto_univ_cumulSpec. @@ -1878,7 +1815,9 @@ Section CheckEnv. Qed. Next Obligation. epose proof (get_wt_indices X_ext wfar wfpars _ _ _ hnth heq Hcs). + destruct HX. specialize_Σ H. + pose proof (abstract_env_ext_wf _ (H1 _ H2)) as wfΣ. unfold check_constructor_spec in Hcs; simpl in *. sq. solve_all. eapply All2_impl; eauto. simpl. @@ -1888,7 +1827,11 @@ Section CheckEnv. - rewrite eq. rewrite it_mkProd_or_LetIn_app. autorewrite with len. lia_f_equal. rewrite /cstr_concl /=. f_equal. rewrite /cstr_concl_head. lia_f_equal. - - now destruct wtinds. + - destruct wtinds as (? & ci). + apply PCUICEnvTyping.ctx_inst_impl with (1 := ci) => t T HT. + split; auto. + unshelve eapply validity in HT as []; eauto. + apply wfΣ. - destruct lets_in_constructor_types; eauto. Qed. @@ -2009,15 +1952,15 @@ End monad_Alli_nth_forall. destruct ind_projs => //. Qed. - Definition checkb_constructors_smaller X_ext (cs : list constructor_univs) (ind_sort : Universe.t) := - List.forallb (List.forallb (fun argsort => abstract_env_leq X_ext argsort ind_sort)) cs. + Definition checkb_constructors_smaller X_ext (cs : list constructor_univs) (ind_sort : sort) := + List.forallb (List.forallb (fun argsort => abstract_env_compare_sort X_ext Cumul argsort ind_sort)) cs. Definition wf_cs_sorts X_ext cs := - forall Σ, abstract_env_ext_rel X_ext Σ -> Forall (Forall (wf_universe Σ)) cs. + forall Σ, abstract_env_ext_rel X_ext Σ -> Forall (Forall (wf_sort Σ)) cs. Lemma check_constructors_smallerP X_ext cs ind_sort : wf_cs_sorts X_ext cs -> - (forall Σ, abstract_env_ext_rel X_ext Σ -> wf_universe Σ ind_sort) -> + (forall Σ, abstract_env_ext_rel X_ext Σ -> wf_sort Σ ind_sort) -> forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ reflect (check_constructors_smaller Σ cs ind_sort) (checkb_constructors_smaller X_ext cs ind_sort) ∥. Proof using Type. unfold check_constructors_smaller, checkb_constructors_smaller. @@ -2026,7 +1969,7 @@ End monad_Alli_nth_forall. eapply forallbP_cond; eauto. clear wfcs. simpl; intros c wfc. eapply forallbP_cond; eauto. simpl. intros x wfx. specialize_Σ H. - apply iff_reflect. apply (abstract_env_compare_universe_correct _ H Cumul); eauto. + apply iff_reflect. apply (abstract_env_compare_sort_correct _ H Cumul); eauto. Qed. Program Definition do_check_ind_sorts X_ext {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X_ext -> NormalizationIn Σ} (params : context) @@ -2034,18 +1977,18 @@ End monad_Alli_nth_forall. (kelim : allowed_eliminations) (indices : context) (cs : list constructor_univs) (wfcs : wf_cs_sorts X_ext cs) - (ind_sort : Universe.t) - (wfi : forall Σ, abstract_env_ext_rel X_ext Σ -> wf_universe Σ ind_sort): + (ind_sort : sort) + (wfi : forall Σ, abstract_env_ext_rel X_ext Σ -> wf_sort Σ ind_sort): typing_result (forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ check_ind_sorts (lift_typing typing) Σ params kelim indices cs ind_sort ∥) := match ind_sort with - | Universe.lSProp => + | sSProp => check_eq_true (allowed_eliminations_subset kelim (elim_sort_sprop_ind cs)) (Msg "Incorrect allowed_elimination for inductive") ;; ret _ - | Universe.lProp => + | sProp => check_eq_true (allowed_eliminations_subset kelim (elim_sort_prop_ind cs)) (Msg "Incorrect allowed_elimination for inductive") ;; ret _ - | Universe.lType u => + | sType u => check_eq_true (checkb_constructors_smaller X_ext cs ind_sort) (Msg ("Incorrect inductive sort: The constructor arguments universes are not smaller than the declared inductive sort")) ;; match indices_matter with @@ -2060,13 +2003,13 @@ End monad_Alli_nth_forall. Next Obligation. unfold check_ind_sorts. simpl. - pose proof (check_constructors_smallerP X_ext cs (Universe.lType u) wfcs wfi). + pose proof (check_constructors_smallerP X_ext cs (sType u) wfcs wfi). rewrite -Heq_anonymous. specialize_Σ H. sq. split => //. match goal with [ H : reflect _ _ |- _ ] => destruct H => // end. Qed. Next Obligation. unfold check_ind_sorts. simpl. - pose proof (check_constructors_smallerP X_ext cs (Universe.lType u) wfcs wfi). + pose proof (check_constructors_smallerP X_ext cs (sType u) wfcs wfi). specialize_Σ H. sq. split. match goal with [ H : reflect _ _ |- _ ] => destruct H => // end. rewrite -Heq_anonymous; auto. @@ -2074,13 +2017,14 @@ End monad_Alli_nth_forall. Lemma sorts_local_ctx_wf_sorts (Σ : global_env_ext) {wfX : wf Σ} Γ Δ s : sorts_local_ctx (lift_typing typing) Σ Γ Δ s -> - Forall (wf_universe Σ) s. + Forall (wf_sort Σ) s. Proof using Type. induction Δ in s |- *; destruct s; simpl; auto. destruct a as [na [b|] ty]. - intros []. eauto. - intros []; eauto. constructor; eauto. - now eapply typing_wf_universe in t0. + destruct l as (_ & ? & t0 & <-). + now eapply typing_wf_sort in t0. Qed. Obligation Tactic := Program.Tactics.program_simplify. @@ -2155,7 +2099,7 @@ End monad_Alli_nth_forall. unfold check_variance in mdeclvar. rewrite -eqvar in mdeclvar. destruct (ind_universes mdecl) as [|[inst cstrs]] => //. - now eapply leq_context_cumul_context. + now eapply PCUICContextConversionTyp.eq_context_cumul_Spec_rel. Qed. Next Obligation. @@ -2176,7 +2120,7 @@ End monad_Alli_nth_forall. eapply All_sigma in indtys as [indus Hinds]. eapply All2_sq in Hinds; eauto. sq. red. - solve_all. now exists y. + solve_all. now eapply has_sort_isType. Qed. Program Definition check_one_ind_body X X_ext {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X_ext -> NormalizationIn Σ} @@ -2242,7 +2186,7 @@ End monad_Alli_nth_forall. destruct pf as [pf ?]; eauto. specialize_Σ H. specialize_Σ wfΣ0. pose proof (abstract_env_ext_wf _ pf). sq. eapply nth_error_all in wfars; eauto; simpl in wfars. - destruct wfars as [s Hs]. + destruct wfars as (_ & s & Hs & _). clear X0; rewrite p in Hs. eapply PCUICSpine.inversion_it_mkProd_or_LetIn in Hs; eauto. eapply inversion_Sort in Hs as [_ [? _]]; eauto. @@ -2258,7 +2202,7 @@ End monad_Alli_nth_forall. sq. clear onprojs onsorts X0. red in wfars. eapply nth_error_all in wfars; eauto; simpl in wfars. - destruct wfars as [s Hs]. + destruct wfars as (_ & s & Hs & _). apply eqb_eq in eqpars; apply eqb_eq in eqindices; subst. rewrite split_at_firstn_skipn in Heq_anonymous. noconf Heq_anonymous. @@ -2284,9 +2228,7 @@ End monad_Alli_nth_forall. rewrite split_at_firstn_skipn in Heq_anonymous. cbn in *. noconf Heq_anonymous. now rewrite {1}H0 {1}H1 /app_context -it_mkProd_or_LetIn_app firstn_skipn. - - red. red. - eapply nth_error_all in wfars; eauto; simpl in wfars. - destruct wfars as [s Hs]. now exists s. + - eapply nth_error_all in wfars; eauto; simpl in wfars. assumption. - unfold check_projections_type in onprojs. destruct (ind_projs idecl) => //. - now apply eqb_eq in eqsort; subst. @@ -2318,8 +2260,12 @@ End monad_Alli_nth_forall. ret (Build_on_inductive_sq check_bodies check_pars check_npars _) end. Next Obligation. - specialize_Σ H. sq. unfold on_constant_decl; rewrite <- Heq_anonymous. - eassumption. + destruct pf. + specialize_Σ H. + pose proof (abstract_env_ext_wf _ (H0 _ H1)) as wfΣ. + sq. unfold on_constant_decl; rewrite <- Heq_anonymous. + unshelve eapply validity in y as HT. 1: apply wfΣ. destruct HT as (_ & ?). + now split. Qed. Next Obligation. specialize_Σ H. sq. unfold on_constant_decl. rewrite <- Heq_anonymous; tea. diff --git a/safechecker/theories/PCUICSafeConversion.v b/safechecker/theories/PCUICSafeConversion.v index e2f8e5be6..b72f18a2c 100644 --- a/safechecker/theories/PCUICSafeConversion.v +++ b/safechecker/theories/PCUICSafeConversion.v @@ -160,16 +160,13 @@ Section Conversion. eapply H0. eapply r; eauto. Qed. - Import PCUICAlpha. - - Definition eqt u v := ∥ u ≡α v ∥. + Definition eqt u v := ∥ u = v :> term ∥. Lemma eqt_eqterm {Σ} {wfΣ : abstract_env_ext_rel X Σ} {u v} : - u ≡α v -> eq_term Σ Σ u v. + u = v -> eq_term Σ Σ u v. Proof using Type. - intros eq. - eapply upto_names_eq_term_refl; tc. - exact eq. + intros <-. + reflexivity. Qed. Local Instance eqt_refl : RelationClasses.Reflexive eqt. @@ -526,11 +523,13 @@ Section Conversion. eapply R_aux_stateR. all: simpl. all: auto. Qed. - Definition abstract_env_compare_global_instance := compare_global_instance (abstract_env_lookup X) (abstract_env_eq X). + Definition abstract_env_compare_global_instance := compare_global_instance (abstract_env_lookup X) (abstract_env_compare_universe X). - Notation eqb_ctx := (eqb_ctx_gen (abstract_env_eq X) abstract_env_compare_global_instance). - Notation eqb_term := (eqb_term_upto_univ (abstract_env_eq X) (abstract_env_eq X) abstract_env_compare_global_instance). - Notation leqb_term := (eqb_term_upto_univ (abstract_env_eq X) (abstract_env_leq X) abstract_env_compare_global_instance). + Notation eqb_ctx := (eqb_ctx_upto (abstract_env_compare_universe X) (abstract_env_compare_sort X) abstract_env_compare_global_instance Conv). + Notation cmpb_term_napp := (eqb_term_upto_univ_napp (abstract_env_compare_universe X) (abstract_env_compare_sort X) abstract_env_compare_global_instance). + Notation cmpb_term pb := (eqb_term_upto_univ (abstract_env_compare_universe X) (abstract_env_compare_sort X) abstract_env_compare_global_instance pb). + Notation eqb_term := (cmpb_term Conv). + Notation leqb_term := (cmpb_term Cumul). Definition eqb_term_stack t1 π1 t2 π2 := eqb_ctx (stack_context π1) (stack_context π2) && @@ -550,10 +549,16 @@ Section Conversion. symmetry; apply reflect_iff. eapply wf_universe_reflect. Qed. - Definition wf_universe_instance_iff Σ u : - wf_universeb_instance Σ u <-> wf_universe_instance Σ u. + Definition wf_sort_iff Σ s : + wf_sortb Σ s <-> wf_sort Σ s. + Proof using Type. + symmetry; apply reflect_iff. eapply wf_sort_reflect. + Qed. + + Definition wf_instance_iff Σ u : + wf_instanceb Σ u <-> wf_instance Σ u. Proof using Type. - symmetry; apply reflect_iff. eapply wf_universe_instanceP. + symmetry; apply reflect_iff. eapply wf_instanceP. Qed. Notation conv_stack_ctx Γ π1 π2 := @@ -1341,62 +1346,47 @@ Section Conversion. eapply typing_wf_local in X0; eauto. Qed. - Definition eqb_universe_instance_gen eq u v := - forallb2 eq (map Universe.make u) (map Universe.make v). - Definition eqb_universe_instance := - eqb_universe_instance_gen (abstract_env_eq X). + compare_universe_instance (abstract_env_compare_universe X Conv). Lemma eqb_universe_instance_spec : forall u v Σ (wfΣ : abstract_env_ext_rel X Σ), - forallb (wf_universeb Σ) (map Universe.make u) -> - forallb (wf_universeb Σ) (map Universe.make v) -> + forallb (wf_universeb Σ) (map Universe.make' u) -> + forallb (wf_universeb Σ) (map Universe.make' v) -> eqb_universe_instance u v -> - R_universe_instance (eq_universe (global_ext_constraints Σ)) u v. + cmp_universe_instance (eq_universe (global_ext_constraints Σ)) u v. Proof using Type. - intros u v Σ wfΣ Hu Hv e. - unfold eqb_universe_instance in e. - eapply forallb2_Forall2 in e. - eapply forallb_Forall in Hu. - eapply forallb_Forall in Hv. - eapply Forall_Forall2_and in e; try exact Hu; clear Hu. - eapply Forall_Forall2_and' in e; try exact Hv; clear Hv. - eapply Forall2_impl. 1: eassumption. - intros. cbn in H. destruct H as [[Hx H] Hy]. - eapply (abstract_env_compare_universe_correct _ _ Conv); eauto; now eapply wf_universe_iff. - Unshelve. eauto. + intros u v Σ wfΣ Hu Hv. + eapply elimT. + eapply reflect_reflectT, reflect_cmp_universe_instance; tea. + intros. + apply iff_reflect. + eapply abstract_env_compare_universe_correct with (conv_pb := Conv); tas. + all: now eapply wf_universe_iff. Qed. Arguments LevelSet.mem : simpl never. - Definition abstract_conv_pb_relb `{cf : checker_flags} - (pb : conv_pb) := - match pb with - | Conv => abstract_env_eq X - | Cumul => abstract_env_leq X - end. - - Lemma compare_universeb_complete Σ (wfΣ : abstract_env_ext_rel X Σ) leq u u' : + Lemma compare_universeb_complete Σ (wfΣ : abstract_env_ext_rel X Σ) pb u u' : wf_universe Σ u -> wf_universe Σ u' -> - compare_universe leq (global_ext_constraints Σ) u u' -> - abstract_conv_pb_relb leq u u'. + compare_universe (global_ext_constraints Σ) pb u u' -> + abstract_env_compare_universe X pb u u'. Proof using Type. intros all1 all2 conv. destruct (heΣ _ wfΣ). - destruct leq; eapply (abstract_env_compare_universe_correct _ _ _); eauto. - Unshelve. all: eauto. + eapply abstract_env_compare_universe_correct; eauto. Qed. Lemma get_level_make l : LevelExpr.get_level (LevelExpr.make l) = l. Proof using Type. now destruct l. Qed. - Lemma compare_universeb_make_complete Σ (wfΣ : abstract_env_ext_rel X Σ) leq x y : - wf_universe_level Σ x -> - wf_universe_level Σ y -> - compare_universe leq (global_ext_constraints Σ) (Universe.make x) (Universe.make y) -> - abstract_conv_pb_relb leq (Universe.make x) (Universe.make y). + Lemma compare_universeb_make_complete Σ (wfΣ : abstract_env_ext_rel X Σ) pb x y : + wf_level Σ x -> + wf_level Σ y -> + compare_universe (global_ext_constraints Σ) pb (Universe.make' x) (Universe.make' y) -> + abstract_env_compare_universe X pb (Universe.make' x) (Universe.make' y). Proof using Type. intros wfx wfy r. eapply compare_universeb_complete; eauto. @@ -1405,64 +1395,52 @@ Qed. Qed. Lemma eqb_universe_instance_complete Σ (wfΣ : abstract_env_ext_rel X Σ) u u' : - wf_universe_instance Σ u -> - wf_universe_instance Σ u' -> - R_universe_instance (eq_universe (global_ext_constraints Σ)) u u' -> + wf_instance Σ u -> + wf_instance Σ u' -> + cmp_universe_instance (eq_universe (global_ext_constraints Σ)) u u' -> eqb_universe_instance u u'. Proof using Type. - intros memu memu' r. - induction u in u', memu, memu', r |- *. - - now destruct u'. - - destruct u'; [easy|]. - depelim memu. - depelim memu'. - depelim r. - cbn in *. - apply Bool.andb_true_iff. - split. - + eapply (compare_universeb_make_complete _ _ Conv); eauto. - + apply IHu; eauto. - Unshelve. all:eauto. - Qed. - - Lemma compare_universe_variance_complete Σ (wfΣ : abstract_env_ext_rel X Σ) leq v u u' : - wf_universe_level Σ u -> - wf_universe_level Σ u' -> - R_universe_variance (eq_universe Σ) (compare_universe leq Σ) v u u' -> - compare_universe_variance (abstract_env_eq X) (abstract_conv_pb_relb leq) v u u'. + intros memu memu'. + eapply introT. + unfold wf_instance in *. + eapply reflect_reflectT, reflect_cmp_universe_instance with (p := wf_universeb Σ); tea. + 1: intros ????; eapply iff_reflect, abstract_env_compare_universe_correct with (conv_pb := Conv); tea. + 1,2: now eapply wf_universe_iff. + all: solve_all; eapply wf_universe_iff; intros ? ->%LevelExprSet.singleton_spec; auto. + Qed. + + Lemma compare_universe_variance_complete Σ (wfΣ : abstract_env_ext_rel X Σ) pb v u u' : + wf_level Σ u -> + wf_level Σ u' -> + cmp_universe_variance (compare_universe Σ) pb v u u' -> + compare_universe_variance (abstract_env_compare_universe X) pb v u u'. Proof using Type. intros memu memu' r. destruct v; cbn in *; eauto. - eapply compare_universeb_make_complete; eauto. - - eapply (compare_universeb_make_complete _ _ Conv); eauto. - Unshelve. eauto. + - eapply compare_universeb_make_complete with (pb := Conv); eauto. Qed. - Lemma compare_universe_instance_variance_complete Σ (wfΣ : abstract_env_ext_rel X Σ) leq v u u' : - wf_universe_instance Σ u -> - wf_universe_instance Σ u' -> - R_universe_instance_variance (eq_universe Σ) (compare_universe leq Σ) v u u' -> - compare_universe_instance_variance (abstract_env_eq X) (abstract_conv_pb_relb leq) v u u'. + Lemma compare_universe_instance_variance_complete Σ (wfΣ : abstract_env_ext_rel X Σ) pb v u u' : + wf_instance Σ u -> + wf_instance Σ u' -> + cmp_universe_instance_variance (compare_universe Σ) pb v u u' -> + compare_universe_instance_variance (abstract_env_compare_universe X) pb v u u'. Proof using Type. intros memu memu' r. - induction u in v, u', memu, memu', r |- *. - - now destruct u'. - - destruct u'; [easy|]. - depelim memu. + induction r in memu, memu' |- *; cbnr. + - depelim memu. depelim memu'. - cbn in *. - destruct v; auto. apply Bool.andb_true_iff. - destruct r. split. + eapply compare_universe_variance_complete; eauto. - + now apply IHu. + + now apply IHr. Qed. Lemma consistent_instance_ext_wf Σ udecl u : consistent_instance_ext Σ udecl u -> - wf_universe_instance Σ u. + wf_instance Σ u. Proof using Type. intros cons. unfold consistent_instance_ext, consistent_instance in *. @@ -1679,7 +1657,7 @@ Qed. eapply declared_constant_inj in decl1; eauto; subst. apply consistent_instance_ext_wf in cons1. apply consistent_instance_ext_wf in cons2. - eapply eqb_universe_instance_complete in r; auto. + eapply eqb_universe_instance_complete in c1; auto. Qed. (* Why Solve All Obligations is not working here ??? *) Next Obligation. solve_unfold_constants aux eq1 eq2 Σ wfΣ. Defined. @@ -1906,7 +1884,7 @@ Qed. [× forall Σ (wfΣ : abstract_env_ext_rel X Σ), declared_inductive Σ ci mdecl idecl, #|pparams p| = ind_npars mdecl, #|pparams p'| = ind_npars mdecl, - eq_context_gen eq eq br.(bcontext) br'.(bcontext), + eq_context_upto_names br.(bcontext) br'.(bcontext), test_context_k (fun k : nat => on_free_vars (closedP k (fun _ : nat => true))) #|pparams p| br.(bcontext) & test_context_k (fun k : nat => on_free_vars (closedP k (fun _ : nat => true))) @@ -1966,7 +1944,7 @@ Qed. - intros. erewrite (abstract_env_ext_irr _ _ wfΣ); eauto. - eapply (wf_predicate_length_pars wf_pred). - eapply (wf_predicate_length_pars wf_pred0). - - eapply alpha_eq_context_gen. etransitivity; tea. + - etransitivity; tea. now symmetry. - eapply PCUICConfluence.eq_context_upto_names_on_free_vars. 2:symmetry; exact p0. @@ -2661,10 +2639,11 @@ Qed. intros typ. destruct (hΣ _ wfΣ). apply PCUICValidity.inversion_mkApps in typ as (?&[typ_prod typ_args]). - apply inversion_Prod in typ_prod as (?&?&?&?&?); [|easy]. + apply inversion_Prod in typ_prod as (?&?&h1&?&?); [|easy]. eapply PCUICSpine.typing_spine_strengthen in typ_args; eauto. 2:{ eapply PCUICArities.isType_Sort. 2:pcuic. - eapply wf_universe_product; now eapply typing_wf_universe. } + eapply unlift_TypUniv in h1. + eapply wf_sort_product; now eapply typing_wf_sort. } clear -typ_args. depelim typ_args. - easy. @@ -2689,21 +2668,37 @@ Qed. Qed. Definition compare_global_instance_correct {Σ} - (wfΣ : abstract_env_ext_rel X Σ) R leq ref n l l' : - (forall u u', wf_universeb Σ u -> wf_universeb Σ u' -> reflect (R u u') (leq u u')) -> - wf_universe_instance Σ l -> - wf_universe_instance Σ l' -> - R_global_instance Σ (eq_universe Σ) R ref n l l' <-> - abstract_env_compare_global_instance leq ref n l l'. + (wfΣ : abstract_env_ext_rel X Σ) pb ref n l l' : + wf_instance Σ l -> + wf_instance Σ l' -> + cmp_global_instance Σ (compare_universe Σ) pb ref n l l' <-> + abstract_env_compare_global_instance pb ref n l l'. Proof. - intros hle hl hl'. apply reflect_iff. eapply reflect_R_global_instance; eauto. - all: try rewrite wf_universeb_instance_forall. - - intros ? ? Hu Hu'; apply iff_reflect; apply (abstract_env_compare_universe_correct _ wfΣ Conv). - + revert Hu. apply reflect_iff. apply wf_universe_reflect. - + revert Hu'. apply reflect_iff. apply wf_universe_reflect. + intros hl hl'. apply reflect_iff. eapply reflect_cmp_global_instance; eauto. + all: try rewrite -> wf_universeb_instance_forall. + - intros ? ? Hu Hu'; apply iff_reflect; apply (abstract_env_compare_universe_correct _ wfΣ). + all: now apply wf_universe_iff. + - intros ? ? Hu Hu'; apply iff_reflect; apply (abstract_env_compare_universe_correct _ wfΣ). + all: now apply wf_universe_iff. - intros; now eapply abstract_env_lookup_correct'. - - revert hl. apply reflect_iff, wf_universe_instanceP. - - revert hl'. apply reflect_iff, wf_universe_instanceP. + - revert hl. apply reflect_iff, wf_instanceP. + - revert hl'. apply reflect_iff, wf_instanceP. + Qed. + + Lemma cmpb_term_correct Σ (wfΣ : abstract_env_ext_rel X Σ) pb napp t u : + wf_universes Σ t -> + wf_universes Σ u -> + cmpb_term_napp pb napp t u <~> + compare_term_napp Σ Σ pb napp t u. + Proof. + intros ht hu. + eassert _ as H. + 2: split; [eapply elimT|eapply introT]; apply H. + eapply reflect_eq_term_upto_univ; tea. + 1,2: intros ????; apply iff_reflect, abstract_env_compare_universe_correct; tea; now apply wf_universe_iff. + 1,2: intros ????; apply iff_reflect, abstract_env_compare_sort_correct; tea; now apply wf_sort_iff. + 1,2: intros ??????; apply iff_reflect, compare_global_instance_correct; tea. + all: rewrite -> wf_universeb_instance_forall in *; now eapply wf_instance_iff. Qed. Lemma reduced_case_discriminee_whne Σ (wfΣ : abstract_env_ext_rel X Σ) Γ π ci p c brs h : @@ -2716,13 +2711,7 @@ Qed. intros eq ir. pose proof (heΣ _ wfΣ) as [[]]. pose proof (hΣ _ wfΣ). destruct ir as (_&[wh]); eauto. - eapply eqb_term_upto_univ_impl with (p := wf_universeb Σ) (q := closedu) in eq; tea. - 2-3: intros; apply iff_reflect; eapply (abstract_env_compare_universe_correct _ wfΣ Conv) ; now eapply wf_universe_iff. - 2:{ intros. rewrite !wf_universeb_instance_forall in H0, H1. - apply wf_universe_instance_iff in H0. - apply wf_universe_instance_iff in H1. - apply compare_global_instance_correct; eauto. - } + eapply cmpb_term_correct in eq; tea. - epose proof (reduce_term_complete _ _ _ _ _ _) as [wh']. eapply whnf_eq_term in eq; [|exact wh']. rewrite zipp_as_mkApps in wh. @@ -2813,12 +2802,7 @@ Qed. intros eq%eq_sym ir. destruct ir as (_&[wh]); eauto. pose proof (hΣ _ wfΣ). - eapply eqb_term_upto_univ_impl in eq; tea. - 2-3: intros; apply iff_reflect; eapply (abstract_env_compare_universe_correct _ wfΣ Conv) ; now eapply wf_universe_iff. - 2:{ intros. rewrite !wf_universeb_instance_forall in H0, H1. - apply wf_universe_instance_iff in H0. - apply wf_universe_instance_iff in H1. - eapply (compare_global_instance_correct wfΣ); eauto. } + eapply cmpb_term_correct in eq; tea. - epose proof (reduce_term_complete _ _ _ _ _ _) as [wh']. eapply whnf_eq_term in eq; [|exact wh']. rewrite zipp_as_mkApps in wh. @@ -3056,11 +3040,11 @@ Qed. (hp : ∥ ws_cumul_pb_terms Σ (Γ,,, stack_context π) (pparams p) (pparams p') ∥) : ∥ ∑ mdecl idecl, [× declared_inductive Σ ci mdecl idecl, - forallb (wf_universeb Σ) (map Universe.make (puinst p)), - forallb (wf_universeb Σ) (map Universe.make (puinst p')), + forallb (wf_universeb Σ) (map Universe.make' (puinst p)), + forallb (wf_universeb Σ) (map Universe.make' (puinst p')), #|pparams p| = ind_npars mdecl, #|pparams p'| = ind_npars mdecl, - eq_context_gen eq eq p.(pcontext) p'.(pcontext), + eq_context_upto_names p.(pcontext) p'.(pcontext), test_context_k (fun k : nat => on_free_vars (closedP k (fun _ : nat => true))) #|pparams p| p.(pcontext) & test_context_k (fun k : nat => on_free_vars (closedP k (fun _ : nat => true))) @@ -3098,13 +3082,13 @@ Qed. exists mdecl, idecl. destruct hcase. destruct hcase'. split; tea. - - eapply Forall_forallb; try eapply consistent_instance_wf_universe; eauto. + - eapply Forall_forallb; try eapply consistent_instance_wf_sort; eauto. intros; apply wf_universe_iff; eauto. - - eapply Forall_forallb; try eapply consistent_instance_wf_universe; eauto. + - eapply Forall_forallb; try eapply consistent_instance_wf_sort; eauto. intros; apply wf_universe_iff; eauto. - eapply (wf_predicate_length_pars wf_pred). - eapply (wf_predicate_length_pars wf_pred0). - - eapply alpha_eq_context_gen. etransitivity; tea. + - etransitivity; tea. now symmetry. - now rewrite <- test_context_k_closed_on_free_vars_ctx. - rewrite test_context_k_closed_on_free_vars_ctx. @@ -3165,7 +3149,7 @@ Qed. | Error ce not_conv_params => no ce; | Success conv_params - with inspect (eqb_universe_instance_gen (abstract_env_eq X) p1.(puinst) p2.(puinst)) := { + with inspect (eqb_universe_instance p1.(puinst) p2.(puinst)) := { | exist false not_eq_insts => no (CasePredUnequalUniverseInstances (Γ,,, stack_context π1) ci1 p1 c1 brs1 @@ -3255,8 +3239,7 @@ Qed. apply consistent_instance_ext_wf in cons. apply consistent_instance_ext_wf in cons0. specialize_Σ wfΣ; destruct H as [[]]. - apply eqb_universe_instance_complete in r; eauto. - unfold eqb_universe_instance in r. + apply eqb_universe_instance_complete in c; eauto. congruence. Qed. Next Obligation. @@ -3405,7 +3388,7 @@ Equations (noeqns) isconv_array_values_aux | prog_view_App _ _ _ _ := False_rect _ _; | prog_view_Const c u c' u' with eq_dec c c' := { - | left eq1 with inspect (eqb_universe_instance_gen (abstract_env_eq X) u u') := { + | left eq1 with inspect (eqb_universe_instance u u') := { | @exist true eq2 with isconv_args_raw leq (tConst c u) π1 (tConst c' u') π2 aux := { | Success h := yes ; (* Unfold both constants at once *) @@ -3604,7 +3587,7 @@ Equations (noeqns) isconv_array_values_aux { | @exist true eqf := yes | @exist false neqf := no (DistinctPrimValues (Γ ,,, stack_context π1) p (Γ ,,, stack_context π2) p') } | (primArray; primArrayModel a) | (primArray; primArrayModel a') - with inspect (abstract_env_eq X (Universe.make (array_level a)) (Universe.make (array_level a'))) := + with inspect (abstract_env_compare_universe X Conv (Universe.make' (array_level a)) (Universe.make' (array_level a'))) := { | @exist false neql := no (ArrayNotConvertibleLevels (Γ ,,, stack_context π1) a (Γ ,,, stack_context π2) a') | @exist true eql with isconv_red_raw Conv (array_type a) (PrimArray_ty a.(array_level) a.(array_value) a.(array_default) :: π1) (array_type a') (PrimArray_ty a'.(array_level) a'.(array_value) a'.(array_default) :: π2) aux := { @@ -3645,10 +3628,10 @@ Equations (noeqns) isconv_array_values_aux - destruct h. eapply welltyped_zipc_zipp in h1; auto. fvs. - constructor. eapply eqb_universe_instance_spec; eauto. + eapply welltyped_zipc_tConst_inv in h1 as (?&?&?); eauto; - eapply Forall_forallb; try eapply consistent_instance_wf_universe; eauto. + eapply Forall_forallb; try eapply consistent_instance_wf_sort; eauto. intros; apply wf_universe_iff; eauto. + eapply welltyped_zipc_tConst_inv in h2 as (?&?&?); eauto; - eapply Forall_forallb; try eapply consistent_instance_wf_universe; eauto. + eapply Forall_forallb; try eapply consistent_instance_wf_sort; eauto. intros; apply wf_universe_iff; eauto. Qed. Next Obligation. @@ -3731,7 +3714,6 @@ Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]; right; split; [easy|]. - unfold eqb_universe_instance. now rewrite <- uneq_u. Qed. @@ -3996,20 +3978,12 @@ Qed. destruct (reduce_stack_Req f _ X _ wfΣ Γ c' [] h) as [e' | hr] end. 1:{ - exfalso. Transparent reduce_term. - unfold reduce_term in eq4. - rewrite e' in eq4. cbn in eq4. - epose proof (eqb_term_upto_univ_refl Σ _ _ _ _ 0 c' _ _). - erewrite H in eq4. - - discriminate. - - intros. apply iff_reflect. eapply abstract_env_compare_universe_correct with (conv_pb := Conv) ; eauto. - - intros. rewrite !wf_universeb_instance_forall in H0, H1. - apply wf_universe_instance_iff in H0. - apply wf_universe_instance_iff in H1. - eapply compare_global_instance_correct; eauto. - clear -X0. intros ? ? Hu Hu'; apply X0. - + revert Hu. apply reflect_iff. apply wf_universe_reflect. - + revert Hu'. apply reflect_iff. apply wf_universe_reflect. + enough false by auto. rewrite eq4. + Transparent reduce_term. + unfold reduce_term. + rewrite e'. cbn. + assert (wf_universes Σ c'). + 2: eapply cmpb_term_correct; tea; cbnr. - pose proof h2 as Hc. specialize_Σ wfΣ. pose proof (hΣ _ wfΣ); sq. apply welltyped_zipc_inv in Hc; eauto. apply welltyped_wf in Hc; eauto. @@ -4017,22 +3991,13 @@ Qed. } dependent destruction hr. 2:{ - exfalso. + enough false by auto. rewrite eq4. destruct y'. simpl in H0. inversion H0. subst. - unfold reduce_term in eq4. - rewrite <- H2 in eq4. - cbn in eq4. - epose proof (eqb_term_upto_univ_refl Σ _ _ _ _ 0 c' _ _). - erewrite H1 in eq4. - - discriminate. - - intros. apply iff_reflect. eapply abstract_env_compare_universe_correct with (conv_pb := Conv); eauto. - - intros. rewrite !wf_universeb_instance_forall in H3, H4. - apply wf_universe_instance_iff in H3. - apply wf_universe_instance_iff in H4. - eapply compare_global_instance_correct; eauto. - clear -X0. intros ? ? Hu Hu'; apply X0. - + revert Hu. apply reflect_iff. apply wf_universe_reflect. - + revert Hu'. apply reflect_iff. apply wf_universe_reflect. + unfold reduce_term. + rewrite <- H2. + cbn. + assert (wf_universes Σ c'). + 2: eapply cmpb_term_correct; tea; cbnr. - pose proof h2 as Hc. specialize_Σ wfΣ. pose proof (hΣ _ wfΣ); sq. apply welltyped_zipc_inv in Hc; eauto. apply welltyped_wf in Hc; eauto. @@ -4042,7 +4007,7 @@ Qed. all: try reflexivity. simpl. intros. eapply cored_zipc. eapply cored_case. erewrite (abstract_env_ext_irr _ _ wfΣ); eauto. - Unshelve. all: eauto. all: intros; eapply abstract_env_compare_universe_correct with (conv_pb := Conv); eauto; reflexivity. + Unshelve. all: eauto. all: intros; eapply abstract_env_compare_sort_correct with (conv_pb := Conv); eauto; reflexivity. Qed. Next Obligation. @@ -4100,20 +4065,12 @@ Qed. destruct (reduce_stack_Req f _ X _ wfΣ Γ c [] h) as [e' | hr] end. 1:{ - exfalso. - unfold reduce_term in eq3. - rewrite e' in eq3. cbn in eq3. - epose proof (eqb_term_upto_univ_refl Σ _ _ _ _ 0 c _ _). - erewrite H in eq3. - - discriminate. - - intros. apply iff_reflect. eapply abstract_env_compare_universe_correct with (conv_pb := Conv); eauto. - - intros. rewrite !wf_universeb_instance_forall in H0, H1. - apply wf_universe_instance_iff in H0. - apply wf_universe_instance_iff in H1. - eapply compare_global_instance_correct; eauto. - clear -X0. intros ? ? Hu Hu'; apply X0. - + revert Hu. apply reflect_iff. apply wf_universe_reflect. - + revert Hu'. apply reflect_iff. apply wf_universe_reflect. + enough false by auto. rewrite eq3. + unfold reduce_term. + rewrite e'. + cbn. + assert (wf_universes Σ c). + 2: eapply cmpb_term_correct; tea; cbnr. - pose proof h1 as Hc. specialize_Σ wfΣ. pose proof (hΣ _ wfΣ); sq. apply welltyped_zipc_inv in Hc; eauto. apply welltyped_wf in Hc; eauto. @@ -4121,22 +4078,13 @@ Qed. } dependent destruction hr. 2:{ - exfalso. + enough false by auto. rewrite eq3. destruct y'. simpl in H0. inversion H0. subst. - unfold reduce_term in eq3. - rewrite <- H2 in eq3. - cbn in eq3. - epose proof (eqb_term_upto_univ_refl Σ _ _ _ _ 0 c _ _). - erewrite H1 in eq3. - - discriminate. - - intros. apply iff_reflect. eapply abstract_env_compare_universe_correct with (conv_pb := Conv); eauto. - - intros. rewrite !wf_universeb_instance_forall in H3, H4. - apply wf_universe_instance_iff in H3. - apply wf_universe_instance_iff in H4. - eapply compare_global_instance_correct; eauto. - clear -X0. intros ? ? Hu Hu'; apply X0. - + revert Hu. apply reflect_iff. apply wf_universe_reflect. - + revert Hu'. apply reflect_iff. apply wf_universe_reflect. + unfold reduce_term. + rewrite <- H2. + cbn. + assert (wf_universes Σ c). + 2: eapply cmpb_term_correct; tea; cbnr. - pose proof h1 as Hc. specialize_Σ wfΣ. pose proof (hΣ _ wfΣ); sq. apply welltyped_zipc_inv in Hc; eauto. apply welltyped_wf in Hc; eauto. @@ -4145,7 +4093,7 @@ Qed. unshelve eapply R_cored. simpl. intros; eapply cored_zipc. eapply cored_case. erewrite (abstract_env_ext_irr _ _ wfΣ); eauto. - Unshelve. all: eauto. all: intros; eapply abstract_env_compare_universe_correct with (conv_pb := Conv); eauto; reflexivity. + Unshelve. all: eauto. all: intros; eapply abstract_env_compare_sort_correct with (conv_pb := Conv); eauto; reflexivity. Qed. Next Obligation. rename H into wfΣ; specialize_Σ wfΣ. @@ -4251,42 +4199,25 @@ Qed. destruct (reduce_stack_Req f _ X _ wfΣ Γ c' [] h) as [e' | hr] end. 1:{ - exfalso. - unfold reduce_term in eq4. - rewrite e' in eq4. cbn in eq4. - epose proof (eqb_term_upto_univ_refl Σ _ _ _ _ 0 c' _ _). - erewrite H in eq4. - - discriminate. - - intros. apply iff_reflect. eapply abstract_env_compare_universe_correct with (conv_pb := Conv); eauto. - - intros. rewrite !wf_universeb_instance_forall in H0, H1. - apply wf_universe_instance_iff in H0. - apply wf_universe_instance_iff in H1. - eapply compare_global_instance_correct; eauto. - clear -X0. intros ? ? Hu Hu'; apply X0. - + revert Hu. apply reflect_iff. apply wf_universe_reflect. - + revert Hu'. apply reflect_iff. apply wf_universe_reflect. + enough false by auto. rewrite eq4. + unfold reduce_term. + rewrite e'. + cbn. + assert (wf_universes Σ c'). + 2: eapply cmpb_term_correct; tea; cbnr. - pose proof h2 as Hc. specialize_Σ wfΣ. pose proof (hΣ _ wfΣ); sq. apply welltyped_zipc_inv in Hc; eauto. apply welltyped_wf in Hc; eauto. } dependent destruction hr. 2:{ - exfalso. + enough false by auto. rewrite eq4. destruct y'. simpl in H0. inversion H0. subst. - unfold reduce_term in eq4. - rewrite <- H2 in eq4. - cbn in eq4. - epose proof (eqb_term_upto_univ_refl Σ _ _ _ _ 0 c' _ _). - erewrite H1 in eq4. - - discriminate. - - intros. apply iff_reflect. eapply abstract_env_compare_universe_correct with (conv_pb := Conv); eauto. - - intros. rewrite !wf_universeb_instance_forall in H3, H4. - apply wf_universe_instance_iff in H3. - apply wf_universe_instance_iff in H4. - eapply compare_global_instance_correct; eauto. - clear -X0. intros ? ? Hu Hu'; apply X0. - + revert Hu. apply reflect_iff. apply wf_universe_reflect. - + revert Hu'. apply reflect_iff. apply wf_universe_reflect. + unfold reduce_term. + rewrite <- H2. + cbn. + assert (wf_universes Σ c'). + 2: eapply cmpb_term_correct; tea; cbnr. - pose proof h2 as Hc. specialize_Σ wfΣ. pose proof (hΣ _ wfΣ); sq. apply welltyped_zipc_inv in Hc; eauto. apply welltyped_wf in Hc; eauto. @@ -4295,7 +4226,7 @@ Qed. all: try reflexivity. simpl. intros; eapply cored_zipc. eapply cored_proj. erewrite (abstract_env_ext_irr _ _ wfΣ); eassumption. - Unshelve. all: eauto. all: intros; eapply abstract_env_compare_universe_correct with (conv_pb := Conv); eauto; reflexivity. + Unshelve. all: eauto. all: intros; eapply abstract_env_compare_sort_correct with (conv_pb := Conv); eauto; reflexivity. Qed. Next Obligation. rename H into wfΣ. specialize_Σ wfΣ. @@ -4356,42 +4287,25 @@ Qed. destruct (reduce_stack_Req f _ X _ wfΣ Γ c [] h) as [e' | hr] end. 1:{ - exfalso. - unfold reduce_term in eq3. - rewrite e' in eq3. cbn in eq3. - epose proof (eqb_term_upto_univ_refl Σ _ _ _ _ 0 c _ _). - erewrite H in eq3. - - discriminate. - - intros. apply iff_reflect. eapply abstract_env_compare_universe_correct with (conv_pb := Conv); eauto. - - intros. rewrite !wf_universeb_instance_forall in H0, H1. - apply wf_universe_instance_iff in H0. - apply wf_universe_instance_iff in H1. - eapply compare_global_instance_correct; eauto. - clear -X0. intros ? ? Hu Hu'; apply X0. - + revert Hu. apply reflect_iff. apply wf_universe_reflect. - + revert Hu'. apply reflect_iff. apply wf_universe_reflect. + enough false by auto. rewrite eq3. + unfold reduce_term. + rewrite e'. + cbn. + assert (wf_universes Σ c). + 2: eapply cmpb_term_correct; tea; cbnr. - pose proof h1 as Hc. specialize_Σ wfΣ. pose proof (hΣ _ wfΣ); sq. apply welltyped_zipc_inv in Hc; eauto. apply welltyped_wf in Hc; eauto. } dependent destruction hr. 2:{ - exfalso. + enough false by auto. rewrite eq3. destruct y'. simpl in H0. inversion H0. subst. - unfold reduce_term in eq3. - rewrite <- H2 in eq3. - cbn in eq3. - epose proof (eqb_term_upto_univ_refl Σ _ _ _ _ 0 c _ _). - erewrite H1 in eq3. - - discriminate. - - intros. apply iff_reflect. eapply abstract_env_compare_universe_correct with (conv_pb := Conv); eauto. - - intros. rewrite !wf_universeb_instance_forall in H3, H4. - apply wf_universe_instance_iff in H3. - apply wf_universe_instance_iff in H4. - eapply compare_global_instance_correct; eauto. - clear -X0. intros ? ? Hu Hu'; apply X0. - + revert Hu. apply reflect_iff. apply wf_universe_reflect. - + revert Hu'. apply reflect_iff. apply wf_universe_reflect. + unfold reduce_term. + rewrite <- H2. + cbn. + assert (wf_universes Σ c). + 2: eapply cmpb_term_correct; tea; cbnr. - pose proof h1 as Hc. specialize_Σ wfΣ. pose proof (hΣ _ wfΣ); sq. apply welltyped_zipc_inv in Hc; eauto. apply welltyped_wf in Hc; eauto. @@ -4399,7 +4313,7 @@ Qed. unshelve eapply R_cored. simpl. intros; eapply cored_zipc. eapply cored_proj. erewrite (abstract_env_ext_irr _ _ wfΣ); eassumption. - Unshelve. all: eauto. all: intros; eapply abstract_env_compare_universe_correct with (conv_pb := Conv); eauto; reflexivity. + Unshelve. all: eauto. all: intros; eapply abstract_env_compare_sort_correct with (conv_pb := Conv); eauto; reflexivity. Qed. Next Obligation. rename H into wfΣ. specialize_Σ wfΣ. @@ -4902,8 +4816,8 @@ Qed. specialize (convdiscrdef _ H) as [convdef]. specialize (convdiscrval _ H) as [convval]. cbn in convty, convdef. - symmetry in eql. unfold abstract_env_eq in eql. - eapply abstract_env_compare_universe_correct in eql; eauto. + symmetry in eql. + eapply abstract_env_compare_universe_correct with (conv_pb := Conv) in eql; eauto. 2:{ destruct h1 as [? ty]; eapply typing_wf_universes in ty; eauto. move/andP: ty => []. rewrite H0 /=. cbn -[wf_universeb]. rtoProp; intuition auto. @@ -4988,8 +4902,8 @@ Qed. specialize (H _ wfΣ). symmetry in neql. destruct H as [H]; eapply invert_cumul_Prim in H; depelim H. + enough false by auto. rewrite <- neql. unshelve eapply (abstract_env_compare_universe_correct _ _ Conv) in e0; eauto. - - unfold abstract_env_eq in neql. now erewrite neql in e0. - rewrite H0 in h1. destruct h1 as [? wt]. eapply typing_wf_universes in wt; eauto. move/andP: wt => []. cbn -[wf_universeb wf_universe]. rtoProp; intuition auto. @@ -5724,8 +5638,7 @@ Qed. isconv_prog leq t1 π1 rt2' (θ2' ++ θ2) aux } } ; - | exist None nored2 with inspect (eqb_termp_napp_gen leq (abstract_env_eq X) (abstract_env_leq X) - abstract_env_compare_global_instance #|(decompose_stack π1).1| t1 t2) := { + | exist None nored2 with inspect (cmpb_term_napp leq #|(decompose_stack π1).1| t1 t2) := { | exist true eq1 := isconv_args leq t1 π1 t2 π2 aux; | exist false noteq := no ( @@ -5982,22 +5895,14 @@ Qed. rewrite Nat.add_0_r. apply All2_length in a. rewrite a in eq1. - eapply eqb_term_upto_univ_impl with (q := closedu); eauto. - + intros. eapply iff_reflect. - eapply (abstract_env_compare_universe_correct _ H Conv) ; now eapply wf_universe_iff. - + intros. eapply iff_reflect. destruct leq. - * eapply (abstract_env_compare_universe_correct _ H Conv) ; now eapply wf_universe_iff. - * eapply (abstract_env_compare_universe_correct _ H Cumul) ; now eapply wf_universe_iff. - + intros. rewrite !wf_universeb_instance_forall in H0, H1. - apply wf_universe_instance_iff in H0. - apply wf_universe_instance_iff in H1. - eapply compare_global_instance_correct; eauto. + eapply cmpb_term_correct; tea. + pose proof h1 as Hc. specialize_Σ H. pose proof (hΣ _ H); sq. apply welltyped_zipc_inv in Hc; eauto. apply welltyped_wf in Hc; eauto. + pose proof h2 as Hc. specialize_Σ H. pose proof (hΣ _ H); sq. apply welltyped_zipc_inv in Hc; eauto. apply welltyped_wf in Hc; eauto. + + now symmetry. Qed. Next Obligation. apply h; clear h. intros Σ wfΣ. @@ -6069,9 +5974,6 @@ Qed. apply consistent_instance_ext_wf in c0. apply consistent_instance_ext_wf in c. eapply compare_global_instance_correct in H3; eauto. - 2: { intros; apply iff_reflect. eapply (abstract_env_compare_universe_correct _ _ leq); apply wf_universe_iff; eauto. - all: apply wf_universe_iff; eauto. - } rewrite eqb_refl in noteq. apply All2_length in rargs1. rewrite <- rargs1 in H3. @@ -6097,14 +5999,6 @@ Qed. rewrite <- rargs1 in H4. apply ssrbool.not_false_is_true. rewrite noteq. cbn. eapply compare_global_instance_correct; eauto. - intros ? ? Hu Hu'. apply iff_reflect. - destruct leq; - [ apply abstract_env_compare_universe_correct with (conv_pb := Conv); eauto | - apply abstract_env_compare_universe_correct with (conv_pb := Cumul); eauto]. - + revert Hu. apply reflect_iff. apply wf_universe_reflect. - + revert Hu'. apply reflect_iff. apply wf_universe_reflect. - + revert Hu. apply reflect_iff. apply wf_universe_reflect. - + revert Hu'. apply reflect_iff. apply wf_universe_reflect. } all: apply conv_cum_alt in conv_hds as [(?&?&[r1 r2 ?])]; auto. all: eapply whnf_red_inv in r1; auto. @@ -6129,7 +6023,7 @@ Qed. simpl in h2. apply inversion_Sort in h2 as (_&h2&_); auto. apply inversion_Sort in h1 as (_&h1&_); auto. - eapply compare_universeb_complete in H0; eauto. + eapply abstract_env_compare_sort_correct in H0; eauto. destruct leq; cbn in *; easy. Unshelve. all:eauto. Qed. @@ -6165,10 +6059,8 @@ Qed. - cbn in *. specialize_Σ wfΣ. eapply cored'_postpone in H as [u' [cor eq]]. eapply cored_welltyped in cor; tea. - destruct eq as [eq]. - eapply welltyped_alpha; tea. symmetry. exact eq. - - simpl in *. destruct e. eapply welltyped_alpha; tea. - now symmetry. + now destruct eq as [<-]. + - simpl in *. destruct e. now rewrite e. Qed. Equations(noeqns) isconv_full (s : state) (Γ : context) diff --git a/safechecker/theories/PCUICSafeReduce.v b/safechecker/theories/PCUICSafeReduce.v index 8cb42f7c8..a7ebc08cc 100644 --- a/safechecker/theories/PCUICSafeReduce.v +++ b/safechecker/theories/PCUICSafeReduce.v @@ -1308,7 +1308,7 @@ Corollary R_Acc_aux : exfalso; eapply invert_cumul_sort_ind; eauto. - apply inversion_Prod in typ as (?&?&?&?&?); auto. exfalso; eapply invert_cumul_sort_ind; eauto. - - apply inversion_Lambda in typ as (?&?&?&?&?); auto. + - apply inversion_Lambda in typ as (?&?&?&?); auto. exfalso; eapply invert_cumul_prod_ind; eauto. - unfold isConstruct_app in ctor. now rewrite decompose_app_mkApps in ctor. @@ -1344,7 +1344,7 @@ Corollary R_Acc_aux : exfalso; eapply invert_cumul_sort_ind; eauto. - apply inversion_Prod in typ as (?&?&?&?&?); auto. exfalso; eapply invert_cumul_sort_ind; eauto. - - apply inversion_Lambda in typ as (?&?&?&?&?); auto. + - apply inversion_Lambda in typ as (?&?&?&?); auto. exfalso; eapply invert_cumul_prod_ind; eauto. - unfold isConstruct_app in ctor. now rewrite decompose_app_mkApps in ctor. @@ -2025,7 +2025,7 @@ Section ReduceFns. (* Definition of assumption-only arities (without lets) *) Definition arity_ass := aname * term. - Fixpoint mkAssumArity (l : list arity_ass) (s : Universe.t) : term := + Fixpoint mkAssumArity (l : list arity_ass) (s : sort) : term := match l with | [] => tSort s | (na, A) :: l => tProd na A (mkAssumArity l s) @@ -2046,7 +2046,7 @@ Section ReduceFns. constructor; auto. Qed. - Lemma mkAssumArity_it_mkProd_or_LetIn (l : list arity_ass) (s : Universe.t) : + Lemma mkAssumArity_it_mkProd_or_LetIn (l : list arity_ass) (s : sort) : mkAssumArity l s = it_mkProd_or_LetIn (arity_ass_context l) (tSort s). Proof using Type. unfold arity_ass_context. @@ -2066,7 +2066,7 @@ Section ReduceFns. Record conv_arity {Γ T} : Type := build_conv_arity { conv_ar_context : list arity_ass; - conv_ar_univ : Universe.t; + conv_ar_univ : sort; conv_ar_red : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ Σ ;;; Γ ⊢ T ⇝ mkAssumArity conv_ar_context conv_ar_univ ∥ }. diff --git a/safechecker/theories/PCUICSafeRetyping.v b/safechecker/theories/PCUICSafeRetyping.v index 385dabc26..467e12876 100644 --- a/safechecker/theories/PCUICSafeRetyping.v +++ b/safechecker/theories/PCUICSafeRetyping.v @@ -99,7 +99,7 @@ Lemma inductive_cumulative_indices_smash {cf : checker_flags} {Σ : global_env_e on_udecl_prop Σ (ind_universes mdecl) -> consistent_instance_ext Σ (ind_universes mdecl) u -> consistent_instance_ext Σ (ind_universes mdecl) u' -> - PCUICEquality.R_global_instance Σ (eq_universe Σ) (leq_universe Σ) (IndRef ind) napp u u' -> + PCUICEquality.cmp_ind_universes Σ ind napp u u' -> forall Γ pars pars', spine_subst Σ Γ pars (List.rev pars) (smash_context [] (subst_instance u (ind_params mdecl))) -> spine_subst Σ Γ pars' (List.rev pars') (smash_context [] (subst_instance u' (ind_params mdecl))) -> @@ -167,10 +167,10 @@ Lemma welltyped_subterm {Σ Γ t} : wellinferred Σ Γ t -> on_subterm (wellinferred Σ) (well_sorted Σ) Γ t. Proof using Type. destruct t; simpl; auto; intros [T HT]; sq. - now inversion HT ; auto; split; do 2 econstructor. - now inversion HT ; auto; split; econstructor ; [econstructor|..]. - now inversion HT ; inversion X1 ; auto; - split; [split|]; econstructor ; [econstructor|..]. + - inversion HT ; subst. apply unlift_TypUniv in X0. split; now do 2 econstructor. + - inversion HT ; subst. destruct X0 as (_ & ? & ? & _); cbn in *. split; econstructor ; [econstructor|..]; eassumption. + - inversion HT ; subst. destruct X0 as (X0' & ? & ? & _); cbn in *. + inversion X0'. split; [split|]; econstructor ; [econstructor|..]; eassumption. Qed. #[local] Notation ret t := (t; _). @@ -181,10 +181,10 @@ Qed. ∑ u, forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ Σ ;;; Γ |- T ▹□ u ∥. #[local] Definition principal_type_type {Γ t} (wt : principal_type Γ t) : term := projT1 wt. - #[local] Definition principal_sort_sort {Γ T} (ps : principal_sort Γ T) : Universe.t + #[local] Definition principal_sort_sort {Γ T} (ps : principal_sort Γ T) : sort := projT1 ps. #[local] Coercion principal_type_type : principal_type >-> term. - #[local] Coercion principal_sort_sort : principal_sort >-> Universe.t. + #[local] Coercion principal_sort_sort : principal_sort >-> sort. Program Definition infer_as_sort {Γ T} (wfΓ : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ wf_local Σ Γ ∥) @@ -199,7 +199,7 @@ Qed. destruct (wf _ wfΣ) as [[]], (hΣ _ wfΣ) as [wΣ]. specialize_Σ wfΣ. sq. - eapply infering_typing, validity in s as []; eauto. + eapply infering_typing, validity in s as (_ & s & Hs & _); eauto. now eexists. Defined. Next Obligation. @@ -303,7 +303,7 @@ Qed. infer Γ wfΓ (tVar n) wt := !; infer Γ wfΓ (tEvar ev args) wt := !; - infer Γ wfΓ (tSort s) wt := ret (tSort (Universe.super s)); + infer Γ wfΓ (tSort s) wt := ret (tSort (Sort.super s)); infer Γ wfΓ (tProd n ty b) wt := let wfΓ' : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ wf_local Σ (Γ ,, vass n ty) ∥ := _ in @@ -311,7 +311,7 @@ Qed. let s1 := infer_as_sort wfΓ (fun a b => (welltyped_subterm (wt a b)).1) ty1 in let ty2 := infer (Γ ,, vass n ty) wfΓ' b (fun a b => (welltyped_subterm (wt a b)).2) in let s2 := infer_as_sort wfΓ' (fun a b => (welltyped_subterm (wt a b)).2) ty2 in - ret (tSort (Universe.sort_of_product s1 s2)); + ret (tSort (Sort.sort_of_product s1 s2)); infer Γ wfΓ (tLambda n t b) wt := let t2 := infer (Γ ,, vass n t) _ b (fun a b => (welltyped_subterm (wt a b)).2) in @@ -402,19 +402,20 @@ Qed. sq. constructor ; tea. inversion X0. - eapply infering_sort_isType; eauto. + now eapply isTypebd_isType in X1. Defined. Next Obligation. cbn ; intros. destruct s1, s2. cbn. specialize_Σ wfΣ. sq. - now constructor. + constructor; eauto. + repeat (eexists; tea). Defined. Next Obligation. pose (hΣ _ wfΣ). specialize_Σ wfΣ. inversion wt. sq. inversion X0 ; subst. constructor ; tea. - now eapply infering_sort_isType. + now eapply lift_sorting_lift_typing. Defined. Next Obligation. case t2 as []. intros; cbn. specialize_Σ wfΣ. @@ -428,9 +429,7 @@ Qed. pose (hΣ _ wfΣ). specialize_Σ wfΣ. inversion wt. sq. inversion X0 ; subst. constructor ; tea. - 1: now eapply infering_sort_isType. - apply checking_typing ; eauto. - now eapply infering_sort_isType. + now eapply lift_sorting_lift_typing. Defined. Next Obligation. cbn; intros; case b'_ty as []. cbn. @@ -580,7 +579,7 @@ Qed. cbn in *. pose proof wt. specialize_Σ wfΣ. destruct infer. pose (hΣ _ wfΣ). cbn. specialize_Σ wfΣ. sq. - eapply infering_typing, validity in s as [] ; eauto. + eapply infering_typing, validity in s as (_ & ? & ? & _) ; eauto. now eexists. Defined. @@ -659,7 +658,7 @@ Qed. specialize_Σ wfΣ. pose (hΣ _ wfΣ); sq. cbn. - eapply infering_typing, validity in s' as []; eauto. + eapply infering_typing, validity in s' as (_ & ? & ? & _); eauto. now eexists. Defined. Next Obligation. @@ -829,7 +828,7 @@ Qed. Opaque type_of_typing. Equations? sort_of_type (Γ : context) (t : PCUICAst.term) (wt : forall Σ : global_env_ext, abstract_env_ext_rel X Σ -> ∥ isType Σ Γ t ∥) : - (∑ u : Universe.t, forall Σ : global_env_ext, abstract_env_ext_rel X Σ -> + (∑ u : sort, forall Σ : global_env_ext, abstract_env_ext_rel X Σ -> ∥ Σ ;;; Γ |- t : tSort u ∥) := sort_of_type Γ t wt with (@type_of_typing Γ t _) := { | T with inspect (reduce_to_sort (X:=X) Γ T.π1 _) := @@ -856,7 +855,7 @@ Qed. cbn in ns. clear ns. specialize (wt _ wfΣ). destruct T as [T HT]. cbn in *. destruct (HT _ wfΣ) as [[hty hp]]. - eapply validity in hty. destruct wt as [[s Hs]]. + eapply validity in hty. destruct wt as [(_ & s & Hs & _)]. red in hp. specialize (hp _ Hs). eapply ws_cumul_pb_Sort_r_inv in hp as [s' [hs' _]]. eapply (H s' hs'). diff --git a/safechecker/theories/PCUICTypeChecker.v b/safechecker/theories/PCUICTypeChecker.v index 396ba2d79..31594a770 100644 --- a/safechecker/theories/PCUICTypeChecker.v +++ b/safechecker/theories/PCUICTypeChecker.v @@ -196,7 +196,7 @@ Lemma isType_mkApps_Ind_smash_inv {cf:checker_flags} {Σ Γ ind u args} (wfΣ : consistent_instance_ext Σ (ind_universes mdecl) u. Proof. move=> wfΓ isType. - destruct isType as [s Hs]. + destruct isType as (_ & s & Hs & _). eapply invert_type_mkApps_ind in Hs as [tyargs cu]; eauto. set (decli' := on_declared_inductive declm). rename declm into decli. @@ -215,13 +215,13 @@ Qed. (* Lemma compare_global_instance_sound {cf Σ} (wfΣ : wf_ext Σ) gr napp (Hφ : on_udecl Σ.1 Σ.2) (G : universes_graph) (HG : is_graph_of_uctx G (global_ext_uctx Σ)) : - subrelation (compare_global_instance Σ (check_eqb_universe G) (check_leqb_universe G) gr napp) - (R_global_instance Σ (eq_universe Σ) (leq_universe Σ) gr napp). -Proof. eapply reflect_R_global_instance. compare_global_instance_impl; tc; intros x y. - - eapply (check_eqb_universe_spec' _ (global_ext_uctx Σ)) => //. + subrelation (compare_global_instance Σ (check_eqb_sort G) (check_leqb_sort G) gr napp) + (cmp_global_instance Σ (eq_sort Σ) (leq_sort Σ) gr napp). +Proof. eapply reflect_cmp_global_instance. compare_global_instance_impl; tc; intros x y. + - eapply (check_eqb_sort_spec' _ (global_ext_uctx Σ)) => //. now eapply wf_ext_global_uctx_invariants. cbn. eapply wfΣ. - - eapply (check_leqb_universe_spec' _ (global_ext_uctx Σ)) => //. + - eapply (check_leqb_sort_spec' _ (global_ext_uctx Σ)) => //. now eapply wf_ext_global_uctx_invariants. cbn. eapply wfΣ. Qed. *) @@ -240,19 +240,18 @@ Lemma substitution_wf_local_rel `{checker_flags} {Σ} {wfΣ : wf Σ} {Γ Γ' s constructor ; cbn. + eapply IHΔ ; tea. + rewrite Nat.add_0_r. - eapply isType_substitution ; tea. - now rewrite -app_context_assoc. - + rewrite Nat.add_0_r. - eapply substitution ; tea. - now rewrite -app_context_assoc. + rewrite app_context_assoc in l. + apply lift_typing_f_impl with (1 := l) => // ?? HT. + now eapply substitution. - cbn. rewrite subst_context_snoc. rewrite app_context_cons in Ht ; depelim Ht. constructor ; cbn. + eapply IHΔ ; tea. + rewrite Nat.add_0_r. - eapply isType_substitution ; tea. - now rewrite -app_context_assoc. + rewrite app_context_assoc in l. + apply lift_typing_f_impl with (1 := l) => // ?? HT. + now eapply substitution. Qed. Section Typecheck. @@ -292,12 +291,12 @@ Section Typecheck. Notation hnf := (hnf (X := X)). - Definition conv_pb_relb_gen_proper pb equ equ' eqlu eqlu' : + Definition conv_pb_relb_gen_proper {T} pb equ equ' eqlu eqlu' : (forall u u', equ u u' = equ' u u') -> (forall u u', eqlu u u' = eqlu' u u') -> - forall u u', - conv_pb_relb_gen pb equ eqlu u u' = - conv_pb_relb_gen pb equ' eqlu' u u'. + forall (u u' : T), + conv_pb_relb_gen equ eqlu pb u u' = + conv_pb_relb_gen equ' eqlu' pb u u'. Proof using Type. now destruct pb. Qed. @@ -316,7 +315,7 @@ Section Typecheck. (hu : forall Σ (wfΣ : abstract_env_ext_rel X Σ), welltyped Σ Γ u) : typing_result_comp (forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ Σ ;;; Γ ⊢ t ≤[le] u ∥) := convert le Γ t u ht hu - with inspect (eqb_termp_napp_gen le (abstract_env_eq X) (abstract_env_leq X) (abstract_env_compare_global_instance _ X) 0 t u) := { + with inspect (eqb_term_upto_univ (abstract_env_compare_universe X) (abstract_env_compare_sort X) (abstract_env_compare_global_instance _ X) le t u) := { | @exist true He := ret _ ; | @exist false He with inspect (isconv_term _ X Γ le t ht u hu) := { @@ -327,19 +326,9 @@ Section Typecheck. raise (NotCumulSmaller false X Γ t u t' u' e) }}. Next Obligation. - unfold eqb_termp_napp_gen in He. pose (heΣ _ wfΣ) as heΣ; sq. + unfold eqb_term_upto_univ_napp in He. pose (heΣ _ wfΣ) as heΣ; sq. constructor; fvs. specialize_Σ wfΣ. - eapply eqb_term_upto_univ_impl; eauto. - - intros. eapply iff_reflect. - eapply (abstract_env_compare_universe_correct _ wfΣ Conv); - try eassumption; apply wf_universe_iff; eauto. - - intros. eapply iff_reflect. destruct le; - eapply (abstract_env_compare_universe_correct _ wfΣ _); - try eassumption; apply wf_universe_iff; eauto. - - intros. rewrite wf_universeb_instance_forall in H. rewrite wf_universeb_instance_forall in H0. - apply wf_universe_instance_iff in H. - apply wf_universe_instance_iff in H0. - eapply (compare_global_instance_correct _ X wfΣ); eauto. + eapply cmpb_term_correct; eauto. - destruct ht as [? ht]. eapply typing_wf_universes in ht; eauto. pose proof ht as [? ?]%andb_and; eassumption. - destruct hu as [? hu]. eapply typing_wf_universes in hu; eauto. @@ -369,7 +358,7 @@ Section Typecheck. typing_result_comp ({ A : term & forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ Σ ;;; Γ |- t ▹ A ∥ })). Equations infer_type Γ (HΓ : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ wf_local Σ Γ ∥) t - : typing_result_comp ({u : Universe.t & forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ Σ ;;; Γ |- t ▹□ u ∥}) := + : typing_result_comp ({u : sort & forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ Σ ;;; Γ |- t ▹□ u ∥}) := infer_type Γ HΓ t := tx <- infer Γ HΓ t ;; s <- reduce_to_sort (X := X) Γ tx.π1 _ ;; @@ -490,18 +479,22 @@ Section Typecheck. Next Obligation. pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. econstructor ; tea. + destruct s as (_ & s). + split; cbn; tas. now eapply checking_typing. Qed. Next Obligation. eapply absurd. intros. pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. inversion H ; subst. - intros. now eapply typing_checking. + apply unlift_TermTyp in X1. + now eapply typing_checking. Qed. Next Obligation. eapply absurd. intros. pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. - now inversion H ; subst. + inversion H ; subst. + destruct X1 as (_ & ?); split; cbn; auto. Qed. Next Obligation. eapply absurd. intros. specialize_Σ wfΣ. sq. @@ -509,8 +502,8 @@ Section Typecheck. Qed. Next Obligation. pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. econstructor; tas. - eexists. - now eapply infering_sort_typing. + eapply has_sort_isType. + now eapply infering_sort_typing. Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. @@ -528,7 +521,7 @@ Section Typecheck. Lemma sq_wf_local_app {Γ Δ} : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ wf_local Σ Γ ∥ -> ∥ wf_local_rel Σ Γ Δ ∥ -> ∥ wf_local Σ (Γ ,,, Δ) ∥. Proof using Type. - intros. sq. now apply wf_local_app. + intros. sq. now apply All_local_env_app. Qed. Equations check_context_rel Γ (wfΓ : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ wf_local Σ Γ ∥) (Δ : context) : @@ -551,19 +544,20 @@ Section Typecheck. Qed. Next Obligation. pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. constructor ; auto. + destruct wfA as (_ & ?); split; cbn; tas. eapply checking_typing ; pcuic. Qed. Next Obligation. apply absurd. intros. pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. inversion H ; subst ; cbn in *. + apply unlift_TermTyp in X1. now eapply typing_checking. Qed. Next Obligation. apply absurd. intros. specialize_Σ wfΣ. sq. inversion H ; subst ; cbn in *. - destruct X1 as [s ?]. - now exists s. + destruct X1. split; cbn; auto. Qed. Next Obligation. apply absurd. intros. specialize_Σ wfΣ. sq. @@ -693,7 +687,8 @@ Section Typecheck. wf_local Σ Γ * wt_decl Σ Γ d. Proof using Type. intros wfd; depelim wfd; split; simpl; pcuic. - now exists t. + all: destruct l as (Hb & s & Ht & _); cbn in Hb, Ht. + all: now eexists. Qed. Lemma cumul_ctx_rel_cons {Pcmp Σ Γ Δ Δ' d d'} (c : cumul_ctx_rel Pcmp Σ Γ Δ Δ') @@ -748,8 +743,8 @@ Section Typecheck. Next Obligation. specialize_Σ wfΣ. sq. depelim wfΔ; simpl. - destruct l; eexists; eauto. - destruct l; split; eexists; eauto. + all: destruct l as (Hb & s' & Ht & _); cbn in Hb, Ht. + all: repeat (eexists; eauto). Qed. Next Obligation. pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. @@ -780,20 +775,20 @@ Section Typecheck. Qed. Equations check_alpha_ws_cumul_ctx Δ Δ' - : typing_result_comp (∥ eq_context_gen eq eq Δ Δ' ∥) := + : typing_result_comp (∥ eq_context_upto_names Δ Δ' ∥) := check_alpha_ws_cumul_ctx Δ Δ' with - inspect (forallb2 (bcompare_decls eqb eqb) Δ Δ') := { + inspect (eqb_context_upto_names Δ Δ') := { | @exist true e := ret _ ; | @exist false e' := raise (Msg "While checking alpha-conversion of contexts: contexts differ") }. Next Obligation. move: e. - elim: reflect_eq_ctx => //. + elim: reflect_eq_context_upto_names => //. Qed. Next Obligation. sq. move: e'. - elim: reflect_eq_ctx => //. + elim: reflect_eq_context_upto_names => //. Qed. (* Equations infer_terms Γ (wfΓ : ∥ wf_local Σ Γ ∥) ts @@ -865,7 +860,7 @@ Section Typecheck. 1: constructor. rewrite subst_empty. apply checking_typing ; auto. - apply wf_local_rel_app_inv in wfΔ as [wt _]. + apply All_local_rel_app_inv in wfΔ as [wt _]. now depelim wt. Qed. Next Obligation. @@ -1044,22 +1039,22 @@ Section Typecheck. Qed. Equations check_is_allowed_elimination - (u : Universe.t) (wfu : forall Σ (wfΣ : abstract_env_ext_rel X Σ), wf_universe Σ u) + (u : sort) (wfu : forall Σ (wfΣ : abstract_env_ext_rel X Σ), wf_sort Σ u) (al : allowed_eliminations) : typing_result_comp (forall Σ (wfΣ : abstract_env_ext_rel X Σ), is_allowed_elimination Σ al u) := check_is_allowed_elimination u wfu IntoSProp - with inspect (Universe.is_sprop u) := { + with inspect (Sort.is_sprop u) := { | @exist true e := ret _ ; | @exist false _ := raise (Msg "Cannot eliminate over this sort") } ; check_is_allowed_elimination u wfu IntoPropSProp - with inspect (is_propositional u) := { + with inspect (Sort.is_propositional u) := { | @exist true _ := ret _ ; | @exist false _ := raise (Msg "Cannot eliminate over this sort") }; check_is_allowed_elimination u wfu IntoSetPropSProp - with inspect (is_propositional u || abstract_env_eq X u Universe.type0) := { + with inspect (Sort.is_propositional u || abstract_env_compare_sort X Conv u Sort.type0) := { | @exist true _ := ret _ ; | @exist false _ := raise (Msg "Cannot eliminate over this sort") } ; @@ -1075,7 +1070,7 @@ Section Typecheck. Next Obligation. symmetry in e; toProp e; destruct e as [-> | e]; [auto|right]. specialize_Σ wfΣ; pose proof (heΣ _ wfΣ) as [heΣ]. - eapply abstract_env_compare_universe_correct with (conv_pb := Conv) in e; eauto using wf_universe_type0. + eapply abstract_env_compare_sort_correct with (conv_pb := Conv) in e; eauto using wf_sort_type0. Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]; specialize_Σ wfΣ; @@ -1084,7 +1079,7 @@ Section Typecheck. symmetry in e0; toProp e0; destruct e0 as [e1 e0]. destruct H as [H|H]; [rewrite H in e1; discriminate e1 | clear e1]. apply diff_false_true. rewrite -e0. - eapply abstract_env_compare_universe_correct with (conv_pb := Conv); eauto using wf_universe_type0. + eapply abstract_env_compare_sort_correct with (conv_pb := Conv); eauto using wf_sort_type0. Qed. Notation wt_brs Γ ci mdecl idecl p ptm ctors brs n := @@ -1096,7 +1091,7 @@ Section Typecheck. Section check_brs. Context (infer : forall (Γ : context) (HΓ : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ wf_local Σ Γ ∥) (t : term), typing_result_comp ({ A : term & forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ Σ ;;; Γ |- t ▹ A ∥ })) - (Γ : context) (wfΓ : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ wf_local Σ Γ ∥) (ps : Universe.t) + (Γ : context) (wfΓ : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ wf_local Σ Γ ∥) (ps : sort) (ci : case_info) (mdecl : mutual_inductive_body) (idecl : one_inductive_body) (p : predicate term) (args : list term). @@ -1111,7 +1106,7 @@ Section Typecheck. Lemma branch_helper n cdecl ctors br (isdecl' : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ Alli (fun i cdecl => declared_constructor Σ (ci, i) mdecl idecl cdecl) n (cdecl :: ctors) ∥) : - ∥ eq_context_gen eq eq (bcontext br) (cstr_branch_context ci mdecl cdecl) ∥ -> + ∥ eq_context_upto_names (bcontext br) (cstr_branch_context ci mdecl cdecl) ∥ -> forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ wf_branch cdecl br × let brctxty := case_branch_type ci.(ci_ind) mdecl idecl p br ptm n cdecl in wf_local Σ (Γ,,, brctxty.1) × Σ;;; Γ,,, brctxty.1 |- brctxty.2 ◃ tSort ps ∥. @@ -1120,7 +1115,7 @@ Section Typecheck. have wfbr : wf_branch cdecl br. { sq. specialize_Σ wfΣ. do 2 red. unfold cstr_branch_context, expand_lets_ctx, expand_lets_k_ctx in H. - move/eq_context_gen_binder_annot: H. + move/eq_context_upto_names_binder_annot: H. now do 3 move/eq_annots_fold. } assert (wfpret' : ∥Σ ;;; Γ ,,, predctx |- preturn p : tSort ps∥). { specialize_Σ wfΣ. sq. eapply infering_sort_typing ; eauto. @@ -1173,7 +1168,7 @@ Section Typecheck. eapply branch_helper in i; tea. pose proof (heΣ _ wfΣ) as [heΣ]. specialize_Σ wfΣ; sq. destruct i as [? []]. - exists ps. + eapply has_sort_isType. apply checking_typing ; eauto. eapply isType_Sort ; eauto. apply infering_sort_typing, validity, isType_Sort_inv in wfpret ; eauto. @@ -1186,7 +1181,7 @@ Section Typecheck. eapply branch_helper in i; tea. specialize_Σ wfΣ; sq. constructor ; tea. split. - * now eapply All2_fold_All2 in check_eq_bcontext. + * assumption. * now destruct i as [? []]. Qed. Next Obligation. @@ -1202,19 +1197,18 @@ Section Typecheck. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]; specialize_Σ wfΣ; sq. depelim H. - apply All2_fold_All2. - now inversion p0. + apply p0. Qed. End check_brs. Lemma eq_context_gen_wf_predicate p ci mdecl idecl : #|pparams p| = ind_npars mdecl -> - eq_context_gen eq eq (pcontext p) (ind_predicate_context ci mdecl idecl) -> + eq_context_upto_names (pcontext p) (ind_predicate_context ci mdecl idecl) -> wf_predicate mdecl idecl p. Proof using Type. intros eqp e. do 2 red. split => //. - eapply eq_context_gen_binder_annot in e. + eapply eq_context_upto_names_binder_annot in e. rewrite /ind_predicate_context in e. rewrite /idecl_binder. destruct (forget_types (pcontext p)); depelim e; cbn in *. @@ -1223,12 +1217,12 @@ Section Typecheck. Qed. Lemma eq_context_gen_wf_branch ci mdecl cdecl br : - eq_context_gen eq eq (bcontext br) (cstr_branch_context ci mdecl cdecl) -> + eq_context_upto_names (bcontext br) (cstr_branch_context ci mdecl cdecl) -> wf_branch cdecl br. Proof using Type. intros e. do 2 red. - eapply eq_context_gen_binder_annot in e. + eapply eq_context_upto_names_binder_annot in e. rewrite /cstr_branch_context in e. now do 3 eapply (proj1 (eq_annots_fold _ _ _)) in e. Qed. @@ -1251,7 +1245,7 @@ Section Typecheck. Next Obligation. pose proof (heΣ _ wfΣ) as [heΣ]. specialize_Σ wfΣ; sq. constructor ; tea. - exists s. + eapply has_sort_isType. now apply infering_sort_typing. Qed. Next Obligation. @@ -1317,15 +1311,15 @@ Section Typecheck. | primInt | decl := check_eq_true (eqb decl.(cst_body) None) (Msg "primitive type is registered to a defined constant") ;; check_eq_true (eqb decl.(cst_universes) Monomorphic_ctx) (Msg "primitive type is registered to a non-monomorphic constant") ;; - check_eq_true (eqb decl.(cst_type) (tSort Universe.type0)) (Msg "primitive type for integers is registered to an axiom whose type is not the sort Set") ;; + check_eq_true (eqb decl.(cst_type) (tSort Sort.type0)) (Msg "primitive type for integers is registered to an axiom whose type is not the sort Set") ;; ret _ | primFloat | decl := check_eq_true (eqb decl.(cst_body) None) (Msg "primitive type is registered to a defined constant") ;; check_eq_true (eqb decl.(cst_universes) Monomorphic_ctx) (Msg "primitive type for floats is registered to non-monomorphic constant") ;; - check_eq_true (eqb decl.(cst_type) (tSort Universe.type0)) (Msg "primitive type for floats is registered to an axiom whose type is not the sort Set") ;; + check_eq_true (eqb decl.(cst_type) (tSort Sort.type0)) (Msg "primitive type for floats is registered to an axiom whose type is not the sort Set") ;; ret _ | primArray | decl := - let s := Universe.make (Level.lvar 0) in + let s := sType (Universe.make' (Level.lvar 0)) in check_eq_true (eqb decl.(cst_body) None) (Msg "primitive type is registered to a defined constant") ;; check_eq_true (eqb decl.(cst_universes) (Polymorphic_ctx array_uctx)) (Msg "primitive type is registered to a monomorphic constant") ;; check_eq_true (eqb decl.(cst_type) (tImpl (tSort s) (tSort s))) (Msg "primitive type for arrays is registered to an axiom whose type is not of shape Type -> Type") ;; @@ -1369,8 +1363,8 @@ Section Typecheck. | (primInt; primIntModel i) := ret _ | (primFloat; primFloatModel f) := ret _ | (primArray; primArrayModel a) := - check_eq_true (abstract_env_ext_wf_universeb X (Universe.make a.(array_level))) (Msg "primitive array level is not well-formed") ;; - check_type <- bdcheck infer Γ wfΓ a.(array_type) (tSort (Universe.make a.(array_level))) _ ;; + check_eq_true (abstract_env_ext_wf_universeb X (Universe.make' a.(array_level))) (Msg "primitive array level is not well-formed") ;; + check_type <- bdcheck infer Γ wfΓ a.(array_type) (tSort (sType (Universe.make' a.(array_level)))) _ ;; check_default <- bdcheck infer Γ wfΓ a.(array_default) a.(array_type) _ ;; check_values <- make_All (fun x => bdcheck infer Γ wfΓ x a.(array_type) _) a.(array_value) ;; ret _. @@ -1381,24 +1375,24 @@ Section Typecheck. 1-2:do 2 constructor. - eauto. - sq. erewrite <- abstract_env_ext_wf_universeb_correct in i; tea. - eexists; eapply type_Sort; eauto. + eapply has_sort_isType; eapply type_Sort; eauto. now move/@wf_universe_reflect: i. - specialize (check_type _ wfΣ) as []. - sq; eapply checking_typing in X0; eauto. now eexists. + sq; eapply checking_typing in X0; eauto. now eapply has_sort_isType. erewrite <- abstract_env_ext_wf_universeb_correct in i; tea. - eexists; eapply type_Sort; eauto. + eapply has_sort_isType; eapply type_Sort; eauto. now move/@wf_universe_reflect: i. - specialize (check_type _ wfΣ) as []. - sq; eapply checking_typing in X0; eauto. now eexists. + sq; eapply checking_typing in X0; eauto. now eapply has_sort_isType. erewrite <- abstract_env_ext_wf_universeb_correct in i; tea. - eexists; eapply type_Sort; eauto. + eapply has_sort_isType; eapply type_Sort; eauto. now move/@wf_universe_reflect: i. - specialize (check_type _ wfΣ) as []. specialize (check_default _ wfΣ) as []. - assert (∥ Σ;;; Γ |- array_type a : tSort (Universe.make (array_level a)) ∥) as []. + assert (∥ Σ;;; Γ |- array_type a : tSort (sType (Universe.make' (array_level a))) ∥) as []. { sq. eapply checking_typing in X0; eauto. erewrite <- abstract_env_ext_wf_universeb_correct in i; tea. - eexists; eapply type_Sort; eauto. now move/@wf_universe_reflect: i. } + eapply has_sort_isType; eapply type_Sort; eauto. now move/@wf_universe_reflect: i. } assert (∥ All (fun x : term => Σ;;; Γ |- x ◃ array_type a) (array_value a) ∥). { induction check_values. - repeat constructor. @@ -1454,15 +1448,15 @@ Section Typecheck. infer Γ HΓ (tEvar ev _) := raise (UnboundEvar ev) ; - infer Γ HΓ (tSort u) with inspect (@abstract_env_ext_wf_universeb _ _ _ _ X u) := { - | exist true _ := ret (tSort (Universe.super u);_) ; - | exist false _ := raise (Msg ("Sort contains an undeclared level " ^ string_of_sort u)) + infer Γ HΓ (tSort s) with inspect (@abstract_env_ext_wf_sortb _ _ _ _ X s) := { + | exist true _ := ret (tSort (Sort.super s);_) ; + | exist false _ := raise (Msg ("Sort contains an undeclared level " ^ string_of_sort s)) } ; infer Γ HΓ (tProd na A B) := - s1 <- infer_type infer Γ HΓ A ;; + s1 <- infer_type infer Γ HΓ A ;; s2 <- infer_type infer (Γ,,vass na A) _ B ;; - Checked_comp (tSort (Universe.sort_of_product s1.π1 s2.π1);_) ; + Checked_comp (tSort (Sort.sort_of_product s1.π1 s2.π1);_) ; infer Γ HΓ (tLambda na A t) := infer_type infer Γ HΓ A ;; @@ -1527,8 +1521,7 @@ Section Typecheck. let chop_args := chop ci.(ci_npar) args in let params := chop_args.1 in let indices := chop_args.2 in cu <- check_consistent_instance (ind_universes mdecl) _ p.(puinst) ;; - check_eq_true (abstract_env_compare_global_instance _ X (abstract_env_leq X) (IndRef ind') - #|args| u p.(puinst)) + check_eq_true (abstract_env_compare_global_instance _ X Cumul (IndRef ind') #|args| u p.(puinst)) (Msg "invalid universe annotation on case, not larger than the discriminee's universes") ;; wt_params <- check_inst infer Γ HΓ (List.rev (smash_context [] (ind_params mdecl))@[p.(puinst)]) _ _ p.(pparams) ;; eq_params <- check_ws_cumul_pb_terms Γ params p.(pparams) _ _ ;; @@ -1608,15 +1601,15 @@ Section Typecheck. Next Obligation. specialize_Σ wfΣ; sq. symmetry in e. - erewrite <- abstract_env_ext_wf_universeb_correct in e; eauto. - eapply (elimT wf_universe_reflect) in e. + erewrite <- abstract_env_ext_wf_sortb_correct in e; eauto. + move: e => /wf_sort_reflect e. sq; econstructor; tas. Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. specialize_Σ wfΣ; sq. - inversion X1 ; subst. erewrite <- abstract_env_ext_wf_universeb_correct in e0; eauto. - move: H0 e0 => /wf_universe_reflect -> //. + inversion X1 ; subst. erewrite <- abstract_env_ext_wf_sortb_correct in e0; eauto. + move: H0 e0 => /wf_sort_reflect -> //. Qed. (* tProd *) Next Obligation. @@ -1628,6 +1621,7 @@ Section Typecheck. Next Obligation. (* intros Γ HΓ t na A B Heq_t [s1 ?] [s2 ?]; *) cbn. specialize_Σ wfΣ. sq. econstructor; try eassumption. + repeat (eexists; tea). Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. @@ -1638,8 +1632,8 @@ Section Typecheck. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. specialize_Σ wfΣ; sq. inversion X1; subst. apply absurd. - eexists. intros. sq. erewrite (abstract_env_ext_irr _ _ wfΣ); eauto. - Unshelve. all: eauto. + eexists. intros. sq. unshelve erewrite (abstract_env_ext_irr _ _ wfΣ); eauto. + now eapply unlift_TypUniv. Qed. (* tLambda *) Next Obligation. @@ -1651,7 +1645,8 @@ Section Typecheck. Next Obligation. (* intros Γ HΓ t0 na A t Heq_t [s ?] [B ?]; *) cbn; pose proof (heΣ _ wfΣ) as [heΣ]. specialize_Σ wfΣ. - sq; econstructor; eassumption. + sq; econstructor; tas. + repeat (eexists; tea). Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. @@ -1662,24 +1657,27 @@ Section Typecheck. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. specialize_Σ wfΣ; sq. inversion X1; subst. apply absurd. - eexists. intros. sq. erewrite (abstract_env_ext_irr _ _ wfΣ); eauto. - Unshelve. all: eauto. + destruct X2 as (_ & u & Hu & _). + eexists. intros. sq. unshelve erewrite (abstract_env_ext_irr _ _ wfΣ); eauto. Qed. (* tLetIn *) Next Obligation. pose proof (heΣ _ wfΣ) as [heΣ]. specialize_Σ wfΣ. - sq; econstructor; tea. + sq. eapply has_sort_isType. eapply infering_sort_typing ; eauto. Qed. Next Obligation. (* intros Γ HΓ t n b b_ty b' Heq_t [? ?] H0; *) pose proof (heΣ _ wfΣ) as [heΣ]. specialize_Σ wfΣ ; sq. econstructor ; tea. - 2: apply checking_typing ; eauto. - all: now eapply infering_sort_isType. + repeat (eexists; tea); cbn. + 1: apply checking_typing ; eauto. + 1: eapply has_sort_isType. + all: now eapply infering_sort_typing. Qed. Next Obligation. - cbn; specialize_Σ wfΣ; sq; econstructor; eauto. + cbn; specialize_Σ wfΣ; sq; econstructor; tas. + repeat (eexists; tea). Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. @@ -1690,14 +1688,14 @@ Section Typecheck. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. specialize_Σ wfΣ; sq. inversion X1; subst. apply absurd. - eexists. intros. sq. erewrite (abstract_env_ext_irr _ _ wfΣ); eauto. - Unshelve. all: eauto. + destruct X3 as (Ht & u & Hu & _). + eexists. intros. sq. unshelve erewrite (abstract_env_ext_irr _ _ wfΣ); eauto. Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. specialize_Σ wfΣ; sq. inversion X1; subst. apply absurd. - eexists. intros. sq. erewrite (abstract_env_ext_irr _ _ wfΣ); eauto. - Unshelve. all: eauto. + destruct X2 as (Ht & u & Hu & _). + eexists. intros. sq. unshelve erewrite (abstract_env_ext_irr _ _ wfΣ); eauto. Qed. (* tApp *) Next Obligation. @@ -1712,10 +1710,10 @@ Section Typecheck. specialize_Σ wfΣ ; sq. eapply infering_typing, type_reduction_closed, validity in X3. 2-4: eauto. - destruct X3 as [s HH]. + destruct X3 as (_ & s & HH & _). eapply inversion_Prod in HH ; auto. destruct HH as [s1 [_ [HH _]]]. - eexists. eassumption. + eapply lift_sorting_forget_univ. eassumption. Qed. Next Obligation. cbn in *; specialize_Σ wfΣ ; sq. @@ -1877,7 +1875,7 @@ Section Typecheck. (* tCase *) Next Obligation. cbn in *. pose proof (heΣ _ wfΣ) as [heΣ]. specialize_Σ wfΣ ; sq. - eapply infering_typing, validity in X0 as []; eauto. + eapply infering_typing, validity in X0 as (_ & ? & ? & _); eauto. eexists; eauto using validity_wf. Qed. Next Obligation. @@ -1929,7 +1927,7 @@ Section Typecheck. unshelve epose proof (ctx_inst_spine_subst _ wt_params). { eapply weaken_wf_local; eauto. eapply on_minductive_wf_params; eauto. now destruct X1. } - eexists; eapply isType_mkApps_Ind_smash; tea. + eapply has_sort_isType, isType_mkApps_Ind_smash; tea. rewrite subst_instance_app List.rev_app_distr smash_context_app_expand. have wf_ctx : wf_local Σ (Γ,,, smash_context [] (ind_params d)@[puinst p],,, @@ -1974,9 +1972,6 @@ Section Typecheck. rewrite -(spine_subst_inst_subst X4). rewrite - !smash_context_subst /= !subst_context_nil. erewrite <- compare_global_instance_correct in i1; eauto. - 2: intros; eapply iff_reflect; - eapply abstract_env_compare_universe_correct with (conv_pb := Cumul); - try eassumption; apply wf_universe_iff; eauto. 2: { eapply consistent_instance_ext_wf; eauto. } 2: { eapply consistent_instance_ext_wf; eauto. } eapply (inductive_cumulative_indices X1); tea. @@ -2033,10 +2028,7 @@ Section Typecheck. Next Obligation. intros. cbn in *. - destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. - pose proof (heΣ _ wfΣ) as [heΣ]. - cbn in *. specialize_Σ wfΣ ; sq. - now eapply All2_fold_All2 in check_wfpctx_conv. + assumption. Qed. Next Obligation. @@ -2067,23 +2059,18 @@ Section Typecheck. induction check_brs; constructor; auto. destruct r0. solve_all. - eapply eq_context_gen_wf_branch. - now eapply All2_fold_All2. + now eapply eq_context_gen_wf_branch. } econstructor ; eauto. - econstructor ; tea. now apply closed_red_red. - - now eapply All2_fold_All2 in check_wfpctx_conv. - - now eapply wf_local_rel_wf_local_bd, wf_local_app_inv, wf_case_predicate_context. + - now eapply wf_local_rel_wf_local_bd, All_local_env_app_inv, wf_case_predicate_context. - eapply ctx_inst_typing_bd ; eauto. eapply ctx_inst_smash. now rewrite subst_instance_smash /= in wt_params. - now eapply negbTE. - erewrite <- compare_global_instance_correct in i1; eauto. - 1: intros; eapply iff_reflect; - eapply abstract_env_compare_universe_correct with (conv_pb := Cumul); - try eassumption; apply wf_universe_iff; eauto. - 1: { apply/wf_universe_instanceP. + 1: { apply/wf_instanceP. rewrite -wf_universeb_instance_forall. assert (tyu : isType Σ Γ (mkApps (tInd ci u) args)). { @@ -2106,10 +2093,9 @@ Section Typecheck. eapply wf_case_branches_types' ; tea. - apply infering_sort_typing ; eauto. now eapply wf_case_predicate_context. - - now eapply All2_fold_All2. } cbn ; intros ? ? ? [? []] ; intuition auto. - now eapply wf_local_rel_wf_local_bd, wf_local_app_inv. + now eapply wf_local_rel_wf_local_bd, All_local_env_app_inv. Qed. Next Obligation. @@ -2192,7 +2178,7 @@ Section Typecheck. subst. apply absurd. sq. - now eapply All2_fold_All2. + assumption. Qed. Next Obligation. @@ -2273,10 +2259,7 @@ Section Typecheck. erewrite All2_length. 2: eassumption. erewrite <- All2_length ; tea. - - intros. eapply iff_reflect. - eapply abstract_env_compare_universe_correct with (conv_pb := Cumul); - try eassumption; apply wf_universe_iff; eauto. - - apply/wf_universe_instanceP. + - apply/wf_instanceP. rewrite -wf_universeb_instance_forall. assert (tyu : isType Σ Γ (mkApps (tInd ind' u) args)). { @@ -2290,7 +2273,7 @@ Section Typecheck. - apply infering_typing, typing_wf_universes in ty ; auto. move: ty => /andP []. rewrite {1}/wf_universes /= wf_universeb_instance_forall => - /andP [] /wf_universe_instanceP; eauto. + /andP [] /wf_instanceP; eauto. Qed. Next Obligation. @@ -2582,9 +2565,19 @@ Section Typecheck. unfold abstract_env_fixguard in guarded. erewrite <- abstract_env_guard_correct in guarded; eauto. constructor; auto. - eapply All_impl ; tea. + eapply All_impl ; tea; cbn. intros. - now apply isType_infering_sort. + eapply lift_sorting_ex_it_impl_gen with X0 => //= H. + eapply typing_infering_sort in H as (? & ? & _). now eexists. + apply All_mix with (2 := wf_types) in wf_bodies. + apply All_mfix_wf in wf_types; eauto. + eapply All_impl with (1 := wf_bodies) ; cbn. + intros d (X0 & (_ & s & X1 & _)). + assert (Σ ;;; Γ ,,, fix_context mfix |- lift0 #|fix_context mfix| (dtype d) : lift0 #|fix_context mfix| (tSort s)). + { apply weakening with (Γ' := fix_context mfix); eauto. } + eapply checking_typing in X0 as X0'; eauto. 2: now eapply has_sort_isType. + eapply typing_infering_sort in X2 as (? & ? & _). + repeat (eexists; tea). erewrite abstract_wf_fixpoint in wffix; eauto. Qed. Next Obligation. @@ -2604,17 +2597,17 @@ Section Typecheck. apply absurd; intros. pose proof (heΣ _ wfΣ) as [heΣ]. cbn in *. specialize_Σ wfΣ ; sq. - now inversion X1. + inversion X1. + apply All_impl with (1 := X3) => d [] //. Qed. Next Obligation. apply absurd; intros. pose proof (heΣ _ wfΣ) as [heΣ]. cbn in *. specialize_Σ wfΣ ; sq. inversion X1 ; subst. - eapply All_impl. - 1: eexact X2. - intros. - now eapply einfering_sort_isType. + eapply All_impl with (1 := X2). + intros d (_ & ? & ? & _). + now eapply infering_sort_isType. Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. @@ -2634,9 +2627,19 @@ Section Typecheck. unfold abstract_env_cofixguard in guarded. erewrite <- abstract_env_guard_correct in guarded; eauto. constructor; auto. - eapply All_impl ; tea. + eapply All_impl ; tea; cbn. intros. - now apply isType_infering_sort. + eapply lift_sorting_ex_it_impl_gen with X0 => //= H. + eapply typing_infering_sort in H as (? & ? & _). now eexists. + apply All_mix with (2 := wf_types) in wf_bodies. + apply All_mfix_wf in wf_types; eauto. + eapply All_impl with (1 := wf_bodies) ; cbn. + intros d (X0 & (_ & s & X1 & _)). + assert (Σ ;;; Γ ,,, fix_context mfix |- lift0 #|fix_context mfix| (dtype d) : lift0 #|fix_context mfix| (tSort s)). + { apply weakening with (Γ' := fix_context mfix); eauto. } + eapply checking_typing in X0 as X0'; eauto. 2: now eapply has_sort_isType. + eapply typing_infering_sort in X2 as (? & ? & _). + repeat (eexists; tea). erewrite abstract_wf_cofixpoint in wfcofix; eauto. Qed. Next Obligation. @@ -2656,17 +2659,17 @@ Section Typecheck. apply absurd; intros. pose proof (heΣ _ wfΣ) as [heΣ]. cbn in *. specialize_Σ wfΣ ; sq. - now inversion X1. + inversion X1. + apply All_impl with (1 := X3) => d [] //. Qed. Next Obligation. apply absurd; intros. pose proof (heΣ _ wfΣ) as [heΣ]. cbn in *. specialize_Σ wfΣ ; sq. inversion X1 ; subst. - eapply All_impl. - 1: eexact X2. - intros. - now eapply einfering_sort_isType. + eapply All_impl with (1 := X2). + intros d (_ & ? & ? & _). + now eapply infering_sort_isType. Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. diff --git a/safechecker/theories/PCUICWfEnv.v b/safechecker/theories/PCUICWfEnv.v index 97278d1a5..d46644153 100644 --- a/safechecker/theories/PCUICWfEnv.v +++ b/safechecker/theories/PCUICWfEnv.v @@ -48,17 +48,13 @@ Definition abstract_env_fixguard {cf:checker_flags} {abstract_env_impl abstract Definition abstract_env_cofixguard {cf:checker_flags} {abstract_env_impl abstract_env_ext_impl : Type} `{!abstract_env_struct abstract_env_impl abstract_env_ext_impl} (X:abstract_env_ext_impl) := abstract_env_guard X CoFix. -Definition abstract_env_conv_pb_relb {cf:checker_flags} {abstract_env_impl abstract_env_ext_impl : Type} `{!abstract_env_struct abstract_env_impl abstract_env_ext_impl} +Definition abstract_env_compare_universe {cf:checker_flags} {abstract_env_impl abstract_env_ext_impl : Type} `{!abstract_env_struct abstract_env_impl abstract_env_ext_impl} (X:abstract_env_ext_impl) : conv_pb -> Universe.t -> Universe.t -> bool := - fun conv_pb => conv_pb_relb_gen conv_pb - (check_eqb_universe_gen (abstract_env_leqb_level_n X)) - (check_leqb_universe_gen (abstract_env_leqb_level_n X)). + check_cmpb_universe_gen (abstract_env_leqb_level_n X). -Definition abstract_env_eq {cf:checker_flags} {abstract_env_impl abstract_env_ext_impl : Type} `{!abstract_env_struct abstract_env_impl abstract_env_ext_impl} - (X:abstract_env_ext_impl) := abstract_env_conv_pb_relb X Conv. - -Definition abstract_env_leq {cf:checker_flags} {abstract_env_impl abstract_env_ext_impl : Type} `{!abstract_env_struct abstract_env_impl abstract_env_ext_impl} - (X:abstract_env_ext_impl) := abstract_env_conv_pb_relb X Cumul. +Definition abstract_env_compare_sort {cf:checker_flags} {abstract_env_impl abstract_env_ext_impl : Type} `{!abstract_env_struct abstract_env_impl abstract_env_ext_impl} + (X:abstract_env_ext_impl) : conv_pb -> sort -> sort -> bool := + check_cmpb_sort_gen (abstract_env_leqb_level_n X). Definition abstract_env_check_constraints {cf:checker_flags} {abstract_env_impl abstract_env_ext_impl : Type} `{!abstract_env_struct abstract_env_impl abstract_env_ext_impl} (X:abstract_env_ext_impl) : ConstraintSet.t -> bool := @@ -66,11 +62,11 @@ Definition abstract_env_check_constraints {cf:checker_flags} {abstract_env_impl Definition abstract_env_ext_wf_universeb {cf:checker_flags} {abstract_env_impl abstract_env_ext_impl : Type} `{!abstract_env_struct abstract_env_impl abstract_env_ext_impl} (X:abstract_env_ext_impl) : Universe.t -> bool := - fun (s : Universe.t) => - match s with - | Universe.lType l => LevelExprSet.for_all (fun l => abstract_env_level_mem X (LevelExpr.get_level l)) l - | _ => true - end. + fun (u : Universe.t) => LevelExprSet.for_all (fun l => abstract_env_level_mem X (LevelExpr.get_level l)) u. + +Definition abstract_env_ext_wf_sortb {cf:checker_flags} {abstract_env_impl abstract_env_ext_impl : Type} `{!abstract_env_struct abstract_env_impl abstract_env_ext_impl} +(X:abstract_env_ext_impl) : sort -> bool := + Sort.on_sort (abstract_env_ext_wf_universeb X) true. Definition abstract_env_level_mem' {cf:checker_flags} {abstract_env_impl abstract_env_ext_impl : Type} `{!abstract_env_struct abstract_env_impl abstract_env_ext_impl} (X:abstract_env_ext_impl) : LevelSet.t -> Level.t -> bool := @@ -221,25 +217,35 @@ Definition abstract_env_is_consistent_empty {cf:checker_flags} {X_type : abstrac Lemma abstract_env_compare_universe_correct {cf:checker_flags} {X_type : abstract_env_impl} (X:X_type.π2.π1) {Σ} (wfΣ : abstract_env_ext_rel X Σ) conv_pb u u' : wf_universe Σ u -> wf_universe Σ u' -> - compare_universe conv_pb Σ u u' <-> - abstract_env_conv_pb_relb X conv_pb u u'. + compare_universe Σ conv_pb u u' <-> + abstract_env_compare_universe X conv_pb u u'. Proof. intros wfu wfu'. pose proof (abstract_env_ext_wf X wfΣ). sq. pose (Hleq := abstract_env_leqb_level_n_correct X wfΣ). erewrite uctx'_eq in Hleq. - destruct conv_pb; split; cbn; intro. - - eapply (check_eqb_universe_complete_gen _ (global_ext_levels Σ, global_ext_constraints Σ)); eauto. - + eapply wf_ext_global_uctx_invariants; eauto. - + eapply wf_ext_consistent; eauto. - - eapply (check_eqb_universe_spec_gen' _ (global_ext_levels Σ, global_ext_constraints Σ)); eauto. - + eapply wf_ext_global_uctx_invariants; eauto. - + eapply wf_ext_consistent; eauto. - - eapply (check_leqb_universe_complete_gen _ (global_ext_levels Σ, global_ext_constraints Σ)); eauto. - + eapply wf_ext_global_uctx_invariants; eauto. - + eapply wf_ext_consistent; eauto. - - eapply (check_leqb_universe_spec_gen' _ (global_ext_levels Σ, global_ext_constraints Σ)); eauto. - + eapply wf_ext_global_uctx_invariants; eauto. - + eapply wf_ext_consistent; eauto. + eapply compare_universeP_gen with (pb := conv_pb) in Hleq. + apply reflect_reflectT in Hleq. + split. + 1: now eapply introT. + 1: now eapply elimT. + all: tea. +Qed. + +Lemma abstract_env_compare_sort_correct {cf:checker_flags} {X_type : abstract_env_impl} + (X:X_type.π2.π1) {Σ} (wfΣ : abstract_env_ext_rel X Σ) conv_pb s s' : + wf_sort Σ s -> wf_sort Σ s' -> + compare_sort Σ conv_pb s s' <-> + abstract_env_compare_sort X conv_pb s s'. +Proof. + intros wfu wfu'. pose proof (abstract_env_ext_wf X wfΣ). sq. + pose (Hleq := abstract_env_leqb_level_n_correct X wfΣ). + erewrite uctx'_eq in Hleq. + eapply compare_sortP_gen with (pb := conv_pb) in Hleq. + apply reflect_reflectT in Hleq. + split. + 1: now eapply introT. + 1: now eapply elimT. + all: tea. Qed. Lemma check_constraints_spec {cf:checker_flags} {X_type : abstract_env_impl} @@ -286,7 +292,7 @@ Lemma abstract_env_ext_wf_universeb_correct {cf:checker_flags} {X_type : abstrac ( X:X_type.π2.π1) {Σ} (wfΣ : abstract_env_ext_rel X Σ) u : wf_universeb Σ u = abstract_env_ext_wf_universeb X u. Proof. - destruct u => //; destruct t as [ [] ?]; cbn. clear is_ok t_ne. + destruct u as [ [] ?]; cbn. clear is_ok t_ne. induction this => //; cbn. set (b := LevelSet.Raw.mem _ _). set (b' := abstract_env_level_mem _ _). assert (Hbb' : b = b'). { unfold b'. apply eq_true_iff_eq. rewrite <- (abstract_env_level_mem_correct X wfΣ (LevelExpr.get_level a)). @@ -294,6 +300,14 @@ Proof. destruct Hbb', b => //. Qed. +Lemma abstract_env_ext_wf_sortb_correct {cf:checker_flags} {X_type : abstract_env_impl} +( X:X_type.π2.π1) {Σ} (wfΣ : abstract_env_ext_rel X Σ) s : + wf_sortb Σ s = abstract_env_ext_wf_sortb X s. +Proof. + destruct s as [| |u]; cbnr. + now apply abstract_env_ext_wf_universeb_correct. +Qed. + Lemma abstract_env_level_mem_correct' {cf:checker_flags} {X_type : abstract_env_impl} ( X:X_type.π2.π1) {Σ} (wfΣ : abstract_env_ext_rel X Σ) levels u : LevelSet.mem u (LevelSet.union levels (global_ext_levels Σ)) = abstract_env_level_mem' X levels u. diff --git a/safechecker/theories/PCUICWfReduction.v b/safechecker/theories/PCUICWfReduction.v index 7f0b46c5d..1452b7933 100644 --- a/safechecker/theories/PCUICWfReduction.v +++ b/safechecker/theories/PCUICWfReduction.v @@ -144,7 +144,7 @@ Section fix_sigma. Qed. Lemma term_subterm_redp {Γ s s' t} {ts : term_subterm s t} : - redp Σ (Γ ,,, term_subterm_context ts) s s' -> + redp Σ (Γ ,,, term_subterm_context ts : context) s s' -> exists t', ∥ redp Σ Γ t t' × ∑ ts' : term_subterm s' t', term_subterm_context ts' = term_subterm_context ts ∥. Proof using Type. intros r. diff --git a/template-coq/src/ast_denoter.ml b/template-coq/src/ast_denoter.ml index a102c550a..9dd801273 100644 --- a/template-coq/src/ast_denoter.ml +++ b/template-coq/src/ast_denoter.ml @@ -17,7 +17,7 @@ struct type quoted_name = name type quoted_aname = name binder_annot type quoted_relevance = relevance - type quoted_sort = Universes0.Universe.t + type quoted_sort = Universes0.Sort.t type quoted_cast_kind = cast_kind type quoted_kernel_name = kername type quoted_inductive = inductive @@ -211,19 +211,20 @@ struct let unquote_level_expr (trm : Universes0.Level.t * quoted_int) : Univ.Universe.t = let l = unquote_level (fst trm) in let u = Univ.Universe.make l in - let n = unquote_int (snd trm) in - if n > 0 then Univ.Universe.super u else u + Caml_nat.iter_nat Univ.Universe.super u (snd trm) - let unquote_universe evd (trm : Universes0.Universe.t) = + let unquote_universe (trm : Universes0.Universe.t) = + let u = Universes0.t_set trm in + let ux_list = Universes0.LevelExprSet.elements u in + let l = List.map unquote_level_expr ux_list in + let u = List.fold_left Univ.Universe.sup (List.hd l) (List.tl l) in + u + + let unquote_sort evd trm = match trm with - | Universes0.Universe.Coq_lSProp -> evd, Sorts.sprop - | Universes0.Universe.Coq_lProp -> evd, Sorts.prop - | Universes0.Universe.Coq_lType u -> - let u = Universes0.t_set u in - let ux_list = Universes0.LevelExprSet.elements u in - let l = List.map unquote_level_expr ux_list in - let u = List.fold_left Univ.Universe.sup (List.hd l) (List.tl l) in - evd, Sorts.sort_of_univ u + | Universes0.Sort.Coq_sSProp -> evd, Sorts.sprop + | Universes0.Sort.Coq_sProp -> evd, Sorts.prop + | Universes0.Sort.Coq_sType u -> evd, Sorts.sort_of_univ (unquote_universe u) let unquote_universe_level evm l = evm, unquote_level l let unquote_universe_instance(evm: Evd.evar_map) (l: quoted_univ_instance): Evd.evar_map * Univ.Instance.t diff --git a/template-coq/src/ast_quoter.ml b/template-coq/src/ast_quoter.ml index 14462a188..2a987dfb5 100644 --- a/template-coq/src/ast_quoter.ml +++ b/template-coq/src/ast_quoter.ml @@ -18,7 +18,7 @@ struct type quoted_name = BasicAst.name type quoted_aname = BasicAst.aname type quoted_relevance = BasicAst.relevance - type quoted_sort = Universes0.Universe.t + type quoted_sort = Universes0.Sort.t type quoted_cast_kind = cast_kind type quoted_kernel_name = Kernames.kername type quoted_inductive = inductive @@ -79,32 +79,26 @@ struct let quote_float64 x = x - (* NOTE: fails if it hits Prop or SProp *) - let quote_nonprop_level (l : Univ.Level.t) : Universes0.Level.t = + let quote_level (l : Univ.Level.t) : Universes0.Level.t = if Univ.Level.is_set l then Universes0.Level.Coq_lzero else match Univ.Level.var_index l with | Some x -> Universes0.Level.Coq_lvar (quote_int x) | None -> Universes0.Level.Coq_level (quote_string (Univ.Level.to_string l)) - let quote_level (l : Univ.Level.t) : (Universes0.PropLevel.t, Universes0.Level.t) Datatypes.sum = - try Coq_inr (quote_nonprop_level l) - with e -> assert false + let quote_level_expr (lvl, shft) : Universes0.LevelExpr.t = quote_level lvl, quote_int shft - let quote_universe s : Universes0.Universe.t = - match Univ.Universe.level s with - Some l -> Universes0.Universe.of_levels (quote_level l) + let quote_universe u : Universes0.Universe.t = + match Univ.Universe.level u with + Some l -> Universes0.Universe.make' (quote_level l) | _ -> - let univs = List.map (fun (l,i) -> - match quote_level l with - | Coq_inl lprop -> assert false - | Coq_inr ql -> (ql, i > 0)) (Univ.Universe.repr s) in - Universes0.Universe.from_kernel_repr (List.hd univs) (List.tl univs) + let levels = Univ.Universe.repr u |> List.map quote_level_expr in + Universes0.Universe.from_kernel_repr (List.hd levels) (List.tl levels) let quote_sort s = match s with - | Sorts.Set -> quote_universe Univ.Universe.type0 - | Sorts.SProp -> Universes0.Universe.of_levels (Coq_inl Universes0.PropLevel.Coq_lSProp) - | Sorts.Prop -> Universes0.Universe.of_levels (Coq_inl Universes0.PropLevel.Coq_lProp) - | Sorts.Type u -> quote_universe u + | Sorts.Set -> Universes0.Sort.type0 + | Sorts.SProp -> Universes0.Sort.Coq_sSProp + | Sorts.Prop -> Universes0.Sort.Coq_sProp + | Sorts.Type u -> Universes0.Sort.Coq_sType (quote_universe u) let quote_sort_family s = match s with @@ -154,14 +148,13 @@ struct | _ -> false let quote_univ_constraint ((l, ct, l') : Univ.univ_constraint) : quoted_univ_constraint = - try ((quote_nonprop_level l, quote_constraint_type ct), quote_nonprop_level l') + try ((quote_level l, quote_constraint_type ct), quote_level l') with e -> assert false - let quote_univ_level = quote_nonprop_level + let quote_univ_level = quote_level let quote_univ_instance (i : Univ.Instance.t) : quoted_univ_instance = let arr = Univ.Instance.to_array i in - (* we assume that valid instances do not contain [Prop] or [SProp] *) - try CArray.map_to_list quote_nonprop_level arr + try CArray.map_to_list quote_level arr with e -> assert false (* (Prop, Le | Lt, l), (Prop, Eq, Prop) -- trivial, (l, c, Prop) -- unsatisfiable *) @@ -188,12 +181,7 @@ struct (names, (quote_univ_instance levels, quote_univ_constraints constraints)) let quote_univ_contextset (uctx : Univ.ContextSet.t) : quoted_univ_contextset = - (* CHECKME: is is safe to assume that there will be no Prop or SProp? *) - let levels = filter_map - (fun l -> match quote_level l with - | Coq_inl _ -> None - | Coq_inr l -> Some l) - (Univ.Level.Set.elements (Univ.ContextSet.levels uctx)) in + let levels = List.map quote_level (Univ.Level.Set.elements (Univ.ContextSet.levels uctx)) in let constraints = Univ.ContextSet.constraints uctx in (Universes0.LevelSetProp.of_list levels, quote_univ_constraints constraints) diff --git a/template-coq/src/caml_nat.ml b/template-coq/src/caml_nat.ml index 5a78380a1..729361a24 100644 --- a/template-coq/src/caml_nat.ml +++ b/template-coq/src/caml_nat.ml @@ -1,11 +1,11 @@ -let nat_of_caml_int i = +let nat_of_caml_int i = let rec aux acc i = if i < 0 then acc else aux (Datatypes.S acc) (i - 1) in aux Datatypes.O (i - 1) -let rec caml_int_of_nat_aux n acc = - match n with +let rec iter_nat f acc = function | Datatypes.O -> acc - | Datatypes.S x -> caml_int_of_nat_aux x (succ acc) -let caml_int_of_nat n = caml_int_of_nat_aux n 0 + | Datatypes.S x -> iter_nat f (f acc) x + +let caml_int_of_nat n = iter_nat succ 0 n diff --git a/template-coq/src/caml_nat.mli b/template-coq/src/caml_nat.mli index b11cab111..72d7f29e5 100644 --- a/template-coq/src/caml_nat.mli +++ b/template-coq/src/caml_nat.mli @@ -1,3 +1,4 @@ val caml_int_of_nat : Datatypes.nat -> int -val nat_of_caml_int : int -> Datatypes.nat \ No newline at end of file +val iter_nat : ('a -> 'a) -> 'a -> Datatypes.nat -> 'a +val nat_of_caml_int : int -> Datatypes.nat diff --git a/template-coq/src/constr_denoter.ml b/template-coq/src/constr_denoter.ml index 675be7ec1..6e4c36987 100644 --- a/template-coq/src/constr_denoter.ml +++ b/template-coq/src/constr_denoter.ml @@ -213,55 +213,66 @@ struct else not_supported_verb trm "unquote_level" + let unquote_universe_level = unquote_level + + let iter_int f a n = + if n < 0 then failwith "Negative number of iterations" + else + let rec loop f a = function 0 -> a | n -> loop f (f a) (n-1) + in + loop f a n + + let unquote_univ_expr evm trm (* of type LevelExpr.t *) : Evd.evar_map * Univ.Universe.t = let (h,args) = app_full trm [] in if constr_equall h c_pair then let l, b = unquote_pair trm in let evm, l' = unquote_level evm l in let u = Univ.Universe.make l' in - evm, if unquote_nat b > 0 then Univ.Universe.super u else u + evm, iter_int Univ.Universe.super u (unquote_nat b) else not_supported_verb trm "unquote_univ_expr" - let unquote_universe evm trm (* of type universe *) = + let unquote_universe evm trm (* of type universe *) = let (h,args) = app_full trm [] in - if constr_equall h lSProp then - match args with - | [] -> evm, Sorts.sprop - | _ -> bad_term_verb trm "unquote_univ_expr" - else if constr_equall h lProp then - match args with - | [] -> evm, Sorts.prop - | _ -> bad_term_verb trm "unquote_univ_expr" - else if constr_equall h lnpe then - match args with - | [x] -> - let (h,args) = app_full x [] in - if constr_equall h tBuild_Universe then - (match args with - | x :: _ :: [] -> - (let (h,args) = app_full x [] in - if constr_equall h tMktLevelExprSet then - match args with - | x :: _ :: [] -> - (match unquote_list x with - | [] -> assert false - | e::q -> - let evm, u = 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 - in evm, Sorts.sort_of_univ u) - | _ -> bad_term_verb trm "unquote_universe 0" - else - not_supported_verb trm "unquote_universe 0") - | _ -> bad_term_verb trm "unquote_universe 1") - else not_supported_verb trm "unquote_universe 2" - | _ -> bad_term_verb trm "unquote_universe 3" - else bad_term_verb trm "unquote_universe 4" - - - let unquote_universe_level evm trm = unquote_level evm trm + if constr_equall h tBuild_Universe then + (match args with + | x :: _ :: [] -> + (let (h,args) = app_full x [] in + if constr_equall h tMktLevelExprSet then + match args with + | x :: _ :: [] -> + (match unquote_list x with + | [] -> assert false + | e::q -> + let evm, u = 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 + in evm, u) + | _ -> bad_term_verb trm "unquote_universe 0" + else + not_supported_verb trm "unquote_universe 0") + | _ -> bad_term_verb trm "unquote_universe 1") + else not_supported_verb trm "unquote_universe 2" + + let unquote_sort evm trm (* of type sort *) = + let (h,args) = app_full trm [] in + if constr_equall h sSProp then + match args with + | [_univ] -> evm, Sorts.sprop + | _ -> bad_term_verb trm "unquote_sort_sprop_args" + else if constr_equall h sProp then + match args with + | [_univ] -> evm, Sorts.prop + | _ -> bad_term_verb trm "unquote_sort_prop_args" + else if constr_equall h sType then + match args with + | [_univ; x] -> + let evm, u = unquote_universe evm x in + evm, Sorts.sort_of_univ u + | _ -> bad_term_verb trm "unquote_sort_type_too_many_args" + else bad_term_verb trm "unquote_sort_type_no_args" let unquote_universe_instance evm trm (* of type universe_instance *) = let l = unquote_list trm in diff --git a/template-coq/src/constr_quoter.ml b/template-coq/src/constr_quoter.ml index b007b6655..1bab1577f 100644 --- a/template-coq/src/constr_quoter.ml +++ b/template-coq/src/constr_quoter.ml @@ -194,27 +194,23 @@ struct let string_of_level s = to_string (Univ.Level.to_string s) - let quote_nonprop_level l = + let quote_level l = if Level.is_set l then Lazy.force lzero else match Level.var_index l with | Some x -> constr_mkApp (tLevelVar, [| quote_int x |]) | None -> constr_mkApp (tLevel, [| string_of_level l |]) - let quote_level l = - Tm_util.debug (fun () -> str"quote_level " ++ Level.pr l); - constr_mkApp (cInr, [|Lazy.force tproplevel;Lazy.force tlevel; quote_nonprop_level l |]) - let quote_universe s = match Univ.Universe.level s with - Some l -> constr_mkApp (tof_levels, [| quote_level l |]) - | _ -> let levels = List.map (fun (l,i) -> pairl tlevel bool_type (quote_nonprop_level l) (quote_bool (i > 0))) (Universe.repr s) in + Some l -> constr_mkApp (tof_level, [| quote_level l |]) + | _ -> let levels = List.map (fun (l,i) -> pairl tlevel tnat (quote_level l) (quote_int i)) (Universe.repr s) in let hd = List.hd levels in - let tl = to_coq_list (prodl tlevel bool_type) (List.tl levels) in + let tl = to_coq_list (prodl tlevel tnat) (List.tl levels) in constr_mkApp (tfrom_kernel_repr, [| hd ; tl |]) let quote_levelset s = let levels = Univ.Level.Set.elements s in - let levels' = to_coq_listl tlevel (List.map quote_nonprop_level levels) in + let levels' = to_coq_listl tlevel (List.map quote_level levels) in constr_mkApp (tLevelSet_of_list, [|levels'|]) let quote_constraint_type (c : Univ.constraint_type) = @@ -224,17 +220,17 @@ struct | Eq -> Lazy.force tunivEq let quote_univ_constraint ((l1, ct, l2) : Univ.univ_constraint) = - let l1 = quote_nonprop_level l1 in - let l2 = quote_nonprop_level l2 in + let l1 = quote_level l1 in + let l2 = quote_level l2 in let ct = quote_constraint_type ct in constr_mkApp (tmake_univ_constraint, [| l1; ct; l2 |]) - let quote_univ_level u = quote_nonprop_level u + let quote_univ_level u = quote_level u (* todo : can be deduced from quote_level, hence shoud be in the Reify module *) let quote_univ_instance u = let arr = Univ.Instance.to_array u in (* we assume that valid instances do not contain [Prop] or [SProp] *) - to_coq_listl tlevel (CArray.map_to_list quote_nonprop_level arr) + to_coq_listl tlevel (CArray.map_to_list quote_level arr) let is_Lt = function | Univ.Lt -> true @@ -326,22 +322,22 @@ struct constr_mkApp (tadd_global_constraints, [|constr_mkApp (cMonomorphic_ctx, [| uctx |]); Lazy.force tinit_graph|]) - let sprop_level = - lazy (constr_mkApp (cInl, [|Lazy.force tproplevel;Lazy.force tlevel;Lazy.force tlevelSProp |])) - let prop_level = - lazy (constr_mkApp (cInl, [|Lazy.force tproplevel;Lazy.force tlevel;Lazy.force tlevelProp |])) + let sprop = + lazy (constr_mkApp (sSProp, [|Lazy.force tuniverse |])) + let prop = + lazy (constr_mkApp (sProp, [|Lazy.force tuniverse |])) let quote_sort s = match s with - | Sorts.Set -> quote_universe Universe.type0 - | Sorts.Prop -> constr_mkApp (tof_levels, [| Lazy.force prop_level |]) - | Sorts.SProp -> constr_mkApp (tof_levels, [| Lazy.force sprop_level |]) - | Sorts.Type u -> quote_universe u + | Sorts.Set -> constr_mkApp (sType, [| Lazy.force tuniverse; quote_universe Universe.type0 |]) + | Sorts.Prop -> Lazy.force prop + | Sorts.SProp -> Lazy.force sprop + | Sorts.Type u -> constr_mkApp (sType, [| Lazy.force tuniverse; quote_universe u |]) let quote_sort_family = function | Sorts.InProp -> Lazy.force sfProp | Sorts.InSet -> Lazy.force sfSet | Sorts.InType -> Lazy.force sfType - | Sorts.InSProp -> Lazy.force sfProp (* FIXME SProp *) + | Sorts.InSProp -> Lazy.force sfSProp let quote_context_decl na b t = constr_mkApp (tmkdecl, [| Lazy.force tTerm; na; quote_optionl tTerm b; t |]) diff --git a/template-coq/src/constr_reification.ml b/template-coq/src/constr_reification.ml index dc323027f..ff8c9f665 100644 --- a/template-coq/src/constr_reification.ml +++ b/template-coq/src/constr_reification.ml @@ -125,9 +125,9 @@ struct let kVmCast = ast "VmCast" let kNative = ast "NativeCast" let kCast = ast "Cast" - let lSProp = ast "universe.lsprop" - let lProp = ast "universe.lprop" - let lnpe = ast "universe.lnpe" + let sSProp = ast "sort.sprop" + let sProp = ast "sort.prop" + let sType = ast "sort.type" let lzero = ast "level.lzero" let tsort_family = ast "sort_family" let lfresh_universe = ast "fresh_universe" @@ -163,6 +163,7 @@ struct let tMPdot = ast "MPdot" let tfresh_evar_id = ast "fresh_evar_id" + let tuniverse = ast "universe.t" let tproplevel = ast "level.prop_level_type" let tlevelSProp = ast "level.lsprop" let tlevelProp = ast "level.lprop" @@ -177,13 +178,12 @@ struct let tBuild_Universe = ast "universe.build0" let tfrom_kernel_repr = ast "universe.from_kernel_repr" (* let tto_kernel_repr = ast "universe.to_kernel_repr" *) - let tof_levels = ast "universe.of_levels" + let tof_level = ast "universe.make_of_level" let tLevelSet_of_list = ast "universe.of_list" let noprop_tSet = ast "noproplevel.lzero" let noprop_tLevel = ast "noproplevel.Level" let noprop_tLevelVar = ast "noproplevel.Var" let univexpr_lProp = ast "levelexpr.prop" - let univexpr_npe = ast "levelexpr.npe" (* let tunivcontext = resolve_symbol pkg_univ "universe_context" *) let tVariance = ast "variance.t" diff --git a/template-coq/src/denoter.ml b/template-coq/src/denoter.ml index 880033ac2..bfb34c032 100644 --- a/template-coq/src/denoter.ml +++ b/template-coq/src/denoter.ml @@ -22,7 +22,7 @@ sig val unquote_inductive : quoted_inductive -> Names.inductive (*val unquote_univ_instance : quoted_univ_instance -> Univ.Instance.t *) val unquote_proj : quoted_proj -> (quoted_inductive * quoted_int * quoted_int) - val unquote_universe : Evd.evar_map -> quoted_sort -> Evd.evar_map * Sorts.t + val unquote_sort : Evd.evar_map -> quoted_sort -> Evd.evar_map * Sorts.t val unquote_universe_level : Evd.evar_map -> quoted_univ_level -> Evd.evar_map * Univ.Level.t val unquote_universe_instance: Evd.evar_map -> quoted_univ_instance -> Evd.evar_map * Univ.Instance.t (* val representsIndConstuctor : quoted_inductive -> Term.constr -> bool *) @@ -72,7 +72,7 @@ struct | ACoq_tEvar (n, l) -> let evm, l' = map_evm (aux env) evm l in D.unquote_evar env evm n l' - | ACoq_tSort x -> let evm, u = D.unquote_universe evm x in evm, Constr.mkSort u + | ACoq_tSort x -> let evm, s = D.unquote_sort evm x in evm, Constr.mkSort s | ACoq_tCast (t,c,ty) -> let evm, t = aux env evm t in let evm, ty = aux env evm ty in evm, Constr.mkCast (t, D.unquote_cast_kind c, ty) diff --git a/template-coq/theories/Ast.v b/template-coq/theories/Ast.v index dbc367c9d..e0cf8a9dc 100644 --- a/template-coq/theories/Ast.v +++ b/template-coq/theories/Ast.v @@ -400,7 +400,7 @@ 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.t) +| tSort (s : sort) | tCast (t : term) (kind : cast_kind) (v : term) | tProd (na : aname) (ty : term) (body : term) | tLambda (na : aname) (ty : term) (body : term) @@ -569,7 +569,7 @@ Fixpoint noccur_between k n (t : term) : bool := | tArray u' arr def ty => tArray (subst_instance_level u u') (List.map (subst_instance_constr u) arr) (subst_instance_constr u def) (subst_instance_constr u ty) | tEvar ev args => tEvar ev (List.map (subst_instance_constr u) args) - | tSort s => tSort (subst_instance_univ u s) + | tSort s => tSort (subst_instance_sort 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') @@ -595,7 +595,7 @@ Fixpoint noccur_between k n (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 k univ + | tSort univ => closedu_sort k univ | tInd _ u => closedu_instance k u | tConstruct _ _ u => closedu_instance k u | tConst _ u => closedu_instance k u diff --git a/template-coq/theories/AstUtils.v b/template-coq/theories/AstUtils.v index 07e3c0d16..8dd36445b 100644 --- a/template-coq/theories/AstUtils.v +++ b/template-coq/theories/AstUtils.v @@ -380,13 +380,15 @@ Definition mkCase_old (Σ : global_env) (ci : case_info) (p : term) (c : term) ( tt oib.(ind_ctors) brs ;; ret (tCase ci p' c brs'). -Definition default_sort_family (u : Universe.t) : allowed_eliminations := - if Universe.is_sprop u then IntoAny - else if Universe.is_prop u then IntoPropSProp - else IntoAny. +Definition default_sort_family (s : sort) : allowed_eliminations := + match s with + | sSProp => IntoSProp + | sProp => IntoPropSProp + | _ => IntoAny + end. -Definition default_relevance (u : Universe.t) : relevance := - if Universe.is_sprop u then Irrelevant +Definition default_relevance (s : sort) : relevance := + if Sort.is_sprop s then Irrelevant else Relevant. (** Convenience functions for building constructors and inductive declarations *) @@ -411,15 +413,15 @@ Definition make_constructor_body (id : ident) (indrel : nat) derived from the universe (i.e. does not handle inductives with singleton elimination, or impredicate set eliminations). *) Definition make_inductive_body (id : ident) (params : context) (indices : context) - (u : Universe.t) (ind_ctors : list constructor_body) : one_inductive_body := + (s : sort) (ind_ctors : list constructor_body) : one_inductive_body := {| ind_name := id; ind_indices := indices; - ind_sort := u; - ind_type := it_mkProd_or_LetIn (params ,,, indices) (tSort u); - ind_kelim := default_sort_family u; + ind_sort := s; + ind_type := it_mkProd_or_LetIn (params ,,, indices) (tSort s); + ind_kelim := default_sort_family s; ind_ctors := ind_ctors; ind_projs := []; - ind_relevance := default_relevance u |}. + ind_relevance := default_relevance s |}. Ltac change_Sk := repeat match goal with diff --git a/template-coq/theories/Checker.v b/template-coq/theories/Checker.v index 8a8574408..6b5c7b1fc 100644 --- a/template-coq/theories/Checker.v +++ b/template-coq/theories/Checker.v @@ -368,7 +368,7 @@ Fixpoint eq_term `{checker_flags} (φ : universes_graph) (t u : term) {struct t} | tRel n, tRel n' => Nat.eqb n n' | tEvar ev args, tEvar ev' args' => Nat.eqb ev ev' && forallb2 (eq_term φ) args args' | tVar id, tVar id' => eqb id id' - | tSort s, tSort s' => check_eqb_universe φ s s' + | tSort s, tSort s' => check_eqb_sort φ s s' | tCast f k T, tCast f' k' T' => eq_term φ f f' && eq_term φ T T' | tApp f args, tApp f' args' => eq_term φ f f' && forallb2 (eq_term φ) args args' | tConst c u, tConst c' u' => eq_constant c c' && eqb_univ_instance φ u u' @@ -400,7 +400,7 @@ Fixpoint leq_term `{checker_flags} (φ : universes_graph) (t u : term) {struct t | tRel n, tRel n' => Nat.eqb n n' | tEvar ev args, tEvar ev' args' => Nat.eqb ev ev' && forallb2 (eq_term φ) args args' | tVar id, tVar id' => eqb id id' - | tSort s, tSort s' => check_leqb_universe φ s s' + | tSort s, tSort s' => check_leqb_sort φ s s' | tApp f args, tApp f' args' => eq_term φ f f' && forallb2 (eq_term φ) args args' | tCast f k T, tCast f' k' T' => eq_term φ f f' && eq_term φ T T' | tConst c u, tConst c' u' => eq_constant c c' && eqb_univ_instance φ u u' @@ -631,7 +631,7 @@ Section Typecheck. | None => raise (NotEnoughFuel fuel) end. - Definition reduce_to_sort Γ (t : term) : typing_result Universe.t := + Definition reduce_to_sort Γ (t : term) : typing_result sort := match t with | tSort s => ret s | _ => @@ -723,7 +723,7 @@ Section Typecheck. | tVar n => raise (UnboundVar n) | tEvar ev args => raise (UnboundEvar ev) - | tSort s => ret (tSort (Universe.super s)) + | tSort s => ret (tSort (Sort.super s)) | tCast c k t => infer_type infer Γ t ;; @@ -733,7 +733,7 @@ Section Typecheck. | tProd n t b => s1 <- infer_type infer Γ t ;; s2 <- infer_type infer (Γ ,, vass n t) b ;; - ret (tSort (Universe.sort_of_product s1 s2)) + ret (tSort (Sort.sort_of_product s1 s2)) | tLambda n t b => infer_type infer Γ t ;; diff --git a/template-coq/theories/Constants.v b/template-coq/theories/Constants.v index 1c7834eca..62c1a06a7 100644 --- a/template-coq/theories/Constants.v +++ b/template-coq/theories/Constants.v @@ -106,8 +106,9 @@ Register MetaCoq.Common.Universes.ConstraintType.Lt as metacoq.ast.constraints.L Register MetaCoq.Common.Universes.ConstraintType.Le0 as metacoq.ast.constraints.Le0. Register MetaCoq.Common.Universes.ConstraintType.Le as metacoq.ast.constraints.Le. Register MetaCoq.Common.Universes.ConstraintType.Eq as metacoq.ast.constraints.Eq. +Register MetaCoq.Common.Universes.Universe.t as metacoq.ast.universe.t. +Register MetaCoq.Common.Universes.Universe.make' as metacoq.ast.universe.make_of_level. Register MetaCoq.Common.Universes.Universe.from_kernel_repr as metacoq.ast.universe.from_kernel_repr. -Register MetaCoq.Common.Universes.Universe.of_levels as metacoq.ast.universe.of_levels. Register MetaCoq.Common.Universes.LevelSetProp.of_list as metacoq.ast.universe.of_list. Register MetaCoq.Common.Universes.Level.t as metacoq.ast.level.t. Register MetaCoq.Common.Universes.Level.level as metacoq.ast.level.Level. @@ -116,13 +117,12 @@ Register MetaCoq.Common.Universes.PropLevel.lProp as metacoq.ast.level.lprop. Register MetaCoq.Common.Universes.PropLevel.lSProp as metacoq.ast.level.lsprop. Register MetaCoq.Common.Universes.Level.lzero as metacoq.ast.level.lzero. Register MetaCoq.Common.Universes.Level.lvar as metacoq.ast.level.Var. -Register MetaCoq.Common.Universes.Universe.lType as metacoq.ast.levelexpr.npe. Register MetaCoq.Common.Universes.LevelExprSet.Mkt as metacoq.ast.levelexprset.mkt. Register MetaCoq.Common.Universes.Build_nonEmptyLevelExprSet as metacoq.ast.universe.build0. -Register MetaCoq.Common.Universes.Universe.lSProp as metacoq.ast.universe.lsprop. -Register MetaCoq.Common.Universes.Universe.lProp as metacoq.ast.universe.lprop. -Register MetaCoq.Common.Universes.Universe.lType as metacoq.ast.universe.lnpe. +Register MetaCoq.Common.Universes.Sort.sSProp as metacoq.ast.sort.sprop. +Register MetaCoq.Common.Universes.Sort.sProp as metacoq.ast.sort.prop. +Register MetaCoq.Common.Universes.Sort.sType as metacoq.ast.sort.type. Register MetaCoq.Common.Universes.Variance.t as metacoq.ast.variance.t. diff --git a/template-coq/theories/EtaExpand.v b/template-coq/theories/EtaExpand.v index 96b1acc7f..2142e448c 100644 --- a/template-coq/theories/EtaExpand.v +++ b/template-coq/theories/EtaExpand.v @@ -255,7 +255,7 @@ Inductive expanded (Γ : list nat): term -> Prop := | expanded_tRel_app (n : nat) args m : nth_error Γ n = Some m -> m <= #|args| -> Forall (expanded Γ) args -> expanded Γ (tApp (tRel n) args) | expanded_tVar (id : ident) : expanded Γ (tVar id) | expanded_tEvar (ev : nat) (args : list term) : Forall (expanded Γ) args -> expanded Γ (tEvar ev args) -| expanded_tSort (s : Universe.t) : expanded Γ (tSort s) +| expanded_tSort (s : sort) : expanded Γ (tSort s) | expanded_tCast (t : term) (kind : cast_kind) (v : term) : expanded Γ t -> expanded Γ v -> expanded Γ (tCast t kind v) | expanded_tProd (na : aname) (ty : term) (body : term) : (* expanded Γ ty -> expanded Γ body -> *) expanded Γ (tProd na ty body) | expanded_tLambda (na : aname) (ty : term) (body : term) : (* expanded Γ ty -> *) expanded (0 :: Γ) body -> expanded Γ (tLambda na ty body) @@ -308,7 +308,7 @@ forall (Σ : global_env) (P : list nat -> term -> Prop), (forall Γ, forall n : nat, forall args, forall m, nth_error Γ n = Some m -> forall Heq : m <= #|args|, Forall (expanded Σ Γ) args -> Forall (P Γ) args -> P Γ (tApp (tRel n) args)) -> (forall Γ, forall id : ident, P Γ (tVar id)) -> (forall Γ, forall (ev : nat) (args : list term), Forall (expanded Σ Γ) args -> Forall (P Γ) args -> P Γ (tEvar ev args)) -> -(forall Γ, forall s : Universe.t, P Γ (tSort s)) -> +(forall Γ, forall s : sort, P Γ (tSort s)) -> (forall Γ, forall (t : term) (kind : cast_kind) (v : term), expanded Σ Γ t -> P Γ t -> expanded Σ Γ v -> P Γ v -> P Γ (tCast t kind v)) -> (forall Γ, forall (na : aname) (ty body : term), P Γ (tProd na ty body)) -> @@ -986,11 +986,23 @@ Proof. | Some (n, _) => n | None => 0 end) Γ') (eta_expand Σg Γ' t)) + (Pj := fun (Σ : global_env_ext) Γ j => option_default (fun t => forall Γ', Forall2 (fun (x : option (nat × term)) (y : context_decl) => + match x with + | Some (_, t0) => decl_type y = t0 /\ _ + | None => True + end) Γ' Γ -> + forall Σg, repr_decls Σg Σ.1 -> + expanded Σ (map (fun x : option (nat × term) => + match x with + | Some (n, _) => n + | None => 0 + end) Γ') (eta_expand Σg Γ' t)) (j_term j) True) (PΓ := fun _ _ _ => True); repeat match goal with | [ |- repr_decls _ _ -> _ ] => intros hrepr | _ => intro end; try now (cbn; eauto). + - apply fst in X. destruct j_term => //. apply snd in X. simpl. apply X. - cbn. eapply Forall2_nth_error_Some_r in H1 as (? & ? & ?); eauto. rewrite H1. destruct x as [[] | ]. @@ -1010,7 +1022,7 @@ Proof. all: econstructor; rewrite nth_error_app1; revgoals; [eapply nth_error_repeat; lia | rewrite repeat_length; lia]. + econstructor. now rewrite nth_error_map H1. - cbn. econstructor. eapply (H1 (up Γ')); tea; econstructor; eauto. - - cbn. econstructor. eauto. eapply (H2 (up Γ')); tea; econstructor; eauto. + - cbn. econstructor. eauto. eapply (H1 (up Γ')); tea; econstructor; eauto. - specialize (H _ H2). assert (Forall(fun t : term => expanded Σ0 (map (fun x : option (nat × term) => @@ -1205,11 +1217,10 @@ Proof. { apply andb_and in H2. destruct H2 as [isl _]. solve_all. } solve_all. { now eapply isLambda_lift, isLambda_eta_expand. } - destruct a as (? & ? & ?). - destruct a0 as (? & ?). + cbn in *. rewrite !firstn_length. rewrite -> !Nat.min_l; try lia. eapply expanded_lift'. - 5: eapply e0; eauto. 2: reflexivity. 2: now len. + 5: eapply a2; eauto. 2: reflexivity. 2: now len. 2: now len. { rewrite map_app. f_equal. rewrite map_rev. f_equal. now rewrite !mapi_map map_mapi. } eapply Forall2_app; solve_all. @@ -1227,7 +1238,7 @@ Proof. len; lia. rewrite repeat_length. len; lia. + cbn - [rev_map seq]. rewrite rev_map_spec. cbn. rewrite Nat.sub_0_r. cbn. destruct List.rev; cbn; congruence. - - cbn. econstructor; eauto. eapply All_Forall, All_map, All_impl. eapply (All_mix X X0). intros ? ((? & ? & ?) & ? & ?). cbn. + - cbn. econstructor; eauto. eapply All_Forall, All_map, All_impl with (1 := X2). cbn. intros ? e0. specialize (e0 (repeat None #|mfix| ++ Γ'))%list. rewrite map_app map_repeat in e0. len. eapply e0; auto. eapply Forall2_app; eauto. unfold types. @@ -1435,8 +1446,9 @@ Proof. cbn. constructor. cbn. constructor. len. rewrite app_nil_r. - red in t1, t2. - forward (eta_expand_expanded (Σ := Σ) Γ (repeat None #|Γ|) _ _ wfΣ t2). + destruct t1 as (t1 & s & t2 & _). + cbn in t1, t2. + forward (eta_expand_expanded (Σ := Σ) Γ (repeat None #|Γ|) _ _ wfΣ t1). clear. induction Γ; cbn; constructor; auto. intros. specialize (H _ hrepr). move: H. @@ -1455,8 +1467,8 @@ Proof. cbn in hs. destruct a as [na [b|] ty]; try destruct hs as [hs ?]. specialize (IHctx' cunivs hs). constructor; auto. constructor. len. rewrite repeat_app. - destruct p as [[s Hs] ?]. - epose proof (eta_expand_expanded (Σ := Σ) _ (repeat None (#|ctx'| + #|ctx|)) _ _ wfΣ t0). + destruct l as [Htm [s Hs]]. + epose proof (eta_expand_expanded (Σ := Σ) _ (repeat None (#|ctx'| + #|ctx|)) _ _ wfΣ Htm). forward H. clear. rewrite -app_context_length. induction (_ ,,, _); cbn; constructor; auto. @@ -1481,7 +1493,7 @@ Proof. - unfold on_constant_decl in ond. destruct c as [na body ty rel]; cbn in *. destruct body. constructor => //; cbn. - apply (eta_expand_expanded (Σ := Σ) [] [] t0 na wf ond). constructor. + apply (eta_expand_expanded (Σ := Σ) [] [] t0 na wf ond.1). constructor. apply hrepr. destruct ond as [s Hs]. constructor => //. - destruct ond as [onI onP onN onV]. @@ -1497,7 +1509,7 @@ Proof. pose proof onc.(on_cargs). eapply eta_expand_context_sorts in X0. now len in X0. exact hrepr. len. len. - pose proof onc.(on_ctype). destruct X0. + pose proof onc.(on_ctype). destruct X0 as (_ & s & t0 & _). epose proof (eta_expand_expanded (Σ := Σ) _ (repeat None #|ind_bodies m|) _ _ wf t0). forward H. rewrite -arities_context_length. clear. induction (arities_context _); constructor; auto. diff --git a/template-coq/theories/Pretty.v b/template-coq/theories/Pretty.v index 694ce8383..2a16201f6 100644 --- a/template-coq/theories/Pretty.v +++ b/template-coq/theories/Pretty.v @@ -131,11 +131,11 @@ Module PrintTermTree. let ctx' := fix_context defs in print_list (print_def (print_term Γ true) (print_term (fresh_names Σ Γ ctx') true)) (nl ^ " with ") defs. - Definition print_universe u := - match u with - | Universe.lProp => "Prop" - | Universe.lSProp => "SProp" - | Universe.lType l => + Definition print_sort (s : sort) := + match s with + | sProp => "Prop" + | sSProp => "SProp" + | sType l => if with_universes then ("Type(" ++ MCString.string_of_list string_of_level_expr (LevelExprSet.elements l) ++ @@ -153,7 +153,7 @@ Module PrintTermTree. end | tVar n => "Var(" ^ n ^ ")" | tEvar ev args => "Evar(" ^ string_of_nat ev ^ "[]" (* TODO *) ^ ")" - | tSort s => print_universe s + | tSort s => print_sort s | tCast c k t => parens top (print_term Γ true c ^ ":" ^ print_term Γ true t) | tProd na dom codom => let na' := (fresh_name Σ Γ na.(binder_name) (Some dom)) in diff --git a/template-coq/theories/ReflectAst.v b/template-coq/theories/ReflectAst.v index 712f4ebd0..f2decc81a 100644 --- a/template-coq/theories/ReflectAst.v +++ b/template-coq/theories/ReflectAst.v @@ -24,7 +24,7 @@ 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.t, u' : Universe.t |- _ => fcase (eq_dec u u') + | u : sort, u' : sort |- _ => fcase (eq_dec u u') | x : Instance.t, y : Instance.t |- _ => fcase (eq_dec x y) | x : list Level.t, y : Instance.t |- _ => diff --git a/template-coq/theories/TemplateMonad/Core.v b/template-coq/theories/TemplateMonad/Core.v index f06d419a8..56452ab8b 100644 --- a/template-coq/theories/TemplateMonad/Core.v +++ b/template-coq/theories/TemplateMonad/Core.v @@ -204,12 +204,18 @@ Definition TypeInstanceOptimized : Common.TMInstance := Definition TypeInstance : Common.TMInstance := Eval hnf in TypeInstanceUnoptimized. -Definition tmQuoteUniverse@{U t u} : TemplateMonad@{t u} Universe.t - := u <- @tmQuote Prop (Type@{U} -> True);; - match u with - | tProd _ (tSort u) _ => ret u +Definition tmQuoteSort@{U t u} : TemplateMonad@{t u} sort + := p <- @tmQuote Prop (Type@{U} -> True);; + match p with + | tProd _ (tSort s) _ => ret s | _ => tmFail "Anomaly: tmQuote (Type -> True) should be (tProd _ (tSort _) _)!"%bs end. +Definition tmQuoteUniverse@{U t u} : TemplateMonad@{t u} Universe.t + := s <- @tmQuoteSort@{U t u};; + match s with + | sType u => ret u + | _ => tmFail "Sort does not carry a universe (is not Type)"%bs + end. Definition tmQuoteLevel@{U t u} : TemplateMonad@{t u} Level.t := u <- tmQuoteUniverse@{U t u};; match Universe.get_is_level u with diff --git a/template-coq/theories/TermEquality.v b/template-coq/theories/TermEquality.v index 653533f9a..e390f61fa 100644 --- a/template-coq/theories/TermEquality.v +++ b/template-coq/theories/TermEquality.v @@ -9,8 +9,8 @@ From Equations.Prop Require Import DepElim. From Equations Require Import Equations. Set Equations With UIP. -Definition R_universe_instance R := - fun u u' => Forall2 R (List.map Universe.make u) (List.map Universe.make u'). +Definition cmp_universe_instance (cmp_univ : Universe.t -> Universe.t -> Prop) : Instance.t -> Instance.t -> Prop := + Forall2 (on_rel cmp_univ Universe.make'). (** Cumulative inductive types: @@ -19,26 +19,15 @@ Definition R_universe_instance R := even on ill-formed terms. It corresponds to the right notion on well-formed terms. *) -Definition R_universe_variance (Re Rle : Universe.t -> Universe.t -> Prop) v u u' := +Definition cmp_universe_variance (cmp_univ : conv_pb -> Universe.t -> Universe.t -> Prop) pb v u u' := match v with | Variance.Irrelevant => True - | Variance.Covariant => Rle (Universe.make u) (Universe.make u') - | Variance.Invariant => Re (Universe.make u) (Universe.make u') + | Variance.Covariant => on_rel (cmp_univ pb) Universe.make' u u' + | Variance.Invariant => on_rel (cmp_univ Conv) Universe.make' u u' end. -Fixpoint R_universe_instance_variance Re Rle v u u' := - match u, u' return Prop with - | u :: us, u' :: us' => - match v with - | [] => R_universe_instance_variance Re Rle v us us' - (* Missing variance stands for irrelevance, we still check that the instances have - the same length. *) - | v :: vs => R_universe_variance Re Rle v u u' /\ - R_universe_instance_variance Re Rle vs us us' - end - | [], [] => True - | _, _ => False - end. +Definition cmp_universe_instance_variance cmp_univ pb v u u' := + Forall3 (cmp_universe_variance cmp_univ pb) v u u'. Definition global_variance_gen lookup gr napp := match gr with @@ -46,64 +35,128 @@ Definition global_variance_gen lookup gr napp := match lookup_inductive_gen lookup ind with | Some (mdecl, idecl) => match destArity [] idecl.(ind_type) with - | Some (ctx, _) => if (context_assumptions ctx) <=? napp then mdecl.(ind_variance) - else None - | None => None + | Some (ctx, _) => if (context_assumptions ctx) <=? napp then + match mdecl.(ind_variance) with + | Some var => Variance var + | None => AllEqual + end + else AllEqual + | None => AllEqual end - | None => None + | None => AllEqual end | ConstructRef ind k => match lookup_constructor_gen lookup ind k with | Some (mdecl, idecl, cdecl) => if (cdecl.(cstr_arity) + mdecl.(ind_npars))%nat <=? napp then (** Fully applied constructors are always compared at the same supertype, - which implies that no universe ws_cumul_pb needs to be checked here. *) - Some [] - else None - | _ => None + which implies that no universe ws_cumul_pb needs to be checked here. + We will still check that the instances have the same length. *) + AllIrrelevant + else AllEqual + | _ => AllEqual end - | _ => None + | _ => AllEqual end. Notation global_variance Σ := (global_variance_gen (lookup_env Σ)). -Definition R_opt_variance Re Rle v := +Definition cmp_opt_variance cmp_univ pb v := match v with - | Some v => R_universe_instance_variance Re Rle v - | None => R_universe_instance Re + | AllEqual => cmp_universe_instance (cmp_univ Conv) + | AllIrrelevant => fun l l' => #|l| = #|l'| + | Variance v => fun u u' => cmp_universe_instance (cmp_univ Conv) u u' \/ cmp_universe_instance_variance cmp_univ pb v u u' end. -Definition R_global_instance_gen Σ Re Rle gr napp := - R_opt_variance Re Rle (global_variance_gen Σ gr napp). +Lemma cmp_instance_opt_variance cmp_univ pb v : + RelationClasses.subrelation (cmp_opt_variance cmp_univ pb AllEqual) (cmp_opt_variance cmp_univ pb v). +Proof. + intros u u' H. + destruct v as [| |v] => //=. + 1: now apply Forall2_length in H. + now left. +Qed. + +Lemma cmp_universe_universe_variance (cmp_univ : conv_pb -> Universe.t -> Universe.t -> Prop) pb v u u' : + RelationClasses.subrelation (cmp_univ Conv) (cmp_univ pb) -> + on_rel (cmp_univ Conv) Universe.make' u u' -> cmp_universe_variance cmp_univ pb v u u'. +Proof. + destruct v => //=. + intros H H1; apply H, H1. +Qed. + +Lemma cmp_instance_variance cmp_univ pb v u u' : + RelationClasses.subrelation (cmp_univ Conv) (cmp_univ pb) -> + #|v| = #|u| -> + cmp_universe_instance (cmp_univ Conv) u u' -> cmp_universe_instance_variance cmp_univ pb v u u'. +Proof. + intros Hsub Hlen Hu. + apply Forall2_triv in Hlen. + eapply Forall2_Forall2_Forall3 in Hu; tea. + apply Forall3_impl with (1 := Hu) => v1 u1 u1' [] _ H1. + now apply cmp_universe_universe_variance. +Qed. + +Lemma cmp_opt_variance_var_dec cmp_univ pb v ui ui' : + RelationClasses.subrelation (cmp_univ Conv) (cmp_univ pb) -> + cmp_universe_instance (cmp_univ Conv) ui ui' \/ cmp_universe_instance_variance cmp_univ pb v ui ui' -> + { cmp_universe_instance (cmp_univ Conv) ui ui' } + { cmp_universe_instance_variance cmp_univ pb v ui ui' }. +Proof. + intros subr H. + elim:(eq_dec #|v| #|ui|). + - right. + destruct H as [|]; tas. + now eapply cmp_instance_variance. + - left. + destruct H as [|]; tas. + eapply @Forall3_Forall2_left with (Q := fun _ _ => True) in H => //. + apply Forall2_length in H. + now exfalso. +Qed. + +Definition cmp_global_instance_gen Σ cmp_universe pb gr napp := + cmp_opt_variance cmp_universe pb (global_variance_gen Σ gr napp). -Notation R_global_instance Σ := (R_global_instance_gen (lookup_env Σ)). +Notation cmp_global_instance Σ := (cmp_global_instance_gen (lookup_env Σ)). -Lemma R_universe_instance_impl R R' : +Lemma cmp_universe_instance_impl R R' : RelationClasses.subrelation R R' -> - RelationClasses.subrelation (R_universe_instance R) (R_universe_instance R'). + RelationClasses.subrelation (cmp_universe_instance R) (cmp_universe_instance R'). Proof. - intros H x y xy. eapply Forall2_impl ; tea. + intros H x y xy. eapply Forall2_impl; tea; unfold on_rel; auto. Qed. -Lemma R_universe_instance_impl' R R' : +Lemma cmp_universe_instance_impl' R R' : RelationClasses.subrelation R R' -> - forall u u', R_universe_instance R u u' -> R_universe_instance R' u u'. + forall u u', cmp_universe_instance R u u' -> cmp_universe_instance R' u u'. Proof. - intros H x y xy. eapply Forall2_impl ; tea. + intros H x y xy. eapply Forall2_impl; tea; unfold on_rel; auto. Qed. -Inductive compare_decls (eq_term leq_term : term -> term -> Type) : context_decl -> context_decl -> Type := - | compare_vass na T na' T' : eq_binder_annot na na' -> - leq_term T T' -> - compare_decls eq_term leq_term (vass na T) (vass na' T') - | compare_vdef na b T na' b' T' : eq_binder_annot na na' -> - eq_term b b' -> leq_term T T' -> - compare_decls eq_term leq_term (vdef na b T) (vdef na' b' T'). +Lemma cmp_universe_instance_variance_impl R R' pb pb' v : + RelationClasses.subrelation (R Conv) (R' Conv) -> + RelationClasses.subrelation (R pb) (R' pb') -> + RelationClasses.subrelation (cmp_universe_instance_variance R pb v) (cmp_universe_instance_variance R' pb' v). +Proof. + intros HConv Hpb x y xy. + eapply Forall3_impl; tea. clear v x y xy. + intros v u u'. + destruct v => //=. + all: unfold on_rel; now auto. +Qed. + + +Inductive eq_decl_upto_names : context_decl -> context_decl -> Type := + | compare_vass {na na' T} : + eq_binder_annot na na' -> eq_decl_upto_names (vass na T) (vass na' T) + | compare_vdef {na na' b T} : + eq_binder_annot na na' -> eq_decl_upto_names (vdef na b T) (vdef na' b T). -Derive Signature NoConfusion for compare_decls. +Derive Signature NoConfusion for eq_decl_upto_names. +Notation eq_context_upto_names := (All2 eq_decl_upto_names). Lemma alpha_eq_context_assumptions {Γ Δ} : - All2 (compare_decls eq eq) Γ Δ -> + eq_context_upto_names Γ Δ -> context_assumptions Γ = context_assumptions Δ. Proof. induction 1 in |- *; cbn; auto. @@ -111,7 +164,7 @@ Proof. Qed. Lemma alpha_eq_extended_subst {Γ Δ k} : - All2 (compare_decls eq eq) Γ Δ -> + eq_context_upto_names Γ Δ -> extended_subst Γ k = extended_subst Δ k. Proof. induction 1 in k |- *; cbn; auto. @@ -120,7 +173,7 @@ Proof. Qed. Lemma expand_lets_eq {Γ Δ t} : - All2 (compare_decls eq eq) Γ Δ -> + eq_context_upto_names Γ Δ -> expand_lets Γ t = expand_lets Δ t. Proof. intros. rewrite /expand_lets /expand_lets_k. @@ -128,16 +181,16 @@ Proof. Qed. Lemma alpha_eq_subst_context {Γ Δ s k} : - All2 (compare_decls eq eq) Γ Δ -> - All2 (compare_decls eq eq) (subst_context s k Γ) (subst_context s k Δ). + eq_context_upto_names Γ Δ -> + eq_context_upto_names (subst_context s k Γ) (subst_context s k Δ). Proof. intros. rewrite /subst_context. induction X. - cbn; auto. - rewrite !fold_context_k_snoc0. constructor; auto. + rewrite -(All2_length X). destruct r; subst; constructor; cbn; auto. - all:now rewrite (All2_length X). Qed. (** ** Syntactic equality up-to universes @@ -147,129 +200,137 @@ Qed. that surround the current term, used to implement cumulativity of inductive types correctly (only fully applied constructors and inductives benefit from it). *) -Inductive eq_term_upto_univ_napp Σ (Re Rle : Universe.t -> Universe.t -> Prop) (napp : nat) : term -> term -> Type := +Inductive eq_term_upto_univ_napp Σ + (cmp_universe : conv_pb -> Universe.t -> Universe.t -> Prop) + (cmp_sort : conv_pb -> sort -> sort -> Prop) + (pb : conv_pb) (napp : nat) : term -> term -> Type := | eq_Rel n : - eq_term_upto_univ_napp Σ Re Rle napp (tRel n) (tRel n) + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp (tRel n) (tRel n) | eq_Evar e args args' : - All2 (eq_term_upto_univ_napp Σ Re Re 0) args args' -> - eq_term_upto_univ_napp Σ Re Rle napp (tEvar e args) (tEvar e args') + All2 (eq_term_upto_univ_napp Σ cmp_universe cmp_sort Conv 0) args args' -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp (tEvar e args) (tEvar e args') | eq_Var id : - eq_term_upto_univ_napp Σ Re Rle napp (tVar id) (tVar id) + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp (tVar id) (tVar id) | eq_Sort s s' : - Rle s s' -> - eq_term_upto_univ_napp Σ Re Rle napp (tSort s) (tSort s') + cmp_sort pb s s' -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp (tSort s) (tSort s') | eq_App t t' u u' : - eq_term_upto_univ_napp Σ Re Rle (#|u| + napp) t t' -> - All2 (eq_term_upto_univ_napp Σ Re Re 0) u u' -> - eq_term_upto_univ_napp Σ Re Rle napp (tApp t u) (tApp t' u') + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb (#|u| + napp) t t' -> + All2 (eq_term_upto_univ_napp Σ cmp_universe cmp_sort Conv 0) u u' -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp (tApp t u) (tApp t' u') | eq_Const c u u' : - R_universe_instance Re u u' -> - eq_term_upto_univ_napp Σ Re Rle napp (tConst c u) (tConst c u') + cmp_universe_instance (cmp_universe Conv) u u' -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp (tConst c u) (tConst c u') | eq_Ind i u u' : - R_global_instance Σ Re Rle (IndRef i) napp u u' -> - eq_term_upto_univ_napp Σ Re Rle napp (tInd i u) (tInd i u') + cmp_global_instance Σ cmp_universe pb (IndRef i) napp u u' -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp (tInd i u) (tInd i u') | eq_Construct i k u u' : - R_global_instance Σ Re Rle (ConstructRef i k) napp u u' -> - eq_term_upto_univ_napp Σ Re Rle napp (tConstruct i k u) (tConstruct i k u') + cmp_global_instance Σ cmp_universe pb (ConstructRef i k) napp u u' -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp (tConstruct i k u) (tConstruct i k u') | eq_Lambda na na' ty ty' t t' : eq_binder_annot na na' -> - eq_term_upto_univ_napp Σ Re Re 0 ty ty' -> - eq_term_upto_univ_napp Σ Re Rle 0 t t' -> - eq_term_upto_univ_napp Σ Re Rle napp (tLambda na ty t) (tLambda na' ty' t') + eq_term_upto_univ_napp Σ cmp_universe cmp_sort Conv 0 ty ty' -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb 0 t t' -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp (tLambda na ty t) (tLambda na' ty' t') | eq_Prod na na' a a' b b' : eq_binder_annot na na' -> - eq_term_upto_univ_napp Σ Re Re 0 a a' -> - eq_term_upto_univ_napp Σ Re Rle 0 b b' -> - eq_term_upto_univ_napp Σ Re Rle napp (tProd na a b) (tProd na' a' b') + eq_term_upto_univ_napp Σ cmp_universe cmp_sort Conv 0 a a' -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb 0 b b' -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp (tProd na a b) (tProd na' a' b') | eq_LetIn na na' t t' ty ty' u u' : eq_binder_annot na na' -> - eq_term_upto_univ_napp Σ Re Re 0 t t' -> - eq_term_upto_univ_napp Σ Re Re 0 ty ty' -> - eq_term_upto_univ_napp Σ Re Rle 0 u u' -> - eq_term_upto_univ_napp Σ Re Rle napp (tLetIn na t ty u) (tLetIn na' t' ty' u') + eq_term_upto_univ_napp Σ cmp_universe cmp_sort Conv 0 t t' -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort Conv 0 ty ty' -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb 0 u u' -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp (tLetIn na t ty u) (tLetIn na' t' ty' u') | eq_Case ind p p' c c' brs brs' : - All2 (eq_term_upto_univ_napp Σ Re Re 0) p.(pparams) p'.(pparams) -> - R_universe_instance Re p.(puinst) p'.(puinst) -> - eq_term_upto_univ_napp Σ Re Re 0 p.(preturn) p'.(preturn) -> + All2 (eq_term_upto_univ_napp Σ cmp_universe cmp_sort Conv 0) p.(pparams) p'.(pparams) -> + cmp_universe_instance (cmp_universe Conv) p.(puinst) p'.(puinst) -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort Conv 0 p.(preturn) p'.(preturn) -> All2 eq_binder_annot p.(pcontext) p'.(pcontext) -> - eq_term_upto_univ_napp Σ Re Re 0 c c' -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort Conv 0 c c' -> All2 (fun x y => All2 (eq_binder_annot) (bcontext x) (bcontext y) * - eq_term_upto_univ_napp Σ Re Re 0 (bbody x) (bbody y) + eq_term_upto_univ_napp Σ cmp_universe cmp_sort Conv 0 (bbody x) (bbody y) ) brs brs' -> - eq_term_upto_univ_napp Σ Re Rle napp (tCase ind p c brs) (tCase ind p' c' brs') + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp (tCase ind p c brs) (tCase ind p' c' brs') | eq_Proj p c c' : - eq_term_upto_univ_napp Σ Re Re 0 c c' -> - eq_term_upto_univ_napp Σ Re Rle napp (tProj p c) (tProj p c') + eq_term_upto_univ_napp Σ cmp_universe cmp_sort Conv 0 c c' -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp (tProj p c) (tProj p c') | eq_Fix mfix mfix' idx : All2 (fun x y => - eq_term_upto_univ_napp Σ Re Re 0 x.(dtype) y.(dtype) * - eq_term_upto_univ_napp Σ Re Re 0 x.(dbody) y.(dbody) * + eq_term_upto_univ_napp Σ cmp_universe cmp_sort Conv 0 x.(dtype) y.(dtype) * + eq_term_upto_univ_napp Σ cmp_universe cmp_sort Conv 0 x.(dbody) y.(dbody) * (x.(rarg) = y.(rarg)) * eq_binder_annot x.(dname) y.(dname) )%type mfix mfix' -> - eq_term_upto_univ_napp Σ Re Rle napp (tFix mfix idx) (tFix mfix' idx) + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp (tFix mfix idx) (tFix mfix' idx) | eq_CoFix mfix mfix' idx : All2 (fun x y => - eq_term_upto_univ_napp Σ Re Re 0 x.(dtype) y.(dtype) * - eq_term_upto_univ_napp Σ Re Re 0 x.(dbody) y.(dbody) * + eq_term_upto_univ_napp Σ cmp_universe cmp_sort Conv 0 x.(dtype) y.(dtype) * + eq_term_upto_univ_napp Σ cmp_universe cmp_sort Conv 0 x.(dbody) y.(dbody) * (x.(rarg) = y.(rarg)) * eq_binder_annot x.(dname) y.(dname) ) mfix mfix' -> - eq_term_upto_univ_napp Σ Re Rle napp (tCoFix mfix idx) (tCoFix mfix' idx) + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp (tCoFix mfix idx) (tCoFix mfix' idx) | eq_Cast t1 c t2 t1' c' t2' : - eq_term_upto_univ_napp Σ Re Re 0 t1 t1' -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort Conv 0 t1 t1' -> eq_cast_kind c c' -> - eq_term_upto_univ_napp Σ Re Re 0 t2 t2' -> - eq_term_upto_univ_napp Σ Re Rle napp (tCast t1 c t2) (tCast t1' c' t2') + eq_term_upto_univ_napp Σ cmp_universe cmp_sort Conv 0 t2 t2' -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp (tCast t1 c t2) (tCast t1' c' t2') -| eq_Int i : eq_term_upto_univ_napp Σ Re Rle napp (tInt i) (tInt i) -| eq_Float f : eq_term_upto_univ_napp Σ Re Rle napp (tFloat f) (tFloat f) +| eq_Int i : eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp (tInt i) (tInt i) +| eq_Float f : eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp (tFloat f) (tFloat f) | eq_Array u u' arr arr' def def' ty ty' : - Re (Universe.make u) (Universe.make u') -> - All2 (eq_term_upto_univ_napp Σ Re Re 0) arr arr' -> - eq_term_upto_univ_napp Σ Re Re 0 def def' -> - eq_term_upto_univ_napp Σ Re Re 0 ty ty' -> - eq_term_upto_univ_napp Σ Re Rle napp (tArray u arr def ty) (tArray u' arr' def' ty'). + cmp_universe_instance (cmp_universe Conv) [u] [u'] -> + All2 (eq_term_upto_univ_napp Σ cmp_universe cmp_sort Conv 0) arr arr' -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort Conv 0 def def' -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort Conv 0 ty ty' -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp (tArray u arr def ty) (tArray u' arr' def' ty'). -Notation eq_term_upto_univ Σ Re Rle := (eq_term_upto_univ_napp Σ Re Rle 0). +Notation eq_term_upto_univ Σ cmp_universe cmp_sort pb := (eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb 0). (* ** Syntactic conversion/cumulativity up-to universes *) -Definition compare_term `{checker_flags} (pb : conv_pb) Σ φ := - eq_term_upto_univ Σ (eq_universe φ) (compare_universe pb φ). +Definition compare_term `{checker_flags} Σ φ (pb : conv_pb) := + eq_term_upto_univ Σ (compare_universe φ) (compare_sort φ) pb. -Notation eq_term := (compare_term Conv). -Notation leq_term := (compare_term Cumul). +Notation eq_term Σ φ := (compare_term Σ φ Conv). +Notation leq_term Σ φ := (compare_term Σ φ Cumul). -Lemma R_global_instance_refl Σ Re Rle gr napp u : - RelationClasses.Reflexive Re -> - RelationClasses.Reflexive Rle -> - R_global_instance Σ Re Rle gr napp u u. +Lemma cmp_universe_instance_refl cmp_universe : + RelationClasses.Reflexive cmp_universe -> + forall u, cmp_universe_instance cmp_universe u u. +Proof. + intros refl_univ u. + apply Forall2_same; reflexivity. +Qed. + +Lemma cmp_global_instance_refl Σ cmp_universe pb gr napp : + RelationClasses.Reflexive (cmp_universe Conv) -> + RelationClasses.Reflexive (cmp_universe pb) -> + forall u, cmp_global_instance Σ cmp_universe pb gr napp u u. Proof. intros rRE rRle. - rewrite /R_global_instance_gen. - destruct global_variance_gen as [v|] eqn:lookup. - - induction u in v |- *; simpl; auto; - unfold R_opt_variance in IHu; destruct v; simpl; auto. - split; auto. - destruct t; simpl; auto. - - apply Forall2_same; eauto. + unfold cmp_global_instance_gen. + destruct global_variance_gen as [| |v] => //= u. + - now apply cmp_universe_instance_refl. + - left. now apply cmp_universe_instance_refl. Qed. #[global] Instance eq_binder_annot_equiv {A} : RelationClasses.Equivalence (@eq_binder_annot A A). @@ -295,13 +356,15 @@ Proof. intros ? ? ? ? ?. now etransitivity. Qed. -Lemma eq_term_upto_univ_refl Σ Re Rle : - RelationClasses.Reflexive Re -> - RelationClasses.Reflexive Rle -> - forall napp t, eq_term_upto_univ_napp Σ Re Rle napp t t. +Lemma eq_term_upto_univ_refl Σ cmp_universe cmp_sort pb : + RelationClasses.Reflexive (cmp_universe Conv) -> + RelationClasses.Reflexive (cmp_universe pb) -> + RelationClasses.Reflexive (cmp_sort Conv) -> + RelationClasses.Reflexive (cmp_sort pb) -> + forall napp t, eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp t t. Proof. - intros hRe hRle napp. - induction t in napp, Rle, hRle |- * using term_forall_list_rect; simpl; + intros refl_univ refl_univ' refl_sort refl_sort' napp t. + induction t in napp, pb, refl_univ', refl_sort' |- * using term_forall_list_rect; simpl; try constructor; try apply Forall_Forall2; try apply All_All2 ; try easy; try now (try apply Forall_All ; apply Forall_True). - eapply All_All2. 1: eassumption. @@ -309,8 +372,8 @@ Proof. - destruct c; constructor. - eapply All_All2. 1: eassumption. intros. easy. - - now apply R_global_instance_refl. - - now apply R_global_instance_refl. + - now apply cmp_global_instance_refl. + - now apply cmp_global_instance_refl. - destruct X as [Ppars Preturn]. eapply All_All2. 1:eassumption. intros; easy. - destruct X as [Ppars Preturn]. now apply Preturn. @@ -325,16 +388,20 @@ Qed. Lemma eq_term_refl `{checker_flags} Σ φ t : eq_term Σ φ t t. Proof. apply eq_term_upto_univ_refl. - - intro; apply eq_universe_refl. - - intro; apply eq_universe_refl. + - apply eq_universe_refl. + - apply eq_universe_refl. + - apply eq_sort_refl. + - apply eq_sort_refl. Qed. Lemma leq_term_refl `{checker_flags} Σ φ t : leq_term Σ φ t t. Proof. apply eq_term_upto_univ_refl. - - intro; apply eq_universe_refl. - - intro; apply leq_universe_refl. + - apply eq_universe_refl. + - apply leq_universe_refl. + - intro; apply eq_sort_refl. + - intro; apply leq_sort_refl. Qed. (* Lemma eq_term_leq_term `{checker_flags} Σ φ napp t u : @@ -343,120 +410,82 @@ Proof. induction t in u |- * using term_forall_list_rect; simpl; inversion 1; subst; constructor; try (now unfold eq_term, leq_term in * ); try eapply Forall2_impl' ; try eapply All2_impl' ; try easy. - now apply eq_universe_leq_universe. - all: try (apply Forall_True, eq_universe_leq_universe). + now apply eq_sort_leq_sort. + all: try (apply Forall_True, eq_sort_leq_sort). apply IHt. Qed. *) -#[global] Instance R_global_instance_impl_same_napp Σ Re Re' Rle Rle' gr napp : - RelationClasses.subrelation Re Re' -> - RelationClasses.subrelation Rle Rle' -> - subrelation (R_global_instance Σ Re Rle gr napp) (R_global_instance Σ Re' Rle' gr napp). -Proof. - intros he hle t t'. - rewrite /R_global_instance_gen /R_opt_variance. - destruct global_variance_gen as [v|] eqn:glob. - induction t in v, t' |- *; destruct v, t'; simpl; auto. - intros []; split; auto. - destruct t0; simpl; auto. - now eapply R_universe_instance_impl'. -Qed. - -Lemma eq_term_upto_univ_morphism0 Σ (Re Re' : _ -> _ -> Prop) - (Hre : forall t u, Re t u -> Re' t u) - : forall t u napp, eq_term_upto_univ_napp Σ Re Re napp t u -> eq_term_upto_univ_napp Σ Re' Re' napp t u. -Proof. - fix aux 4. - destruct 1; constructor; eauto. - all: unfold R_universe_instance in *. - all: try solve[ match goal with - | H : All2 _ _ _ |- _ => clear -H aux; induction H; constructor; eauto - | H : Forall2 _ _ _ |- _ => induction H; constructor; eauto - end]. - - eapply R_global_instance_impl_same_napp; eauto. - - eapply R_global_instance_impl_same_napp; eauto. - - induction a1; constructor; auto. intuition auto. - - induction a; constructor; auto. intuition auto. - - induction a; constructor; auto. intuition auto. -Qed. - -Lemma eq_term_upto_univ_morphism Σ (Re Re' Rle Rle' : _ -> _ -> Prop) - (Hre : forall t u, Re t u -> Re' t u) - (Hrle : forall t u, Rle t u -> Rle' t u) - : forall t u napp, eq_term_upto_univ_napp Σ Re Rle napp t u -> eq_term_upto_univ_napp Σ Re' Rle' napp t u. +#[global] Instance cmp_global_instance_impl_same_napp Σ cmp_universe cmp_universe' pb pb' gr napp : + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe' Conv) -> + RelationClasses.subrelation (cmp_universe pb) (cmp_universe' pb') -> + subrelation (cmp_global_instance Σ cmp_universe pb gr napp) (cmp_global_instance Σ cmp_universe' pb' gr napp). Proof. - fix aux 4. - destruct 1; constructor; eauto using eq_term_upto_univ_morphism0. - all: unfold R_universe_instance in *. - all: try solve [match goal with - | H : Forall2 _ _ _ |- _ => induction H; constructor; - eauto using eq_term_upto_univ_morphism0 - | H : All2 _ _ _ |- _ => induction H; constructor; - eauto using eq_term_upto_univ_morphism0 - end]. - - clear X. induction a; constructor; eauto using eq_term_upto_univ_morphism0. - - eapply R_global_instance_impl_same_napp; eauto. - - eapply R_global_instance_impl_same_napp; eauto. - - clear X1 X2. induction a1; constructor; eauto using eq_term_upto_univ_morphism0. - destruct r0. split; eauto using eq_term_upto_univ_morphism0. - - induction a; constructor; eauto using eq_term_upto_univ_morphism0. - destruct r as [[[? ?] ?] ?]. - repeat split; eauto using eq_term_upto_univ_morphism0. - - induction a; constructor; eauto using eq_term_upto_univ_morphism0. - destruct r as [[[? ?] ?] ?]. - repeat split; eauto using eq_term_upto_univ_morphism0. + intros sub_conv sub_pb u u'. + unfold cmp_global_instance_gen, cmp_opt_variance. + destruct global_variance_gen as [| |v] => //. + 1: now apply cmp_universe_instance_impl. + + intros [H | H]; [left | right]. + 1: eapply cmp_universe_instance_impl; tea. + + eapply Forall3_impl; tea. clear v u u' H. + intros v u u'. + destruct v => //=. + all: unfold on_rel; now auto. Qed. - -Lemma global_variance_napp_mon {Σ gr napp napp' v} : +Lemma global_variance_napp_mon Σ gr {napp napp'} : napp <= napp' -> - global_variance Σ gr napp = Some v -> - global_variance Σ gr napp' = Some v. + match global_variance Σ gr napp with + | Variance v => global_variance Σ gr napp' = Variance v + | AllEqual => True + | AllIrrelevant => global_variance Σ gr napp' = AllIrrelevant + end. Proof. intros hnapp. rewrite /global_variance_gen. - destruct gr; try congruence. - - destruct lookup_inductive_gen as [[mdecl idec]|] => //. - destruct destArity as [[ctx s]|] => //. - elim: Nat.leb_spec => // cass indv. + destruct gr => //=. + - destruct lookup_inductive_gen as [[mdecl idec]|] => //=. + destruct destArity as [[ctx s]|] => //=. + elim: Nat.leb_spec => // cass. + destruct ind_variance => //=. elim: Nat.leb_spec => //. lia. - - destruct lookup_constructor_gen as [[[mdecl idecl] cdecl]|] => //. - elim: Nat.leb_spec => // cass indv. + - destruct lookup_constructor_gen as [[[mdecl idecl] cdecl]|] => //=. + elim: Nat.leb_spec => // cass. elim: Nat.leb_spec => //. lia. Qed. -#[global] Instance R_global_instance_impl Σ Re Re' Rle Rle' gr napp napp' : - RelationClasses.subrelation Re Re' -> - RelationClasses.subrelation Re Rle' -> - RelationClasses.subrelation Rle Rle' -> +#[global] +Instance cmp_global_instance_impl Σ cmp_universe cmp_universe' pb pb' gr napp napp' : + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe' Conv) -> + RelationClasses.subrelation (cmp_universe pb) (cmp_universe' pb') -> napp <= napp' -> - subrelation (R_global_instance Σ Re Rle gr napp) (R_global_instance Σ Re' Rle' gr napp'). + subrelation (cmp_global_instance Σ cmp_universe pb gr napp) (cmp_global_instance Σ cmp_universe' pb' gr napp'). Proof. - intros he hle hele hnapp t t'. - rewrite /R_global_instance_gen /R_opt_variance. - destruct global_variance_gen as [v|] eqn:glob. - rewrite (global_variance_napp_mon hnapp glob). - induction t in v, t' |- *; destruct v, t'; simpl; auto. - intros []; split; auto. - destruct t0; simpl; auto. - destruct (global_variance _ _ napp') as [v|] eqn:glob'; eauto using R_universe_instance_impl'. - induction t in v, t' |- *; destruct v, t'; simpl; auto; intros H; inv H. - eauto. - split; auto. - destruct t0; simpl; auto. + intros sub_conv sub_pb lenapp u u'. + unfold cmp_global_instance_gen. + move: (global_variance_napp_mon Σ gr lenapp). + destruct global_variance_gen as [| |v] => //. + all: [> intros _ | intros ->; cbn ..]; auto. + 1: intro H; apply: cmp_instance_opt_variance; move: H => /=. + - now apply cmp_universe_instance_impl. + - intros [H | H]; [left | right]. + 1: eapply cmp_universe_instance_impl; tea. + eapply cmp_universe_instance_variance_impl; eassumption. Qed. -#[global] Instance eq_term_upto_univ_impl Σ Re Re' Rle Rle' napp napp' : - RelationClasses.subrelation Re Re' -> - RelationClasses.subrelation Rle Rle' -> - RelationClasses.subrelation Re Rle' -> +#[global] Instance eq_term_upto_univ_impl Σ cmp_universe cmp_universe' cmp_sort cmp_sort' pb pb' napp napp' : + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe' Conv) -> + RelationClasses.subrelation (cmp_universe pb) (cmp_universe' pb') -> + RelationClasses.subrelation (cmp_sort Conv) (cmp_sort' Conv) -> + RelationClasses.subrelation (cmp_sort pb) (cmp_sort' pb') -> napp <= napp' -> - subrelation (eq_term_upto_univ_napp Σ Re Rle napp) (eq_term_upto_univ_napp Σ Re' Rle' napp'). + subrelation (eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp) (eq_term_upto_univ_napp Σ cmp_universe' cmp_sort' pb' napp'). Proof. - intros he hle hele hnapp t t'. - induction t in napp, napp', hnapp, t', Rle, Rle', hle, hele |- * using term_forall_list_rect; + intros univ_sub_conv univ_sub_pb sort_sub_conv sort_sub_pb hnapp t t'. + induction t in napp, napp', hnapp, pb, pb', univ_sub_pb, sort_sub_pb, t' |- * using term_forall_list_rect; try (inversion 1; subst; constructor; - eauto using R_universe_instance_impl'; fail). + eauto using cmp_universe_instance_impl'; fail). - inversion 1; subst; constructor. eapply All2_impl'; tea. eapply All_impl; eauto. @@ -464,14 +493,14 @@ Proof. eapply IHt. 4:eauto. all:auto with arith. eauto. solve_all. - inversion 1; subst; constructor. - eapply R_global_instance_impl. 5:eauto. all:auto. + eapply cmp_global_instance_impl. 4:eauto. all:auto. - inversion 1; subst; constructor. - eapply R_global_instance_impl. 5:eauto. all:eauto. + eapply cmp_global_instance_impl. 4:eauto. all:eauto. - destruct X as [IHpars IHret]. inversion 1; subst; constructor; eauto. eapply All2_impl'; tea. eapply All_impl; eauto. - eapply R_universe_instance_impl; eauto. + eapply cmp_universe_instance_impl; eauto. eapply All2_impl'; eauto. cbn. eapply All_impl; eauto. @@ -485,20 +514,44 @@ Proof. eapply All_impl; eauto. cbn. intros x [? ?] y [[[? ?] ?] ?]. repeat split; eauto. - inversion 1; subst; constructor; eauto. + 1: eapply cmp_universe_instance_impl; eauto. eapply All2_impl'; tea. eapply All_impl; eauto. Qed. +Lemma eq_term_upto_univ_morphism0 Σ cmp_universe cmp_universe' cmp_sort cmp_sort' : + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe' Conv) -> + RelationClasses.subrelation (cmp_sort Conv) (cmp_sort' Conv) -> + forall t u napp, eq_term_upto_univ_napp Σ cmp_universe cmp_sort Conv napp t u -> eq_term_upto_univ_napp Σ cmp_universe' cmp_sort' Conv napp t u. +Proof. + intros univ_sub sort_sub t u napp. + apply eq_term_upto_univ_impl. + 5: auto with arith. all: auto. +Qed. + +Lemma eq_term_upto_univ_morphism Σ cmp_universe cmp_universe' cmp_sort cmp_sort' pb pb' : + RelationClasses.subrelation (cmp_universe Conv) (cmp_universe' Conv) -> + RelationClasses.subrelation (cmp_universe pb) (cmp_universe' pb') -> + RelationClasses.subrelation (cmp_sort Conv) (cmp_sort' Conv) -> + RelationClasses.subrelation (cmp_sort pb) (cmp_sort' pb') -> + forall t u napp, eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp t u -> eq_term_upto_univ_napp Σ cmp_universe' cmp_sort' pb' napp t u. +Proof. + intros univ_sub univ_sub_pb sort_sub sort_sub_pb t u napp. + apply eq_term_upto_univ_impl. + 5: auto with arith. all: auto. +Qed. Lemma eq_term_leq_term `{checker_flags} Σ φ t u : eq_term Σ φ t u -> leq_term Σ φ t u. Proof. - eapply eq_term_upto_univ_morphism. auto. - intros. - now apply eq_universe_leq_universe. + eapply eq_term_upto_univ_morphism. + - reflexivity. + - apply eq_leq_universe. + - reflexivity. + - apply eq_leq_sort. Qed. -Lemma eq_term_upto_univ_App `{checker_flags} Σ Re Rle napp f f' : - eq_term_upto_univ_napp Σ Re Rle napp f f' -> +Lemma eq_term_upto_univ_App `{checker_flags} Σ cmp_universe cmp_sort pb napp f f' : + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp f f' -> isApp f = isApp f'. Proof. inversion 1; reflexivity. @@ -511,14 +564,14 @@ Proof. inversion 1; reflexivity. Qed. -Lemma eq_term_upto_univ_mkApps `{checker_flags} Σ Re Rle napp f l f' l' : - eq_term_upto_univ_napp Σ Re Rle (#|l| + napp) f f' -> - All2 (eq_term_upto_univ Σ Re Re) l l' -> - eq_term_upto_univ_napp Σ Re Rle napp (mkApps f l) (mkApps f' l'). +Lemma eq_term_upto_univ_mkApps `{checker_flags} Σ cmp_universe cmp_sort pb napp f l f' l' : + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb (#|l| + napp) f f' -> + All2 (eq_term_upto_univ Σ cmp_universe cmp_sort Conv) l l' -> + eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp (mkApps f l) (mkApps f' l'). Proof. induction l in f, f' |- *; intro e; inversion_clear 1. - assumption. - - pose proof (eq_term_upto_univ_App _ _ _ _ _ _ e). + - pose proof (eq_term_upto_univ_App _ _ _ _ _ _ _ e). case_eq (isApp f). + intro X; rewrite X in H0. destruct f; try discriminate. @@ -540,10 +593,9 @@ Lemma leq_term_mkApps `{checker_flags} Σ φ f l f' l' : leq_term Σ φ (mkApps f l) (mkApps f' l'). Proof. intros. - eapply eq_term_upto_univ_mkApps. - eapply eq_term_upto_univ_impl. 5:eauto. 4:auto with arith. - 1-3:typeclasses eauto. - apply X0. + eapply eq_term_upto_univ_mkApps. 2: assumption. + eapply eq_term_upto_univ_impl. 6:eauto. 5:auto with arith. + all:typeclasses eauto. Qed. Lemma leq_term_App `{checker_flags} Σ φ f f' : @@ -551,4 +603,4 @@ Lemma leq_term_App `{checker_flags} Σ φ f f' : isApp f = isApp f'. Proof. inversion 1; reflexivity. -Qed. \ No newline at end of file +Qed. diff --git a/template-coq/theories/Typing.v b/template-coq/theories/Typing.v index 7fa6c4555..2a313d667 100644 --- a/template-coq/theories/Typing.v +++ b/template-coq/theories/Typing.v @@ -521,7 +521,7 @@ Reserved Notation " Σ ;;; Γ |- t <=[ pb ] u " (at level 50, Γ, t, u at next l *) Inductive cumul_gen `{checker_flags} (Σ : global_env_ext) (Γ : context) (pb : conv_pb) : term -> term -> Type := - | cumul_refl t u : compare_term pb Σ (global_ext_constraints Σ) t u -> Σ ;;; Γ |- t <=[pb] u + | cumul_refl t u : compare_term Σ (global_ext_constraints Σ) pb t u -> Σ ;;; Γ |- t <=[pb] u | cumul_red_l t u v : red1 Σ.1 Γ t v -> Σ ;;; Γ |- v <=[pb] u -> Σ ;;; Γ |- t <=[pb] u | cumul_red_r t u v : Σ ;;; Γ |- t <=[pb] v -> red1 Σ.1 Γ u v -> Σ ;;; Γ |- t <=[pb] u where "Σ ;;; Γ |- t <=[ pb ] u " := (cumul_gen Σ Γ pb t u). @@ -772,29 +772,28 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> | type_Sort s : wf_local Σ Γ -> - wf_universe Σ s -> - Σ ;;; Γ |- tSort s : tSort (Universe.super s) + wf_sort Σ s -> + Σ ;;; Γ |- tSort s : tSort (Sort.super s) | type_Cast c k t s : Σ ;;; Γ |- t : tSort s -> Σ ;;; Γ |- c : t -> Σ ;;; Γ |- tCast c k t : t -| type_Prod n t b s1 s2 : - Σ ;;; Γ |- t : tSort s1 -> - Σ ;;; Γ ,, vass n t |- b : tSort s2 -> - Σ ;;; Γ |- tProd n t b : tSort (Universe.sort_of_product s1 s2) +| type_Prod na t b s1 s2 : + lift_typing typing Σ Γ (j_vass_s na t s1) -> + Σ ;;; Γ ,, vass na t |- b : tSort s2 -> + Σ ;;; Γ |- tProd na t b : tSort (Sort.sort_of_product s1 s2) -| type_Lambda n t b s1 bty : - Σ ;;; Γ |- t : tSort s1 -> - Σ ;;; Γ ,, vass n t |- b : bty -> - Σ ;;; Γ |- tLambda n t b : tProd n t bty +| type_Lambda na t b bty : + lift_typing typing Σ Γ (j_vass na t) -> + Σ ;;; Γ ,, vass na t |- b : bty -> + Σ ;;; Γ |- tLambda na t b : tProd na t bty -| type_LetIn n b b_ty b' s1 b'_ty : - Σ ;;; Γ |- b_ty : tSort s1 -> - Σ ;;; Γ |- b : b_ty -> - Σ ;;; Γ ,, vdef n b b_ty |- b' : b'_ty -> - Σ ;;; Γ |- tLetIn n b b_ty b' : tLetIn n b b_ty b'_ty +| type_LetIn na b b_ty b' b'_ty : + lift_typing typing Σ Γ (j_vdef na b b_ty) -> + Σ ;;; Γ ,, vdef na b b_ty |- b' : b'_ty -> + Σ ;;; Γ |- tLetIn na b b_ty b' : tLetIn na b b_ty b'_ty | type_App t l t_ty t' : Σ ;;; Γ |- t : t_ty -> @@ -827,8 +826,8 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> context_assumptions mdecl.(ind_params) = #|p.(pparams)| -> consistent_instance_ext Σ (ind_universes mdecl) p.(puinst) -> let predctx := case_predicate_context ci.(ci_ind) mdecl idecl p in - ctx_inst typing Σ Γ (p.(pparams) ++ indices) - (List.rev (ind_params mdecl ,,, ind_indices idecl)@[p.(puinst)]) -> + ctx_inst (typing Σ) Γ (p.(pparams) ++ indices) + (List.rev (ind_params mdecl ,,, ind_indices idecl : context)@[p.(puinst)]) -> Σ ;;; Γ ,,, predctx |- p.(preturn) : tSort ps -> is_allowed_elimination Σ idecl.(ind_kelim) ps -> Σ ;;; Γ |- c : mkApps (tInd ci.(ci_ind) p.(puinst)) (p.(pparams) ++ indices) -> @@ -852,8 +851,8 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> fix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> wf_local Σ Γ -> - All (fun d => {s & Σ ;;; Γ |- d.(dtype) : tSort s}) mfix -> - All (fun d => (Σ ;;; Γ ,,, fix_context mfix |- d.(dbody) : lift0 #|fix_context mfix| d.(dtype))) mfix -> + All (on_def_type (lift_typing1 (typing Σ)) Γ) mfix -> + All (on_def_body (lift_typing1 (typing Σ)) (fix_context mfix) Γ) mfix -> wf_fixpoint Σ mfix -> Σ ;;; Γ |- tFix mfix n : decl.(dtype) @@ -861,8 +860,8 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> cofix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> wf_local Σ Γ -> - All (fun d => {s & Σ ;;; Γ |- d.(dtype) : tSort s}) mfix -> - All (fun d => Σ ;;; Γ ,,, fix_context mfix |- d.(dbody) : lift0 #|fix_context mfix| d.(dtype)) mfix -> + All (on_def_type (lift_typing1 (typing Σ)) Γ) mfix -> + All (on_def_body (lift_typing1 (typing Σ)) (fix_context mfix) Γ) mfix -> wf_cofixpoint Σ mfix -> Σ ;;; Γ |- tCoFix mfix n : decl.(dtype) @@ -881,15 +880,15 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> Σ ;;; Γ |- tFloat p : tConst prim_ty [] | type_Array prim_ty cdecl u arr def ty : - wf_local Σ Γ -> - primitive_constant Σ primArray = Some prim_ty -> - declared_constant Σ prim_ty cdecl -> - primitive_invariants primArray cdecl -> - let s := Universe.make u in - Σ ;;; Γ |- ty : tSort s -> - Σ ;;; Γ |- def : ty -> - All (fun t => Σ ;;; Γ |- t : ty) arr -> - Σ ;;; Γ |- tArray u arr def ty : tApp (tConst prim_ty [u]) [ty] + wf_local Σ Γ -> + primitive_constant Σ primArray = Some prim_ty -> + declared_constant Σ prim_ty cdecl -> + primitive_invariants primArray cdecl -> + let s := sType (Universe.make' u) in + Σ ;;; Γ |- ty : tSort s -> + Σ ;;; Γ |- def : ty -> + All (fun t => Σ ;;; Γ |- t : ty) arr -> + Σ ;;; Γ |- tArray u arr def ty : tApp (tConst prim_ty [u]) [ty] | type_Conv t A B s : Σ ;;; Γ |- t : A -> @@ -897,7 +896,7 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> Σ ;;; Γ |- A <= B -> Σ ;;; Γ |- t : B where " Σ ;;; Γ |- t : T " := (typing Σ Γ t T) : type_scope - and "'wf_local' Σ Γ " := (All_local_env (lift_typing typing Σ) Γ) : type_scope + and "'wf_local' Σ Γ " := (All_local_env (lift_typing1 (typing Σ)) Γ) : type_scope (* Typing of "spines", currently just the arguments of applications *) @@ -921,13 +920,10 @@ Definition unlift_opt_pred (P : global_env_ext -> context -> option term -> term (global_env_ext -> context -> term -> term -> Type) := fun Σ Γ t T => P Σ Γ (Some t) T. -Definition infer_sorting {cf: checker_flags} Σ Γ T := { s : Universe.t & typing Σ Γ T (tSort s) }. - Module TemplateTyping <: Typing TemplateTerm Env TemplateTermUtils TemplateEnvTyping TemplateConversion TemplateConversionPar. Definition typing := @typing. - Definition infer_sorting := @infer_sorting. End TemplateTyping. @@ -962,7 +958,7 @@ Section CtxInstSize. Context (typing : global_env_ext -> context -> term -> term -> Type) (typing_size : forall {Σ Γ t T}, typing Σ Γ t T -> size). - Fixpoint ctx_inst_size {Σ Γ args Δ} (c : ctx_inst typing Σ Γ args Δ) : size := + Fixpoint ctx_inst_size {Σ Γ args Δ} (c : ctx_inst (typing Σ) Γ args Δ) : size := match c with | ctx_inst_nil => 0 | ctx_inst_ass na t i inst Δ ty ctxi => ((typing_size _ _ _ _ ty) + (ctx_inst_size ctxi))%nat @@ -987,25 +983,33 @@ Proof. | H : Alli _ _ _ |- _ => idtac | H : typing_spine _ _ _ _ _ |- _ => idtac | H : _ + _ |- _ => idtac + | H : lift_typing0 _ _ |- _ => idtac | H1 : size, H2 : size, H3 : size |- _ => exact (S (Nat.max H1 (Nat.max H2 H3))) | H1 : size, H2 : size |- _ => exact (S (Nat.max H1 H2)) | H1 : size |- _ => exact (S H1) | _ => exact 1 end. - - exact (S (All_local_env_size typing_size _ _ a)). - - exact (S (All_local_env_size typing_size _ _ a)). + - exact (S (All_local_env_size (typing_size _) _ a)). + - exact (S (All_local_env_size (typing_size _) _ a)). + - exact (S (Nat.max d (lift_typing_size (typing_size _ _) _ l))). + - exact (S (Nat.max d (lift_typing_size (typing_size _ _) _ l))). + - exact (S (Nat.max d (lift_typing_size (typing_size _ _) _ l))). - exact (S (Nat.max d (typing_spine_size typing_size _ _ _ _ _ t0))). - - exact (S (S (All_local_env_size typing_size _ _ a))). - - exact (S (S (All_local_env_size typing_size _ _ a))). - - exact (S (S (All_local_env_size typing_size _ _ a))). + - exact (S (S (All_local_env_size (typing_size _) _ a))). + - exact (S (S (All_local_env_size (typing_size _) _ a))). + - exact (S (S (All_local_env_size (typing_size _) _ a))). - exact (S (Nat.max d1 (Nat.max d2 (Nat.max (ctx_inst_size _ typing_size c1) (all2i_size _ (fun _ x y p => Nat.max (typing_size _ _ _ _ p.1.2) (typing_size _ _ _ _ p.2)) a0))))). - - exact (S (Nat.max (Nat.max (All_local_env_size typing_size _ _ a) (all_size _ (fun x p => infer_sort_size (typing_sort_size typing_size) Σ _ _ p) a0)) (all_size _ (fun x p => typing_size Σ _ _ _ p) a1))). - - exact (S (Nat.max (Nat.max (All_local_env_size typing_size _ _ a) (all_size _ (fun x p => infer_sort_size (typing_sort_size typing_size) Σ _ _ p) a0)) (all_size _ (fun x p => typing_size Σ _ _ _ p) a1))). - - exact (S (All_local_env_size typing_size _ _ a)). - - exact (S (All_local_env_size typing_size _ _ a)). - - exact (S (Nat.max d2 (Nat.max d3 (all_size _ (fun t => typing_size Σ Γ t ty) a0)))). + - exact (S (Nat.max + (Nat.max (All_local_env_size (typing_size _) _ a) (all_size _ (fun x p => on_def_type_size (typing_size Σ) _ _ p) a0)) + (all_size (on_def_body (lift_typing typing _) _ _) (fun x p => on_def_body_size (typing_size Σ) _ _ _ p) a1))). + - exact (S (Nat.max + (Nat.max (All_local_env_size (typing_size _) _ a) (all_size _ (fun x p => on_def_type_size (typing_size Σ) _ _ p) a0)) + (all_size (on_def_body (lift_typing typing _) _ _) (fun x p => on_def_body_size (typing_size Σ) _ _ _ p) a1))). + - exact (S (All_local_env_size (typing_size _) _ a)). + - exact (S (All_local_env_size (typing_size _) _ a)). + - exact (S (Nat.max (All_local_env_size (typing_size _) _ a) (Nat.max d2 (Nat.max d3 (all_size _ (fun t p => typing_size Σ Γ t ty p) a0))))). Defined. Lemma typing_size_pos `{checker_flags} {Σ Γ t T} (d : Σ ;;; Γ |- t : T) : typing_size d > 0. @@ -1033,41 +1037,44 @@ Arguments lexprod [A B]. Definition wf `{checker_flags} Σ := on_global_env cumul_gen (lift_typing typing) Σ. Definition wf_ext `{checker_flags} := on_global_env_ext cumul_gen (lift_typing typing). -Lemma typing_wf_local `{checker_flags} {Σ} {Γ t T} : +Lemma typing_wf_local `{cf : checker_flags} {Σ} {Γ t T} : Σ ;;; Γ |- t : T -> wf_local Σ Γ. Proof. - induction 1; eauto. + revert Σ Γ t T. + fix f 5. + destruct 1; eauto. + all: exact (f _ _ _ _ l.2.π2.1). Defined. #[global] Hint Resolve typing_wf_local : wf. Lemma type_Prop `{checker_flags} Σ : - Σ ;;; [] |- tSort Universe.lProp : tSort Universe.type1. - change ( Σ ;;; [] |- tSort (Universe.lProp) : tSort Universe.type1); + Σ ;;; [] |- tSort sProp : tSort Sort.type1. + change ( Σ ;;; [] |- tSort (sProp) : tSort Sort.type1); constructor;auto. constructor. constructor. Defined. Lemma type_Prop_wf `{checker_flags} Σ Γ : wf_local Σ Γ -> - Σ ;;; Γ |- tSort Universe.lProp : tSort Universe.type1. + Σ ;;; Γ |- tSort sProp : tSort Sort.type1. Proof. constructor;auto. constructor. Defined. -Definition env_prop `{checker_flags} (P : forall Σ Γ t T, Type) (PΓ : forall Σ Γ (wfΓ : wf_local Σ Γ), Type):= +Definition env_prop `{checker_flags} (P : forall Σ Γ t T, Type) (Pj : forall Σ Γ j, Type) (PΓ : forall Σ Γ (wfΓ : wf_local Σ Γ), Type):= forall (Σ : global_env_ext) (wfΣ : wf Σ) Γ (wfΓ : wf_local Σ Γ) t T (ty : Σ ;;; Γ |- t : T), - on_global_env cumul_gen (lift_typing P) Σ * (PΓ Σ Γ (typing_wf_local ty) * P Σ Γ t T). + on_global_env cumul_gen Pj Σ * (PΓ Σ Γ (typing_wf_local ty) * P Σ Γ t T). -Lemma env_prop_typing `{checker_flags} {P PΓ} : env_prop P PΓ -> +Lemma env_prop_typing `{checker_flags} {P Pj PΓ} : env_prop P Pj PΓ -> forall (Σ : global_env_ext) (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) (t T : term), Σ ;;; Γ |- t : T -> P Σ Γ t T. Proof. intros. now apply X. Qed. -Lemma env_prop_wf_local `{checker_flags} {P PΓ} : env_prop P PΓ -> +Lemma env_prop_wf_local `{checker_flags} {P Pj PΓ} : env_prop P Pj PΓ -> forall (Σ : global_env_ext) (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ), PΓ Σ Γ wfΓ. Proof. intros. red in X. now apply (X _ wfΣ _ wfΓ _ _ (type_Prop_wf Σ Γ wfΓ)). Qed. -Lemma env_prop_sigma `{checker_flags} {P PΓ} : env_prop P PΓ -> - forall Σ (wfΣ : wf Σ), on_global_env cumul_gen (lift_typing P) Σ. +Lemma env_prop_sigma `{checker_flags} {P Pj PΓ} : env_prop P Pj PΓ -> + forall Σ (wfΣ : wf Σ), on_global_env cumul_gen Pj Σ. Proof. intros. eapply (X (empty_ext Σ)). apply wfΣ. constructor. @@ -1081,52 +1088,34 @@ Proof. Defined. #[global] Hint Immediate wf_local_app_l : wf. -Lemma typing_wf_local_size `{checker_flags} {Σ} {Γ t T} +Lemma typing_wf_local_size `{cf : checker_flags} {Σ} {Γ t T} (d :Σ ;;; Γ |- t : T) : - All_local_env_size (@typing_size _) _ _ (typing_wf_local d) < typing_size d. + All_local_env_size (@typing_size _ _) _ (typing_wf_local d) < typing_size d. Proof. + revert Σ Γ t T d. fix f 5. induction d; simpl; - change (fun (x : global_env_ext) (x0 : context) (x1 x2 : term) - (x3 : x;;; x0 |- x1 : x2) => typing_size x3) with (@typing_size H); try lia. + change (fun Γ t T (Hty : Σ;;; Γ |- t : T) => typing_size Hty) with (@typing_size cf Σ); try lia. + 1-3: pose proof (f _ _ _ _ l.2.π2.1); unfold lift_sorting_size; cbn in *; try lia. + Qed. Lemma wf_local_inv `{checker_flags} {Σ Γ'} (w : wf_local Σ Γ') : forall d Γ, Γ' = d :: Γ -> - ∑ w' : wf_local Σ Γ, - match d.(decl_body) with - | Some b => - ∑ u (ty : Σ ;;; Γ |- b : d.(decl_type)), - { ty' : Σ ;;; Γ |- d.(decl_type) : tSort u | - All_local_env_size (@typing_size _) Σ _ w' < - All_local_env_size (@typing_size _) _ _ w /\ - typing_size ty <= All_local_env_size (@typing_size _) _ _ w /\ - typing_size ty' <= All_local_env_size (@typing_size _) _ _ w } - - | None => - ∑ u, - { ty : Σ ;;; Γ |- d.(decl_type) : tSort u | - All_local_env_size (@typing_size _) Σ _ w' < - All_local_env_size (@typing_size _) _ _ w /\ - typing_size ty <= All_local_env_size (@typing_size _) _ _ w } - end. + ∑ (w' : wf_local Σ Γ) u, + { ty : lift_typing1 (typing Σ) Γ (Judge d.(decl_body) d.(decl_type) (Some u)) | + All_local_env_size (@typing_size _ Σ) _ w' < + All_local_env_size (@typing_size _ _) _ w /\ + lift_typing_size (@typing_size _ _ _) _ ty <= All_local_env_size (@typing_size _ _) _ w }. Proof. - intros d Γ. - destruct w. - - simpl. congruence. - - intros [=]. subst d Γ0. - exists w. simpl. destruct l. exists x. exists t0. pose (typing_size_pos t0). - cbn. split. - + lia. - + auto with arith. - - intros [=]. subst d Γ0. - exists w. simpl. simpl in l. destruct l as [u h]. - simpl in l0. - exists u, l0, h. cbn. - pose (typing_size_pos h). - pose (typing_size_pos l0). - intuition eauto. - all: try lia. + intros d Γ ->. + depelim w; cbn. + all: exists w, l.2.π1, (lift_sorting_extract l). + all: pose proof (typing_size_pos l.2.π2.1). + 2: pose proof (typing_size_pos l.1). + all: simpl in *. + all: simpl; split. + all: try lia. Qed. @@ -1150,48 +1139,48 @@ Implicit Types (Σ : global_env_ext). Lemma typing_ind_env `{cf : checker_flags} : forall (P : global_env_ext -> context -> term -> term -> Type) + (Pj : global_env_ext -> context -> judgment -> Type) (Pdecl := fun Σ Γ wfΓ t T tyT => P Σ Γ t T) (PΓ : forall Σ Γ, wf_local Σ Γ -> Type), + (forall Σ (wfΣ : wf Σ) (Γ : context) (* (wfΓ : wf_local Σ Γ) *) j, + lift_typing_conj (typing Σ) (P Σ) Γ j -> Pj Σ Γ j) -> + (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ), - All_local_env_over typing Pdecl Σ Γ wfΓ -> PΓ Σ Γ wfΓ) -> + All_local_env_over (typing Σ) (Pdecl Σ) Γ wfΓ -> All_local_env (Pj Σ) Γ -> PΓ Σ Γ wfΓ) -> (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) (n : nat) decl, nth_error Γ n = Some decl -> PΓ Σ Γ wfΓ -> P Σ Γ (tRel n) (lift0 (S n) decl.(decl_type))) -> - (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) (u : Universe.t), + (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) (u : sort), PΓ Σ Γ wfΓ -> - wf_universe Σ u -> - P Σ Γ (tSort u) (tSort (Universe.super u))) -> + wf_sort Σ u -> + P Σ Γ (tSort u) (tSort (Sort.super u))) -> (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) (c : term) (k : cast_kind) - (t : term) (s : Universe.t), + (t : term) (s : sort), Σ ;;; Γ |- t : tSort s -> P Σ Γ t (tSort s) -> Σ ;;; Γ |- c : t -> P Σ Γ c t -> P Σ Γ (tCast c k t) t) -> - (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) (n : aname) (t b : term) (s1 s2 : Universe.t), + (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (t b : term) (s1 s2 : sort), PΓ Σ Γ wfΓ -> - Σ ;;; Γ |- t : tSort s1 -> - P Σ Γ t (tSort s1) -> - Σ ;;; Γ,, vass n t |- b : tSort s2 -> - P Σ (Γ,, vass n t) b (tSort s2) -> P Σ Γ (tProd n t b) (tSort (Universe.sort_of_product s1 s2))) -> + lift_typing typing Σ Γ (j_vass_s na t s1) -> + Pj Σ Γ (j_vass_s na t s1) -> + Σ ;;; Γ,, vass na t |- b : tSort s2 -> + P Σ (Γ,, vass na t) b (tSort s2) -> P Σ Γ (tProd na t b) (tSort (Sort.sort_of_product s1 s2))) -> - (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) (n : aname) (t b : term) - (s1 : Universe.t) (bty : term), + (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (t b : term) (bty : term), PΓ Σ Γ 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)) -> + lift_typing typing Σ Γ (j_vass na t) -> + Pj Σ Γ (j_vass na t) -> + Σ ;;; Γ,, vass na t |- b : bty -> P Σ (Γ,, vass na t) b bty -> P Σ Γ (tLambda na t b) (tProd na t bty)) -> - (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) (n : aname) (b b_ty b' : term) - (s1 : Universe.t) (b'_ty : term), + (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (b b_ty b' b'_ty : term), PΓ Σ Γ wfΓ -> - Σ ;;; Γ |- b_ty : tSort s1 -> - P Σ Γ b_ty (tSort s1) -> - Σ ;;; Γ |- b : b_ty -> - P Σ Γ b b_ty -> - Σ ;;; Γ,, vdef n b b_ty |- b' : b'_ty -> - P Σ (Γ,, vdef n b b_ty) b' b'_ty -> P Σ Γ (tLetIn n b b_ty b') (tLetIn n b b_ty b'_ty)) -> + lift_typing typing Σ Γ (j_vdef na b b_ty) -> + Pj Σ Γ (j_vdef na b b_ty) -> + Σ ;;; Γ,, vdef na b b_ty |- b' : b'_ty -> + P Σ (Γ,, vdef na b b_ty) b' b'_ty -> P Σ Γ (tLetIn na b b_ty b') (tLetIn na b b_ty b'_ty)) -> (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) (t : term) (l : list term) (t_ty t' : term), Σ ;;; Γ |- t : t_ty -> P Σ Γ t t_ty -> @@ -1201,7 +1190,7 @@ Lemma typing_ind_env `{cf : checker_flags} : P Σ Γ (tApp t l) t') -> (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) cst u (decl : constant_body), - on_global_env cumul_gen (lift_typing P) Σ.1 -> + on_global_env cumul_gen Pj Σ.1 -> PΓ Σ Γ wfΓ -> declared_constant Σ.1 cst decl -> consistent_instance_ext Σ decl.(cst_universes) u -> @@ -1209,14 +1198,14 @@ Lemma typing_ind_env `{cf : checker_flags} : (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) (ind : inductive) u mdecl idecl (isdecl : declared_inductive Σ.1 ind mdecl idecl), - on_global_env cumul_gen (lift_typing P) Σ.1 -> + on_global_env cumul_gen Pj Σ.1 -> PΓ Σ Γ wfΓ -> consistent_instance_ext Σ mdecl.(ind_universes) u -> P Σ Γ (tInd ind u) (subst_instance u (ind_type idecl))) -> (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) (ind : inductive) (i : nat) u mdecl idecl cdecl (isdecl : declared_constructor Σ.1 (ind, i) mdecl idecl cdecl), - on_global_env cumul_gen (lift_typing P) Σ.1 -> + on_global_env cumul_gen Pj Σ.1 -> PΓ Σ Γ wfΓ -> consistent_instance_ext Σ mdecl.(ind_universes) u -> P Σ Γ (tConstruct ind i u) (type_of_constructor mdecl cdecl (ind, i) u)) -> @@ -1224,15 +1213,15 @@ Lemma typing_ind_env `{cf : checker_flags} : (forall (Σ : global_env_ext) (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ), forall (ci : case_info) p c brs indices ps mdecl idecl (isdecl : declared_inductive Σ.1 ci.(ci_ind) mdecl idecl), - on_global_env cumul_gen (lift_typing P) Σ.1 -> + on_global_env cumul_gen Pj Σ.1 -> PΓ Σ Γ wfΓ -> mdecl.(ind_npars) = ci.(ci_npar) -> wf_nactx p.(pcontext) (ind_predicate_context ci.(ci_ind) mdecl idecl) -> context_assumptions mdecl.(ind_params) = #|p.(pparams)| -> consistent_instance_ext Σ (ind_universes mdecl) p.(puinst) -> let predctx := case_predicate_context ci.(ci_ind) mdecl idecl p in - ctx_inst (Prop_conj typing P) Σ Γ (p.(pparams) ++ indices) - (List.rev (ind_params mdecl ,,, ind_indices idecl)@[p.(puinst)]) -> + ctx_inst (Prop_conj typing P Σ) Γ (p.(pparams) ++ indices) + (List.rev (ind_params mdecl ,,, ind_indices idecl : context)@[p.(puinst)]) -> forall pret : Σ ;;; Γ ,,, predctx |- p.(preturn) : tSort ps, P Σ (Γ ,,, predctx) p.(preturn) (tSort ps) -> PΓ Σ (Γ ,,, predctx) (typing_wf_local pret) -> @@ -1251,7 +1240,7 @@ Lemma typing_ind_env `{cf : checker_flags} : (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) (p : projection) (c : term) u mdecl idecl cdecl pdecl (isdecl : declared_projection Σ.1 p mdecl idecl cdecl pdecl) args, - on_global_env cumul_gen (lift_typing P) Σ.1 -> + on_global_env cumul_gen Pj Σ.1 -> PΓ Σ Γ wfΓ -> Σ ;;; Γ |- c : mkApps (tInd p.(proj_ind) u) args -> P Σ Γ c (mkApps (tInd p.(proj_ind) u) args) -> @@ -1263,9 +1252,11 @@ Lemma typing_ind_env `{cf : checker_flags} : fix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> PΓ Σ Γ wfΓ -> - All (on_def_type (lift_typing2 typing P Σ) Γ) mfix -> - All (on_def_body (lift_typing2 typing P Σ) types Γ) mfix -> - wf_fixpoint Σ.1 mfix -> + All (on_def_type (lift_typing typing Σ) Γ) mfix -> + All (on_def_type (Pj Σ) Γ) mfix -> + All (on_def_body (lift_typing typing Σ) types Γ) mfix -> + All (on_def_body (Pj Σ) types Γ) mfix -> + wf_fixpoint Σ.1 mfix -> P Σ Γ (tFix mfix n) decl.(dtype)) -> (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) (mfix : list (def term)) (n : nat) decl, @@ -1273,9 +1264,11 @@ Lemma typing_ind_env `{cf : checker_flags} : cofix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> PΓ Σ Γ wfΓ -> - All (on_def_type (lift_typing2 typing P Σ) Γ) mfix -> - All (on_def_body (lift_typing2 typing P Σ) types Γ) mfix -> - wf_cofixpoint Σ.1 mfix -> + All (on_def_type (lift_typing typing Σ) Γ) mfix -> + All (on_def_type (Pj Σ) Γ) mfix -> + All (on_def_body (lift_typing typing Σ) types Γ) mfix -> + All (on_def_body (Pj Σ) types Γ) mfix -> + wf_cofixpoint Σ.1 mfix -> P Σ Γ (tCoFix mfix n) decl.(dtype)) -> (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) p prim_ty cdecl, @@ -1297,7 +1290,7 @@ Lemma typing_ind_env `{cf : checker_flags} : primitive_constant Σ primArray = Some prim_ty -> declared_constant Σ prim_ty cdecl -> primitive_invariants primArray cdecl -> - let s := Universe.make u in + let s := sType (Universe.make' u) in Σ ;;; Γ |- ty : tSort s -> P Σ Γ ty (tSort s) -> Σ ;;; Γ |- def : ty -> @@ -1314,10 +1307,10 @@ Lemma typing_ind_env `{cf : checker_flags} : Σ ;;; Γ |- A <= B -> P Σ Γ t B) -> - env_prop P PΓ. + env_prop P Pj PΓ. Proof. - intros P Pdecl PΓ; unfold env_prop. - intros XΓ. + intros P Pj Pdecl PΓ; unfold env_prop. + intros Xj XΓ. intros X X0 Xcast X1 X2 X3 X4 X5 X6 X7 X8 X9 X10 X11 Xint Xfloat Xarr X12 Σ wfΣ Γ wfΓ t T H. (* NOTE (Danil): while porting to 8.9, I had to split original "pose" into 2 pieces, otherwise it takes forever to execure the "pose", for some reason *) @@ -1352,61 +1345,59 @@ Proof. destruct Xg. rename udecl0 into udecl. rename on_global_decl_d0 into Xg. constructor; auto; try constructor; auto. - - unshelve eset (IH' := IH ((Σ', udecl); wfΣ; []; tSort Universe.lProp; _; _)). + - unshelve eset (IH' := IH ((Σ', udecl); wfΣ; []; tSort sProp; _; _)). shelve. simpl. apply type_Prop. forward IH'. constructor 1; cbn. lia. apply IH'; auto. - simpl. simpl in *. destruct d. - + destruct c; simpl in *. - destruct cst_body0; apply lift_typing_impl with (1 := Xg); intros ? Hs. - all: specialize (IH ((Σ', udecl); wfΣ; _; _; _; Hs)). - all: forward IH; [constructor 1; simpl; subst Σ' Σg; cbn; lia|]. - all: apply IH. + + apply Xj; tas. + apply lift_typing_impl with (1 := Xg); intros ?? Hs. split; tas. + specialize (IH ((Σ', udecl); wfΣ; _; _; _; Hs)). + forward IH; [constructor 1; simpl; subst Σ' Σg; cbn; lia|]. + apply IH. + red in Xg. - destruct Xg as [onI onP onnp]; constructor; eauto. + destruct Xg as [onI onP onNP onV]; constructor; eauto. * unshelve eset (IH' := fun p => IH ((Σ', udecl); wfΣ; p) _). constructor. cbn; subst Σ' Σg; lia. clearbody IH'. cbn in IH'. clear IH; rename IH' into IH. - eapply Alli_impl; eauto. cbn in IH. clear onI onP onnp. intros n x Xg. + eapply Alli_impl; eauto. cbn in IH. clear onI onP onNP onV. intros n x Xg. refine {| ind_arity_eq := Xg.(ind_arity_eq); ind_cunivs := Xg.(ind_cunivs) |}. -- apply onArity in Xg. - apply lift_typing_impl with (1 := Xg); intros ? Hs. + apply Xj; tas. + apply lift_typing_impl with (1 := Xg); intros ?? Hs. split; tas. apply (IH (_; _; _; Hs)). -- pose proof Xg.(onConstructors) as Xg'. - eapply All2_impl; eauto. intros. - destruct X13 as [cass tyeq onctyp oncargs oncind]. - unshelve econstructor; eauto. - { apply lift_typing_impl with (1 := onctyp); intros ? Hs. + eapply All2_impl; eauto. + intros ? ? [cass tyeq onctyp oncargs oncind]; unshelve econstructor; eauto. + { apply Xj; tas. + apply lift_typing_impl with (1 := onctyp); intros ?? Hs. split; tas. + apply (IH (_; _; _; Hs)). } + { apply sorts_local_ctx_impl with (1 := oncargs); intros Γ' j' Hj'. + apply Xj; tas. + apply lift_typing_impl with (1 := Hj'); intros ?? Hs. split; tas. apply (IH (_; _; _; Hs)). } - { apply sorts_local_ctx_impl with (1 := oncargs); intros Γ' t' T' HT'. - apply lift_typing_impl with (1 := HT'); intros ? Hs. + { eapply ctx_inst_impl with (1 := oncind); intros ?? Hj. + apply Xj; tas. + apply lift_typing_impl with (1 := Hj); intros ?? Hs. split; tas. apply (IH (_; _; _; Hs)). } - { clear -IH oncind. - revert oncind. - generalize (List.rev (lift_context #|cstr_args x0| 0 (ind_indices x))). - generalize (cstr_indices x0). induction 1; constructor; auto. - do 2 red in t0 |- *. - apply (IH (_; (_; (_; t0)))). } -- pose proof (onProjections Xg); auto. - -- destruct Xg. simpl. unfold check_ind_sorts in *. - destruct Universe.is_prop; auto. - destruct Universe.is_sprop; auto. - split. apply ind_sorts0. destruct indices_matter; auto. - eapply type_local_ctx_impl. eapply ind_sorts0. intros ??? HT. - apply lift_typing_impl with (1 := HT); intros ? Hs. + -- pose proof (ind_sorts Xg) as Xg'. unfold check_ind_sorts in *. + destruct Sort.to_family; auto. + split. apply Xg'. destruct indices_matter; auto. + eapply type_local_ctx_impl. eapply Xg'. auto. intros ?? Hj. apply Xj; tas. + apply lift_typing_impl with (1 := Hj); intros ?? Hs. split; tas. apply (IH (_; _; _; Hs)). -- apply (onIndices Xg). - * red in onP |- *. - apply All_local_env_impl with (1 := onP); intros ??? HT. - apply lift_typing_impl with (1 := HT); intros ? Hs. + * apply All_local_env_impl with (1 := onP); intros ?? Hj. apply Xj; tas. + apply lift_typing_impl with (1 := Hj); intros ?? Hs. split; tas. apply (IH ((Σ', udecl); (wfΣ; _; _; _; Hs))). constructor 1. simpl. subst Σ' Σg; cbn; lia. - assert (forall Γ t T (Hty : Σ ;;; Γ |- t : T), typing_size Hty < typing_size H -> - on_global_env cumul_gen (lift_typing P) Σ.1 * P Σ Γ t T). + on_global_env cumul_gen Pj Σ.1 * P Σ Γ t T). { intros. specialize (IH (Σ; wfΣ; _; _; _; Hty)). simpl in IH. @@ -1416,22 +1407,37 @@ Proof. rename X13 into X14. + assert (Xj' : forall Γ j (Hj : lift_typing typing Σ Γ j), + lift_typing_size (fun t T H => @typing_size cf Σ Γ t T H) _ Hj < typing_size H -> + Pj Σ Γ j). + { intros ?? Hj Hlt. + apply Xj; tas. + eapply lift_typing_size_impl with (Psize := fun t T H => @typing_size _ _ _ t T H) (tu := Hj). + intros ?? Hty Hlt'; split; tas. + apply X14 with (Hty := Hty). set (size_mid := lift_typing_size _ _ Hj) in *. lia. } + clear Xj; rename Xj' into Xj. + assert (forall Γ' t T (Hty : Σ ;;; Γ' |- t : T), typing_size Hty <= typing_size H -> PΓ Σ Γ' (typing_wf_local Hty)). - { intros. apply XΓ; auto. - clear -Pdecl wfΣ X14 H0. - pose proof (typing_wf_local_size Hty). + { intros. + pose proof (typing_wf_local_size Hty) as Hlt0. set (foo := typing_wf_local Hty) in *. - clearbody foo. assert (All_local_env_size (@typing_size cf) Σ Γ' foo < typing_size H) by lia. - clear H1 H0 Hty. - revert foo H2. - induction foo; simpl in *; try destruct t3 as [u Hu]; cbn in *; constructor. + clearbody foo. assert (All_local_env_size (@typing_size cf Σ) Γ' foo < typing_size H) as Hlt by lia. + clear -wfΣ X14 Xj Hlt XΓ. + apply XΓ; tas. + + revert foo Hlt. + induction foo; simpl in *; cbn in *; constructor. - simpl in *. apply IHfoo. lia. - - red. apply (X14 _ _ _ Hu). lia. + - red. apply (X14 _ _ _ t2.2.π2.1). cbn. lia. - simpl in *. apply IHfoo. lia. - - red. apply (X14 _ _ _ t4). lia. - - red. simpl. apply (X14 _ _ _ Hu). lia. } + - red. apply (X14 _ _ _ t2.1). cbn. lia. + - red. apply (X14 _ _ _ t2.2.π2.1). cbn. lia. + + revert foo Hlt. + induction foo; cbn in *; constructor. + 1,3: apply IHfoo; lia. + all: eapply Xj with (Hj := t2); simpl; cbn. + all: lia. } clear IH. assert (pΓ : PΓ Σ Γ (typing_wf_local H)). @@ -1445,24 +1451,17 @@ Proof. end; eauto; unshelve eapply X14; simpl; auto with arith]. - -- match reverse goal with - H : _ |- _ => eapply H - end; eauto; - unshelve eapply X14; simpl; eauto with arith wf. + all: try solve [ match reverse goal with H : _ |- _ => eapply H end; + eauto; [ unshelve eapply Xj | unshelve eapply X14 ]; simpl; auto with arith ]. - -- clear X X0 Xcast X1 X2 X3 X5 X6 X7 X8 X9 X10 X11 X12 X13. + -- clear X X0 Xcast X1 X2 X3 X5 X6 X7 X8 X9 X10 X11 X12 X13 Xj. eapply X4 with t_ty t0; eauto. clear X4. unshelve eapply X14; simpl; auto with arith. simpl in X14. - assert(X: forall Γ0 : context, - wf_local Σ Γ0 -> - forall (t1 T : term) (Hty : Σ;;; Γ0 |- t1 : T), - typing_size Hty < - S - ((typing_spine_size - (fun x (x0 : context) (x1 x2 : term) (x3 : x;;; x0 |- x1 : x2) => - typing_size x3) Σ Γ t_ty l t' t0)) -> - on_global_env cumul_gen (lift_typing P) Σ.1 * P Σ Γ0 t1 T). { + assert(X: forall Γ0, wf_local Σ Γ0 -> + forall t1 T (Hty : Σ;;; Γ0 |- t1 : T), + typing_size Hty < S ((typing_spine_size (fun _ _ _ _ H => typing_size H) Σ Γ t_ty l t' t0)) -> + on_global_env cumul_gen Pj Σ.1 * P Σ Γ0 t1 T). { intros. unshelve eapply X14; eauto. lia. } clear X14. simpl in pΓ. clear n e H pΓ. induction t0; constructor. @@ -1486,16 +1485,15 @@ Proof. eapply (X8 Σ wfΣ Γ (typing_wf_local H0) ci); eauto. ++ eapply (X14 _ _ _ H0); eauto. simpl; auto with arith. lia. ++ clear -c1 X14. - assert (forall (Γ' : context) (t T : term) (Hty : Σ;;; Γ' |- t : T), + assert (forall Γ' t T (Hty : Σ;;; Γ' |- t : T), typing_size Hty <= ctx_inst_size _ (@typing_size _) c1 -> P Σ Γ' t T). { intros. eapply (X14 _ _ _ Hty). simpl. - change (fun (x : global_env_ext) (x0 : context) (x1 x2 : term) - (x3 : x;;; x0 |- x1 : x2) => typing_size x3) with (@typing_size cf). + change (fun _ _ _ _ H => typing_size H) with (@typing_size cf). lia. } clear -X c1. revert c1 X. - generalize (List.rev (subst_instance (puinst p) (ind_params mdecl ,,, ind_indices idecl))). + generalize (List.rev (subst_instance (puinst p) (ind_params mdecl ,,, ind_indices idecl : context))). generalize (pparams p ++ indices). intros l c ctxi IH; induction ctxi; constructor; eauto. * split; tas. eapply (IH _ _ _ t0); simpl; auto. lia. @@ -1506,7 +1504,7 @@ Proof. ++ simpl in *. eapply (X13 _ _ _ H); eauto. simpl. subst predctx. lia. ++ eapply (X14 _ _ _ H0); simpl. lia. - ++ clear X13. revert a X14. clear. intros. + ++ clear X13 Xj. revert a X14. clear. intros. subst ptm predctx. Opaque case_branch_type. simpl in X14. @@ -1528,248 +1526,70 @@ Proof. unshelve eapply X14; eauto. -- clear X X0 Xcast X1 X2 X3 X4 X5 X6 X7 X8 X9 X11 X12. - eapply X10; eauto; clear X10. simpl in *. - * assert(forall (t T : term) (Hty : Σ;;; Γ |- t : T), - typing_size Hty < - S (all_size (fun x : def term => - ∑ s : Universe.t, Σ;;; Γ |- dtype x : tSort s) - (fun (x : def term) '(existT s d) => typing_size d) a0) -> - on_global_env cumul_gen (lift_typing P) Σ.1 * P Σ Γ t T). - intros; eauto. eapply (X14 _ _ _ Hty); eauto. unfold infer_sort. lia. - clear X13 X14 a pΓ. + eapply X10; eauto; clear X10. + * assert (forall j Hj, lift_typing_size (fun t T H => @typing_size cf Σ Γ t T H) j Hj < + S (all_size _ (fun x p => on_def_type_size (@typing_size _ Σ) _ _ p) a0) -> + Pj Σ Γ j). + { intros. eapply Xj with (Hj := Hj); eauto. simpl. lia. } clear -a0 X. induction a0; constructor. - destruct p as [s Hs]. exists s; split; auto. - apply (X (dtype x) (tSort s) Hs). simpl. lia. - apply IHa0. intros. eapply (X _ _ Hty); eauto. - simpl. lia. + ++ eapply X with (Hj := p). cbn. + unfold on_def_type_sorting_size, lift_sorting_size, option_default_size in *. + cbn in *. lia. + ++ apply IHa0. intros. eapply X with (Hj := Hj); eauto. + simpl. lia. * simpl in X14. - assert(forall Γ0 : context, - wf_local Σ Γ0 -> - forall (t T : term) (Hty : Σ;;; Γ0 |- t : T), - typing_size Hty < - S - (all_size (fun x : def term => (Σ;;; Γ ,,, fix_context mfix |- dbody x : lift0 #|fix_context mfix| (dtype x))%type - )%type - (fun (x : def term) p => typing_size p) a1) -> - on_global_env cumul_gen (lift_typing P) Σ.1 * P Σ Γ0 t T). - {intros. eapply (X14 _ _ _ Hty); eauto. lia. } - clear X14 X13. - clear e decl i a0 i0 pΓ. - remember (fix_context mfix) as mfixcontext. clear Heqmfixcontext. - - induction a1; econstructor; eauto. - ++ split; auto. - eapply (X _ (typing_wf_local p) _ _ p). simpl. lia. - ++ eapply IHa1. intros. - eapply (X _ X0 _ _ Hty). simpl; lia. + assert (forall Γ, wf_local Σ Γ -> + forall j Hj, lift_typing_size (fun t T H => @typing_size cf Σ Γ t T H) j Hj < + S (all_size (on_def_body (lift_typing typing _) _ _) (fun x p => on_def_body_size (fun Γ t T H => @typing_size cf Σ Γ t T H) _ _ _ p) a1) -> + Pj Σ Γ j). + { intros. eapply Xj with (Hj := Hj); eauto. simpl. lia. } + clear -a1 X. + remember (fix_context mfix) as mfixcontext eqn:e. clear e. + induction a1; constructor. + ++ eapply X with (Hj := p). 1: eapply typing_wf_local, p.1. cbn. lia. + ++ apply IHa1. intros. eapply X with (Hj := Hj); eauto. + simpl. lia. -- clear X X0 Xcast X1 X2 X3 X4 X5 X6 X7 X8 X9 X10 X12. eapply X11; eauto; clear X11. simpl in *. - * assert(forall (t T : term) (Hty : Σ;;; Γ |- t : T), - typing_size Hty < - S (all_size (fun x : def term => - ∑ s : Universe.t, Σ;;; Γ |- dtype x : tSort s) - (fun (x : def term) '(existT s d) => typing_size d) a0) -> - on_global_env cumul_gen (lift_typing P) Σ.1 * P Σ Γ t T). - intros; eauto. eapply (X14 _ _ _ Hty); eauto. unfold infer_sort. lia. - clear X13 X14 a pΓ. - clear -a0 X. - induction a0; constructor. - destruct p as [s Hs]. exists s; split; auto. - apply (X (dtype x) (tSort s) Hs). simpl. lia. - apply IHa0. intros. eapply (X _ _ Hty); eauto. - simpl. lia. - * simpl in X14. - assert(forall Γ0 : context, - wf_local Σ Γ0 -> - forall (t T : term) (Hty : Σ;;; Γ0 |- t : T), - typing_size Hty < - S - (all_size (fun x : def term => (Σ;;; Γ ,,, fix_context mfix |- dbody x : lift0 #|fix_context mfix| (dtype x))%type) - (fun (x : def term) p => typing_size p) a1) -> - on_global_env cumul_gen (lift_typing P) Σ.1 * P Σ Γ0 t T). - { intros. eapply (X14 _ _ _ Hty); eauto. lia. } - clear X14 X13. - clear e decl a0 i i0 pΓ. - remember (fix_context mfix) as mfixcontext. clear Heqmfixcontext. - - induction a1; econstructor; eauto. - ++ split; auto. - eapply (X _ (typing_wf_local p) _ _ p). simpl. lia. - ++ eapply IHa1. intros. - eapply (X _ X0 _ _ Hty). simpl; lia. + * assert (forall j Hj, lift_typing_size (fun t T H => @typing_size cf Σ Γ t T H) j Hj < + S (all_size _ (fun x p => on_def_type_size (@typing_size _ Σ) _ _ p) a0) -> + Pj Σ Γ j). + { intros. eapply Xj with (Hj := Hj); eauto. simpl. lia. } + clear -a0 X. + induction a0; constructor. + ++ eapply X with (Hj := p). cbn. + unfold on_def_type_sorting_size, lift_sorting_size, option_default_size in *. + cbn in *. lia. + ++ apply IHa0. intros. eapply X with (Hj := Hj); eauto. + simpl. lia. + * simpl in X14. + assert (forall Γ, wf_local Σ Γ -> + forall j Hj, lift_typing_size (fun t T H => @typing_size cf Σ Γ t T H) j Hj < + S (all_size (on_def_body (lift_typing typing _) _ _) (fun x p => on_def_body_size (fun Γ t T H => @typing_size cf Σ Γ t T H) _ _ _ p) a1) -> + Pj Σ Γ j). + { intros. eapply Xj with (Hj := Hj); eauto. simpl. lia. } + clear -a1 X. + remember (fix_context mfix) as mfixcontext eqn:e. clear e. + induction a1; constructor. + ++ eapply X with (Hj := p). 1: eapply typing_wf_local, p.1. cbn. lia. + ++ apply IHa1. intros. eapply X with (Hj := Hj); eauto. + simpl. lia. -- eapply Xarr; tea. * eapply (X14 _ _ _ H). simpl. subst s; lia. * eapply (X14 _ _ _ H0). simpl. subst s; lia. - * clear -a0 X14. - assert (forall (Γ0 : context) (t T : term) (Hty : Σ;;; Γ0 |- t : T), - typing_size Hty < S (all_size _ (fun t p => typing_size p) a0) -> - on_global_env cumul_gen (lift_typing P) Σ.1 × P Σ Γ0 t T). - { intros ??? Hty ?; eapply (X14 _ _ _ Hty). simpl. lia. } - clear X14. clear -X a0. induction a0; constructor; eauto. - split => //. eapply (X _ _ _ p). cbn. lia. - eapply IHa0. intros. eapply (X _ _ _ Hty). simpl. lia. -Qed. - -(** * Lemmas about All_local_env *) - -Lemma nth_error_All_local_env {P Γ n} (isdecl : n < #|Γ|) : - All_local_env P Γ -> - on_some (on_local_decl P (skipn (S n) Γ)) (nth_error Γ n). -Proof. - induction 1 in n, isdecl |- *. red; simpl. - - destruct n; simpl; inv isdecl. - - destruct n. red; simpl. red. simpl. apply t0. - simpl. apply IHX. simpl in isdecl. lia. - - destruct n. auto. - apply IHX. simpl in *. lia. + * assert (Xi : forall t (Hty : Σ;;; Γ |- t : ty), typing_size Hty < S (all_size _ (fun t p => typing_size p) a0) -> + P Σ Γ t ty). + { intros ? Hty ?; eapply (X14 _ _ _ Hty). simpl. lia. } + clear -Xi a0. + induction a0; constructor; eauto. + + split => //. eapply (Xi _ p). cbn. lia. + + eapply IHa0. intros. eapply (Xi _ Hty). simpl. lia. Qed. Arguments on_global_env {cf} Pcmp P !g. -Lemma lookup_on_global_env `{checker_flags} {Pcmp P} {Σ : global_env} {c decl} : - on_global_env Pcmp P Σ -> - lookup_env Σ c = Some decl -> - { Σ' : global_env_ext & on_global_env Pcmp P Σ' × strictly_extends_decls Σ' Σ × on_global_decl Pcmp P Σ' c decl }. -Proof. - unfold on_global_env. - destruct Σ as [univs Σ retro]; cbn. intros [cu ond]. - induction ond; cbn in * => //. - destruct o. rename udecl0 into udecl. - case: eqb_specT => [-> [= <-]| ne]. - - exists ({| universes := univs; declarations := Σ; retroknowledge := retro |}, udecl). - split; try constructor; tas. - cbn. split => //=. now exists [(kn, d)]. - - intros hl. - destruct (IHond hl) as [[Σ' udecl'] [ong [[equ ext extretro] ond']]]. - exists (Σ', udecl'). cbn in equ |- *. subst univs. split; cbn; auto; try apply ong. - split; cbn; auto. split; cbn; auto. - cbn in ext. destruct ext as [Σ'' ->]. cbn in *. - now exists ((kn, d) :: Σ''). -Qed. - -Lemma All_local_env_app_inv - (P : context -> term -> typ_or_sort -> Type) l l' : - All_local_env P (l ,,, l') -> - All_local_env P l * All_local_env (fun Γ t T => P (l ,,, Γ) t T) l'. -Proof. - induction l'; simpl; split; auto. - - constructor. - - intros. unfold app_context in X. - inv X. - + intuition auto. - + auto. apply IHl'. auto. - - inv X. - + eapply localenv_cons_abs. - * apply IHl'. apply X0. - * apply X1. - + eapply localenv_cons_def. - * apply IHl'. apply X0. - * apply X1. - * eapply X2. -Qed. - -Lemma All_local_env_lookup {P Γ n} {decl} : - All_local_env P Γ -> - nth_error Γ n = Some decl -> - on_local_decl P (skipn (S n) Γ) decl. -Proof. - induction 1 in n, decl |- *. - - simpl. destruct n; simpl; congruence. - - destruct n. - + red. simpl. intro HH; apply some_inj in HH; rewrite <- HH; tas. - + apply IHX. - - destruct n. - + red. simpl. intro HH; apply some_inj in HH; rewrite <- HH; tas. - + apply IHX. -Qed. - -Lemma All_local_env_app `{checker_flags} (P : context -> term -> typ_or_sort -> Type) l l' : - All_local_env P l * All_local_env (fun Γ t T => P (l ,,, Γ) t T) l' -> - All_local_env P (l ,,, l'). -Proof. - induction l'; simpl; auto. intuition. - intuition. destruct a. destruct decl_body. - inv b. econstructor; eauto. inv b; econstructor; eauto. Qed. - -Lemma All_local_env_map `{checker_flags} (P : context -> term -> typ_or_sort -> Type) f l : - (forall u, f (tSort u) = tSort u) -> - All_local_env (fun Γ t T => P (map (map_decl f) Γ) (f t) (typ_or_sort_map f T)) l -> All_local_env P (map (map_decl f) l). -Proof. - intros Hf. induction 1; econstructor; eauto. -Qed. - -Definition property `{checker_flags} := - forall (Σ : global_env_ext) (Γ : context), - wf_local Σ Γ -> forall t T : term, typing Σ Γ t T -> Type. - -Definition lookup_wf_local {Γ P} (wfΓ : All_local_env P Γ) (n : nat) (isdecl : n < #|Γ|) : - All_local_env P (skipn (S n) Γ). -Proof. - induction wfΓ in n, isdecl |- *; simpl. constructor. - cbn -[skipn] in *. destruct n. - simpl. exact wfΓ. - apply IHwfΓ. auto with arith. - destruct n. exact wfΓ. - apply IHwfΓ. auto with arith. -Defined. - -Definition lookup_wf_local_decl {Γ P} (wfΓ : All_local_env P Γ) (n : nat) - {decl} (eq : nth_error Γ n = Some decl) : - ∑ Pskip : All_local_env P (skipn (S n) Γ), - on_local_decl P (skipn (S n) Γ) decl. -Proof. - induction wfΓ in n, decl, eq |- *; simpl. - - elimtype False. destruct n; inversion eq. - - destruct n. - + simpl. exists wfΓ. injection eq; intros <-. apply t0. - + apply IHwfΓ. auto with arith. - - destruct n. - + exists wfΓ. injection eq; intros <-. apply t1. - + apply IHwfΓ. apply eq. -Defined. - -Definition on_wf_local_decl `{checker_flags} {Σ Γ} - (P : forall Σ Γ (wfΓ : wf_local Σ Γ) t T, Σ ;;; Γ |- t : T -> Type) - (wfΓ : wf_local Σ Γ) {d} - (H : on_local_decl (lift_typing typing Σ) Γ d) := - match d as d' return (on_local_decl (lift_typing typing Σ) Γ d') -> Type with - | {| decl_name := na; decl_body := Some b; decl_type := ty |} - => fun H => P Σ Γ wfΓ b ty H - | {| decl_name := na; decl_body := None; decl_type := ty |} - => fun H => P Σ Γ wfΓ _ _ (projT2 H) - end H. - - -Lemma nth_error_All_local_env_over `{checker_flags} {P Σ Γ n decl} (eq : nth_error Γ n = Some decl) {wfΓ : wf_local Σ Γ} : - All_local_env_over typing P Σ Γ wfΓ -> - let Γ' := skipn (S n) Γ in - let p := lookup_wf_local_decl wfΓ n eq in - (All_local_env_over typing P Σ Γ' (projT1 p) * on_wf_local_decl P (projT1 p) (projT2 p))%type. -Proof. - induction 1 in n, decl, eq |- *. simpl. - - destruct n; simpl; elimtype False; discriminate eq. - - destruct n; simpl. - + cbn [skipn]. simpl in *. split; tas. - destruct (f_equal (fun e => match e with - | Some c => c - | None => {| decl_name := na; - decl_body := None; decl_type := t |} - end) eq). - assumption. - + apply IHX. - - destruct n. - + cbn [skipn]. simpl in *. split; tas. - destruct (f_equal (fun e => match e with - | Some c => c - | None => {| decl_name := na; - decl_body := Some b; decl_type := t |} - end) eq). - assumption. - + apply IHX. -Defined. - Definition wf_ext_wf `{checker_flags} Σ : wf_ext Σ -> wf Σ := fst. #[global] diff --git a/template-coq/theories/TypingWf.v b/template-coq/theories/TypingWf.v index affb93cc9..a122f6b0d 100644 --- a/template-coq/theories/TypingWf.v +++ b/template-coq/theories/TypingWf.v @@ -27,65 +27,27 @@ Proof. - destruct a as [na [body|] ty]. + econstructor. * apply IHΓ. inv X; eauto. - * red. inv X. split. - -- apply X0. - -- constructor. - * red. inv X. eauto. + * red. inv X. + apply X0. + econstructor. * apply IHΓ. inv X; eauto. - * red. inv X. split. - -- apply X0. - -- constructor. + * red. inv X. + apply X0. Qed. +(* Still needed ? *) Lemma on_global_decl_impl `{checker_flags} Σ P Q kn d : - (forall Σ Γ t T, on_global_env cumul_gen P Σ.1 -> P Σ Γ t T -> Q Σ Γ t T) -> + (forall Σ Γ j, on_global_env cumul_gen P Σ.1 -> P Σ Γ j -> Q Σ Γ j) -> on_global_env cumul_gen P Σ.1 -> on_global_decl cumul_gen P Σ kn d -> on_global_decl cumul_gen Q Σ kn d. Proof. - unfold on_global_env. - intros X X0 o. - destruct d; simpl. - - destruct c; simpl. destruct cst_body0; simpl in *. - red in o |- *. simpl in *. now eapply X. - red in o |- *. simpl in *. now eapply X. - - simpl in *. - destruct o as [onI onP onNP]. - constructor; auto. - -- eapply Alli_impl. exact onI. eauto. intros. - - refine {| ind_arity_eq := X1.(ind_arity_eq); - ind_cunivs := X1.(ind_cunivs) |}. - --- apply onArity in X1. unfold on_type in *; simpl in *. - now eapply X. - --- pose proof X1.(onConstructors) as X11. red in X11. - eapply All2_impl; eauto. - simpl. intros. destruct X2 as [? ? ? ?]; unshelve econstructor; eauto. - * apply X; eauto. - * clear -X0 X on_cargs. revert on_cargs. - generalize (cstr_args x0), y. - induction c; destruct y0; simpl; auto; - destruct a as [na [b|] ty]; simpl in *; auto; - split; intuition eauto. - * clear -X0 X on_cindices. - revert on_cindices. - generalize (List.rev (lift_context #|cstr_args x0| 0 (ind_indices x))). - generalize (cstr_indices x0). - induction 1; simpl; constructor; auto. - --- simpl; intros. apply (onProjections X1). - --- destruct X1. simpl. unfold check_ind_sorts in *. - destruct Universe.is_prop; auto. - destruct Universe.is_sprop; auto. - split. apply ind_sorts. destruct indices_matter; auto. - eapply type_local_ctx_impl. eapply ind_sorts. auto. - --- apply (onIndices X1). - -- red in onP. red. - eapply All_local_env_impl. eauto. - intros. now apply X. + intro HPQ. + eapply on_global_decl_impl_simple. + apply HPQ. Qed. Lemma on_global_env_impl `{checker_flags} Σ P Q : - (forall Σ Γ t T, on_global_env cumul_gen P Σ.1 -> P Σ Γ t T -> Q Σ Γ t T) -> + (forall Σ Γ j, on_global_env cumul_gen P Σ.1 -> P Σ Γ j -> Q Σ Γ j) -> on_global_env cumul_gen P Σ -> on_global_env cumul_gen Q Σ. Proof. destruct Σ as [univs Σ]; cbn. @@ -168,7 +130,7 @@ Proof. intro H. revert n. induction H; constructor; eauto. Qed. Ltac wf := intuition try (eauto with wf || congruence || solve [constructor]). #[global] -Hint Unfold wf_decl vass vdef : wf. +Hint Unfold wf_decl wf_decl_pred vass vdef j_term j_typ : wf. #[global] Hint Extern 10 => progress simpl : wf. #[global] @@ -251,7 +213,7 @@ Qed. Lemma wf_decl_extends {cf} {Σ : global_env} T {Σ' : global_env} P : on_global_env cumul_gen P Σ' -> wf_decl Σ T -> extends Σ Σ' -> wf_decl Σ' T. Proof. - intros wf [] ext. red. destruct decl_body; split; eauto using wf_extends. + intros wf [] ext. red. destruct decl_body; split; cbn in *; eauto using wf_extends. Qed. Arguments lookup_on_global_env {H} {Pcmp P Σ c decl}. @@ -296,12 +258,27 @@ Proof. eapply nth_error_alli in Hidecl; eauto. pose proof (onArity Hidecl). rewrite Hidecl.(ind_arity_eq) in X0. - destruct X0 as [s Hs]; wf. - eapply wf_it_mkProd_or_LetIn in s as [? H]. + destruct X0 as [_ Hs]; cbn in Hs; wf. + eapply wf_it_mkProd_or_LetIn in Hs as [? H]. eapply wf_it_mkProd_or_LetIn in H as []. solve_all. eapply wf_decl_extends; tea; typeclasses eauto. Qed. +Lemma sorts_local_ctx_All_wf_decl {cf:checker_flags} {Σ} {mdecl} {u: list sort} {args} : + sorts_local_ctx (fun Σ : global_env_ext => wf_decl_pred Σ) Σ + (arities_context (ind_bodies mdecl),,, ind_params mdecl) + args u -> + All (wf_decl Σ) args. +Proof. + induction args as [|[na [b|] ty] args] in u |- * ; + constructor. + - constructor; now destruct X as (?&?&?). + - eapply IHargs; now apply X. + - destruct u => //; constructor; cbnr; now destruct X as (?&?&?). + - destruct u => //; eapply IHargs; now apply X. +Qed. + + Lemma declared_inductive_wf_ctors {cf:checker_flags} {Σ} {ind} {mdecl idecl} : on_global_env cumul_gen wf_decl_pred Σ -> declared_inductive Σ ind mdecl idecl -> @@ -315,15 +292,10 @@ Proof. apply onInductives in prf. eapply nth_error_alli in Hidecl; eauto. pose proof (onConstructors Hidecl). red in X0. - solve_all. destruct X0. - clear -X ext on_cargs. - induction (cstr_args x) as [|[na [b|] ty] args] in on_cargs, y |- * ; - try destruct on_cargs; - constructor; unfold wf_decl in *; cbn in *; intuition eauto using wf_extends; simpl in *. - destruct b0. intuition eauto using wf_extends with typeclass_instances. - destruct a. intuition eauto using wf_extends with typeclass_instances. - destruct y => //. destruct on_cargs. destruct w; eauto using wf_extends with typeclass_instances. - destruct y => //. eapply IHargs; intuition eauto. + solve_all. + apply on_cargs in X0. + eapply sorts_local_ctx_All_wf_decl in X0. + solve_all. eapply wf_decl_extends; tea; typeclasses eauto. Qed. Lemma All_local_env_wf_decls Σ ctx : @@ -331,7 +303,6 @@ Lemma All_local_env_wf_decls Σ ctx : All (wf_decl Σ) ctx. Proof. induction 1; constructor; auto. - destruct t0 as [s Hs]. split; simpl; intuition auto. Qed. Lemma on_global_inductive_wf_params {cf:checker_flags} {Σ : global_env_ext} {kn mdecl} : @@ -340,8 +311,7 @@ Lemma on_global_inductive_wf_params {cf:checker_flags} {Σ : global_env_ext} {kn Proof. intros prf. apply onParams in prf. red in prf. - apply All_local_env_wf_decls in prf. - solve_all. + now apply All_local_env_wf_decls in prf. Qed. Lemma destArity_spec ctx T : @@ -514,7 +484,7 @@ Section WfAst. apply wf_subst_instance_context. rewrite /ind_predicate_context. constructor. - simpl; split; auto. simpl. auto. simpl. + split; try constructor. simpl. eapply wf_mkApps. now econstructor. apply wf_reln. constructor. eapply wf_subst_context. @@ -522,14 +492,32 @@ Section WfAst. now apply wf_extended_subst. Qed. + Lemma lift_typing_wf_pred Σ' Γ j : + lift_typing (fun Σ _ t T => WfAst.wf Σ t * WfAst.wf Σ T) Σ' Γ j -> wf_decl_pred Σ' Γ j. + Proof. + intros (Xtm & s & (Xty & _) & _). + split; tas. + destruct j_term => //. + now destruct Xtm. + Qed. + + Lemma lift_typing2_wf_pred P Γ j : + lift_typing_conj P (fun _ t T => WfAst.wf Σ t * WfAst.wf Σ T) Γ j -> wf_decl_pred Σ Γ j. + Proof. + intro H. + eapply lift_typing_wf_pred. + apply lift_typing_impl with (1 := H). + move => ??[]//. + Qed. + Lemma Forall_decls_on_global_wf : Forall_decls_typing (fun (Σ : global_env_ext) (_ : context) (t T : term) => WfAst.wf Σ t * WfAst.wf Σ T) Σ -> on_global_env cumul_gen wf_decl_pred Σ. Proof using Type. - apply on_global_env_impl => Σ' Γ t []; simpl; unfold wf_decl_pred; - intros; auto. destruct X0 as [s []]; intuition auto. + apply on_global_env_impl; intros ??? _. + apply lift_typing_wf_pred. Qed. (* Hint Resolve on_global_wf_Forall_decls : wf. *) @@ -674,18 +662,12 @@ Section WfLookup. { unfold on_constructors in onConstructors. clear -onConstructors. induction onConstructors; constructor; auto. - destruct r. - clear -on_cargs. - revert on_cargs. revert y. generalize (cstr_args x). - induction c as [|[? [] ?] ?]; simpl; - destruct y; intuition auto; - constructor; - try red; simpl; try red in a0, b0; intuition eauto. - now red in b. } + apply on_cargs in r. + eapply sorts_local_ctx_All_wf_decl; tea. } split => //. - now destruct onArity. - rewrite ind_arity_eq in onArity . - destruct onArity as [ona _]. + destruct onArity as [_ ona]. eapply wf_it_mkProd_or_LetIn in ona as [_ ona]. now eapply wf_it_mkProd_or_LetIn in ona as []. - unfold on_constructors in onConstructors. @@ -698,7 +680,7 @@ Section WfLookup. induction onConstructors; constructor; auto. destruct r. rewrite cstr_eq in on_ctype. - destruct on_ctype as [wf _]. + destruct on_ctype as [_ wf]. eapply wf_it_mkProd_or_LetIn in wf as [_ wf]. eapply wf_it_mkProd_or_LetIn in wf as [_ wf]. rewrite /cstr_concl in wf. @@ -786,8 +768,8 @@ Section WfRed. 2: destruct wfΣ; now eapply NoDup_on_global_decls. eapply lookup_on_global_env in H as [Σ' [onΣ' [ext prf]]]; eauto. destruct decl; simpl in *. - subst cst_body0; simpl in *; unfold on_constant_decl in prf; cbn in prf. - unfold wf_decl_pred in prf. intuition eauto using wf_extends with typeclass_instances. + subst cst_body0; simpl in *; unfold on_constant_decl in prf; cbn in prf; destruct prf as [prfbod prftyp]. + unfold j_term in prfbod. intuition eauto using wf_extends with typeclass_instances. - apply wf_mkApps_inv in X. eapply nth_error_all in X; eauto. - simpl in *. econstructor; eauto. cbn. @@ -885,8 +867,7 @@ Section WfRed. Lemma declared_constant_wf cst decl : on_global_env cumul_gen wf_decl_pred Σ -> declared_constant Σ cst decl -> - WfAst.wf Σ decl.(cst_type) * - on_some_or_none (WfAst.wf Σ) decl.(cst_body). + option_default (WfAst.wf Σ) (cst_body decl) unit * WfAst.wf Σ (cst_type decl). Proof using Type. intros wΣ h. unfold declared_constant in h. @@ -894,15 +875,15 @@ Section WfRed. 2: destruct wΣ; now eapply NoDup_on_global_decls. destruct (lookup_on_global_env wΣ h) as [Σ' [wΣ' [ext h']]]. simpl in h'. - destruct decl as [ty [bo|]]. all: cbn in *. - - destruct h'. intuition eauto using wf_extends with typeclass_instances. - - destruct h'. intuition eauto using wf_extends with typeclass_instances. + destruct h'. + split. 1: destruct cst_body => //=; cbn in o. + all: intuition eauto using wf_extends with typeclass_instances. Qed. Lemma wf_it_mkProd_or_LetIn_inv (Σ' : global_env_ext) Γ (wfΓ : wf_local Σ' Γ) - : All_local_env_over typing - (fun (Σ : global_env_ext) (Γ : context) (_ : wf_local Σ Γ) - (t T : term) (_ : Σ;;; Γ |- t : T) => WfAst.wf Σ t * WfAst.wf Σ T) Σ' + : All_local_env_over (typing Σ') + (fun (Γ : context) (_ : wf_local Σ' Γ) + (t T : term) (_ : Σ';;; Γ |- t : T) => WfAst.wf Σ' t * WfAst.wf Σ' T) Γ wfΓ -> forall t, WfAst.wf Σ' t -> WfAst.wf Σ' (it_mkProd_or_LetIn Γ t). Proof using Type. @@ -966,14 +947,18 @@ Section TypingWf. Lemma typing_wf_gen : env_prop (fun Σ Γ t T => WfAst.wf Σ t * WfAst.wf Σ T) + (fun Σ Γ j => wf_decl_pred Σ Γ j) (fun Σ Γ wfΓ => All (wf_decl Σ) Γ). Proof using Type. apply typing_ind_env; intros; auto with wf; - specialize_goal; + specialize_goal; unfold wf_decl_pred in *; try solve [split; try constructor; intuition auto with wf]. - - eapply All_local_env_wf_decls. - induction X; constructor; auto; red; intuition auto. + - now eapply lift_typing2_wf_pred. + - apply All_local_env_over_2 in X. + eapply All_local_env_wf_decls. + eapply All_local_env_impl with (1 := X) => Γ' j. + apply lift_typing2_wf_pred. - split; wf. apply wf_lift. apply (nth_error_all H X). - split. constructor; auto. wf. @@ -983,23 +968,18 @@ Section TypingWf. clear X H H0. induction X1; auto. apply IHX1. apply wf_subst. now destruct p0. destruct p. now inv w. - - split. wf. apply wf_subst_instance. wf. - eapply lookup_global_Some_iff_In_NoDup in H; eauto. - 2: destruct wfΣ; now eapply NoDup_on_global_decls. - destruct (lookup_on_global_env X H) as [Σ' [wfΣ' [ext prf]]]; eauto. - red in prf. destruct decl; destruct cst_body0; red in prf; simpl in *; wf. - destruct prf as [s []]. wf. + - split. wf. apply wf_subst_instance. + apply declared_constant_wf in H; tas. + apply H. - split. wf. apply wf_subst_instance. eapply declared_inductive_wf; eauto. - now eapply Forall_decls_on_global_wf. - split. wf. unfold type_of_constructor. apply wf_subst; auto with wf. apply wf_inds. apply wf_subst_instance. eapply declared_constructor_wf; eauto. - now eapply Forall_decls_on_global_wf. - destruct X3 as [wfret wps]. destruct X6 as [wfc wfapps]. @@ -1015,23 +995,19 @@ Section TypingWf. apply wf_mkApps_inv in b. apply All_rev. solve_all. eapply declared_projection_wf in isdecl; eauto. now eapply wf_subst_instance. - now eapply Forall_decls_on_global_wf. - subst types. clear H. split. + constructor. - solve_all; destruct a, b. - all: intuition. - + eapply All_nth_error in X0; eauto. - destruct X0 as [s ?]; intuition. + solve_all; destruct a0, b; cbn in *; assumption. + + eapply All_nth_error in X1 as []; eauto. - subst types. split. + constructor. - solve_all; destruct a, b. - all: intuition. - + eapply All_nth_error in X0; eauto. destruct X0 as [s ?]; intuition. + solve_all; destruct a0, b; cbn in *; assumption. + + eapply All_nth_error in X1 as []; eauto. - split => //. + constructor; intuition auto. solve_all. @@ -1048,11 +1024,7 @@ Section TypingWf. Lemma typing_wf_sigma Σ (wfΣ : wf Σ) : on_global_env cumul_gen wf_decl_pred Σ. Proof using Type. - intros. - pose proof (env_prop_sigma typing_wf_gen _ wfΣ). red in X. - do 2 red in wfΣ. - eapply on_global_env_impl; eauto; simpl; intros. - destruct T. red. apply X1. red. destruct X1 as [x [a wfs]]. split; auto. + eapply (env_prop_sigma typing_wf_gen _ wfΣ). Qed. Lemma typing_wf Σ (wfΣ : wf Σ.1) Γ t T : @@ -1068,7 +1040,6 @@ Section TypingWf. Proof using Type. intros declm. pose proof (typing_wf_gen (Env.empty_ext Σ) wfΣ _ localenv_nil _ _ (type_Prop _)) as [X _]. - eapply Forall_decls_on_global_wf in X. eapply lookup_global_Some_iff_In_NoDup in declm; eauto. 2: destruct X; now eapply NoDup_on_global_decls. destruct (lookup_on_global_env X declm) as [? [? [ext ?]]]; eauto. diff --git a/template-coq/theories/WfAst.v b/template-coq/theories/WfAst.v index a258e0466..76433aed0 100644 --- a/template-coq/theories/WfAst.v +++ b/template-coq/theories/WfAst.v @@ -186,17 +186,10 @@ Proof. - inv X2; auto. eapply Harr; eauto using lift_to_wf_list. Qed. -Definition wf_decl Σ d := - match decl_body d with - | Some b => wf Σ b - | None => True - end * wf Σ (decl_type d). - -Definition wf_decl_pred Σ : context -> term -> typ_or_sort -> Type := - (fun _ t T => wf Σ t * match T with - | Typ T => wf Σ T - | Sort => True - end). +Definition wf_decl Σ d := option_default (wf Σ) (decl_body d) unit × wf Σ (decl_type d). + +Definition wf_decl_pred Σ : context -> judgment -> Type := + (fun _ j => option_default (wf Σ) (j_term j) unit × wf Σ (j_typ j)). Lemma wf_mkApp Σ u a : wf Σ u -> wf Σ a -> wf Σ (mkApp u a). Proof. diff --git a/template-pcuic/theories/PCUICToTemplateCorrectness.v b/template-pcuic/theories/PCUICToTemplateCorrectness.v index c651532b0..9ef998cae 100644 --- a/template-pcuic/theories/PCUICToTemplateCorrectness.v +++ b/template-pcuic/theories/PCUICToTemplateCorrectness.v @@ -74,7 +74,7 @@ Proof. reflexivity. Qed. -Lemma trans_local_app Γ Δ : trans_local (SE.app_context Γ Δ) = trans_local Γ ,,, trans_local Δ. +Lemma trans_local_app Γ Δ : trans_local (Γ ,,, Δ) = trans_local Γ ,,, trans_local Δ. Proof. now rewrite /trans_local map_app. Qed. @@ -633,7 +633,7 @@ Qed. Lemma alpha_eq_trans {Γ Δ} : eq_context_upto_names Γ Δ -> - All2 (TermEquality.compare_decls eq eq) (trans_local Γ) (trans_local Δ). + TermEquality.eq_context_upto_names (trans_local Γ) (trans_local Δ). Proof. intros. eapply All2_map, All2_impl; tea. @@ -644,7 +644,7 @@ Definition wt {cf} Σ Γ t := ∑ T, Σ ;;; Γ |- t : T. Lemma isType_wt {cf} Σ Γ T : isType Σ Γ T -> wt Σ Γ T. Proof. - intros [s hs]; now exists (PCUICAst.tSort s). + intros (_ & s & hs & _); now exists (PCUICAst.tSort s). Qed. Coercion isType_wt : isType >-> wt. @@ -665,7 +665,7 @@ Section wtsub. ci.(ci_npar) = mdecl.(ind_npars), consistent_instance_ext Σ (PCUICEnvironment.ind_universes mdecl) (PCUICAst.puinst p), wf_predicate mdecl idecl p, - All2 (PCUICEquality.compare_decls eq eq) (PCUICCases.ind_predicate_context ci mdecl idecl) (PCUICAst.pcontext p), + eq_context_upto_names (PCUICCases.ind_predicate_context ci mdecl idecl) (PCUICAst.pcontext p), wf_local_rel Σ (Γ ,,, smash_context [] (ind_params mdecl)@[p.(puinst)]) p.(pcontext)@[p.(puinst)], PCUICSpine.spine_subst Σ Γ (PCUICAst.pparams p) (List.rev (PCUICAst.pparams p)) (PCUICEnvironment.smash_context [] (PCUICEnvironment.ind_params mdecl)@[PCUICAst.puinst p]), @@ -673,9 +673,9 @@ Section wtsub. wt Γ c & All2i (fun i cdecl br => [× wf_branch cdecl br, - All2 (PCUICEquality.compare_decls eq eq) (bcontext br) (PCUICCases.cstr_branch_context ci mdecl cdecl), + eq_context_upto_names (bcontext br) (PCUICCases.cstr_branch_context ci mdecl cdecl), wf_local_rel Σ (Γ ,,, smash_context [] (ind_params mdecl)@[p.(puinst)]) br.(bcontext)@[p.(puinst)], - All2 (PCUICEquality.compare_decls eq eq) + eq_context_upto_names (Γ ,,, PCUICCases.case_branch_context ci mdecl p (forget_types br.(bcontext)) cdecl) (Γ ,,, inst_case_branch_context p br) & wt (Γ ,,, PCUICCases.case_branch_context ci mdecl p (forget_types br.(bcontext)) cdecl) br.(bbody)]) 0 idecl.(ind_ctors) brs] @@ -696,16 +696,26 @@ Section wtsub. eexists; tea. eexists; tea. Qed. + Lemma lift_wt_inv {Γ j} : lift_typing typing Σ Γ j -> lift_wf_term (wt Σ Γ) j. + Proof. + intros (? & ? & ? & _). + split. 1: destruct j_term => //. + all: eexists; eassumption. + Qed. + Lemma wt_inv {Γ t} : wt Σ Γ t -> wt_subterm Γ t. Proof. destruct t; simpl; intros [T h]; try exact tt. - now eapply inversion_Evar in h. - - eapply inversion_Prod in h as (?&?&?&?&?); tea. - split; eexists; eauto. - - eapply inversion_Lambda in h as (?&?&?&?&?); tea. - split; eexists; eauto. - - eapply inversion_LetIn in h as (?&?&?&?&?&?); tea. - repeat split; eexists; eauto. + - eapply inversion_Prod in h as (?&?&h1&?&?); tea. + apply lift_wt_inv in h1 as []. + split; tas; eexists; eauto. + - eapply inversion_Lambda in h as (?&h1&?&?); tea. + apply lift_wt_inv in h1 as []. + split; tas; eexists; eauto. + - eapply inversion_LetIn in h as (?&h1&?&?); tea. + apply lift_wt_inv in h1 as []. + repeat split; tas; eexists; eauto. - eapply inversion_App in h as (?&?&?&?&?&?); tea. split; eexists; eauto. - eapply inversion_Case in h as (mdecl&idecl&decli&indices&[]&?); tea. @@ -715,7 +725,7 @@ Section wtsub. split. eapply spine_subst_wt in s. now eapply All_rev_inv in s. exists mdecl, idecl. split; auto. now symmetry. - * eapply wf_local_app_inv. eapply wf_local_alpha. + * eapply All_local_env_app_inv. eapply wf_local_alpha. eapply All2_app; [|reflexivity]. eapply alpha_eq_subst_instance. symmetry; tea. eapply wf_ind_predicate; tea. pcuic. @@ -725,7 +735,7 @@ Section wtsub. eapply All2i_All2_mix_left in brs_ty; tea. eapply All2i_nth_hyp in brs_ty. solve_all. split; auto. - { eapply wf_local_app_inv. eapply wf_local_alpha. + { eapply All_local_env_app_inv. eapply wf_local_alpha. eapply All2_app; [|reflexivity]. eapply alpha_eq_subst_instance in a2. symmetry; tea. @@ -744,12 +754,12 @@ Section wtsub. eexists; eauto. - eapply inversion_Fix in h as (?&?&?&?&?&?&?); tea. eapply All_prod. - eapply (All_impl a). intros ? h; exact h. - eapply (All_impl a0). intros ? h; eexists; tea. + eapply (All_impl a). intros ? (_ & ? & ? & _); now eexists. + eapply (All_impl a0). intros ? (h & _); cbn in h; eexists; tea. - eapply inversion_CoFix in h as (?&?&?&?&?&?&?); tea. eapply All_prod. - eapply (All_impl a). intros ? h; exact h. - eapply (All_impl a0). intros ? h; eexists; tea. + eapply (All_impl a). intros ? (_ & ? & ? & _); now eexists. + eapply (All_impl a0). intros ? (h & _); cbn in h; eexists; tea. - eapply inversion_Prim in h as (?&?&[]); eauto. destruct prim as [? []]; cbn in *; eauto; try exact tt. depelim p0. repeat split. now eexists. now eexists. solve_all. now eexists. @@ -824,8 +834,8 @@ Qed. Lemma map2_set_binder_name_alpha (nas : list aname) (Δ Δ' : context) : All2 (fun x y => eq_binder_annot x (decl_name y)) nas Δ -> - All2 (TermEquality.compare_decls eq eq) Δ Δ' -> - All2 (TermEquality.compare_decls eq eq) (map2 set_binder_name nas Δ) Δ'. + TermEquality.eq_context_upto_names Δ Δ' -> + TermEquality.eq_context_upto_names (map2 set_binder_name nas Δ) Δ'. Proof. intros hl. induction 1 in nas, hl |- *; cbn; auto. destruct nas; cbn; auto. @@ -856,7 +866,7 @@ Qed. Lemma map2_set_binder_name_alpha_eq (nas : list aname) (Δ Δ' : context) : All2 (fun x y => x = (decl_name y)) nas Δ' -> - All2 (TermEquality.compare_decls eq eq) Δ Δ' -> + TermEquality.eq_context_upto_names Δ Δ' -> (map2 set_binder_name nas Δ) = Δ'. Proof. intros hl. induction 1 in nas, hl |- *; cbn; auto. @@ -925,26 +935,26 @@ Proof. Qed. #[global] -Instance compare_decls_refl : Reflexive (TermEquality.compare_decls eq eq). +Instance compare_decls_refl : Reflexive (TermEquality.eq_decl_upto_names). Proof. intros [na [b|] ty]; constructor; auto. Qed. #[global] -Instance All2_compare_decls_refl : Reflexive (All2 (TermEquality.compare_decls eq eq)). +Instance All2_compare_decls_refl : Reflexive (TermEquality.eq_context_upto_names). Proof. intros x; apply All2_refl; reflexivity. Qed. #[global] -Instance All2_compare_decls_sym : Symmetric (All2 (TermEquality.compare_decls eq eq)). +Instance All2_compare_decls_sym : Symmetric (TermEquality.eq_context_upto_names). Proof. intros x y. eapply All2_symP. clear. intros d d' []; subst; constructor; auto; now symmetry. Qed. #[global] -Instance All2_compare_decls_trans : Transitive (All2 (TermEquality.compare_decls eq eq)). +Instance All2_compare_decls_trans : Transitive (TermEquality.eq_context_upto_names). Proof. intros x y z. eapply All2_trans. clear. intros d d' d'' [] H; depelim H; subst; constructor; auto; now etransitivity. @@ -1039,7 +1049,7 @@ Qed. Lemma red1_alpha_eq Σ Γ Δ T U : TT.red1 Σ Γ T U -> - All2 (TermEquality.compare_decls eq eq) Γ Δ -> + TermEquality.eq_context_upto_names Γ Δ -> TT.red1 Σ Δ T U. Proof. intros r; revert Δ. @@ -1094,7 +1104,7 @@ Proof. Qed. Lemma map2_set_binder_name_eq nas Δ Δ' : - All2 (TermEquality.compare_decls eq eq) Δ Δ' -> + TermEquality.eq_context_upto_names Δ Δ' -> map2 set_binder_name nas Δ = map2 set_binder_name nas Δ'. Proof. induction 1 in nas |- *; cbn; auto. @@ -1135,7 +1145,7 @@ Proof. Qed. Lemma trans_local_set_binder_name_eq nas Γ Δ : - All2 (PCUICEquality.compare_decls eq eq) Γ Δ -> + eq_context_upto_names Γ Δ -> trans_local (map2 PCUICEnvironment.set_binder_name nas Γ) = trans_local (map2 PCUICEnvironment.set_binder_name nas Δ). Proof. @@ -1317,19 +1327,17 @@ Proof. eapply PCUICSafeLemmata.wf_inst_case_context; tea. destruct w2. pcuic. rewrite /inst_case_context. - apply compare_decls_conv. + apply PCUICContextConversion.eq_context_upto_names_conv_context. eapply All2_app. 2:{ reflexivity. } - eapply compare_decls_eq_context. apply (PCUICAlpha.inst_case_predicate_context_eq (ind:=ci) w). - cbn. apply compare_decls_eq_context. now symmetry. } + cbn. now symmetry. } rewrite [trans_local _]map_app. eapply All2_app; [|reflexivity]. symmetry. etransitivity. rewrite (trans_case_predicate_context _ _ _ _ p d c0 s w). reflexivity. eapply alpha_eq_trans. - eapply All2_fold_All2, PCUICAlpha.inst_case_predicate_context_eq => //. - symmetry. now eapply All2_fold_All2. + eapply PCUICAlpha.inst_case_predicate_context_eq => //. - eapply wt_inv in wt as [hpars [mdecl [idecl []]]]. econstructor; eauto. @@ -1350,7 +1358,7 @@ Proof. destruct w4 as [s' Hs]. exists s'. eapply PCUICContextConversionTyp.context_conversion; tea. eapply wf_local_alpha; tea. pcuic. - now eapply compare_decls_conv. + now eapply PCUICContextConversion.eq_context_upto_names_conv_context. - eapply red1_mkApp; auto. eapply trans_wf; tea. - eapply (red1_mkApps_r _ _ _ [_] [_]); auto. @@ -1391,12 +1399,12 @@ Proof. eapply OnOne2_map. solve_all. Qed. -Lemma trans_R_global_instance Σ Re Rle gref napp u u' : - PCUICEquality.R_global_instance Σ Re Rle gref napp u u' -> - TermEquality.R_global_instance (trans_global_env Σ) Re Rle gref napp u u'. +Lemma trans_cmp_global_instance Σ cmp_universe pb gref napp u u' : + PCUICEquality.cmp_global_instance Σ cmp_universe pb gref napp u u' -> + TermEquality.cmp_global_instance (trans_global_env Σ) cmp_universe pb gref napp u u'. Proof. - unfold PCUICEquality.R_global_instance, PCUICEquality.R_global_instance_gen, PCUICEquality.global_variance. - unfold TermEquality.R_global_instance, TermEquality.global_variance. + unfold PCUICEquality.cmp_global_instance, PCUICEquality.cmp_global_instance_gen, PCUICEquality.global_variance. + unfold TermEquality.cmp_global_instance, TermEquality.global_variance. destruct gref; simpl; auto. - unfold S.lookup_inductive, S.lookup_minductive. unfold S.lookup_inductive_gen, S.lookup_minductive_gen. @@ -1421,21 +1429,19 @@ Proof. Qed. Lemma trans_eq_context_gen_eq_binder_annot Γ Δ : - eq_context_gen eq eq Γ Δ -> + eq_context_upto_names Γ Δ -> All2 eq_binder_annot (map decl_name (trans_local Γ)) (map decl_name (trans_local Δ)). Proof. - move/All2_fold_All2. intros; do 2 eapply All2_map. solve_all. destruct X; cbn; auto. Qed. -Lemma trans_eq_term_upto_univ {cf} : - forall Σ Re Rle t u napp, - PCUICEquality.eq_term_upto_univ_napp Σ Re Rle napp t u -> - TermEquality.eq_term_upto_univ_napp (trans_global_env Σ) Re Rle napp (trans t) (trans u). +Lemma trans_eq_term_upto_univ {cf} Σ cmp_universe cmp_sort pb napp t u : + PCUICEquality.eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp t u -> + TermEquality.eq_term_upto_univ_napp (trans_global_env Σ) cmp_universe cmp_sort pb napp (trans t) (trans u). Proof. - intros Σ Re Rle t u napp e. - induction t using term_forall_list_ind in Rle, napp, u, e |- *. + intros e. + induction t using term_forall_list_ind in pb, napp, u, e |- *. all: invs e; cbn. all: try solve [ constructor ; auto ]. all: try solve [ @@ -1445,11 +1451,12 @@ Proof. eapply ih ; auto end ]. - 1,6,7: try solve [ constructor; solve_all ]. - all: try solve [ constructor; now eapply trans_R_global_instance ]. - - eapply (TermEquality.eq_term_upto_univ_mkApps _ _ _ _ _ [_] _ [_]); simpl; eauto. + 1,6,7: try solve [ constructor; unfold eq_mfixpoint in *; solve_all ]. + all: try solve [ constructor; now eapply trans_cmp_global_instance ]. + - eapply (TermEquality.eq_term_upto_univ_mkApps _ _ _ _ _ _ [_] _ [_]); simpl; eauto. - destruct X1 as [Hpars [Huinst [Hctx Hret]]]. destruct X as [IHpars [IHctx IHret]]. + unfold eq_branches, eq_branch in *. constructor; cbn; auto. solve_all. red in X0. eapply trans_eq_context_gen_eq_binder_annot in Hctx. @@ -1458,6 +1465,7 @@ Proof. eapply trans_eq_context_gen_eq_binder_annot in a. now rewrite !map_context_trans. - depelim X0; cbn in X |- *; try econstructor; intuition eauto; solve_all. + repeat (constructor; tas). Qed. Lemma trans_leq_term {cf} Σ ϕ T U : @@ -1511,10 +1519,10 @@ Proof. eapply trans_leq_term in c. now rewrite -trans_global_ext_constraints. - destruct w as [r ht hv]. - apply trans_red1 in r; eauto. 2:destruct ht as [s hs]; now eexists. + apply trans_red1 in r; eauto. 2:destruct ht as (_ & s & hs & _); now eexists. econstructor 2; tea. - destruct w as [r ht hv]. - apply trans_red1 in r; eauto. 2:destruct ht as [s hs]; now eexists. + apply trans_red1 in r; eauto. 2:destruct ht as (_ & s & hs & _); now eexists. econstructor 3; tea. Qed. @@ -1523,25 +1531,25 @@ Definition TTwf_local {cf} Σ Γ := TT.All_local_env (TT.lift_typing TT.typing Lemma trans_wf_local' {cf} : forall (Σ : SE.global_env_ext) Γ (wfΓ : wf_local Σ Γ), let P := - (fun Σ0 Γ0 _ (t T : PCUICAst.term) _ => - TT.typing (trans_global Σ0) (trans_local Γ0) (trans t) (trans T)) + (fun Γ0 _ (t T : PCUICAst.term) _ => + TT.typing (trans_global Σ) (trans_local Γ0) (trans t) (trans T)) in - ST.All_local_env_over ST.typing P Σ Γ wfΓ -> + ST.All_local_env_over (ST.typing Σ) P Γ wfΓ -> TTwf_local (trans_global Σ) (trans_local Γ). Proof. intros. induction X. - simpl. constructor. - - simpl. econstructor. - + eapply IHX. - + simpl. destruct tu. exists x. eapply Hs. - - simpl. constructor; auto. red. destruct tu. exists x. auto. + - simpl. econstructor; auto. + simpl. destruct tu as (? & ? & ? & ?); cbn in *. repeat (eexists; tea). + - simpl. constructor; auto. + simpl. destruct tu as (? & ? & ? & ?); cbn in *. repeat (eexists; tea); cbn. Qed. Lemma trans_wf_local_env {cf} Σ Γ : ST.All_local_env (ST.lift_typing - (fun (Σ : SE.global_env_ext) Γ b ty => + (fun Σ Γ b ty => ST.typing Σ Γ b ty × TT.typing (trans_global Σ) (trans_local Γ) (trans b) (trans ty)) Σ) Γ -> TTwf_local (trans_global Σ) (trans_local Γ). @@ -1549,11 +1557,10 @@ Proof. intros. induction X. - simpl. constructor. - - simpl. econstructor. - + eapply IHX. - + simpl. destruct t0. exists x. eapply p. - - simpl. constructor; auto. red. destruct t0. exists x. intuition auto. - red. red in t1. destruct t1. eapply t2. + - simpl. econstructor; auto. + simpl. destruct t0 as (_ & ? & (? & ?) & ?); cbn in *. repeat (eexists; tea). + - simpl. constructor; auto. + simpl. destruct t0 as ((? & ?) & ? & (? & ?) & ?); cbn in *. repeat (eexists; cbn; tea). Qed. Lemma trans_branches {cf} Σ Γ brs btys ps: @@ -1587,7 +1594,7 @@ Lemma trans_mfix_All {cf} Σ Γ mfix: (Γ : SE.context) (b ty : PCUICAst.term) => ST.typing Σ Γ b ty × TT.typing (trans_global Σ) (trans_local Γ) (trans b) (trans ty)) Σ) - (SE.app_context Γ (SE.fix_context mfix)) -> + (Γ ,,, (SE.fix_context mfix)) -> TTwf_local (trans_global Σ) (trans_local Γ ,,, TT.fix_context (map (map_def trans trans) mfix)). Proof. @@ -1595,63 +1602,46 @@ Proof. rewrite <- trans_fix_context. match goal with |- TTwf_local _ ?A => - replace A with (trans_local (SE.app_context Γ (SE.fix_context mfix))) + replace A with (trans_local (Γ ,,, (SE.fix_context mfix))) end. 2: { - unfold trans_local, SE.app_context. + unfold trans_local, app_context. now rewrite map_app. } induction X;cbn;constructor;auto;cbn in *. - - destruct t0 as (?&?&?). - exists x. - apply t1. - - destruct t0 as (?&?&?). - eexists;eassumption. - - destruct t1. - assumption. + - destruct t0 as (_&?&(?&?)&?). + repeat (eexists; tea); eauto. + - destruct t0 as ((?&?)&?&(?&?)&?). + repeat (eexists; tea); eauto. Qed. Lemma trans_mfix_All2 {cf} Σ Γ mfix xfix: - All - (fun d : def PCUICAst.term => - (ST.typing Σ (SE.app_context Γ (SE.fix_context xfix)) - (dbody d) - (S.lift0 #|SE.fix_context xfix| (dtype d))) - × TT.typing (trans_global Σ) - (trans_local (SE.app_context Γ (SE.fix_context xfix))) - (trans (dbody d)) - (trans - (S.lift0 #|SE.fix_context xfix| - (dtype d)))) mfix -> - All - (fun d : def term => - TT.typing (trans_global Σ) - (trans_local Γ ,,, TT.fix_context (map (map_def trans trans) xfix)) - (dbody d) (T.lift0 #|TT.fix_context (map (map_def trans trans) xfix)| (dtype d))) - (map (map_def trans trans) mfix). + All (on_def_body (lift_typing1 + (fun Γ t T => TT.typing (trans_global Σ) (trans_local Γ) (trans t) (trans T))) + (SE.fix_context xfix) Γ) mfix -> + All (TT.on_def_body (fun Γ => TT.lift_typing0 (TT.typing (trans_global Σ) Γ)) + (TT.fix_context (map (map_def trans trans) xfix)) (trans_local Γ)) (map (map_def trans trans) mfix). Proof. induction 1. - constructor. - - simpl; constructor. - + destruct p as []. - unfold app_context, SE.app_context in *. - unfold trans_local in t0. - rewrite map_app trans_fix_context in t0. - rewrite trans_dbody trans_lift trans_dtype in t0. - replace(#|SE.fix_context xfix|) with - (#|TT.fix_context (map(map_def trans trans) xfix)|) in t0. - 2:now rewrite TT.fix_context_length map_length fix_context_length. - now destruct x. - + auto. + - simpl; constructor; auto. + destruct p as (Hb & s & Ht & _). cbn in Hb, Ht. + unfold app_context in *. + rewrite /trans_local map_app trans_fix_context in Hb, Ht. + rewrite trans_lift in Hb, Ht. + replace(#|SE.fix_context xfix|) with + (#|TT.fix_context (map(map_def trans trans) xfix)|) in Hb, Ht. + 2:now rewrite TT.fix_context_length map_length fix_context_length. + repeat (eexists; cbn; tea); eauto. Qed. Lemma All_over_All {cf} Σ Γ wfΓ : - ST.All_local_env_over ST.typing - (fun (Σ : SE.global_env_ext) (Γ : SE.context) + ST.All_local_env_over (ST.typing Σ) + (fun (Γ : SE.context) (_ : wf_local Σ Γ) (t T : PCUICAst.term) (_ : ST.typing Σ Γ t T) => - TT.typing (trans_global Σ) (trans_local Γ) (trans t) (trans T)) Σ Γ wfΓ -> + TT.typing (trans_global Σ) (trans_local Γ) (trans t) (trans T)) Γ wfΓ -> ST.All_local_env (ST.lift_typing (fun (Σ0 : SE.global_env_ext) (Γ0 : SE.context) @@ -1661,19 +1651,8 @@ Lemma All_over_All {cf} Σ Γ wfΓ : Proof. intros H. induction H;cbn. - - constructor. - - constructor. - + apply IHAll_local_env_over_gen. - + cbn in *. - destruct tu. - eexists;split;eassumption. - - constructor. - + apply IHAll_local_env_over_gen. - + cbn in *. - destruct tu. - eexists;split;eassumption. - + cbn in *. - split;eassumption. + all: constructor; eauto. + all: apply lift_sorting_it_impl with tu => //. Qed. Lemma trans_decompose_prod_assum : @@ -1894,9 +1873,9 @@ Section Typing_Spine_size. Fixpoint typing_spine_size {t T U} (s : typing_spine Σ Γ t T U) : size := match s with | type_spine_eq ty => 0 - | type_spine_nil ty ty' ist cum => typing_size ist.π2 + | type_spine_nil ty ty' ist cum => typing_size ist.2.π2.1 | type_spine_cons hd tl na A B T B' typrod cumul ty s' => - (max (typing_size typrod.π2) (max (typing_size ty) (typing_spine_size s')))%nat + (max (typing_size typrod.2.π2.1) (max (typing_size ty) (typing_spine_size s')))%nat end. End Typing_Spine_size. @@ -1925,11 +1904,11 @@ Lemma typing_spine_weaken_concl_size {cf:checker_flags} {Σ Γ T args S S'} (Hs : Σ ;;; Γ ⊢ S ≤ S') (ist : isType Σ Γ S') : typing_spine_size (typing_spine_weaken_concl sp tyT Hs ist) <= - max (typing_spine_size sp) (typing_size ist.π2). + max (typing_spine_size sp) (typing_size ist.2.π2.1). Proof. induction sp; simpl; auto. lia. rewrite - !Nat.max_assoc. - specialize (IHsp (PCUICArities.isType_apply i t) Hs). lia. + specialize (IHsp (PCUICArities.isType_apply i t) Hs). cbn in *. lia. Qed. Ltac sig := unshelve eexists. @@ -1944,7 +1923,7 @@ Lemma typing_spine_app {cf:checker_flags} Σ Γ ty args na A B arg (sp : typing_spine Σ Γ ty args (tProd na A B)) (argd : Σ ;;; Γ |- arg : A) : ∑ sp' : typing_spine Σ Γ ty (args ++ [arg]) (B {0 := arg}), - typing_spine_size sp' <= max (typing_size isty.π2) (max (typing_spine_size sp) (typing_size argd)). + typing_spine_size sp' <= max (typing_size isty.2.π2.1) (max (typing_spine_size sp) (typing_size argd)). Proof. revert arg argd. depind sp. @@ -1961,7 +1940,7 @@ Proof. specialize (IHsp wf isty _ Harg) as [sp' Hsp']. sig. * econstructor. apply i. auto. auto. exact sp'. - * simpl. lia. + * simpl. cbn in *. lia. Qed. (** Likewise, in.Common-Coq, we can append without re-typing the result *) @@ -2012,7 +1991,7 @@ Proof. exists fty, d'. split. lia. destruct hd' as [sp' Hsp]. - destruct (typing_spine_app _ _ _ _ _ _ _ _ wfΣ (s; d1) sp' d3) as [appsp Happ]. + destruct (typing_spine_app _ _ _ _ _ _ _ _ wfΣ (has_sort_isType s d1) sp' d3) as [appsp Happ]. simpl in Happ. move: appsp Happ. rewrite equ firstn_skipn. intros app happ. exists app. lia. @@ -2021,7 +2000,7 @@ Proof. rewrite PCUICAstUtils.atom_decompose_app. destruct t => /= //. eexists _, d2. split. lia. unshelve eexists. - econstructor. exists s; eauto. reflexivity. assumption. constructor. + econstructor. eapply has_sort_isType; eauto. reflexivity. assumption. constructor. simpl. lia. - destruct (isApp t) eqn:isapp => //. @@ -2033,10 +2012,10 @@ Proof. destruct s0 as [sp Hsp]. unshelve eexists. eapply typing_spine_weaken_concl; eauto. exact (PCUICValidity.validity d'). eapply cumulSpec_cumulAlgo_curry in c; eauto; fvs. - exact (s; d2). cbn. + eapply (has_sort_isType s d2). epose proof (typing_spine_weaken_concl_size wfΣ sp (validity_term wfΣ d') (cumulSpec_cumulAlgo_curry Σ wfΣ Γ A B (typing_closed_context d') (type_is_open_term d1) (subject_is_open_term d2) c) - (s; d2)). simpl in H0. lia. + (has_sort_isType s d2)). simpl in H0. simpl. lia. Qed. (** Finally, for each typing spine built above, assuming we can apply the induction hypothesis @@ -2066,7 +2045,8 @@ Proof. split. eapply cumul_decorate in c; tea. now eapply trans_cumul in c. - specialize (IH _ _ i.π2 H). now exists i.π1. + destruct i as (?&?&Hty&?); simpl in *. + specialize (IH _ _ Hty H). now eexists. - simpl; intros. forward IHsp. @@ -2076,7 +2056,8 @@ Proof. forward IHt by lia. eapply cumul_decorate in c; tea. apply trans_cumul in c. - specialize (IH _ _ i.π2) as IHi. + destruct i as (?&?&Hty&?); simpl in *. + specialize (IH _ _ Hty) as IHi. forward IHi by lia. simpl in IHi. destruct IHsp as [T' [Hsp eq]]. @@ -2217,9 +2198,13 @@ Proof. apply (typing_ind_env_app_size (fun Σ Γ t T => TT.typing (trans_global Σ) (trans_local Γ) (trans t) (trans T) )%type + (fun Σ Γ j => TT.lift_typing TT.typing (trans_global Σ) (trans_local Γ) (judgment_map trans j)) (fun Σ Γ => TT.All_local_env (TT.lift_typing TT.typing (trans_global Σ)) (trans_local Γ)) );intros. + - destruct X as (Hb & s & (_ & IHt) & e). + repeat (eexists; tea); cbn. + destruct j_term => //=. destruct Hb => //. - now eapply trans_wf_local_env, All_over_All. - rewrite trans_lift. rewrite trans_decl_type. @@ -2303,7 +2288,7 @@ Proof. rewrite -[List.rev (trans_local _)]map_rev. clear. unfold app_context. change subst_instance_context with SE.subst_instance_context. unfold context. rewrite -map_rev. set (ctx := map _ (List.rev _)). clearbody ctx. - intro HH; pose proof (ctx_inst_impl _ (fun _ _ _ _ => TT.typing _ _ _ _ ) _ _ _ _ HH (fun _ _ H => H.2)); revert X; clear HH. + intro HH; pose proof (ctx_inst_impl _ (fun _ _ _ => TT.typing _ _ _ _ ) _ _ _ HH (fun _ _ H => H.2)); revert X; clear HH. now move: ctx; induction 1; cbn; constructor; auto; rewrite -(List.rev_involutive (map trans_decl Δ)) subst_telescope_subst_context -map_rev -(trans_subst_context [_]) -map_rev -PCUICSpine.subst_telescope_subst_context List.rev_involutive. @@ -2348,12 +2333,12 @@ Proof. reflexivity. + rewrite /trans_local map_app in X. now eapply TT.All_local_env_app_inv in X as []. - + eapply All_map, (All_impl X0). - intuition auto. destruct X2 as [s [? ?]]. - exists s; intuition auto. + + eapply All_map, (All_impl X1). + intros d (_ & ? & ? & _) => //. + repeat (eexists; tea). + fold trans. subst types. - eapply trans_mfix_All2; eassumption. + eapply trans_mfix_All2. eassumption. + now rewrite trans_wf_fixpoint. - rewrite trans_dtype. simpl. eapply TT.type_CoFix; auto. @@ -2366,8 +2351,9 @@ Proof. + rewrite /trans_local map_app in X. now eapply TT.All_local_env_app_inv in X as []. + fold trans. - eapply All_map, (All_impl X0). - intros x [s ?]; exists s; intuition auto. + eapply All_map, (All_impl X1). + intros d (_ & s & IHt & _). + repeat (eexists; tea); cbn. + fold trans;subst types. now apply trans_mfix_All2. + now rewrite trans_wf_cofixpoint. @@ -2386,7 +2372,7 @@ Proof. eapply ws_cumul_pb_forget in X4. eapply cumul_decorate in X4; tea. 2:eapply validity; tea. - 2:now exists s. + 2:now eapply has_sort_isType. now eapply trans_cumul. Qed. diff --git a/template-pcuic/theories/TemplateToPCUIC.v b/template-pcuic/theories/TemplateToPCUIC.v index 595df18f5..7df5f442f 100644 --- a/template-pcuic/theories/TemplateToPCUIC.v +++ b/template-pcuic/theories/TemplateToPCUIC.v @@ -46,7 +46,7 @@ Section Trans. Context (Σ : global_env_map). Definition dummy_decl : context_decl := - vass {| binder_name := nAnon; binder_relevance := Relevant |} (tSort Universe.type0). + vass {| binder_name := nAnon; binder_relevance := Relevant |} (tSort Sort.type0). Definition trans_predicate ind mdecl idecl pparams puinst pcontext preturn := let pctx := map2_bias_left set_binder_name dummy_decl pcontext (ind_predicate_context ind mdecl idecl) in @@ -90,7 +90,7 @@ Section Trans. still work. *) tCase ci {| pparams := p'.(Ast.pparams); puinst := p'.(Ast.puinst); - pcontext := map (fun na => vass na (tSort Universe.type0)) p'.(Ast.pcontext); + pcontext := map (fun na => vass na (tSort Sort.type0)) p'.(Ast.pcontext); preturn := p'.(Ast.preturn) |} (trans c) [] end diff --git a/template-pcuic/theories/TemplateToPCUICCorrectness.v b/template-pcuic/theories/TemplateToPCUICCorrectness.v index 9c342a166..1f7514b1c 100644 --- a/template-pcuic/theories/TemplateToPCUICCorrectness.v +++ b/template-pcuic/theories/TemplateToPCUICCorrectness.v @@ -278,7 +278,7 @@ Proof. intros wfΣ ext wftΣ wfΣ' H. destruct H. rewrite /trans_one_ind_body; destruct b; cbn in *. f_equal; solve_all. - - rewrite trans_decl_weakening //. + - rewrite trans_weakening //. - rewrite trans_weakening //. - unfold trans_constructor_body; destruct x; cbn in *; f_equal. * rewrite [map _ _]trans_local_weakening //. @@ -312,13 +312,14 @@ Qed. Import TypingWf. -Lemma weaken_wf_decl_pred {cf} (Σ Σ' : Ast.Env.global_env) Γ t T : +Lemma weaken_wf_decl_pred {cf} (Σ Σ' : Ast.Env.global_env) Γ j : Typing.wf Σ -> Ast.Env.extends Σ Σ' -> Typing.wf Σ' -> - WfAst.wf_decl_pred Σ Γ t T -> WfAst.wf_decl_pred Σ' Γ t T. + WfAst.wf_decl_pred Σ Γ j -> WfAst.wf_decl_pred Σ' Γ j. Proof. intros wf ext wf' ong. red in ong |- *. - destruct T; intuition eauto using wf_extends. + destruct j_term; cbn in *. + all: intuition eauto using wf_extends. Qed. Lemma trans_lookup {cf} Σ cst : @@ -739,7 +740,7 @@ Section Trans_Global. End Trans_Global. Lemma on_global_env_impl `{checker_flags} {Σ : Ast.Env.global_env} P Q : - (forall (Σ : Ast.Env.global_env_ext) Γ t T, ST.on_global_env Typing.cumul_gen P Σ -> P Σ Γ t T -> Q Σ Γ t T) -> + (forall (Σ : Ast.Env.global_env_ext) Γ j, ST.on_global_env Typing.cumul_gen P Σ -> P Σ Γ j -> Q Σ Γ j) -> ST.on_global_env Typing.cumul_gen P Σ -> ST.on_global_env Typing.cumul_gen Q Σ. Proof. apply on_global_env_impl. @@ -748,14 +749,16 @@ Qed. Lemma typing_wf_wf {cf}: forall (Σ : Ast.Env.global_env), ST.wf Σ -> - ST.Forall_decls_typing - (fun (Σ : Ast.Env.global_env_ext) (_ : Ast.Env.context) (t T : Ast.term) => WfAst.wf Σ t × WfAst.wf Σ T) Σ. + ST.on_global_env ST.cumul_gen + (fun Σ Γ j => WfAst.wf_decl_pred Σ Γ j) Σ. Proof. intros Σ. eapply on_global_env_impl. clear. - intros Σ' Γ t T ? HT. - apply ST.lift_typing_impl with (1 := HT); intros ? Hs. - now eapply typing_wf. + intros Σ' Γ j ? Hj. + apply lift_typing_wf_pred with (Γ := Γ). + apply ST.lift_typing_impl with (1 := Hj). + eapply typing_wf. + apply X. Qed. Lemma declared_inductive_inj {cf Σ mdecl mdecl' ind idecl idecl'} @@ -791,13 +794,13 @@ Section Trans_Global. Context (wfΣ : Typing.wf Σ). Context (wfΣ' : wf Σ'). - Lemma trans_R_global_instance {Re Rle gref napp u u'} : - SEq.R_global_instance Σ Re Rle gref napp u u' -> - R_global_instance (trans_global_env Σ) Re Rle gref napp u u'. + Lemma trans_cmp_global_instance {cmp_universe pb gref napp u u'} : + SEq.cmp_global_instance Σ cmp_universe pb gref napp u u' -> + cmp_global_instance (trans_global_env Σ) cmp_universe pb gref napp u u'. Proof. - unfold SEq.R_global_instance, SEq.global_variance. + unfold SEq.cmp_global_instance, SEq.global_variance. destruct gref; simpl; auto. - - unfold R_global_instance_gen, R_opt_variance; cbn. + - unfold cmp_global_instance_gen, cmp_opt_variance; cbn. unfold Ast.lookup_inductive_gen, lookup_inductive_gen, Ast.lookup_minductive_gen, lookup_minductive_gen. rewrite trans_lookup. destruct Ast.Env.lookup_env eqn:look => //; simpl. @@ -810,7 +813,7 @@ Section Trans_Global. generalize (trans_destArity Σ [] (Ast.Env.ind_type o) wfty wfΣ'). destruct Ast.destArity as [[ctx ps]|] eqn:eq' => /= // -> //. now rewrite context_assumptions_map. - - unfold R_global_instance_gen, R_opt_variance; cbn. + - unfold cmp_global_instance_gen, cmp_opt_variance; cbn. unfold lookup_constructor, lookup_inductive, lookup_minductive. unfold Ast.lookup_constructor, Ast.lookup_inductive, Ast.lookup_minductive. unfold lookup_constructor_gen, lookup_inductive_gen, lookup_minductive_gen. @@ -824,7 +827,7 @@ Section Trans_Global. Lemma eq_binder_annot_eq_context_gen_set_binder_name Γ Γ' Δ : All2 eq_binder_annot Γ Γ' -> - eq_context_gen eq eq (map2 set_binder_name Γ Δ) (map2 set_binder_name Γ' Δ). + eq_context_upto_names (map2 set_binder_name Γ Δ) (map2 set_binder_name Γ' Δ). Proof. induction 1 in Δ |- *. - constructor. @@ -855,16 +858,14 @@ Section Trans_Global. induction 1 in l' |- *; intros H; depelim H; intros H'; depelim H'; cbn; constructor; auto. Qed. - (* TODO updateTemplate Coq's eq_term to reflect PCUIC's cumulativity *) - Lemma trans_eq_term_upto_univ {Re Rle t u napp} : - RelationClasses.subrelation Re Rle -> + Lemma trans_eq_term_upto_univ {cmp_universe cmp_sort pb napp t u} : WfAst.wf Σ t -> WfAst.wf Σ u -> - SEq.eq_term_upto_univ_napp Σ Re Rle napp t u -> - eq_term_upto_univ_napp (trans_global_env Σ) Re Rle napp (trans (trans_global_env Σ) t) (trans (trans_global_env Σ) u). + SEq.eq_term_upto_univ_napp Σ cmp_universe cmp_sort pb napp t u -> + eq_term_upto_univ_napp (trans_global_env Σ) cmp_universe cmp_sort pb napp (trans (trans_global_env Σ) t) (trans (trans_global_env Σ) u). Proof. - intros sub wt wu e. - induction t using Induction.term_forall_list_rect in sub, Rle, napp, wt, u, wu, e |- *. + intros wt wu e. + induction t using Induction.term_forall_list_rect in pb, napp, wt, u, wu, e |- *. all: invs e; cbn. all: try solve [ constructor ; auto ]. all: repeat (match goal with @@ -874,17 +875,17 @@ Section Trans_Global. all: try solve [ repeat constructor ; auto ; match goal with - | ih : forall Rle (u : Ast.term) (napp : nat), _ |- _ => + | ih : forall pb napp (u : Ast.term), _ |- _ => now eapply ih end ]. - constructor. - solve_all. eapply a; auto. tc. + solve_all. - eapply eq_term_upto_univ_napp_mkApps. + rewrite map_length. now eapply IHt. - + destruct wt, wu. solve_all. eapply a0; auto; tc. - - constructor. apply trans_R_global_instance; auto. - - constructor. apply trans_R_global_instance; auto. + + destruct wt, wu. solve_all. + - constructor. apply trans_cmp_global_instance; auto. + - constructor. apply trans_cmp_global_instance; auto. - red in X, X0. destruct wt as [mdecl' [idecl' [decli hci hpctx lenpar eqpars eqret eqc eqbrs]]]. destruct wu as [mdecl'' [idecl'' [decli' hci' hpctx' lenpars' eqpars' eqret' eqc' eqbrs']]]. @@ -899,7 +900,7 @@ Section Trans_Global. destruct X. constructor. all: try solve [ match goal with - | ih : forall Rle u, _ |- _ => + | ih : forall pb napp u, _ |- _ => now eapply ih end ]. @@ -916,7 +917,8 @@ Section Trans_Global. eapply All2_All2_All2_All3; tea. cbn. intros cdecl br br' [[eq wfb] IH] [eq' wfb'] [eqbs eqbods]. split. - { rewrite map2_map2_bias_left; len. + { unfold trans_branch; cbn. + rewrite map2_map2_bias_left; len. eapply All2_length in eq. now len in eq. rewrite map2_map2_bias_left; len. eapply All2_length in eq'. now len in eq'. @@ -948,7 +950,7 @@ Section Trans_Global. simpl. intros [? ? ? ?] [? ? ? ?] [[[? ?] [[ih1 ih2] [? ?]]] [? ?]]. simpl in *. - intuition eauto. now eapply ih1. now eapply ih2. + intuition eauto. - constructor. assert ( w1 : @@ -969,11 +971,10 @@ Section Trans_Global. simpl. intros [? ? ? ?] [? ? ? ?] [[[? ?] [[ih1 ih2] [? ?]]] [? ?]]. simpl in *. - intuition eauto. now eapply ih1. now eapply ih2. + intuition eauto. - constructor; eauto. intuition auto; constructor; cbn; eauto. - eapply (IHt1 Re); eauto. reflexivity. - eapply (IHt2 Re); eauto. reflexivity. - solve_all. eapply a; eauto. reflexivity. + + inv H4. auto. + + solve_all. Qed. Lemma trans_leq_term ϕ T U : @@ -981,7 +982,7 @@ Section Trans_Global. leq_term (trans_global_env Σ) ϕ (trans Σ' T) (trans Σ' U). Proof. intros HT HU H. - eapply trans_eq_term_upto_univ ; eauto. tc. + eapply trans_eq_term_upto_univ ; eauto. Qed. Lemma trans_eq_term φ t u : @@ -989,7 +990,7 @@ Section Trans_Global. eq_term (trans_global_env Σ) φ (trans Σ' t) (trans Σ' u). Proof. intros HT HU H. - eapply trans_eq_term_upto_univ ; eauto. tc. + eapply trans_eq_term_upto_univ ; eauto. Qed. Lemma trans_eq_term_list {φ l l'} : @@ -1340,7 +1341,7 @@ Section Trans_Global. now rewrite trans_subst_context map_rev trans_local_subst_instance. Qed. - Lemma trans_local_app Γ Δ : trans_local Σ' (Ast.Env.app_context Γ Δ) = trans_local Σ' Γ ,,, trans_local Σ' Δ. + Lemma trans_local_app Γ Δ : trans_local Σ' (Γ ,,, Δ) = trans_local Σ' Γ ,,, trans_local Σ' Δ. Proof using Σ. now rewrite /trans_local map_app. Qed. @@ -1661,7 +1662,7 @@ Section Trans_Global. apply (OnOne2_All_mix_left Hwf) in X. solve_all. red. rewrite <- !map_dtype. rewrite <- !map_dbody. intuition eauto. - + unfold Ast.Env.app_context, trans_local in b. + + unfold app_context, trans_local in b. simpl in a. rewrite -> map_app in b. unfold app_context. unfold ST.fix_context in b. rewrite map_rev map_mapi in b. simpl in b. @@ -1694,7 +1695,7 @@ Section Trans_Global. apply (OnOne2_All_mix_left Hwf) in X. solve_all. red. rewrite <- !map_dtype. rewrite <- !map_dbody. intuition eauto. - + unfold Ast.Env.app_context, trans_local in b. + + unfold app_context, trans_local in b. simpl in a. rewrite -> map_app in b. unfold app_context. unfold ST.fix_context in b. rewrite map_rev map_mapi in b. simpl in b. @@ -1771,42 +1772,34 @@ Proof. eapply IHX2. auto. eapply wf_red1 in r; tea. now eapply typing_wf_sigma; auto. Qed. -Definition Tlift_typing (P : Ast.Env.global_env_ext -> Ast.Env.context -> Ast.term -> Ast.term -> Type) := - fun Σ Γ t T => - match T with - | Some T => P Σ Γ t T - | None => { s : Universe.t & P Σ Γ t (Template.Ast.tSort s) } - end. - Definition TTy_wf_local {cf : checker_flags} Σ Γ := ST.All_local_env (ST.lift_typing ST.typing Σ) Γ. Lemma trans_wf_local {cf}: forall (Σ : Ast.Env.global_env_ext) (Γ : Ast.Env.context) (wfΓ : TTy_wf_local Σ Γ), let P := - (fun Σ0 Γ0 _ (t T : Ast.term) _ => - let Σ' := trans_global Σ0 in + (fun Γ0 _ (t T : Ast.term) _ => + let Σ' := trans_global Σ in wf Σ'.1 -> - trans_global Σ0;;; trans_local Σ' Γ0 |- trans Σ' t : trans Σ' T) + trans_global Σ;;; trans_local Σ' Γ0 |- trans Σ' t : trans Σ' T) in let Σ' := trans_global Σ in wf Σ' -> - ST.All_local_env_over ST.typing P Σ Γ wfΓ -> + ST.All_local_env_over (ST.typing Σ) P Γ wfΓ -> wf_local (trans_global Σ) (trans_local Σ' Γ). Proof. intros. induction X0. - simpl. constructor. - - simpl. econstructor. - + eapply IHX0. - + simpl. destruct tu. exists x. now eapply Hs. - - simpl. constructor; auto. red. destruct tu. exists x; auto. - simpl. eauto. + - simpl. econstructor; auto. + simpl. destruct tu as (? & ? & ? & ?); cbn in *. repeat (eexists; tea). eauto. + - simpl. constructor; auto. + simpl. destruct tu as (? & ? & ? & ?); cbn in *. repeat (eexists; tea); cbn. all: eauto. Qed. Lemma trans_wf_local_env {cf} Σ Γ : ST.All_local_env (ST.lift_typing - (fun (Σ : Ast.Env.global_env_ext) Γ (b ty : Ast.term) => + (fun Σ Γ (b ty : Ast.term) => let Σ' := trans_global Σ in ST.typing Σ Γ b ty × trans_global Σ;;; trans_local Σ' Γ |- trans Σ' b : trans Σ' ty) Σ) Γ -> @@ -1815,12 +1808,10 @@ Proof. intros. induction X. - simpl. constructor. - - simpl. econstructor. - + eapply IHX. - + simpl. destruct t0. exists x. eapply p. - - simpl. constructor; auto. red. destruct t0. exists x. intuition auto. - red in t1. destruct t1. cbn. eapply p. - red; red in t1. destruct t1. eapply t2. + - simpl. econstructor; auto. + simpl. destruct t0 as (_ & ? & (? & ?) & ?); cbn in *. repeat (eexists; tea). + - simpl. constructor; auto. + simpl. destruct t0 as ((? & ?) & ? & (? & ?) & ?); cbn in *. repeat (eexists; cbn; tea). Qed. #[global] @@ -2080,7 +2071,7 @@ Lemma isType_mkApps_Ind_inv_spine {cf:checker_flags} {Σ : global_env_ext} {Γ i Proof. move=> isdecl isty. pose proof (isType_wf_local isty). - destruct isty as [s Hs]. + destruct isty as (_ & s & Hs & _). eapply invert_type_mkApps_ind in Hs as [sp cu]; tea. erewrite ind_arity_eq in sp. 2: eapply PCUICInductives.oib ; eauto. @@ -2213,12 +2204,10 @@ Proof. rewrite trans_reln //. Qed. -#[local] Hint Unfold lift_judgment : core. - -Lemma trans_wf_universe Σ u : S.wf_universe Σ u -> - wf_universe (trans_global Σ) u. +Lemma trans_wf_sort Σ u : S.wf_sort Σ u -> + wf_sort (trans_global Σ) u. Proof. - unfold S.wf_universe, wf_universe. + unfold S.wf_sort, wf_sort, S.wf_universe, wf_universe. now rewrite global_ext_levels_trans. Qed. @@ -2229,9 +2218,9 @@ Proof. induction decls; cbn; auto. Qed. -Local Hint Resolve trans_wf_universe : trans. +Local Hint Resolve trans_wf_sort : trans. Local Hint Transparent Ast.Env.global_env_ext : trans. -Local Hint Transparent Universe.t : trans. +Local Hint Transparent sort : trans. Local Hint Variables Transparent : trans. Ltac trans := try typeclasses eauto with trans. (* bug in Coq, typeclasses eauto tries exact with a quantified hypothesis starting with a let-in *) @@ -2246,6 +2235,11 @@ Theorem template_to_pcuic {cf} : let Σ' := trans_global Σ in wf Σ' -> typing Σ' (trans_local Σ' Γ) (trans Σ' t) (trans Σ' T)) + (fun Σ Γ j => + let Σ' := trans_global Σ in + wf Σ' -> + lift_typing typing Σ' (trans_local Σ' Γ) (judgment_map (trans Σ') j) + ) (fun Σ Γ _ => let Σ' := trans_global Σ in wf Σ' -> @@ -2255,6 +2249,11 @@ Proof. all: intros. all: try solve [ econstructor; trans ]. + - destruct X as (Hb & s & (_ & Ht) & e). unfold judgment_map. + split; cbn. + + destruct j_term => //=. now eapply Hb. + + exists s. split; tas. now eapply Ht. + - eapply trans_wf_local; eauto. - rewrite trans_lift. @@ -2263,19 +2262,21 @@ Proof. f_equal. - (* Casts *) + assert (lift_typing0 (typing Σ' (trans_local (trans_global_env Σ.1) Γ)) (TypUniv (trans (trans_global_env Σ.1) t) s)). + { repeat (eexists; eauto). } eapply refine_type; cbn. * eapply type_App. - 2:{ eapply type_Lambda; eauto. eapply type_Rel. econstructor; eauto. - eapply typing_wf_local; eauto. now eexists. reflexivity. } + 2:{ eapply type_Lambda; eauto. now eapply lift_sorting_forget_univ. eapply type_Rel. econstructor; eauto. + eapply typing_wf_local; eauto. now eapply has_sort_isType. reflexivity. } eapply type_Prod. eauto. instantiate (1 := s). simpl. eapply (weakening _ _ [_] _ (tSort _)); eauto. constructor; eauto. eapply typing_wf_local; eauto. - now eexists. + now eapply has_sort_isType. now eapply X2. * unfold subst1. rewrite simpl_subst; auto. now rewrite lift0_p. - - cbn. econstructor; trans. + - cbn. econstructor; trans. 1: now eapply X1. cbn in *. trans. cbn in *; trans. now apply X3. - econstructor; eauto with trans. @@ -2292,12 +2293,12 @@ Proof. * simpl in p. destruct (TypingWf.typing_wf _ wfΣ _ _ _ typrod) as [wfAB _]. econstructor; eauto. - + exists s; eauto. + + now eapply has_sort_isType. + rewrite -/Σ'. change (tProd na (trans Σ' A) (trans _ B)) with (trans Σ' (Ast.tProd na A B)). eapply trans_cumulSpec_typed; tea. eapply TypingWf.typing_all_wf_decl; eauto. - exists s; eauto. + now eapply has_sort_isType. + eapply typing_wf in ty; eauto. destruct ty as [wfhd _]. rewrite trans_subst in IHX1; eauto with wf. eapply IHX1; cycle 1. @@ -2306,7 +2307,7 @@ Proof. eapply PCUICInversion.inversion_Prod in p as [s1 [s2 [HA [HB Hs]]]]; auto. eapply (PCUICArities.isType_subst (Δ := [vass na (trans Σ' A)])); eauto. eapply subslet_ass_tip. eauto with pcuic. - now exists s2. + now eapply has_sort_isType. - cbn; rewrite trans_subst_instance. pose proof (forall_decls_declared_constant _ _ _ _ _ H). @@ -2432,18 +2433,19 @@ Proof. cbn. econstructor; eauto. eapply fix_guard_trans. assumption. now rewrite nth_error_map H0. - -- eapply All_map, (All_impl X0). - intros x [s [Hs Hts]]. - now exists s. - -- apply All_map. eapply All_impl; eauto. - intuition eauto 3 with wf; cbn. hnf in X3. - rewrite H2. rewrite /trans_local map_length. - rewrite /trans_local map_app in X3. - rewrite <- trans_lift. apply X3; auto. - -- eapply trans_wf_fixpoint => //. - solve_all; destruct a as [s [Hs IH]], b as [Hs' IH']. - now eapply TypingWf.typing_wf in Hs. - now eapply TypingWf.typing_wf in Hs'. + -- eapply All_map, (All_impl X1). + intros x (_ & s & Hts & _). assumption. + repeat (eexists; tea). + -- apply All_map. eapply (All_impl X3); eauto. + intros x (Htb & s & Hts & _). assumption. + unfold on_def_body, types in *; cbn in *. + rewrite H2. rewrite /trans_local !map_length map_app in Htb, Hts |- *. + rewrite <- trans_lift. + repeat (eexists; tea); cbn; eauto. + -- eapply trans_wf_fixpoint => //. clear X1 X3. + solve_all; destruct a as (_ & s & Hs & _), b as (Hs' & _). + now eapply TypingWf.typing_wf in Hs. + now eapply TypingWf.typing_wf in Hs'. -- destruct decl; reflexivity. - eapply refine_type. @@ -2456,15 +2458,17 @@ Proof. cbn; econstructor; eauto. -- now eapply cofix_guard_trans. -- now rewrite nth_error_map H0. - -- eapply All_map, (All_impl X0). - intros x [s [Hs Hts]]. now exists s. - -- apply All_map. eapply All_impl; eauto. - intuition eauto 3 with wf. hnf in X3. - rewrite H2. rewrite /trans_local map_length. - rewrite trans_local_app in X3. - cbn. rewrite <- trans_lift. now apply X3. - -- eapply trans_wf_cofixpoint => //. - solve_all; destruct a as [s [Hs IH]], b as [Hs' IH']. + -- eapply All_map, (All_impl X1). + intros x (_ & s & Hts & _). assumption. + repeat (eexists; tea). + -- apply All_map. eapply (All_impl X3); eauto. + intros x (Htb & s & Hts & _). assumption. + unfold on_def_body, types in *; cbn in *. + rewrite H2. rewrite /trans_local !map_length map_app in Htb, Hts |- *. + rewrite <- trans_lift. + repeat (eexists; tea); cbn; eauto. + -- eapply trans_wf_cofixpoint => //. clear X1 X3. + solve_all; destruct a as (_ & s & Hs & _), b as (Hs' & _). now eapply TypingWf.typing_wf in Hs. now eapply TypingWf.typing_wf in Hs'. -- destruct decl; reflexivity. @@ -2499,7 +2503,7 @@ Proof. now rewrite H1 H2 H3 /= in H0 |- *. + constructor; eauto. cbn [array_level a]. eapply validity in X1; eauto. eapply PCUICWfUniverses.isType_wf_universes in X1. cbn [trans PCUICWfUniverses.wf_universes] in X1. - unfold PCUICWfUniverses.wf_universes in X1. cbn [PCUICWfUniverses.on_universes] in X1. + unfold PCUICWfUniverses.wf_universes in X1. cbn [PCUICWfUniverses.on_universes Sort.on_sort s] in X1. move: X1. case: PCUICWfUniverses.wf_universe_reflect => //; eauto. eauto. cbn [a array_value]. solve_all. - assert (WfAst.wf Σ B). @@ -2510,7 +2514,7 @@ Proof. eapply typing_wf in X0; eauto. destruct X0. auto. specialize (X1 X5). now eapply validity in X1. - specialize (X3 X5). now exists s. + specialize (X3 X5). now eapply has_sort_isType. Qed. Lemma Alli_map {A B} (P : nat -> B -> Type) n (f : A -> B) l : @@ -2642,10 +2646,10 @@ Lemma trans_cumul_ctx_rel {cf} {Σ : Ast.Env.global_env_ext} Γ Δ Δ' : let Σ' := trans_global Σ in Typing.wf Σ -> wf Σ' -> ST.TemplateConversion.cumul_ctx_rel Typing.cumul_gen Σ Γ Δ Δ' -> - closed_ctx (trans_local Σ' (Ast.Env.app_context Γ Δ)) -> - closed_ctx (trans_local Σ' (Ast.Env.app_context Γ Δ')) -> - All (WfAst.wf_decl Σ) (Ast.Env.app_context Γ Δ) -> - All (WfAst.wf_decl Σ) (Ast.Env.app_context Γ Δ') -> + closed_ctx (trans_local Σ' (Γ ,,, Δ)) -> + closed_ctx (trans_local Σ' (Γ ,,, Δ')) -> + All (WfAst.wf_decl Σ) (Γ ,,, Δ) -> + All (WfAst.wf_decl Σ) (Γ ,,, Δ') -> cumul_ctx_rel cumulSpec0 Σ' (trans_local Σ' Γ) (trans_local Σ' Δ) (trans_local Σ' Δ'). Proof. intros Σ' wfΣ wfΣ'. @@ -2753,17 +2757,14 @@ Lemma trans_cstr_respects_variance {cf} Σ mdecl v cdecl : All (WfAst.wf Σ) (Ast.Env.cstr_indices cdecl) -> (closed_ctx (trans_local (trans_global_env Σ) - (Ast.Env.app_context - (Ast.Env.app_context (ST.ind_arities mdecl) - (Ast.Env.smash_context [] (Ast.Env.ind_params mdecl))) + (ST.ind_arities mdecl ,,, Ast.Env.smash_context [] (Ast.Env.ind_params mdecl) ,,, (Ast.Env.expand_lets_ctx (Ast.Env.ind_params mdecl) (Ast.Env.smash_context [] (Ast.Env.cstr_args cdecl)))))) -> All (fun x : Ast.term => closedn (Ast.Env.context_assumptions (Ast.Env.cstr_args cdecl) + Ast.Env.context_assumptions (Ast.Env.ind_params mdecl) + #|ST.ind_arities mdecl|) (trans Σ' (Ast.Env.expand_lets - (Ast.Env.app_context (Ast.Env.ind_params mdecl) - (Ast.Env.cstr_args cdecl)) x))) + (Ast.Env.ind_params mdecl ,,, Ast.Env.cstr_args cdecl) x))) (Ast.Env.cstr_indices cdecl) -> ST.cstr_respects_variance Typing.cumul_gen Σ mdecl v cdecl -> cstr_respects_variance cumulSpec0 Σ' mdecl' v cdecl'. @@ -2818,7 +2819,7 @@ Proof. rewrite -trans_expand_lets_ctx trans_ind_arities. rewrite -!trans_local_app. move/(closed_ctx_on_free_vars xpred0). - rewrite !Ast.Env.app_context_assoc. + rewrite !app_context_assoc. now rewrite on_free_vars_ctx_subst_instance. } { eapply closedn_on_free_vars. len. move: a0. rewrite /ST.ind_arities; len. @@ -2853,8 +2854,7 @@ Lemma trans_ind_respects_variance {cf} Σ mdecl v idecl : wf_inductive_body Σ idecl -> All (WfAst.wf_decl Σ) (Ast.Env.ind_params mdecl) -> closed_ctx (trans_local (trans_global_env Σ) - (Ast.Env.app_context (Ast.Env.ind_params mdecl) - (Ast.Env.ind_indices idecl))) -> + (Ast.Env.ind_params mdecl ,,, Ast.Env.ind_indices idecl)) -> ST.ind_respects_variance Typing.cumul_gen Σ mdecl v (Ast.Env.ind_indices idecl) -> ind_respects_variance cumulSpec0 Σ' mdecl' v (ind_indices idecl'). Proof. @@ -2891,15 +2891,13 @@ Proof. Qed. Lemma trans_type_local_ctx {cf} {Σ : Ast.Env.global_env_ext} Γ cs s (Σ' := trans_global Σ) : - (forall (Σ : Ast.Env.global_env_ext) - (Γ : Ast.Env.context) (t : Ast.term) (T : typ_or_sort_ Ast.term), + (forall (Σ : Ast.Env.global_env_ext) (Γ : Ast.Env.context) j, Typing.wf Σ -> - Typing.lift_typing Typing.typing Σ Γ t T -> + Typing.lift_typing Typing.typing Σ Γ j -> wf (trans_global_env Σ) -> lift_typing typing (trans_global Σ) (trans_local (trans_global_env Σ) Γ) - (trans (trans_global_env Σ) t) - (typ_or_sort_map (trans (trans_global_env Σ)) T)) -> + (judgment_map (trans (trans_global_env Σ)) j)) -> Typing.wf Σ -> wf Σ' -> ST.type_local_ctx (Typing.TemplateEnvTyping.lift_typing Typing.typing) Σ Γ cs s -> @@ -2909,17 +2907,14 @@ Lemma trans_type_local_ctx {cf} {Σ : Ast.Env.global_env_ext} Γ cs s (Σ' := tr (trans_local Σ' cs) s. Proof. intros IH wfΣ wfΣ'. induction cs; simpl; auto. - { now intros; eapply trans_wf_universe. } + { now intros; eapply trans_wf_sort. } destruct a as [na [b|] ty] => //; - intros [? ?]; cbn. destruct p. - intuition auto. - - eapply (IH _ _ _ Sort) in i; auto. + intros [? ?]; cbn. + all: split; auto. + - eapply (IH _ _ (TermTyp _ _)) in l; auto. rewrite -trans_local_app //. - - eapply (IH _ _ _ (Typ _)) in t0 => //. + - eapply (IH _ _ (TypUniv _ _)) in l => //. rewrite -trans_local_app //. - - eapply (IH _ _ _ (Typ _)) in t0 => //. - rewrite -trans_local_app; split => //. - now eapply IHcs. Qed. Lemma trans_check_ind_sorts {cf} Σ udecl kn mdecl n idecl @@ -2929,14 +2924,13 @@ Lemma trans_check_ind_sorts {cf} Σ udecl kn mdecl n idecl Typing.wf Σ -> wf Σ' -> (forall (Σ : Ast.Env.global_env_ext) - (Γ : Ast.Env.context) (t : Ast.term) (T : typ_or_sort_ Ast.term), + (Γ : Ast.Env.context) j, Typing.wf Σ -> - Typing.lift_typing Typing.typing Σ Γ t T -> + Typing.lift_typing Typing.typing Σ Γ j -> wf (trans_global_env Σ) -> lift_typing typing (trans_global Σ) (trans_local (trans_global_env Σ) Γ) - (trans (trans_global_env Σ) t) - (typ_or_sort_map (trans (trans_global_env Σ)) T)) -> + (judgment_map (trans (trans_global_env Σ)) j)) -> forall (oni: Typing.on_ind_body Typing.cumul_gen (Typing.TemplateEnvTyping.lift_typing Typing.typing) (Σ, udecl) kn mdecl n idecl), ST.check_ind_sorts @@ -2954,8 +2948,7 @@ Lemma trans_check_ind_sorts {cf} Σ udecl kn mdecl n idecl Proof. intros wfΣ wfΣ' IH oni. unfold ST.check_ind_sorts, check_ind_sorts. cbn. - destruct Universe.is_prop => //. - destruct Universe.is_sprop => //. + destruct Sort.to_family => //. intros []. split => //. now rewrite -global_ext_constraints_trans in c. destruct indices_matter => //. @@ -2969,10 +2962,9 @@ Lemma on_global_decl_wf {cf} {Σ : Ast.Env.global_env_ext} {kn d} : Proof. intros. eapply TypingWf.on_global_decl_impl; tea. intros. - destruct T. - * eapply typing_wf; tea. - * destruct X2 as [s Hs]. - red. split => //. now eapply typing_wf in Hs; tea. + apply lift_typing_wf_pred. cbn in X2. + eapply Typing.lift_typing_impl with (1 := X2). + now eapply typing_wf. Qed. Lemma Alli_All_mix {A} {P : nat -> A -> Type} {Q} {n l} : @@ -3002,7 +2994,7 @@ Proof. cbn. rewrite map_app closedn_ctx_app /=. len. apply/andP; split. rewrite /test_decl /trans_decl /=. - destruct p as [s hs]. now eapply subject_closed in hs. + destruct p as (_ & s & hs & _). now eapply subject_closed in hs. eapply closedn_ctx_upwards; tea. lia. Qed. @@ -3042,11 +3034,11 @@ Proof. Qed. Lemma trans_on_global_env `{checker_flags} Σ : - (forall (Σ : Ast.Env.global_env_ext) Γ t T, Typing.wf Σ -> - Typing.lift_typing Typing.typing Σ Γ t T -> + (forall (Σ : Ast.Env.global_env_ext) Γ j, Typing.wf Σ -> + Typing.lift_typing Typing.typing Σ Γ j -> let Σ' := trans_global_env Σ in wf Σ' -> - lift_typing typing (trans_global Σ) (trans_local Σ' Γ) (trans Σ' t) (typ_or_sort_map (trans Σ') T)) -> + lift_typing typing (trans_global Σ) (trans_local Σ' Γ) (judgment_map (trans Σ') j)) -> Typing.wf Σ -> wf (trans_global_env Σ). Proof. intros X X0. @@ -3075,11 +3067,7 @@ Proof. { split; rewrite trans_env_env_universes //. } have wfdecl := on_global_decl_wf (Σ := (Σg, udecl)) X0 on_global_decl_d. destruct d eqn:eqd. - * destruct c; simpl. destruct cst_body0; simpl in *. - red in on_udecl_udecl |- *. simpl in *. - eapply (X (Σg, cst_universes0) [] t (Typ cst_type0)); auto. - red in on_global_decl_d |- *. simpl in *. - now apply (X (Σg, cst_universes0) [] cst_type0 Sort). + * eapply (X (Σg, _) [] (TermoptTyp _ _)); auto. * destruct on_global_decl_d as [onI onP onNP]. simpl. change (trans_env_env (trans_global_env Σg), Ast.Env.ind_universes m) with (global_env_ext_map_global_env_ext (trans_global (Σg, Ast.Env.ind_universes m))) in *. @@ -3091,14 +3079,14 @@ Proof. { eapply closed_arities_context => //. clear -onu wfΣg X0 IHond X onI. eapply Alli_All; tea; cbv beta. move=> n x /Typing.onArity => o. - apply (X (Σg, Ast.Env.ind_universes m) [] (Ast.Env.ind_type x) Sort X0 o) => //. } + apply (X (Σg, Ast.Env.ind_universes m) [] _ X0 o) => //. } eapply Alli_All_mix in onI; tea. eapply Alli_map. eapply Alli_impl. exact onI. eauto. intros n idecl [oni wf]. have onarity : on_type (PCUICEnvTyping.lift_typing typing) (trans_global (Σg, Ast.Env.ind_universes m)) [] (ind_type (trans_one_ind_body (trans_global_env Σg) idecl)). { apply ST.onArity in oni. unfold on_type in *; simpl in *. - now apply (X (Σg, Ast.Env.ind_universes m) [] (Ast.Env.ind_type idecl) Sort). } + now apply (X (Σg, Ast.Env.ind_universes m) [] (Typ (Ast.Env.ind_type idecl))). } unshelve refine {| ind_cunivs := oni.(ST.ind_cunivs) |}; tea. --- cbn -[trans_global_env]. rewrite oni.(ST.ind_arity_eq). now rewrite ![trans _ _]trans_it_mkProd_or_LetIn. @@ -3116,7 +3104,7 @@ Proof. (cstr_type (trans_constructor_body (trans_global_env Σg) x)). { unfold cstr_type, Ast.Env.cstr_type in on_ctype |- *; simpl in *. red. move: (X (Σg, Ast.Env.ind_universes m) (Ast.Env.arities_context (Ast.Env.ind_bodies m)) - (Ast.Env.cstr_type x) Sort). + (Typ (Ast.Env.cstr_type x))). rewrite trans_arities_context. intros H'. apply H' => //. } have ceq : cstr_type (trans_constructor_body (trans_global_env Σg) x) = @@ -3140,22 +3128,18 @@ Proof. + simpl. unfold trans_local. rewrite context_assumptions_map. now rewrite cstr_args_length. + clear -X0 wfΣg onu IHond X on_cargs. revert on_cargs. simpl. - have foo := (X (Σg, udecl) _ _ _ X0). + have foo := (X (Σg, udecl) _ _ X0). rewrite -trans_arities_context. clear X. induction (Ast.Env.cstr_args x) in cs |- *; destruct cs; simpl; auto; destruct a as [na [b|] ty]; simpl in *; auto; split; intuition eauto; specialize (foo - (Ast.Env.app_context (Ast.Env.app_context - (Ast.Env.arities_context (Ast.Env.ind_bodies m)) (Ast.Env.ind_params m)) - c)); + (Ast.Env.arities_context (Ast.Env.ind_bodies m) ,,, Ast.Env.ind_params m ,,, c)); rewrite /trans_local !map_app in foo. - now eapply (foo ty Sort). - now apply (foo b (Typ ty)). - now apply (foo ty Sort). - now apply (foo b (Typ ty)). - now apply (foo ty (Typ (Ast.tSort _))). + now apply (foo (TermTyp b ty)). + now apply (foo (TermTyp b ty)). + now apply (foo (TypUniv ty _)). + clear -X0 onu wfΣg IHond X on_cindices. revert on_cindices. rewrite -trans_lift_context /trans_local -map_rev. @@ -3163,12 +3147,11 @@ Proof. generalize (List.rev (Ast.Env.lift_context #|Ast.Env.cstr_args x| 0 (Ast.Env.ind_indices idecl))). rewrite -trans_arities_context. induction 1; simpl; constructor; auto; - have foo := (X (Σg, Ast.Env.ind_universes m) _ _ _ X0); - specialize (foo (Ast.Env.app_context (Ast.Env.app_context - (Ast.Env.arities_context (Ast.Env.ind_bodies m)) - (Ast.Env.ind_params m)) (Ast.Env.cstr_args x))); + have foo := (X (Σg, Ast.Env.ind_universes m) _ _ X0); + specialize (foo + (Ast.Env.arities_context (Ast.Env.ind_bodies m) ,,, Ast.Env.ind_params m ,,, Ast.Env.cstr_args x)); rewrite /trans_local !map_app in foo. - now apply (foo i (Typ t)). + now apply (foo (TermTyp i t)). now rewrite (trans_subst_telescope _ [i] 0) in IHon_cindices. now rewrite (trans_subst_telescope _ [b] 0) in IHon_cindices. + cbn -[trans_global_env] in *. @@ -3203,7 +3186,7 @@ Proof. have wfty : WfAst.wf Σg (Ast.Env.ind_type i). { rewrite nth_error_rev in H2. len. rewrite List.rev_involutive in H2. eapply nth_error_alli in onI; tea. cbn in onI. - destruct onI as [onI _]. eapply Typing.onArity in onI as [s Hs]. + destruct onI as [onI _]. eapply Typing.onArity in onI as (_ & s & Hs & _). now eapply typing_wf in Hs. } rewrite /trans_one_ind_body /= /ind_realargs /= /ST.ind_realargs. generalize (trans_destArity Σg [] _ wfty wfΣg). @@ -3222,13 +3205,13 @@ Proof. specialize (on_ctype_variance _ indv). cbn -[Σg]. eapply trans_cstr_respects_variance => //. - { do 2 red in onty. destruct onty as [s hs]. + { do 2 red in onty. destruct onty as (_ & s & hs & _). rewrite ceq in hs. eapply PCUICClosedTyp.subject_closed in hs. len in hs. rewrite -it_mkProd_or_LetIn_app in hs. rewrite closedn_it_mkProd_or_LetIn in hs. move/andP: hs => [] cl _. move: cl. - rewrite -Ast.Env.app_context_assoc. + rewrite -app_context_assoc. rewrite !trans_local_app !trans_smash_context trans_expand_lets_ctx /=. rewrite !trans_smash_context. rewrite -(smash_context_app_expand []). @@ -3237,7 +3220,7 @@ Proof. now apply closedn_smash_context. } { rewrite ceq in onty. rewrite /cstr_concl in onty. - destruct onty as [s hs]. + destruct onty as (_ & s & hs & _). rewrite -it_mkProd_or_LetIn_app in hs. eapply PCUICSpine.type_it_mkProd_or_LetIn_inv in hs as [Δs [us []]]. eapply typing_expand_lets in t. eapply subject_closed in t; move: t. @@ -3265,7 +3248,7 @@ Proof. unfold ST.on_projection, on_projection. cbn -[inds Σg]. rewrite context_assumptions_map. rewrite -[trans_local _ _ ++ _]trans_local_app -(trans_smash_context _ []) nth_error_map. - rewrite /Ast.Env.app_context. destruct nth_error => // /=. + rewrite /app_context. destruct nth_error => // /=. rewrite /trans_projection_body /=. move=> [] /= hb -> ->. cbn -[Σg]. split => //. rewrite trans_subst trans_inds. cbn -[Σg]. f_equal. @@ -3278,7 +3261,7 @@ Proof. eapply trans_ind_respects_variance; tea. move: onarity; cbn -[Σg]. rewrite oni.(ST.ind_arity_eq). rewrite !trans_it_mkProd_or_LetIn. - move=> [] s /subject_closed. rewrite -it_mkProd_or_LetIn_app closedn_it_mkProd_or_LetIn /=. + move=> [] _ [] s [] /subject_closed. rewrite -it_mkProd_or_LetIn_app closedn_it_mkProd_or_LetIn /=. rewrite andb_true_r trans_local_app //. -- red in onP. red. epose proof (Typing.env_prop_wf_local template_to_pcuic (Σg, Ast.Env.ind_universes m) X0 _ onP). @@ -3300,14 +3283,11 @@ Proof. intros Hu. eapply trans_on_global_env; eauto. simpl; intros. epose proof (ST.env_prop_typing template_to_pcuic _ X Γ). - forward X2. - red in X0. destruct T. - now eapply ST.typing_wf_local. - destruct X0 as [s Hs]. eapply ST.typing_wf_local; eauto. - destruct T; simpl in *. - - eapply X2; eauto. - - destruct X0 as [s Hs]. exists s. - eapply (X2 _ (Ast.tSort s)); eauto. + forward X2. { destruct X0 as (_ & s & X0 & _). now eapply ST.typing_wf_local. } + destruct X0 as (Hb & s & Ht & e). unfold judgment_map. + repeat (eexists; tea); cbn. + - destruct j_term => //. eapply X2; eauto. + - eapply (X2 _ (Ast.tSort s)); eauto. Qed. Lemma template_to_pcuic_env_ext {cf} Σ :Template.Typing.wf_ext Σ -> wf_ext (trans_global Σ). diff --git a/template-pcuic/theories/TemplateToPCUICExpanded.v b/template-pcuic/theories/TemplateToPCUICExpanded.v index cf6f8e115..a1a1d9dd0 100644 --- a/template-pcuic/theories/TemplateToPCUICExpanded.v +++ b/template-pcuic/theories/TemplateToPCUICExpanded.v @@ -259,9 +259,6 @@ Proof. destruct wf. split => //. now depelim o0. eapply typing_wf_wf in wf. depelim wf. cbn in o0. depelim o0. cbn. destruct o1. split => //. - eapply TypingWf.on_global_decl_impl; tea. cbn. - intros. destruct T => //. red. red in X0. destruct X0. intuition auto. - cbn. split => //. Qed. Lemma trans_global_env_cons univs retro (Σ : Ast.Env.global_declarations) decl : @@ -311,9 +308,9 @@ Lemma wf_context_sorts {cf} {Σ ctx ctx' cunivs} {wfΣ : Typing.wf_ext Σ} : Proof. induction ctx' in cunivs |- *; cbn; auto. destruct a as [na [b|] ty]. - intros [? []]. constructor; auto. eauto. + intros [? ?]. constructor; auto. eauto. destruct cunivs => //. - intros [? []]. constructor; eauto. constructor; cbn; eauto. + intros [? ?]. constructor; auto. eauto. Qed. Lemma expanded_trans_global_env {cf} Σ {wfΣ : Typing.wf_ext Σ} : diff --git a/test-suite/bug369.v b/test-suite/bug369.v index 22546bb68..88bfa87f9 100644 --- a/test-suite/bug369.v +++ b/test-suite/bug369.v @@ -13,7 +13,7 @@ Definition mkImplN name (A B : term) : term := Definition one_pt_i : one_inductive_entry := {| mind_entry_typename := "Point"; - mind_entry_arity := tSort Universe.type0; + mind_entry_arity := tSort Sort.type0; mind_entry_consnames := ["mkPoint"]; mind_entry_lc := [ mkImplN "coordx"%bs (tRel 0) (mkImplN "coordy"%bs (tRel 1) (tApp (tRel 3) [tRel 2]))]; @@ -24,7 +24,7 @@ Definition mut_pt_i : mutual_inductive_entry := mind_entry_record := Some (Some "mkPoint" (* Irrelevant *)); mind_entry_finite := BiFinite; mind_entry_params := [{| decl_name := bnamed "A"; decl_body := None; - decl_type := (tSort Universe.type0) |}]; + decl_type := (tSort Sort.type0) |}]; mind_entry_inds := [one_pt_i]; mind_entry_universes := Monomorphic_entry ContextSet.empty; mind_entry_template := false; diff --git a/test-suite/castprop.v b/test-suite/castprop.v index f1514c150..3fc76da7a 100644 --- a/test-suite/castprop.v +++ b/test-suite/castprop.v @@ -18,4 +18,4 @@ Definition setprop : { x : nat | x = 0 } := exist _ 0 eq_refl. MetaCoq Quote Recursively Definition q_setprop := setprop. Notation proof t := - (Ast.tCast t BasicAst.Cast (Ast.tCast _ BasicAst.Cast (Ast.tSort ((Universes.Universe.lProp :: nil)%list; _)))). + (Ast.tCast t BasicAst.Cast (Ast.tCast _ BasicAst.Cast (Ast.tSort ((Universes.sProp :: nil)%list; _)))). diff --git a/test-suite/inferind.v b/test-suite/inferind.v index db164e740..c45953788 100644 --- a/test-suite/inferind.v +++ b/test-suite/inferind.v @@ -12,12 +12,7 @@ Definition qlist := Eval compute in match <% list %> with Definition refresh_sort t := match t with - | tSort s => - match s with - | Universe.lProp => tSort Universe.lProp - | Universe.lSProp => tSort Universe.lSProp - | Universe.lType _ => tSort Universes.fresh_universe - end + | tSort s => tSort (Sort.map (fun _ => fresh_universe) s) | _ => t end. diff --git a/test-suite/tmFix.v b/test-suite/tmFix.v index db8c26c69..ec702551e 100644 --- a/test-suite/tmFix.v +++ b/test-suite/tmFix.v @@ -71,12 +71,18 @@ Module Unquote. Local Unset Universe Minimization ToSet. (* idk why this is needed... *) #[local] Hint Extern 1 (Monad _) => refine TemplateMonad_Monad : typeclass_instances. + Definition tmQuoteSort@{U t u} : TemplateMonad@{t u} sort + := p <- @tmQuote Prop (Type@{U} -> True);; + match p with + | tProd _ (tSort s) _ => ret s + | _ => tmFail "Anomaly: tmQuote (Type -> True) should be (tProd _ (tSort _) _)!"%bs + end. Definition tmQuoteUniverse@{U t u} : TemplateMonad@{t u} Universe.t - := u <- @tmQuote Prop (Type@{U} -> True);; - match u with - | tProd _ (tSort u) _ => ret u - | _ => tmFail "Anomaly: tmQuote (Type -> True) should be (tProd _ (tSort _) _)!"%bs - end. + := s <- @tmQuoteSort@{U t u};; + match s with + | sType u => ret u + | _ => tmFail "Sort does not carry a universe (is not Type)"%bs + end. Definition tmQuoteLevel@{U t u} : TemplateMonad@{t u} Level.t := u <- tmQuoteUniverse@{U t u};; match Universe.get_is_level u with diff --git a/test-suite/univ.v b/test-suite/univ.v index ced8b4b93..bb1d830a8 100644 --- a/test-suite/univ.v +++ b/test-suite/univ.v @@ -17,7 +17,7 @@ MetaCoq Quote Definition a_random_univ := Type. Example a_random_univ_ex : exists l, a_random_univ = - tSort (Universe.from_kernel_repr (Level.level l, false) []). + tSort (sType (Universe.make' (Level.level l))). Proof. eexists. reflexivity. Qed. (* Back and forth *) @@ -30,18 +30,18 @@ Check eq_refl : univ_foo = univ_foo_back. Print univ_foo_back. -Fail MetaCoq Unquote Definition t1 := (tSort (Universe.make (Level.level "Top.400"))). +Fail MetaCoq Unquote Definition t1 := (tSort (sType (Universe.make' (Level.level "Top.400")))). (* Fails with "Level Top. not a declared level and you are in Strict Unquote Universe Mode." *) Unset MetaCoq Strict Unquote Universe Mode. -MetaCoq Unquote Definition t2 := (tSort fresh_universe). -MetaCoq Unquote Definition t3 := (tSort (Universe.make (Level.level "Top.400"))). +MetaCoq Unquote Definition t2 := (tSort (sType fresh_universe)). +MetaCoq Unquote Definition t3 := (tSort (sType (Universe.make' (Level.level "Top.400")))). Monomorphic Universe i j. Set MetaCoq Strict Unquote Universe Mode. MetaCoq Test Quote (Type@{j} -> Type@{i}). -MetaCoq Unquote Definition T'' := (tSort (Universe.make (Level.level "j"))). +MetaCoq Unquote Definition T'' := (tSort (sType (Universe.make' (Level.level "j")))). Unset MetaCoq Strict Unquote Universe Mode. @@ -196,7 +196,7 @@ Definition nNamedR (s : string) := mkBindAnn (nNamed s) Relevant. Definition nAnonR := mkBindAnn nAnon Relevant. Unset MetaCoq Strict Unquote Universe Mode. -MetaCoq Unquote Definition bla' := (tLambda (nNamedR "T") (tSort (Universe.make (Level.level "Top.46"))) (tLambda (nNamedR "T2") (tSort (Universe.make (Level.level "Top.477"))) (tProd nAnonR (tRel 1) (tRel 1)))). +MetaCoq Unquote Definition bla' := (tLambda (nNamedR "T") (tSort (sType (Universe.make' (Level.level "Top.46")))) (tLambda (nNamedR "T2") (tSort (sType (Universe.make' (Level.level "Top.477")))) (tProd nAnonR (tRel 1) (tRel 1)))). Set Printing Universes. Print bla. diff --git a/translations/param_cheap_packed.v b/translations/param_cheap_packed.v index 7f5254513..79c1e4ad2 100644 --- a/translations/param_cheap_packed.v +++ b/translations/param_cheap_packed.v @@ -8,9 +8,14 @@ Local Existing Instance config.default_checker_flags. Local Existing Instance default_fuel. +Definition refresh_universe u := + if Universe.is_level u then u else fresh_universe. + +Definition refresh_sort_universe := Sort.map refresh_universe. + Fixpoint refresh_universes (t : term) {struct t} := match t with - | tSort s => tSort (if Universe.is_level s then s else fresh_universe) + | tSort s => tSort (refresh_sort_universe s) | tProd na b t => tProd na b (refresh_universes t) | tLetIn na b t' t => tLetIn na b t' (refresh_universes t) | _ => t diff --git a/translations/param_generous_packed.v b/translations/param_generous_packed.v index c8211e2b7..281844476 100644 --- a/translations/param_generous_packed.v +++ b/translations/param_generous_packed.v @@ -28,10 +28,14 @@ Definition proj2 (t : term) : term Definition proj (b : bool) (t : term) : term := tProj (mkProjection sigma_ind 2 (if b then 0 else S 0)) t. +Definition refresh_universe u := + if Universe.is_level u then u else fresh_universe. + +Definition refresh_sort_universe := Sort.map refresh_universe. Fixpoint refresh_universes (t : term) {struct t} := match t with - | tSort s => tSort (if Universe.is_level s then s else fresh_universe) + | tSort s => tSort (refresh_sort_universe s) | tProd na b t => tProd na b (refresh_universes t) | tLetIn na b t' t => tLetIn na b t' (refresh_universes t) | _ => t @@ -108,8 +112,8 @@ with tsl_term (fuel : nat) (Σ : global_env_ext) (E : tsl_table) (Γ : context) | tRel n => ret (tRel n) | tSort s => - ret (pair (tSort fresh_universe) - (tLambda (nNamed "A") (tSort fresh_universe) (tProd nAnon (tRel 0) (tSort fresh_universe))) + ret (pair (tSort (sType fresh_universe)) + (tLambda (nNamed "A") (tSort (sType fresh_universe)) (tProd nAnon (tRel 0) (tSort (sType fresh_universe)))) (tSort s) (tLambda (nNamed "A") (tSort s) (tProd nAnon (tRel 0) (tSort s)))) diff --git a/translations/times_bool_fun.v b/translations/times_bool_fun.v index ee85d7df2..41ae75cbe 100644 --- a/translations/times_bool_fun.v +++ b/translations/times_bool_fun.v @@ -207,10 +207,14 @@ Definition tsl_mind_body (ΣE : tsl_context) (mp : modpath) (kn : kername) exact (fun A Γ' => Γ' ,, vass (decl_name A) (tsl Γ' (decl_type A))). Defined. +Definition refresh_universe u := + if Universe.is_level u then u else fresh_universe. + +Definition refresh_sort_universe := Sort.map refresh_universe. Fixpoint refresh_universes (t : term) {struct t} := match t with - | tSort s => tSort (if Universe.is_level s then s else fresh_universe) + | tSort s => tSort (refresh_sort_universe s) | tProd na b t => tProd na b (refresh_universes t) | tLetIn na b t' t => tLetIn na b t' (refresh_universes t) | tCast x x0 x1 => tCast (refresh_universes x) x0 (refresh_universes x1) diff --git a/utils/theories/All_Forall.v b/utils/theories/All_Forall.v index 6960f2987..13c8e97d6 100644 --- a/utils/theories/All_Forall.v +++ b/utils/theories/All_Forall.v @@ -80,7 +80,17 @@ Inductive All3 {A B C : Type} (R : A -> B -> C -> Type) : list A -> list B -> li R x y z -> All3 R l l' l'' -> All3 R (x :: l) (y :: l') (z :: l''). Arguments All3_nil {_ _ _ _}. Arguments All3_cons {_ _ _ _ _ _ _ _ _ _}. -Derive Signature NoConfusionHom for All3. + +Inductive Forall3 {A B C : Type} (R : A -> B -> C -> Type) : list A -> list B -> list C -> Prop := + Forall3_nil : Forall3 R [] [] [] +| Forall3_cons : forall (x : A) (y : B) (z : C) (l : list A) (l' : list B) (l'' : list C), + R x y z -> Forall3 R l l' l'' -> Forall3 R (x :: l) (y :: l') (z :: l''). +Arguments Forall3_nil {_ _ _ _}. +Arguments Forall3_cons {_ _ _ _ _ _ _ _ _ _}. + +#[global] Hint Constructors All3 Forall3 : core. +Derive Signature for All3 Forall3. +Derive NoConfusionHom for All3. Definition invert_Forall2 {A B R l l'} (a : @Forall2 A B R l l') := match a in Forall2 _ l l' @@ -269,6 +279,18 @@ Section Forallb2. End Forallb2. +Section Forallb3. + Context {A B C} (f : A -> B -> C -> bool). + + Fixpoint forallb3 l l' l'' := + match l, l', l'' with + | hd :: tl, hd' :: tl', hd'' :: tl'' => f hd hd' hd'' && forallb3 tl tl' tl'' + | [], [], [] => true + | _, _, _ => false + end. + +End Forallb3. + Lemma forallb2_refl : forall A (R : A -> A -> bool), (forall x, R x x) -> @@ -428,10 +450,28 @@ Proof. rewrite andb_and. intuition auto. Qed. -Lemma All2P {A B : Type} {p : A -> B -> bool} {l l'} : - reflectT (All2 p l l') (forallb2 p l l'). +Lemma forallb3_All3 {A B C : Type} {p : A -> B -> C -> bool} + {l : list A} {l' : list B} {l'' : list C}: + is_true (forallb3 p l l' l'') -> All3 (fun x y z => is_true (p x y z)) l l' l''. +Proof. + induction l in l', l'' |- *; destruct l', l''; simpl; intros; try congruence. + - constructor. + - constructor. revert H; rewrite andb_and; intros [px pl]. auto. + apply IHl. revert H; rewrite andb_and; intros [px pl]. auto. +Qed. + +Lemma All3_forallb3 {A B C : Type} {p : A -> B -> C -> bool} + {l : list A} {l' : list B} {l'' : list C} : + All3 (fun x y z => is_true (p x y z)) l l' l'' -> is_true (forallb3 p l l' l''). Proof. - apply equiv_reflectT. apply All2_forallb2. apply forallb2_All2. + induction 1; simpl; intros; try congruence. + rewrite andb_and. intuition auto. +Qed. + +Lemma All3P {A B C : Type} {p : A -> B -> C -> bool} {l l' l''} : + reflectT (All3 p l l' l'') (forallb3 p l l' l''). +Proof. + apply equiv_reflectT. apply All3_forallb3. apply forallb3_All3. Qed. Lemma All2_refl {A} {P : A -> A -> Type} l : @@ -715,6 +755,18 @@ Proof. eapply All2_All_right; tea. auto. Qed. +Lemma All2P {A B : Type} {P : A -> B -> Type} {p : A -> B -> bool} {l l'} : + (forall x y, reflectT (P x y) (p x y)) -> + reflectT (All2 P l l') (forallb2 p l l'). +Proof. + intro H. + apply equiv_reflectT; intro X. + - eapply All2_forallb2, All2_impl with (1 := X). + move => x y /H //. + - apply forallb2_All2 in X. + eapply All2_impl with (1 := X). + move => x y /H //. +Qed. Lemma All2i_All_left {A B} {P : nat -> A -> B -> Type} {Q : A -> Type} {n l l'} : All2i P n l l' -> @@ -883,6 +935,12 @@ Proof. induction 1; simpl; constructor; auto. Qed. +Lemma Alli_refl {A} (P : nat -> A -> Type) n (l : list A) : + (forall n x, P n x) -> Alli P n l. +Proof. + intros H. induction l in n |- *; constructor; auto. +Defined. + Lemma Alli_rev {A} {P : nat -> A -> Type} k l : Alli P k l -> Alli (fun k' => P (Nat.pred #|l| - k' + k)) 0 (List.rev l). @@ -962,6 +1020,12 @@ Proof. intros H; induction 1; constructor; try inv H; intuition. Qed. +Lemma OnOne2_All_mix_both {A} {P : A -> A -> Type} {Q R : A -> Type} {l l'} : + All Q l -> All R l' -> OnOne2 P l l' -> OnOne2 (fun x y => (P x y × Q x × R y)%type) l l'. +Proof. + intros H H'; induction 1; constructor; try inv H; try inv H'; intuition. +Qed. + Lemma OnOne2_app {A} (P : A -> A -> Type) l tl tl' : OnOne2 P tl tl' -> OnOne2 P (l ++ tl) (l ++ tl'). Proof. induction l; simpl; try constructor; auto. Qed. @@ -1636,6 +1700,49 @@ Proof. induction 1; split; constructor; intuition eauto. Qed. +Lemma All_reflect_reflect_All2 {A B} R (P : A -> B -> Type) (p : A -> B -> bool) l l' : + All (fun x => forall y, R y -> reflectT (P x y) (p x y)) l -> + All R l' -> + reflectT (All2 P l l') (forallb2 p l l'). +Proof. + intros X' XP. + apply equiv_reflectT; intro X. + - apply All2_All_mix_left with (1 := X'), All2_All_mix_right with (1 := XP) in X. + eapply All2_forallb2, All2_impl with (1 := X). + move => x y [] [] H /H //. + - apply forallb2_All2, All2_All_mix_left with (1 := X'), All2_All_mix_right with (1 := XP) in X. + eapply All2_impl with (1 := X). + move => x y [] [] H /H //. +Qed. + +Lemma All2_reflect_reflect_All2 {A B} (P : A -> B -> Type) (p : A -> B -> bool) l l' : + All2 (fun x y => reflectT (P x y) (p x y)) l l' -> + reflectT (All2 P l l') (forallb2 p l l'). +Proof. + intro X'. + apply equiv_reflectT; intro X. + - apply All2_All2_mix with (1 := X') in X. + eapply All2_forallb2, All2_impl with (1 := X). + move => x y [] H /H //. + - apply forallb2_All2, All2_All2_mix with (1 := X') in X. + eapply All2_impl with (1 := X). + move => x y [] H /H //. +Qed. + +Lemma All3_Forall3 {A B C} {P : A -> B -> C -> Prop} {l l' l''} : + All3 P l l' l'' -> Forall3 P l l' l''. +Proof. + induction 1; auto. +Qed. + +(* Bad! Uses template polymorphism. *) +Lemma Forall3_All3 {A B C} {P : A -> B -> C -> Prop} l l' l'' : Forall3 P l l' l'' -> All3 P l l' l''. +Proof. + intros f; induction l in l', l'', f |- *; destruct l', l''; try constructor; auto. + 1-6: elimtype False; inv f. + inv f; auto. + apply IHl. inv f; auto. +Qed. Ltac toAll := match goal with @@ -1651,10 +1758,18 @@ Ltac toAll := | |- Forall2 _ _ _ => apply All2_Forall2 + | H : Forall3 _ _ _ |- _ => apply Forall3_All3 in H + + | |- Forall3 _ _ _ => apply All3_Forall3 + | H : is_true (forallb2 _ _ _) |- _ => apply forallb2_All2 in H | |- is_true (forallb2 _ _ _) => apply All2_forallb2 + | H : is_true (forallb3 _ _ _) |- _ => apply forallb3_All3 in H + + | |- is_true (forallb3 _ _ _) => apply All3_forallb3 + | [ H : All2_dep (fun x y _ => @?R' x y) ?a |- _ ] => apply (@All2_undep _ _ _ R' _ _ a) in H | [ |- All2_dep (fun x y _ => @?R' x y) ?a ] => apply (@All2_undep _ _ _ R' _ _ a) @@ -2088,6 +2203,9 @@ Qed. Lemma Forall2_length {A B} {P : A -> B -> Prop} {l l'} : Forall2 P l l' -> #|l| = #|l'|. Proof. induction 1; simpl; auto. Qed. +Lemma Forall2_triv {A B} {l : list A} {l' : list B} : #|l| = #|l'| -> Forall2 (fun _ _ => True) l l'. +Proof. induction l in l' |- *; destruct l' => //=. auto. Qed. + Lemma Forall2_app_r {A} (P : A -> A -> Prop) l l' r r' : Forall2 P (l ++ [r]) (l' ++ [r']) -> (Forall2 P l l') /\ (P r r'). Proof. @@ -2791,6 +2909,14 @@ Proof. - constructor ; eauto. Qed. +Lemma Forall2_symP {A} (P : A -> A -> Prop) : + RelationClasses.Symmetric P -> + RelationClasses.Symmetric (Forall2 P). +Proof. + intros h l l' hl. + induction hl. all: auto. +Qed. + Lemma All_All2_All2_mix {A B} (P : B -> B -> Type) (Q R : A -> B -> Type) l l' l'' : All (fun x => forall y z, Q x y -> R x z -> ∑ v, P y v * P z v) l -> All2 Q l l' -> All2 R l l'' -> ∑ l''', All2 P l' l''' * All2 P l'' l'''. @@ -3153,6 +3279,33 @@ Proof. constructor; auto. now destruct (Hp a b). Qed. +Lemma forallb3_Forall3 : + forall A B C (p : A -> B -> C -> bool) l l' l'', + forallb3 p l l' l'' -> + Forall3 (fun x y z => p x y z) l l' l''. +Proof. + intros A B C p l l' l'' h. + induction l in l', l'', h |- *. + - destruct l', l''; try discriminate. + constructor. + - destruct l', l''; try discriminate. + simpl in h. move/andb_and: h => [? ?]. + constructor. all: auto. +Qed. + +Lemma forallb3P {A B C} (P : A -> B -> C -> Prop) (p : A -> B -> C -> bool) l l' l'' : + (forall x y z, reflect (P x y z) (p x y z)) -> + reflect (Forall3 P l l' l'') (forallb3 p l l' l''). +Proof. + intros Hp. + apply iff_reflect; change (forallb3 p l l' l'' = true) with (forallb3 p l l' l'' : Prop); split. + - induction 1; rewrite /= // IHForall3 // andb_true_r. + now destruct (Hp x y z). + - induction l in l', l'' |- *; destruct l', l''; rewrite /= //. rewrite andb_and. + intros [pa pl]. + constructor; auto. now destruct (Hp a b c). +Qed. + (** All, All2 and In interactions. *) Lemma All2_In {A B} (P : A -> B -> Type) l l' x : In x l -> @@ -3333,6 +3486,14 @@ Proof. congruence. Qed. +Lemma All2i_All2 {A B} {P : nat -> A -> B -> Type} {Q : A -> B -> Type} n l l' : + All2i P n l l' -> + (forall i x y, P i x y -> Q x y) -> + All2 Q l l'. +Proof. + induction 1; constructor; eauto. +Qed. + Lemma All2i_All2_All2i_All2i {A B C n} {P : nat -> A -> B -> Type} {Q : B -> C -> Type} {R : nat -> A -> C -> Type} {S : nat -> A -> C -> Type} {l l' l''} : All2i P n l l' -> @@ -3559,6 +3720,14 @@ Section All2_fold. induction H; constructor; auto. Qed. + Lemma All2_fold_flip Γ Δ : + All2_fold P Γ Δ -> + All2_fold (fun Γ Γ' x y => P Γ' Γ y x) Δ Γ. + Proof using Type. + intros H. + induction H; constructor; auto. + Qed. + Lemma All2_fold_app_inv Γ Γ' Δ Δ' : #|Δ| = #|Δ'| -> All2_fold P (Δ ++ Γ) (Δ' ++ Γ') -> @@ -3852,6 +4021,90 @@ Proof. induction 1; constructor; auto. Qed. +Lemma Forall3_impl {A B C} {P Q : A -> B -> C -> Prop} {l l' l''} : + Forall3 P l l' l'' -> + (forall x y z, P x y z -> Q x y z) -> + Forall3 Q l l' l''. +Proof. + induction 1; constructor; auto. +Qed. + +Lemma Forall3_Forall2_left {A B C} {P : A -> B -> C -> Prop} {Q : A -> B -> Prop} {l l' l''} : + Forall3 P l l' l'' -> + (forall x y z, P x y z -> Q x y) -> + Forall2 Q l l'. +Proof. + induction 1; constructor; eauto. +Qed. + +Lemma Forall3_Forall2_edges {A B C} {P : A -> B -> C -> Prop} {Q : A -> C -> Prop} {l l' l''} : + Forall3 P l l' l'' -> + (forall x y z, P x y z -> Q x z) -> + Forall2 Q l l''. +Proof. + induction 1; constructor; eauto. +Qed. + +Lemma Forall3_Forall2_right {A B C} {P : A -> B -> C -> Prop} {Q : B -> C -> Prop} {l l' l''} : + Forall3 P l l' l'' -> + (forall x y z, P x y z -> Q y z) -> + Forall2 Q l' l''. +Proof. + induction 1; constructor; eauto. +Qed. + +Lemma Forall2_Forall2_Forall3 {A B C} {P : A -> B -> Prop} {Q : B -> C -> Prop} {l l' l''} : + Forall2 P l l' -> + Forall2 Q l' l'' -> + Forall3 (fun x y z => P x y /\ Q y z) l l' l''. +Proof. + induction 1 in l'' |- *; intro HQ; inv HQ; constructor; eauto. +Qed. + +Lemma Forall3_symP {A B} (P : B -> A -> A -> Prop) : + (forall b, RelationClasses.Symmetric (P b)) -> + forall l, RelationClasses.Symmetric (Forall3 P l). +Proof. + intros h l l' l'' hl. + induction hl; constructor; eauto. + now symmetry. +Qed. + +Lemma Forall3_transP {A B} (P : B -> A -> A -> Prop) : + (forall b, RelationClasses.Transitive (P b)) -> + forall l, RelationClasses.Transitive (Forall3 P l). +Proof. + intros h l0 l l' l'' hl hl'. + induction hl in l'', hl' |- *; inv hl'; econstructor; eauto. + now etransitivity. +Qed. + +Lemma Forall3_antisymP {A B} (P P' : B -> A -> A -> Prop) : + (forall b x y, P b x y -> P b y x -> P' b x y) -> + forall l l' l'', Forall3 P l l' l'' -> Forall3 P l l'' l' -> Forall3 P' l l' l''. +Proof. + intros h l l' l'' hl hl'. + induction hl in hl' |- *; inv hl'; constructor; eauto. +Qed. + +Lemma Forall3_map_inv {A B C A' B' C'} (R : A' -> B' -> C' -> Prop) (f : A -> A') (g : B -> B') (h : C -> C') l l' l'' : + Forall3 R (map f l) (map g l') (map h l'') -> + Forall3 (fun x y z => R (f x) (g y) (h z)) l l' l''. +Proof. + induction l in l', l'' |- * ; destruct l', l''; try solve [ inversion 1 ]. + - constructor. + - constructor. + + inversion H. subst. assumption. + + eapply IHl. inversion H. assumption. +Qed. + +Lemma Forall3_map {A B C A' B' C'} (R : A' -> B' -> C' -> Prop) (f : A -> A') (g : B -> B') (h : C -> C') l l' l'' : + Forall3 (fun x y z => R (f x) (g y) (h z)) l l' l'' -> + Forall3 R (map f l) (map g l') (map h l''). +Proof. + induction 1; constructor; auto. +Qed. + Lemma map2_app {A B C} (f : A -> B -> C) l0 l0' l1 l1' : #|l0| = #|l1| -> #|l0'| = #|l1'| -> map2 f (l0 ++ l0') (l1 ++ l1') = diff --git a/utils/theories/MCList.v b/utils/theories/MCList.v index f7a238fe7..125f70869 100644 --- a/utils/theories/MCList.v +++ b/utils/theories/MCList.v @@ -198,6 +198,11 @@ Lemma map_map_compose : map g (map f l) = map (fun x => g (f x)) l. Proof. apply map_map. Qed. +Lemma map_cst_repeat {A B} (f : B) (l: list A) : map (fun _ => f) l = repeat f #|l|. +Proof. + unfold map. induction l; cbn; auto. now rewrite IHl. +Qed. + Lemma map_id_f {A} (l : list A) (f : A -> A) : (forall x, f x = x) -> map f l = l. diff --git a/utils/theories/MCReflect.v b/utils/theories/MCReflect.v index 86b1b4040..9ec78f43c 100644 --- a/utils/theories/MCReflect.v +++ b/utils/theories/MCReflect.v @@ -35,6 +35,12 @@ Proof. intros [] => //. Qed. Hint View for move/ introT|2. +Lemma reflectT_change_left P Q p : + CRelationClasses.iffT P Q -> reflectT P p -> reflectT Q p. +Proof. + intros [] []; constructor; auto. +Qed. + Lemma reflectT_subrelation {A} {R} {r : A -> A -> bool} : (forall x y, reflectT (R x y) (r x y)) -> CRelationClasses.subrelation R r. Proof. intros. intros x y h. destruct (X x y); auto. diff --git a/utils/theories/MCRelations.v b/utils/theories/MCRelations.v index 59bf5552f..a2fe36e2b 100644 --- a/utils/theories/MCRelations.v +++ b/utils/theories/MCRelations.v @@ -2,6 +2,7 @@ Require Import ssreflect. Require Import Equations.Type.Relation Equations.Type.Relation_Properties. Require Import CRelationClasses. +Require RelationClasses. #[global] Hint Mode Reflexive ! ! : typeclass_instances. @@ -34,6 +35,60 @@ Notation Trel_sig R S := Notation on_Trel_eq R f g := (fun x y => (R (f x) (f y) * (g x = g y)))%type. +#[global] +Instance on_rel_refl {A B} (P : B -> B -> Prop) (f : A -> B) : + RelationClasses.Reflexive P -> + RelationClasses.Reflexive (on_rel P f). +Proof. + intros refl u. + apply refl. +Qed. + +#[global] +Instance on_rel_sym {A B} (P : B -> B -> Prop) (f : A -> B) : + RelationClasses.Symmetric P -> + RelationClasses.Symmetric (on_rel P f). +Proof. + intros sym u u'. + apply sym. +Qed. + +#[global] +Instance on_rel_trans {A B} (P : B -> B -> Prop) (f : A -> B) : + RelationClasses.Transitive P -> + RelationClasses.Transitive (on_rel P f). +Proof. + intros trans u u' u''. + apply trans. +Qed. + +#[global] +Instance on_Trel_refl {A B} (P : B -> B -> Type) (f : A -> B) : + Reflexive P -> + Reflexive (on_Trel P f). +Proof. + intros refl u. + apply refl. +Qed. + +#[global] +Instance on_Trel_sym {A B} (P : B -> B -> Type) (f : A -> B) : + Symmetric P -> + Symmetric (on_Trel P f). +Proof. + intros sym u u'. + apply sym. +Qed. + +#[global] +Instance on_Trel_trans {A B} (P : B -> B -> Type) (f : A -> B) : + Transitive P -> + Transitive (on_Trel P f). +Proof. + intros trans u u' u''. + apply trans. +Qed. + Section Flip. Local Set Universe Polymorphism. Context {A : Type} (R : crelation A). diff --git a/utils/theories/utils.v b/utils/theories/utils.v index 6ea1c27b3..494979635 100644 --- a/utils/theories/utils.v +++ b/utils/theories/utils.v @@ -30,3 +30,5 @@ Notation "A * B" := (prod A B) : type_scope2. Global Open Scope type_scope2. Global Open Scope metacoq_scope. + +#[global] Hint Constructors unit : core.