diff --git a/common/theories/BasicAst.v b/common/theories/BasicAst.v index b3908a723..daaecaffe 100644 --- a/common/theories/BasicAst.v +++ b/common/theories/BasicAst.v @@ -138,6 +138,10 @@ Lemma map_dbody {A B} (f : A -> B) (g : A -> B) (d : def A) : g (dbody d) = dbody (map_def f g d). Proof. destruct d; reflexivity. Qed. +Lemma map_dname {A B} (f : A -> B) (g : A -> B) (d : def A) : + dname d = dname (map_def f g d). +Proof. destruct d; reflexivity. Qed. + Definition mfixpoint term := list (def term). Definition test_def {A} (tyf bodyf : A -> bool) (d : def A) := @@ -217,13 +221,13 @@ Record judgment_ {universe Term} := Judge { j_term : option Term; j_typ : Term; j_univ : option universe; - (* j_rel : option relevance; *) + 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) *). + Judge (option_map f (j_term j)) (f (j_typ j)) (j_univ j) (j_rel j). Section Contexts. Context {term : Type}. @@ -239,17 +243,21 @@ 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) *)) (only parsing). -Notation j_vass_s na ty s := (TypUniv ty s (* na.(binder_relevance) *)) (only parsing). -Notation j_vdef na b ty := (TermTyp b ty (* na.(binder_relevance) *)) (only parsing). -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) *)). +Notation Typ typ := (Judge None typ None None). +Notation TypRel typ rel := (Judge None typ None (Some rel)). +Notation TermTyp tm ty := (Judge (Some tm) ty None None). +Notation TermTypRel tm ty rel := (Judge (Some tm) ty None (Some rel)). +Notation TermoptTyp tm typ := (Judge tm typ None None). +Notation TermoptTypRel tm typ rel := (Judge tm typ None (Some rel)). +Notation TypUniv ty u := (Judge None ty (Some u) None). +Notation TypUnivRel ty u rel := (Judge None ty (Some u) (Some rel)). +Notation TermTypUniv tm ty u := (Judge (Some tm) ty (Some u) None). + +Notation j_vass na ty := (TypRel ty na.(binder_relevance)). +Notation j_vass_s na ty s := (TypUnivRel ty s na.(binder_relevance)). +Notation j_vdef na b ty := (TermTypRel b ty na.(binder_relevance)). +Notation j_decl d := (TermoptTypRel (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 (Some (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); diff --git a/common/theories/Environment.v b/common/theories/Environment.v index 679511cc7..06e81f7af 100644 --- a/common/theories/Environment.v +++ b/common/theories/Environment.v @@ -149,6 +149,8 @@ Module Environment (T : Term). (** Local (de Bruijn) context *) Definition context := list context_decl. + Definition mark_context := list relevance. + Definition marks_of_context : context -> mark_context := fun l => List.map (fun d => d.(decl_name).(binder_relevance)) l. (** Last declaration first *) @@ -850,7 +852,7 @@ Module Environment (T : Term). end. Definition tImpl (dom codom : term) : term := - tProd {| binder_name := nAnon; binder_relevance := Relevant |} + tProd {| binder_name := nAnon; binder_relevance := rel_of_Type |} dom (lift 1 0 codom). Definition array_uctx := ([nAnon], ConstraintSet.empty). @@ -998,12 +1000,30 @@ Module Environment (T : Term). Definition arities_context (l : list one_inductive_body) := rev_map (fun ind => vass (mkBindAnn (nNamed ind.(ind_name)) - (ind.(ind_relevance))) ind.(ind_type)) l. + rel_of_Type) ind.(ind_type)) l. Lemma arities_context_length l : #|arities_context l| = #|l|. Proof. unfold arities_context. now rewrite rev_map_length. Qed. #[global] Hint Rewrite arities_context_length : len. + Lemma nth_error_arities_context idecls i idecl : + nth_error (List.rev idecls) i = Some idecl -> + nth_error (arities_context idecls) i = + Some {| decl_name := {| binder_name := nNamed idecl.(ind_name); binder_relevance := rel_of_Type |}; + decl_body := None; + decl_type := idecl.(ind_type) |}. + Proof using Type. + intros hnth. + epose proof (nth_error_Some_length hnth). autorewrite with len in H. + rewrite nth_error_rev in hnth. now autorewrite with len. + rewrite List.rev_involutive in hnth. autorewrite with len in hnth. + unfold arities_context. + rewrite rev_map_spec. + rewrite nth_error_rev; autorewrite with len; auto. + rewrite List.rev_involutive nth_error_map. + rewrite hnth. simpl. reflexivity. + 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 => diff --git a/common/theories/EnvironmentTyping.v b/common/theories/EnvironmentTyping.v index ff7df2e95..2206b97a9 100644 --- a/common/theories/EnvironmentTyping.v +++ b/common/theories/EnvironmentTyping.v @@ -325,20 +325,20 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). (** Well-formedness of local environments embeds a sorting for each variable *) Notation on_local_decl P Γ d := - (P Γ (j_decl d)). + (P Γ (j_decl d)) (only parsing). Definition on_def_type (P : context -> judgment -> Type) Γ d := - P Γ (Typ d.(dtype)). + P Γ (TypRel d.(dtype) d.(dname).(binder_relevance)). Definition on_def_body (P : context -> judgment -> Type) types Γ d := - P (Γ ,,, types) (TermTyp d.(dbody) (lift0 #|types| d.(dtype))). + P (Γ ,,, types) (TermTypRel d.(dbody) (lift0 #|types| d.(dtype)) d.(dname).(binder_relevance)). (* 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_wfu_term wf_term wf_univ (j : judgment) := option_default wf_term (j_term j) (unit : Type) × wf_term (j_typ j) × option_default wf_univ (j_univ j) (unit : Type). 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 Γ)). @@ -347,7 +347,9 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). 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. + ∑ s, sorting (j_typ j) s × + option_default (fun u => u = s) (j_univ j) True /\ + isSortRelOpt s (j_rel j). Notation lift_sorting1 checking sorting := (fun Γ => lift_sorting (checking Γ) (sorting Γ)). Notation lift_sorting2 checking sorting := (fun Σ Γ => lift_sorting (checking Σ Γ) (sorting Σ Γ)). @@ -365,22 +367,22 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). 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), + Lemma lift_wf_term_it_impl {P Q} {tm tm' : option term} {t t' : term} {u u' r r'} : + forall tu: lift_wf_term P (Judge tm t u r), 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'). + lift_wf_term Q (Judge tm' t' u' r'). 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' : + Lemma lift_wf_term_f_impl P Q tm t u u' r r' : forall f, - lift_wf_term P (Judge tm t u) -> + lift_wf_term P (Judge tm t u r) -> (forall t, P t -> Q (f t)) -> - lift_wf_term Q (Judge (option_map f tm) (f t) u'). + lift_wf_term Q (Judge (option_map f tm) (f t) u' r'). Proof. unfold lift_wf_term; cbn. intros f (Hb & Ht) HPQ. @@ -399,12 +401,12 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). destruct j_term; cbn in *; auto. Defined. - Lemma lift_wfbu_term_f_impl (P Q : term -> bool) tm t u : + Lemma lift_wfbu_term_f_impl (P Q : term -> bool) tm t u r : forall f fu, - lift_wfbu_term P (P ∘ tSort) (Judge tm t u) -> + lift_wfbu_term P (P ∘ tSort) (Judge tm t u r) -> (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)). + lift_wfbu_term Q (Q ∘ tSort) (Judge (option_map f tm) (f t) (option_map fu u) r). Proof. unfold lift_wfbu_term; cbn. intros. rtoProp. @@ -413,58 +415,85 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). destruct u; rewrite //= -H0 //. auto. Defined. - Lemma unlift_TermTyp {Pc Ps tm ty u} : - lift_sorting Pc Ps (Judge (Some tm) ty u) -> + Lemma lift_wf_wfb_term (p : _ -> bool) j : + reflectT (lift_wf_term p j) (lift_wfb_term p j). + Proof. + rewrite /lift_wf_term /lift_wfb_term. + set (HP := @reflectT_pred _ p). + destruct (HP (j_typ j)) => //; + destruct (reflect_option_default HP (j_term j)) => /= //; now constructor. + Qed. + + Lemma lift_wfu_wfbu_term (p : _ -> bool) (pu : _ -> bool) j : + reflectT (lift_wfu_term p pu j) (lift_wfbu_term p pu j). + Proof. + set (HP := @reflectT_pred _ p); set (HPu := @reflectT_pred _ pu). + rewrite /lift_wfu_term /lift_wfbu_term. + destruct (HP (j_typ j)) => //; + destruct (reflect_option_default HP (j_term j)) => /= //; + destruct (reflect_option_default HPu (j_univ j)) => /= //; now constructor. + Qed. + + Lemma unlift_TermTyp {Pc Ps tm ty u r} : + lift_sorting Pc Ps (Judge (Some tm) ty u r) -> Pc tm ty. Proof. apply fst. Defined. - Definition unlift_TypUniv {Pc Ps tm ty u} : - lift_sorting Pc Ps (Judge tm ty (Some u)) -> + Definition unlift_TypUniv {Pc Ps tm ty u r} : + lift_sorting Pc Ps (Judge tm ty (Some u) r) -> Ps ty u - := fun H => eq_rect_r _ H.2.π2.1 H.2.π2.2. + := fun H => eq_rect_r _ H.2.π2.1 H.2.π2.2.p1. - 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)). + Definition lift_sorting_extract {c s tm ty r} (w : lift_sorting c s (Judge tm ty None r)) : + lift_sorting c s (Judge tm ty (Some w.2.π1) r) + := (w.1, (w.2.π1; (w.2.π2.1, (conj eq_refl w.2.π2.2.p2)))). + + Lemma lift_sorting_forget_univ {Pc Ps tm ty u r} : + lift_sorting Pc Ps (Judge tm ty u r) -> + lift_sorting Pc Ps (Judge tm ty None r). + Proof. + intros (? & ? & ? & ? & ?). + repeat (eexists; tea). + Qed. - 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). + Lemma lift_sorting_forget_body {Pc Ps tm ty u r} : + lift_sorting Pc Ps (Judge tm ty u r) -> + lift_sorting Pc Ps (Judge None ty u r). Proof. - intros (? & ? & ? & ?). + 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). + Lemma lift_sorting_forget_rel {Pc Ps tm ty u r} : + lift_sorting Pc Ps (Judge tm ty u r) -> + lift_sorting Pc Ps (Judge tm ty u None). Proof. - intros (? & ? & ? & ?). + 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), + Lemma lift_sorting_ex_it_impl_gen {Pc Qc Ps Qs} {tm tm' : option term} {t t' : term} {r r' : option relevance} : + forall tu: lift_sorting Pc Ps (Judge tm t None r), 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'). + (Ps t s -> isSortRelOpt s r -> ∑ s', Qs t' s' × isSortRelOpt s' r') -> + lift_sorting Qc Qs (Judge tm' t' None r'). Proof. - intros (? & ? & Hs & e) s HPQc HPQs. + intros (? & ? & Hs & e & er) s HPQc HPQs. split. - destruct tm, tm' => //=. now apply HPQc. - - specialize (HPQs Hs) as (s' & H). - eexists. split; eassumption. + - specialize (HPQs Hs er) as (s' & H & He). + eexists. repeat 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), + Lemma lift_sorting_it_impl_gen {Pc Qc Ps Qs} {tm tm' : option term} {t t' : term} {u r} : + forall tu: lift_sorting Pc Ps (Judge tm t u r), 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). + lift_sorting Qc Qs (Judge tm' t' u r). Proof. intros (? & ? & Hs & e) s HPQc HPQs. split. @@ -473,19 +502,23 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). 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), + Lemma lift_sorting_fu_it_impl {Pc Qc Ps Qs} {tm : option term} {t : term} {u r} : + forall f fu, forall tu: lift_sorting Pc Ps (Judge tm t u r), let s := tu.2.π1 in + option_default (fun rel => isSortRel s rel -> isSortRel (fu s) rel) r True -> 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)). + lift_sorting Qc Qs (Judge (option_map f tm) (f t) (option_map fu u) r). Proof. - intros ?? (? & ? & Hs & e) s HPQc HPQs. + intros ?? (? & ? & Hs & e & er) s Hr HPQc HPQs. split. - destruct tm => //=. now apply HPQc. - eexists. split; [now apply HPQs|]. - destruct u => //. - cbn. f_equal => //. + split. + + destruct u => //. + cbn. f_equal => //. + + destruct r => //. + now apply Hr. Qed. Lemma lift_sorting_f_it_impl {Pc Qc Ps Qs} {j : judgment} : @@ -496,9 +529,12 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). 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))). + replace (judgment_map f j) with (Judge (option_map f (j_term j)) (f (j_typ j)) (option_map id (j_univ j)) (j_rel j)). 2: unfold judgment_map; destruct j_univ => //. - apply lift_sorting_fu_it_impl with (fu := id). + intro tu. + apply lift_sorting_fu_it_impl with (fu := id) (tu := tu). + destruct tu as (? & s & ?); cbn; clear. + destruct j_rel => //. Qed. Lemma lift_sorting_it_impl {Pc Qc Ps Qs} {j} : @@ -513,26 +549,29 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). destruct j, j_term => //. Qed. - Lemma lift_sorting_fu_impl {Pc Qc Ps Qs tm t u} : + Lemma lift_sorting_fu_impl {Pc Qc Ps Qs tm t u r} : forall f fu, - lift_sorting Pc Ps (Judge tm t u) -> + lift_sorting Pc Ps (Judge tm t u r) -> + (forall r s, isSortRelOpt s r -> isSortRelOpt (fu s) r) -> (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)). + lift_sorting Qc Qs (Judge (option_map f tm) (f t) (option_map fu u) r). Proof. - intros ?? tu ??. + intros ?? tu Hr ??. apply lift_sorting_fu_it_impl with (tu := tu); auto. + 1: destruct r => //=; apply Hr with (r := Some _). destruct tm => //=. auto. Qed. - Lemma lift_typing_fu_impl {P Q tm t u} : + Lemma lift_typing_fu_impl {P Q tm t u r} : forall f fu, - lift_typing0 P (Judge tm t u) -> + lift_typing0 P (Judge tm t u r) -> (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)). + (forall r s, isSortRelOpt s r -> isSortRelOpt (fu s) r) -> + lift_typing0 Q (Judge (option_map f tm) (f t) (option_map fu u) r). Proof. - intros ?? HT HPQ Hf. + intros ?? HT HPQ Hf Hr. apply lift_sorting_fu_impl with (1 := HT); tas. intros; rewrite -Hf; now apply HPQ. Qed. @@ -570,10 +609,11 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). 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) -> + Lemma lift_typing_mapu {P} f fu {tm ty u r} : + lift_typing0 (fun t T => P (f t) (f T)) (Judge tm ty u r) -> (forall u, f (tSort u) = tSort (fu u)) -> - lift_typing0 P (Judge (option_map f tm) (f ty) (option_map fu u)). + (forall r s, isSortRelOpt s r -> isSortRelOpt (fu s) r) -> + lift_typing0 P (Judge (option_map f tm) (f ty) (option_map fu u) r). Proof. intros HT. eapply lift_typing_fu_impl with (1 := HT) => //. @@ -600,6 +640,30 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). intros; now apply HPQ. Qed. + Lemma lift_typing_subject {P Q j} : + lift_typing0 P j -> + (forall t T, P t T -> Q t) -> + lift_wf_term Q j. + Proof. + intros (? & ? & ? & _) HPQ. + split; eauto. + destruct j_term => //=. + eauto. + Qed. + + Lemma lift_typing_subjtyp {P Q Q' j} : + lift_typing0 P j -> + (forall t T, P t T -> Q t × Q T) -> + (forall u, Q (tSort u) -> Q' u) -> + lift_wfu_term Q Q' j. + Proof. + intros (Htm & s & Hty & e & er) HPQ HQQ'. + repeat split. + - destruct j_term => //=. eapply fst, HPQ, Htm. + - eapply fst, HPQ, Hty. + - destruct j_univ => //=. rewrite e. eapply HQQ', snd, HPQ, Hty. + Qed. + Section TypeLocal. Context (typing : forall (Γ : context), judgment -> Type). @@ -625,7 +689,7 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). 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) + All_local_env typing Γ -> typing Γ (TermoptTypRel bo t na.(binder_relevance)) -> 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) := @@ -647,6 +711,16 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). now apply X0. Defined. + Lemma All_local_env_All_fold P Γ : + All_local_env P Γ <~> + All_fold (fun Γ decl => P Γ (j_decl decl)) Γ. + Proof using Type. + split. + - induction 1 using All_local_env_ind1; constructor; auto. + - induction 1; [constructor|]. + destruct d as [na [b|] ty]; cbn in p; constructor; auto. + Qed. + 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) Γ). @@ -1182,7 +1256,7 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT match Δ with | [] => wf_sort Σ u | {| decl_name := na; decl_body := None; decl_type := t |} :: Δ => - type_local_ctx Σ Γ Δ u × P Σ (Γ ,,, Δ) (TypUniv t u (* na.(binder_relevance) *)) + type_local_ctx Σ Γ Δ u × P Σ (Γ ,,, Δ) (TypUnivRel t u na.(binder_relevance)) | {| decl_body := Some _; |} as d :: Δ => type_local_ctx Σ Γ Δ u × P Σ (Γ ,,, Δ) (j_decl d) end. @@ -1191,7 +1265,7 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT match Δ, us with | [], [] => unit | {| decl_name := na; decl_body := None; decl_type := t |} :: Δ, u :: us => - sorts_local_ctx Σ Γ Δ us × P Σ (Γ ,,, Δ) (TypUniv t u (* na.(binder_relevance) *)) + sorts_local_ctx Σ Γ Δ us × P Σ (Γ ,,, Δ) (TypUnivRel t u na.(binder_relevance)) | {| decl_body := Some _ |} as d :: Δ, us => sorts_local_ctx Σ Γ Δ us × P Σ (Γ ,,, Δ) (j_decl d) | _, _ => False @@ -1200,6 +1274,7 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT Implicit Types (mdecl : mutual_inductive_body) (idecl : one_inductive_body) (cdecl : constructor_body). Definition on_type Σ Γ T := P Σ Γ (Typ T). + Definition on_type_rel Σ Γ T r := P Σ Γ (TypRel T r). Open Scope type_scope. @@ -1441,7 +1516,7 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT arguments ending with a reference to the inductive applied to the (non-lets) parameters and arguments *) - on_ctype : on_type Σ (arities_context mdecl.(ind_bodies)) (cstr_type cdecl); + on_ctype : on_type_rel Σ (arities_context mdecl.(ind_bodies)) (cstr_type cdecl) idecl.(ind_relevance); on_cargs : sorts_local_ctx Σ (arities_context mdecl.(ind_bodies) ,,, mdecl.(ind_params)) cdecl.(cstr_args) cunivs; @@ -1578,7 +1653,7 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT (it_mkProd_or_LetIn idecl.(ind_indices) (tSort idecl.(ind_sort))); (** It must be well-typed in the empty context. *) - onArity : on_type Σ [] idecl.(ind_type); + onArity : on_type_rel Σ [] idecl.(ind_type) rel_of_Type; (** The sorts of the arguments contexts of each constructor *) ind_cunivs : list constructor_univs; @@ -1603,6 +1678,8 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT check_ind_sorts Σ mdecl.(ind_params) idecl.(ind_kelim) idecl.(ind_indices) ind_cunivs idecl.(ind_sort); + ind_relevance_compat : isSortRel idecl.(ind_sort) idecl.(ind_relevance); + onIndices : (* The inductive type respect the variance annotation on polymorphic universes, if any. *) match ind_variance mdecl with @@ -1643,7 +1720,7 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT (** *** Typing of constant declarations *) Definition on_constant_decl Σ d := - P Σ [] (TermoptTyp d.(cst_body) d.(cst_type)). + P Σ [] (TermoptTypRel d.(cst_body) d.(cst_type) d.(cst_relevance)). Definition on_global_decl Σ kn decl := match decl with @@ -1749,6 +1826,7 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT Arguments onConstructors {_ Pcmp P Σ mind mdecl i idecl}. Arguments onProjections {_ Pcmp P Σ mind mdecl i idecl}. Arguments ind_sorts {_ Pcmp P Σ mind mdecl i idecl}. + Arguments ind_relevance_compat {_ Pcmp P Σ mind mdecl i idecl}. Arguments onIndices {_ Pcmp P Σ mind mdecl i idecl}. Arguments onInductives {_ Pcmp P Σ mind mdecl}. @@ -2054,6 +2132,7 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT do 2 destruct indices_matter => //=. 2: now rewrite ?andb_false_r //= in Xcf. intro. eapply type_local_ctx_impl; eauto. + + apply (ind_relevance_compat X). + generalize (X.(onIndices)). destruct ind_variance => //. apply ind_respects_variance_impl. destruct variance_universes as [[[]]|] => //=. @@ -2180,9 +2259,18 @@ Module DeclarationTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E) Import T E L TU ET CT GM CS Ty. + Definition isTypeRel `{checker_flags} (Σ : global_env_ext) (Γ : context) (t : term) (rel : relevance) := + on_type_rel (lift_typing typing) Σ Γ t rel. + Definition isType `{checker_flags} (Σ : global_env_ext) (Γ : context) (t : term) := on_type (lift_typing typing) Σ Γ t. + Lemma isTypeRel_isType `{checker_flags} (Σ : global_env_ext) (Γ : context) (t : term) (r : relevance) : isTypeRel Σ Γ t r -> isType Σ Γ t. + Proof. apply lift_sorting_forget_rel. Defined. + #[export] Hint Resolve isTypeRel_isType : pcuic. + + Coercion isTypeRel_isType : isTypeRel >-> isType. + (** This predicate enforces that there exists typing derivations for every typable term in env. *) Definition Forall_decls_typing `{checker_flags} diff --git a/common/theories/MonadBasicAst.v b/common/theories/MonadBasicAst.v index b1cb40fbb..dad84cab3 100644 --- a/common/theories/MonadBasicAst.v +++ b/common/theories/MonadBasicAst.v @@ -24,7 +24,7 @@ Section with_monad. Definition monad_typ_or_sort_map {univ T' T''} (f: T' -> T T'') (t : judgment_ univ T') := match t with - | Judge tm ty u => ftm <- monad_option_map f tm;; fty <- f ty;; ret (Judge ftm fty u) + | Judge tm ty u rel => ftm <- monad_option_map f tm;; fty <- f ty;; ret (Judge ftm fty u rel) end. Definition monad_map_decl {term term'} (f : term -> T term') (d : context_decl term) : T (context_decl term') := diff --git a/common/theories/Universes.v b/common/theories/Universes.v index eca654b27..b920a79f4 100644 --- a/common/theories/Universes.v +++ b/common/theories/Universes.v @@ -2119,6 +2119,25 @@ Section SortCompare. Proof using Type. apply cmp_sort_subset with (pb := Cumul). Qed. End SortCompare. + + +Definition relevance_of_family (s : Sort.family) := + match s with + | Sort.fSProp => Irrelevant + | _ => Relevant + end. + +#[global] Opaque relevance_of_family. + +Notation rel_of_Type := (relevance_of_family Sort.fType). +Notation relevance_of_sort s := (relevance_of_family (Sort.to_family s)). + +Notation isSortRel s rel := (relevance_of_sort s = rel). +Notation isSortRelOpt s relopt := + (option_default (fun rel => isSortRel s rel) relopt True). + + + (** Elimination restriction *) @@ -2271,6 +2290,35 @@ Tactic Notation "unfold_univ_rel" "eqn" ":"ident(H) := Ltac cong := intuition congruence. +Lemma leq_relevance_eq {cf φ} {s s'} : + leq_sort φ s s' -> relevance_of_sort s = relevance_of_sort s'. +Proof. + now destruct s, s'. +Qed. + +Lemma leq_relevance_opt {cf φ} {s s' rel} : + leq_sort φ s s' -> isSortRelOpt s rel -> isSortRelOpt s' rel. +Proof. + now destruct s, s'. +Qed. + +Lemma leq_relevance {cf φ} {s s' rel} : + leq_sort φ s s' -> isSortRel s rel -> isSortRel s' rel. +Proof. + now destruct s, s'. +Qed. + +Lemma geq_relevance {cf φ} {s s' rel} : + leq_sort φ s' s -> isSortRel s rel -> isSortRel s' rel. +Proof. + now destruct s, s'. +Qed. + +Lemma relevance_super s : relevance_of_sort (Sort.super s) = rel_of_Type. +Proof using Type. + now destruct s. +Qed. + Lemma leq_sort_product_mon {cf} ϕ s1 s1' s2 s2' : leq_sort ϕ s1 s1' -> leq_sort ϕ s2 s2' -> @@ -2412,6 +2460,7 @@ Section no_prop_leq_type. now rewrite ps in H. Qed. + End no_prop_leq_type. @@ -2475,6 +2524,25 @@ Qed. #[global] Instance subst_instance_instance : UnivSubst Instance.t := fun u u' => List.map (subst_instance_level u) u'. + +Theorem relevance_subst_eq {cf} u s : relevance_of_sort (subst_instance_sort u s) = relevance_of_sort s. +Proof. + now destruct s. +Qed. + +Theorem relevance_subst_opt {cf} u rel s : + isSortRelOpt s rel -> isSortRelOpt (subst_instance_sort u s) rel. +Proof. + now destruct s. +Qed. + +Theorem relevance_subst {cf} u rel s : + isSortRel s rel -> isSortRel (subst_instance_sort u s) rel. +Proof. + now destruct s. +Qed. + + (** Tests that the term is closed over [k] universe variables *) Section Closedu. Context (k : nat). diff --git a/erasure/theories/EArities.v b/erasure/theories/EArities.v index f80d881f6..e84914ea9 100644 --- a/erasure/theories/EArities.v +++ b/erasure/theories/EArities.v @@ -77,6 +77,7 @@ Lemma isWfArity_prod_inv (Σ : global_env_ext) (Γ : context) (x : aname) (x0 x1 Proof. intros wfΣ (? & ? & ? & ?). cbn in e. eapply isType_tProd in i as [dom codom]; auto. + apply isTypeRel_isType in dom. split; auto. split; auto. clear dom codom. @@ -214,6 +215,7 @@ Proof. apply (X (subst_context [hd0] 0 Γ0) ltac:(len; reflexivity) _ _ sp). eapply isType_apply in i; tea. eapply (type_ws_cumul_pb (pb:=Conv)); tea. 2:now symmetry. + eapply isTypeRel_isType. now eapply isType_tProd in i as []. Qed. @@ -300,7 +302,7 @@ Proof. eapply invert_cumul_arity_r in c0; eauto. eapply typing_spine_strengthen in t0. 3:eauto. eapply wf_cofixpoint_spine in i0; eauto. - 2-3:eapply nth_error_all in a; eauto; simpl in a; eauto. + 2-3:eapply nth_error_all, isTypeRel_isType in a; eauto; simpl in a; eauto. destruct i0 as (Γ' & T & DA & ind & u & indargs & (eqT & ck) & cum). destruct (Nat.ltb #|x1| (context_assumptions Γ')). eapply invert_cumul_arity_r_gen in c0; eauto. diff --git a/erasure/theories/ESubstitution.v b/erasure/theories/ESubstitution.v index 6ab3417cd..152f48a3e 100644 --- a/erasure/theories/ESubstitution.v +++ b/erasure/theories/ESubstitution.v @@ -233,9 +233,9 @@ Proof. now apply All_local_env_app_inv in X3. - econstructor. + eauto. - + eapply H4; eauto. - + red in H6. - eapply Forall2_All2 in H6. + + eapply H5; eauto. + + red in H7. + eapply Forall2_All2 in H7. eapply All2i_All2_mix_left in X6; tea. clear H6. eapply All2i_nth_hyp in X6. @@ -489,15 +489,15 @@ Proof. + cbn. econstructor; auto. + econstructor. eapply is_type_subst; eauto. - - depelim H7. + - depelim H8. + cbn. econstructor. * eauto. - * eapply H4; eauto. + * eapply H5; eauto. * eapply All2_map. eapply All2_impl_In; eauto. - intros. destruct H11, x, y. cbn in e0. subst. split; eauto. - eapply In_nth_error in H9 as []. - move: H6. rewrite /wf_branches. + intros. destruct H12, x, y. cbn in e0. subst. split; eauto. + eapply In_nth_error in H10 as []. + move: H7. rewrite /wf_branches. move/Forall2_All2 => hbrs. eapply All2_nth_error_Some_r in hbrs; tea. set (br := {| bcontext := _ |}). @@ -514,10 +514,10 @@ Proof. move/(substitution_wf_local X8) => hwf. specialize (e0 _ _ _ t _ hwf X8). len in e0. cbn in e0. - have := PCUICCasesContexts.inst_case_branch_context_eq (p:=p) eqctx => H6. - rewrite /inst_case_branch_context /= in H6. + have := PCUICCasesContexts.inst_case_branch_context_eq (p:=p) eqctx => H7. + rewrite /inst_case_branch_context /= in H7. forward e0. - { move: e. cbn. rewrite /inst_case_branch_context /= -H6. + { move: e. cbn. rewrite /inst_case_branch_context /= -H7. now rewrite app_context_assoc. } forward e0. { now rewrite app_context_assoc. } diff --git a/erasure/theories/ErasureCorrectness.v b/erasure/theories/ErasureCorrectness.v index 4a12486ef..421aa5c04 100644 --- a/erasure/theories/ErasureCorrectness.v +++ b/erasure/theories/ErasureCorrectness.v @@ -60,7 +60,8 @@ 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 (? & h1 & ? & e0). + apply isTypeRel_isType in h1. assert (Σ ;;; [] |- a' : t). { eapply subject_reduction_eval; eauto. eapply PCUICConversion.ws_cumul_pb_Prod_Prod_inv in e0 as [? e1 e2]. @@ -827,7 +828,7 @@ Proof. eapply inversion_CoFix in t; destruct_sigma t; auto. pose proof (wfΣ' := wfΣ.1). eapply PCUICSpine.typing_spine_strengthen in t0; eauto. - 2:{ now eapply nth_error_all in a; tea. } + 2:{ now eapply nth_error_all, isTypeRel_isType in a; tea. } invs He. * edestruct IHeval1 as (? & ? & ?); eauto. now depelim Hed. depelim Hed. diff --git a/erasure/theories/ErasureProperties.v b/erasure/theories/ErasureProperties.v index 125ca26b8..3d6742a43 100644 --- a/erasure/theories/ErasureProperties.v +++ b/erasure/theories/ErasureProperties.v @@ -321,13 +321,13 @@ Proof. - cbn. econstructor. eapply H in X3; eauto. apply X3. cbn. econstructor. eauto. - eapply lift_typing_fu_impl with (1 := X0) => // ?? HT. + eapply lift_typing_fu_impl with (1 := X0) => // ?? HT; eauto using relevance_subst_opt. now apply typing_subst_instance. - cbn. econstructor. now edestruct X1; tea; eauto. eapply H in X3; eauto. exact X3. cbn. econstructor. eauto. - eapply lift_typing_fu_impl with (1 := X0) => // ?? HT. + eapply lift_typing_fu_impl with (1 := X0) => // ?? HT; eauto using relevance_subst_opt. now apply typing_subst_instance. - unfold subst_instance. cbn [subst_instance_constr]. econstructor; eauto. @@ -519,8 +519,8 @@ Section wellscoped. unfold lookup_constructor. rewrite (declared_constructor_lookup_gen isdecl) //. - unshelve eapply declared_inductive_to_gen in isdecl; eauto. unfold lookup_inductive. now rewrite (declared_inductive_lookup_gen isdecl). - - red in H8. eapply Forall2_All2 in H8. - eapply All2i_All2_mix_left in X4; tea. clear H8. + - red in H9. eapply Forall2_All2 in H9. + eapply All2i_All2_mix_left in X4; tea. clear H9. solve_all. - unshelve eapply declared_projection_to_gen in isdecl; eauto. unfold lookup_projection. now rewrite (declared_projection_lookup_gen isdecl). diff --git a/erasure/theories/Prelim.v b/erasure/theories/Prelim.v index 02bbd94d9..ee2dd5e8a 100644 --- a/erasure/theories/Prelim.v +++ b/erasure/theories/Prelim.v @@ -213,7 +213,7 @@ Proof. eapply (substitution (Δ := [])) in a0'; eauto. 2:{ eapply subslet_cofix_subst; pcuic. constructor; eauto. } rewrite PCUICLiftSubst.simpl_subst_k in a0'. now autorewrite with len. - eapply a0'. now eapply nth_error_all in a; tea. + eapply a0'. now eapply nth_error_all, isTypeRel_isType in a; tea. Qed. (** Assumption contexts: constructor arguments/case branches contexts contain only assumptions, no local definitions *) diff --git a/erasure/theories/Typed/Erasure.v b/erasure/theories/Typed/Erasure.v index df57eefa6..5dbd45cd1 100644 --- a/erasure/theories/Typed/Erasure.v +++ b/erasure/theories/Typed/Erasure.v @@ -107,7 +107,7 @@ Proof. intros [(_ & s & typ & _)]. sq. apply inversion_Prod in typ as (s' & ? & ? & ? & ?); [|now eauto]. - now eapply lift_sorting_forget_univ. + now eapply lift_sorting_forget_univ, lift_sorting_forget_rel. Qed. Hint Resolve isType_prod_dom : erase. @@ -1097,7 +1097,8 @@ Proof. sq. destruct r as [?? r]. eapply subject_reduction in r; eauto. - apply inversion_Lambda in r as (?&?&?&?); auto. + apply inversion_Lambda in r as (?&h1&?&?); auto. + eapply isTypeRel_isType in h1; auto. - clear inf. destruct (typ _ wfΣ) as [typ0]. reduce_term_sound. @@ -1168,10 +1169,8 @@ Proof. subst. assert (∥ wf Σ0 ∥) by now apply HΣ. unfold on_constant_decl in wt. - destruct (PCUICEnvironment.cst_body cst); cbn in *. - + sq;eapply validity;eauto. now eapply unlift_TermTyp. - + destruct wt. - eexists; eassumption. + sq. + now apply lift_sorting_forget_body, lift_sorting_forget_rel in wt. - assert (rΣ = Σ). { eapply abstract_env_ext_irr;eauto. } easy. @@ -1180,10 +1179,8 @@ Proof. subst. assert (∥ wf Σ ∥) by now apply HΣ. unfold on_constant_decl in wt. - destruct (PCUICEnvironment.cst_body cst). - + sq. - now eapply unlift_TermTyp, validity in wt. - + assumption. + sq. + now apply lift_sorting_forget_body, lift_sorting_forget_rel in wt. Qed. Import P. @@ -1213,7 +1210,7 @@ Qed. Definition ind_aname (oib : PCUICEnvironment.one_inductive_body) := {| binder_name := nNamed (PCUICEnvironment.ind_name oib); - binder_relevance := PCUICEnvironment.ind_relevance oib |}. + binder_relevance := rel_of_Type |}. Definition arities_contexts (mind : kername) @@ -1343,6 +1340,7 @@ Proof. all: intros;assert (rΣ = Σ) by (eapply abstract_env_ext_irr;eauto);subst. - abstract ( destruct wt as [wt];sq; + eapply isTypeRel_isType; exact (onArity wt.π2)). - abstract ( destruct p; diff --git a/erasure/theories/Typed/TypeAnnotations.v b/erasure/theories/Typed/TypeAnnotations.v index 972518ce5..75e25aec9 100644 --- a/erasure/theories/Typed/TypeAnnotations.v +++ b/erasure/theories/Typed/TypeAnnotations.v @@ -291,11 +291,13 @@ Proof. all: try subst erΓ1. - depelim er0. now apply inversion_Evar in X. - - apply inversion_Lambda in X as (? & ? & ? & ?); auto. - - apply inversion_Lambda in X as (? & ? & ? & ?); auto. + - apply inversion_Lambda in X as (? & h1 & ? & ?); auto. + now apply isTypeRel_isType in h1. + - apply inversion_Lambda in X as (? & h1 & ? & ?); auto. + apply isTypeRel_isType in h1. econstructor; eauto. - - apply inversion_LetIn in X as (?&?&?&?); auto. - constructor; eapply lift_sorting_forget_body; eauto. + - apply inversion_LetIn in X as (?&h1&?&?); auto. + now apply lift_sorting_forget_body, isTypeRel_isType in h1. - apply inversion_LetIn in X as (?&?&?&?); auto. econstructor; eapply unlift_TermTyp; eauto. - apply inversion_LetIn in X as (?&?&?&?); auto. @@ -331,7 +333,7 @@ Proof. apply rev_Forall. apply All_Forall. eapply All_impl. apply a. - intros. cbn in *;now constructor. + intros d X. cbn in *. apply isTypeRel_isType in X. now constructor. - apply inversion_Fix in t0 as (?&?&?&?&?&?&?); auto. depelim er0. clear -a0 X. @@ -352,7 +354,7 @@ Proof. apply rev_Forall. apply All_Forall. eapply All_impl. apply a. - intros. cbn in *;now constructor. + intros d X. cbn in *. apply isTypeRel_isType in X. now constructor. - apply inversion_CoFix in t0 as (?&?&?&?&?&?&?); auto. depelim er0. clear -a0 X. diff --git a/pcuic/theories/Bidirectional/BDFromPCUIC.v b/pcuic/theories/Bidirectional/BDFromPCUIC.v index d92a5a6fb..425910c0b 100644 --- a/pcuic/theories/Bidirectional/BDFromPCUIC.v +++ b/pcuic/theories/Bidirectional/BDFromPCUIC.v @@ -99,16 +99,17 @@ Proof. intro Hj. apply lift_sorting_ex_it_impl_gen with Hj => //=. 1: now destruct b. - intros (u' & ? & _); now eexists. + intros (u' & ? & Hle) er; eexists; split; tea. eapply geq_relevance; tea. 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 *. + intros (Htm & u0 & (u' & Hty & Hle) & <- & Her); cbn in *. exists u'; split; tas. repeat (eexists; tea). + eapply geq_relevance; tea. Qed. Section BDFromPCUIC. @@ -169,7 +170,8 @@ Proof. split. + econstructor. all: eassumption. + apply ws_cumul_pb_Prod ; auto. - by eapply isType_ws_cumul_pb_refl. + eapply isType_ws_cumul_pb_refl. + now eapply isTypeRel_isType. - intros na t A ? ? ? ? CumtA ? (?&?&?). apply conv_lift_judgment with (na := na) in CumtA. @@ -218,7 +220,7 @@ Proof. 1: fvs. now eapply closed_on_free_vars, declared_constructor_closed_type. - - intros ci p c brs indices ps mdecl idecl isdecl wfΣ' wfbΓ epar ? predctx wfpred ? ? ty_p Cump ? ? Hinst ty_c Cumc ? ? ? ty_br. + - intros ci p c brs indices ps mdecl idecl isdecl wfΣ' wfbΓ epar ? predctx wfpred ? ? ty_p Cump ? ? Her Hinst ty_c Cumc ? ? ? ty_br. apply conv_infer_sort in Cump as (?&?&?) ; auto. apply conv_infer_ind in Cumc as (?&?&[]) ; auto. @@ -229,6 +231,7 @@ Proof. * by eapply All_local_app_rel. * eapply is_allowed_elimination_monotone. all: eassumption. + * eapply geq_relevance; eassumption. * 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. @@ -343,6 +346,7 @@ Proof. split. 2:{ apply isType_ws_cumul_pb_refl. + eapply isTypeRel_isType. eapply nth_error_all in Htypes as Hj ; tea. } constructor ; eauto. @@ -354,6 +358,7 @@ Proof. split. 2:{ apply isType_ws_cumul_pb_refl. + eapply isTypeRel_isType. eapply nth_error_all in Htypes as Hj ; tea. } constructor ; eauto. @@ -407,14 +412,34 @@ Proof. now apply typing_infering. Qed. +Lemma lift_typing_lift_sorting `{checker_flags} {Σ} (wfΣ : wf Σ) {Γ} {tm ty rel} : + lift_typing typing Σ Γ (Judge tm ty None rel) -> lift_sorting (checking Σ Γ) (infering_sort Σ Γ) (Judge tm ty None rel). +Proof. + intro Hj. + apply lift_sorting_ex_it_impl_gen with Hj. + - destruct tm => //. apply typing_checking. + - intros Hty Her. + eapply typing_infering_sort in Hty as [? []]; tea. + eexists; split; tea. destruct rel as [rel|] => //=. + now eapply geq_relevance; tea. +Qed. + Lemma isType_infering_sort `{checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} {Γ t} : isType Σ Γ t -> ∑ u', Σ ;;; Γ |- t ▹□ u'. Proof. + move/lift_typing_lift_sorting. intros (_ & s & ty & _). - eapply typing_infering_sort in ty as [? []]; tea. now eexists. Qed. +Theorem isTypeRel_infering_sort `{checker_flags} (Σ : global_env_ext) Γ t r (wfΣ : wf Σ) : + isTypeRel Σ Γ t r -> ∑ u, Σ ;;; Γ |- t ▹□ u × relevance_of_sort u = r. +Proof. + move/lift_typing_lift_sorting. + intros (_ & s & ty & _ & Her). + now eexists; split; tea. +Qed. + Lemma typing_infer_prod `{checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} {Γ t na A B} : Σ ;;; Γ |- t : tProd na A B -> ∑ na' A' B', [× Σ ;;; Γ |- t ▹Π (na',A',B'), @@ -454,7 +479,7 @@ Proof. 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. + - intros Hty Her. eapply typing_infering_sort in Hty as [? []]; tea. eexists; split; tea. eapply geq_relevance; tea. Qed. Theorem ctx_inst_typing_bd `{checker_flags} (Σ : global_env_ext) Γ l Δ (wfΣ : wf Σ) : diff --git a/pcuic/theories/Bidirectional/BDStrengthening.v b/pcuic/theories/Bidirectional/BDStrengthening.v index 40c0c72b5..dc04ce7ca 100644 --- a/pcuic/theories/Bidirectional/BDStrengthening.v +++ b/pcuic/theories/Bidirectional/BDStrengthening.v @@ -750,7 +750,7 @@ Proof using wfΣ. solve_all. now eapply conv_renameP. + by apply rename_wf_branches. - + eapply Forall2_All2 in H6. + + eapply Forall2_All2 in H7. eapply All2i_All2_mix_left in X9; eauto. eapply All2i_All_mix_right in X9 ; eauto. eapply All2i_nth_hyp in X9. diff --git a/pcuic/theories/Bidirectional/BDToPCUIC.v b/pcuic/theories/Bidirectional/BDToPCUIC.v index 2fd27e6d3..f54d0e19c 100644 --- a/pcuic/theories/Bidirectional/BDToPCUIC.v +++ b/pcuic/theories/Bidirectional/BDToPCUIC.v @@ -164,12 +164,13 @@ Section BDToPCUICTyping. intros wfΔ args ctxi ; inversion ctxi. - subst d. subst. - assert (isType Σ Γ t). + assert (isTypeRel Σ Γ t na.(binder_relevance)). { eapply All_local_rel_app_inv in wfΔ as [wfd _]. inversion_clear wfd. eassumption. } + apply isTypeRel_isType in X2 as X2'. constructor ; auto. apply X ; auto. 1: by rewrite subst_telescope_length ; reflexivity. @@ -189,7 +190,7 @@ Section BDToPCUICTyping. { eapply All_local_rel_app_inv in wfΔ as [wfd _]. inversion_clear wfd. - apply lift_sorting_it_impl_gen with X2 => //. + now eapply lift_sorting_forget_body, lift_sorting_forget_rel. } constructor ; auto. apply X ; auto. @@ -211,7 +212,7 @@ Section BDToPCUICTyping. 2: intros [_ IH]; now apply IH. destruct j_term => //=. intros [_ IH]; apply IH; tas. - apply lift_sorting_forget_univ in X. + apply lift_sorting_forget_univ, lift_sorting_forget_rel, lift_sorting_forget_body in X. apply lift_sorting_it_impl_gen with X => //. intros [_ IH']; now apply IH'. @@ -239,7 +240,7 @@ Section BDToPCUICTyping. apply validity in X0 as (_ & s & X0 & _). apply inversion_Prod in X0 as (s1 & s2 & X0 & _). 2: done. - eapply lift_sorting_forget_univ; eassumption. + eapply lift_sorting_forget_rel, lift_sorting_forget_univ; eassumption. - red ; intros ; econstructor ; eauto. @@ -342,7 +343,7 @@ Section BDToPCUICTyping. specialize (Hd X3). apply lift_sorting_it_impl with Hd => //. - + have Htypes : All (fun d => isType Σ Γ (dtype d)) mfix. + + have Htypes : All (fun d => isTypeRel Σ Γ (dtype d) (dname d).(binder_relevance)) mfix. { apply All_impl with (1 := X0) => d Hd. specialize (Hd X3). apply lift_sorting_it_impl with Hd => //. @@ -362,7 +363,7 @@ Section BDToPCUICTyping. specialize (Hd X3). apply lift_sorting_it_impl with Hd => //. - + have Htypes : All (fun d => isType Σ Γ (dtype d)) mfix. + + have Htypes : All (fun d => isTypeRel Σ Γ (dtype d) (dname d).(binder_relevance)) mfix. { apply All_impl with (1 := X0) => d Hd. specialize (Hd X3). apply lift_sorting_it_impl with Hd => //. @@ -439,6 +440,14 @@ Proof. now apply infering_sort_typing. Qed. +Theorem infering_sort_isTypeRel `{checker_flags} (Σ : global_env_ext) Γ t u (wfΣ : wf Σ) : + wf_local Σ Γ -> Σ ;;; Γ |- t ▹□ u -> isTypeRel Σ Γ t (relevance_of_sort u). +Proof. + intros wfΓ Ht. + repeat (eexists; tea). + now apply infering_sort_typing. +Qed. + Theorem einfering_sort_isType `{checker_flags} (Σ : global_env_ext) Γ t (wfΣ : wf Σ) : wf_local Σ Γ -> (∑ u, Σ ;;; Γ |- t ▹□ u) -> isType Σ Γ t. Proof. @@ -510,6 +519,7 @@ Proof. rewrite /= app_context_assoc in wfΓ. eapply All_local_env_app_inv in wfΓ as [wfΓ _]. inversion wfΓ ; subst. + apply isTypeRel_isType in X0. now apply checking_typing. } constructor ; tea. diff --git a/pcuic/theories/Bidirectional/BDTyping.v b/pcuic/theories/Bidirectional/BDTyping.v index a384ca756..1544583e0 100644 --- a/pcuic/theories/Bidirectional/BDTyping.v +++ b/pcuic/theories/Bidirectional/BDTyping.v @@ -82,6 +82,7 @@ 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 -> + isSortRel ps ci.(ci_relevance) -> ctx_inst (checking Σ) Γ (pparams p) (List.rev mdecl.(ind_params)@[p.(puinst)]) -> isCoFinite mdecl.(ind_finite) = false -> @@ -415,6 +416,7 @@ Section BidirectionalInduction. wf_local_bd_rel Σ Γ predctx -> PΓ_rel Γ predctx -> is_allowed_elimination Σ (ind_kelim idecl) ps -> + isSortRel ps ci.(ci_relevance) -> ctx_inst (checking Σ) Γ (pparams p) (List.rev mdecl.(ind_params)@[p.(puinst)]) -> ctx_inst Pcheck Γ p.(pparams) diff --git a/pcuic/theories/Bidirectional/BDUnique.v b/pcuic/theories/Bidirectional/BDUnique.v index 13bb91a57..11ab9e85e 100644 --- a/pcuic/theories/Bidirectional/BDUnique.v +++ b/pcuic/theories/Bidirectional/BDUnique.v @@ -104,11 +104,13 @@ Proof using wfΣ. 1: constructor. rewrite subst_empty. eapply checking_typing ; tea. + eapply isTypeRel_isType. now eapply isType_tProd, validity, infering_prod_typing. + constructor. 1: constructor. rewrite subst_empty. eapply checking_typing ; tea. + eapply isTypeRel_isType. now eapply isType_tProd, validity, infering_prod_typing. - pose proof (declared_constant_to_gen (wfΣ := wfΣ) isdecl). @@ -204,9 +206,9 @@ Proof using wfΣ. - intros ? T' ty_T'. inversion ty_T' ; subst. - unshelve eapply declared_inductive_to_gen in H13; cycle -2; eauto. + unshelve eapply declared_inductive_to_gen in H14; cycle -2; eauto. unshelve eapply declared_inductive_to_gen in H; cycle -2; eauto. - move: (H) => /declared_inductive_inj /(_ H13) [? ?]. + move: (H) => /declared_inductive_inj /(_ H14) [? ?]. subst. assert (op' : is_open_term Γ (mkApps ptm0 (skipn (ci_npar ci) args0 ++ [c]))). by now eapply type_is_open_term, infering_typing. @@ -475,9 +477,14 @@ Proof. Qed. - - - - - +Corollary isTypeRel_unique `{checker_flags} {Σ} (wfΣ : wf Σ) {Γ T rel rel'} : + isTypeRel Σ Γ T rel -> isTypeRel Σ Γ T rel' -> rel = rel'. +Proof. + intros (_ & s1 & H1 & _ & <-) (_ & s2 & H2 & _ & <-). cbn in H1, H2. + apply typing_wf_local in H1 as wfΓ. + apply typing_infering_sort in H1 as (u1 & H1 & Hle), H2 as (u2 & H2 & Hle'). + apply leq_relevance_eq in Hle as <-, Hle' as <-. + eapply infering_sort_sort in H2 as ->; tea. + reflexivity. +Qed. diff --git a/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v b/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v index be8df6415..c8933db39 100644 --- a/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v +++ b/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v @@ -1656,7 +1656,7 @@ Lemma All_local_env_over_subst_instance {cf : checker_flags} Σ Γ (wfΓ : wf_lo wf_local (Σ.1, univs) (subst_instance u Γ). Proof. induction 1; simpl; rewrite /subst_instance /=; constructor; cbn in *; auto. - all: eapply lift_sorting_fu_it_impl with (tu := tu); cbn in *; eauto. + all: eapply lift_sorting_fu_it_impl with (tu := tu); cbn in *; eauto using relevance_subst. Qed. #[global] Hint Resolve All_local_env_over_subst_instance : univ_subst. @@ -2016,15 +2016,13 @@ Section SubstIdentity. subst_instance u Γ = Γ). Proof using Type. 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 lift_typing_subjtyp with (1 := X). 2: intros s Hs; now depelim Hs. + intros ?? []; auto. } { 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. } + eapply All2_eq, All2_map_left. apply All_All2 with (1 := X0) => decl IH. + destruct IH as (? & ? & ?), decl as [na bo ty]; tas; unfold map_decl; cbn in *. + f_equal; tas. destruct bo => //. cbn in *. now f_equal. } all: try ((subst u || subst u0); split; [f_equal|]; intuition eauto). diff --git a/pcuic/theories/PCUICAlpha.v b/pcuic/theories/PCUICAlpha.v index 6a8a6efd0..2b8d53048 100644 --- a/pcuic/theories/PCUICAlpha.v +++ b/pcuic/theories/PCUICAlpha.v @@ -37,24 +37,6 @@ Qed. Section Alpha. Context {cf:checker_flags}. - (* TODO MOVE *) - Lemma wf_local_nth_error_vass : - forall Σ Γ i na ty, - wf Σ.1 -> - wf_local Σ Γ -> - nth_error Γ i = Some (vass na ty) -> - lift_typing typing Σ Γ (Typ (lift0 (S i) ty)). - Proof using Type. - intros Σ Γ i na ty hΣ hΓ e. simpl. - 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 *) Lemma fix_context_nth_error : forall m i d, @@ -397,6 +379,7 @@ Section Alpha. intros a eq. induction eq; depelim a; cbn; try solve [constructor; auto]; depelim r; subst; constructor; auto. + all: rewrite e. 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. @@ -462,13 +445,13 @@ Section Alpha. now eapply eq_context_conversion. Qed. - Lemma lift_judgment_alpha {Σ : global_env_ext} {Γ tm tm' t t' u} : + Lemma lift_judgment_alpha {Σ : global_env_ext} {Γ tm tm' t t' u r} : lift_typing0 (fun t T : term => forall u : term, upto_names' t u -> Σ;;; Γ |- u : T) - (Judge tm t u) -> + (Judge tm t u r) -> 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). + lift_typing typing Σ Γ (Judge tm' t' u r). Proof. intros tu Htm Hty. apply lift_sorting_it_impl_gen with tu => //. @@ -476,7 +459,7 @@ Section Alpha. 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 lift_sorting_forget_rel, 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. @@ -507,6 +490,7 @@ Section Alpha. * intros Δ eq; destruct Δ; depelim eq; constructor. * intros Δ eq; depelim eq. depelim c. all: constructor; auto. + all: rewrite -eqna. all: eapply lift_judgment_alpha with (1 := X _ eq) => //. - intros n decl hnth ih Δ v e eqctx; invs e. @@ -526,27 +510,29 @@ Section Alpha. eapply eq_context_conversion; tea. eapply type_Sort; assumption. - intros na A B s1 s2 ih hA ihA hB ihB Δ v e eqctx; invs e. - econstructor. - + eapply lift_judgment_alpha with (1 := ihA _ eqctx) => //. - + eapply context_conversion. + assert (lift_typing typing Σ Δ (j_vass_s na' a' s1)). + + rewrite -H3. + eapply lift_judgment_alpha with (1 := ihA _ eqctx) => //. + + econstructor; tas. + eapply context_conversion. * eapply ihB. assumption. reflexivity. * constructor; eauto. - eapply lift_sorting_forget_univ. - eapply lift_judgment_alpha with (1 := ihA _ eqctx) => //. + now eapply lift_sorting_forget_univ. * constructor. -- now eapply upto_names_conv_context. -- constructor. assumption. constructor. eapply upto_names_impl_eq_term. assumption. - intros na A t B ih hA ihA hB ihB Δ v e eqctx; invs e. eapply type_Cumul'. - + econstructor. - * eapply lift_judgment_alpha with (1 := ihA _ eqctx) => //. - * eapply eq_context_conversion. + + assert (lift_typing typing Σ Δ (j_vass na' ty')). + * rewrite -H3. + eapply lift_judgment_alpha with (1 := ihA _ eqctx) => //. + * econstructor; eauto. + eapply eq_context_conversion. -- eapply ihB. assumption. reflexivity. -- constructor. 1: assumption. simpl. constructor; auto. -- 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. @@ -560,15 +546,16 @@ Section Alpha. assert (isType Σ Γ (tLetIn na b B A)). { eapply validity. econstructor; eauto. } eapply type_Cumul'. - + econstructor. - * eapply lift_judgment_alpha with (1 := ihbB _ eqctx) => //. - * eapply eq_context_conversion. + + assert (lift_typing typing Σ Δ (j_vdef na' t' ty')). + * rewrite -H4. + eapply lift_judgment_alpha with (1 := ihbB _ eqctx) => //. + * econstructor; auto. + eapply eq_context_conversion. -- eapply ihA; trea. -- constructor. ++ assumption. ++ constructor; auto. -- constructor; auto. - 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. @@ -601,7 +588,7 @@ Section Alpha. 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. + wfcpc kelim Her 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 (1 := HIHctxi). eassert (PCUICEnvTyping.ctx_inst _ _ _ _) as IHctxi. @@ -680,7 +667,7 @@ Section Alpha. eapply ws_cumul_pb_mkApps; trea. eapply ws_cumul_pb_refl => //. eauto with fvs. eapply wf_local_closed_context in wfΓ. - eapply isType_open in X1. + eapply isType_is_open_term in X1. rewrite on_free_vars_mkApps in X1. move/andP: X1 => [] _. rewrite forallb_app => /andP[] hargs hc. eapply All2_app. @@ -759,7 +746,7 @@ 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)]. + destruct r as (Hty & eqty & eqbod & eqrarg & eqann). rewrite -eqann. 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. @@ -795,9 +782,10 @@ 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)]. + destruct r as (Hty & eqty & eqbod & eqrarg & eqann). rewrite /on_def_type -eqann. eapply lift_judgment_alpha with (1 := Hty _ ltac:(reflexivity)) => //. * unfold eq_mfixpoint in *. solve_all. + rewrite /on_def_body -b. specialize (b1 (Γ ,,, types)) as IHb. forward IHb by eapply eq_context_upto_cat; reflexivity. eapply @lift_judgment_alpha with (tm' := Some _) in IHb; tea. @@ -814,7 +802,7 @@ Section Alpha. 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. unfold eq_mfixpoint in *. solve_all. - + eapply All_nth_error in hmfix; tea. + + eapply isTypeRel_isType. eapply All_nth_error in hmfix; tea. + apply eq_term_upto_univ_cumulSpec, eq_term_leq_term, upto_names_impl_eq_term. now symmetry. @@ -827,7 +815,7 @@ 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)]. + destruct r as (Hty & eqty & eqbod & eqrarg & eqann). rewrite -eqann. 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. @@ -863,9 +851,10 @@ 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)]. + destruct r as (Hty & eqty & eqbod & eqrarg & eqann). rewrite /on_def_type -eqann. eapply lift_judgment_alpha with (1 := Hty _ ltac:(reflexivity)) => //. * unfold eq_mfixpoint in *. solve_all. + rewrite /on_def_body -b. specialize (b1 (Γ ,,, types)) as IHb. forward IHb by eapply eq_context_upto_cat; reflexivity. eapply @lift_judgment_alpha with (tm' := Some _) in IHb; tea. @@ -876,7 +865,7 @@ Section Alpha. unfold wf_cofixpoint, wf_cofixpoint_gen. enough (map check_one_cofix mfix = map check_one_cofix mfix') as ->; auto. apply upto_names_check_cofix. unfold eq_mfixpoint in *. solve_all. - + eapply All_nth_error in hmfix; tea. + + eapply isTypeRel_isType. 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'. diff --git a/pcuic/theories/PCUICArities.v b/pcuic/theories/PCUICArities.v index 9e23e62db..4b5f37620 100644 --- a/pcuic/theories/PCUICArities.v +++ b/pcuic/theories/PCUICArities.v @@ -211,7 +211,7 @@ Section WfEnv. Qed. Lemma isType_tProd {Γ} {na A B} : - isType Σ Γ (tProd na A B) <~> (isType Σ Γ A × isType Σ (Γ,, vass na A) B). + isType Σ Γ (tProd na A B) <~> (isTypeRel Σ Γ A na.(binder_relevance) × isType Σ (Γ,, vass na A) B). Proof. split; intro HH. - destruct HH as (_ & s & H & _). @@ -225,27 +225,6 @@ Section WfEnv. econstructor; eassumption. Defined. - Lemma isType_subst {Γ Δ A} s : - subslet Σ Γ s Δ -> - isType Σ (Γ ,,, Δ) A -> - isType Σ Γ (subst0 s A). - Proof using wfΣ. - intros sub HT. - apply lift_typing_f_impl with (1 := HT) => // ?? Hs. - have wf := typing_wf_local Hs. - now eapply (substitution (Δ := [])). - Qed. - - Lemma isType_subst_gen {Γ Δ Δ'} {A} s : - subslet Σ Γ s Δ -> - isType Σ (Γ ,,, Δ ,,, Δ') A -> - isType Σ (Γ ,,, subst_context s 0 Δ') (subst s #|Δ'| A). - Proof using wfΣ. - intros sub HT. - apply lift_typing_f_impl with (1 := HT) => // ?? Hs. - now eapply substitution. - Qed. - Lemma type_ws_cumul_pb {pb Γ t} T {U} : Σ ;;; Γ |- t : T -> isType Σ Γ U -> @@ -287,7 +266,7 @@ Section WfEnv. Lemma wf_local_ass {Γ na A} : wf_local Σ Γ -> - isType Σ Γ A -> + isTypeRel Σ Γ A na.(binder_relevance) -> wf_local Σ (Γ ,, vass na A). Proof using Type. constructor; eauto with pcuic. @@ -295,7 +274,7 @@ Section WfEnv. Lemma wf_local_def {Γ na d ty} : wf_local Σ Γ -> - isType Σ Γ ty -> + isTypeRel Σ Γ ty na.(binder_relevance) -> Σ ;;; Γ |- d : ty -> wf_local Σ (Γ ,, vdef na d ty). Proof using Type. @@ -312,7 +291,7 @@ Section WfEnv. isType Σ Γ (B {0 := t}). Proof using wfΣ. move/isType_tProd => [hA hB] ht. - eapply (isType_subst (Δ:= [vass na A])); eauto with pcuic. + eapply (isType_subst (Δ := [vass na A])); eauto with pcuic. Qed. Hint Resolve isType_wf_local : pcuic. @@ -392,7 +371,7 @@ Section WfEnv. Qed. Lemma type_mkProd_or_LetIn {Γ} d {u t s} : - Σ ;;; Γ |- decl_type d : tSort u -> + lift_typing typing Σ Γ (j_decl_s d (Some u)) -> Σ ;;; Γ ,, d |- t : tSort s -> match decl_body d return Type with | Some b => Σ ;;; Γ |- mkProd_or_LetIn d t : tSort s @@ -408,7 +387,6 @@ Section WfEnv. - 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} : @@ -428,13 +406,13 @@ Section WfEnv. eapply type_Cumul. eapply IHΓ'; 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'. + destruct a as [na [b|] ty]; cbn; destruct equ as (? & Hdecl). { 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.π2.1. - + eapply type_Cumul; eauto. - econstructor; eauto with pcuic. - eapply cumul_Sort. eapply leq_sort_product. } + pose proof (lift_sorting_extract Hdecl). + eapply (type_mkProd_or_LetIn {| decl_body := Some b |}); eauto. + eapply type_Cumul; eauto. + econstructor; eauto with pcuic. + 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. @@ -467,11 +445,10 @@ Section WfEnv. induction Γ'; simpl; auto; move=> Γ us s t equ Ht. - destruct us => //. - destruct a as [na [b|] ty]; intuition auto. - * 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 as (? & _ & s' & Hty & ->). + * eapply IHΓ'; eauto. + pose proof (lift_sorting_extract b0). + eapply (type_mkProd_or_LetIn {| decl_body := Some b |}); eauto. + * destruct us => //. destruct equ. simpl. eapply IHΓ'; eauto. apply (type_mkProd_or_LetIn {| decl_body := None |}) => /=; eauto. diff --git a/pcuic/theories/PCUICCasesHelper.v b/pcuic/theories/PCUICCasesHelper.v index 2ff8c1ec9..b66af2563 100644 --- a/pcuic/theories/PCUICCasesHelper.v +++ b/pcuic/theories/PCUICCasesHelper.v @@ -465,6 +465,7 @@ Lemma type_Case_helper (* consistent_instance_ext Σ (ind_universes mib) puinst -> *) is_allowed_elimination Σ (ind_kelim oib) ps -> + isSortRel ps oib.(ind_relevance) -> isCoFinite (ind_finite mib) = false -> @@ -487,7 +488,7 @@ Lemma type_Case_helper Σ ;;; Γ |- make_case ind mib oib pparams puinst discr rettyp brs : subst_rettyp. Proof. - move=> wfΣ inddecl pparamslen discrtyp p predctx ptyp (* puinstok *) pselimok isfinite brslen ptm brstyp srettyp. + move=> wfΣ inddecl pparamslen discrtyp p predctx ptyp (* puinstok *) pselimok Her isfinite brslen ptm brstyp srettyp. set ci := {| ci_ind := ind ; ci_npar := ind_npars mib ; ci_relevance := ind_relevance oib |}. have wfΣ1 := (wfΣ : wf Σ.1). @@ -581,6 +582,7 @@ Lemma type_Case_subst_helper (wf_local Σ (Γ,,, predctx) -> Σ ;;; Γ,,, predctx |- preturn p : tSort ps) -> is_allowed_elimination Σ (ind_kelim oib) ps -> + isSortRel ps oib.(ind_relevance) -> isCoFinite (ind_finite mib) = false -> @@ -602,7 +604,7 @@ Lemma type_Case_subst_helper Σ ;;; Γ |- make_case ind mib oib pparams puinst discr rettyp brs : subst_rettyp. Proof. - move=> wfΣ inddecl pparamslen discrtyp p predctx predWty elimok finok brslen brsok substrettyp. + move=> wfΣ inddecl pparamslen discrtyp p predctx predWty elimok Her finok brslen brsok substrettyp. set ci := {| ci_ind := ind ; ci_npar := ind_npars mib ; ci_relevance := ind_relevance oib |}. @@ -716,6 +718,7 @@ Lemma type_Case_simple_subst_helper Σ ;;; Γ |- rettyp : tSort ps -> is_allowed_elimination Σ (ind_kelim oib) ps -> + isSortRel ps oib.(ind_relevance) -> isCoFinite (ind_finite mib) = false -> @@ -736,7 +739,7 @@ Lemma type_Case_simple_subst_helper Σ ;;; Γ |- make_case ind mib oib pparams puinst discr rettyp' brs : rettyp. Proof. - move=> wfΣ inddecl pparamslen discrtyp rettyp' p predctx predWty elimok finok brslen brsok. + move=> wfΣ inddecl pparamslen discrtyp rettyp' p predctx predWty elimok Her finok brslen brsok. set ci := {| ci_ind := ind ; ci_npar := ind_npars mib ; ci_relevance := ind_relevance oib |}. diff --git a/pcuic/theories/PCUICClassification.v b/pcuic/theories/PCUICClassification.v index 94ed68cc1..0f6cbe583 100644 --- a/pcuic/theories/PCUICClassification.v +++ b/pcuic/theories/PCUICClassification.v @@ -234,7 +234,7 @@ Section Spines. intros sp. have ty := (typing_spine_isType_dom sp). have r := (red_it_mkProd_or_LetIn_smash Γ Δ T). - specialize (r (wf_local_closed_context (isType_wf_local ty)) (isType_open ty)). + specialize (r (wf_local_closed_context (isType_wf_local ty)) (isType_is_open_term ty)). eapply typing_spine_strengthen; tea. eapply isType_red; tea. eapply r. eapply ws_cumul_pb_eq_le; symmetry. @@ -301,6 +301,7 @@ Qed. now rewrite -(expand_lets_subst_comm _ _ _). eapply isType_apply in i; tea. eapply (type_ws_cumul_pb (pb:=Conv)); tea. 2:now symmetry. + eapply isTypeRel_isType. now eapply isType_tProd in i as []. Qed. @@ -330,6 +331,7 @@ Qed. now len. eapply isType_apply in i; tea. eapply (type_ws_cumul_pb (pb:=Conv)); tea. 2:now symmetry. + eapply isTypeRel_isType. now eapply isType_tProd in i as []. Qed. @@ -649,6 +651,7 @@ Section classification. eapply inversion_CoFix in Hcof as (? & ? & ? & ? & ? & ? & ?); auto. eapply nth_error_all in a; eauto. simpl in a. + apply isTypeRel_isType in a. eapply typing_spine_strengthen in sp; eauto. eapply wf_cofixpoint_spine in sp as (Γ' & concl & da & ?); eauto. eapply decompose_prod_assum_it_mkProd_or_LetIn in da. @@ -705,7 +708,7 @@ Section classification. 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). + now eapply isTypeRel_isType, (declared_inductive_valid_type decl). Qed. Lemma invert_fix_ind Γ mfix i args ind u args' : @@ -718,8 +721,8 @@ Section classification. intros no_arg typ. eapply inversion_mkApps in typ as (?&?&?); eauto. eapply inversion_Fix in t as (? & ? & ? & ? & ? & ? & ?); auto. + eapply nth_error_all in a; eauto. apply isTypeRel_isType in a. eapply PCUICSpine.typing_spine_strengthen in t0; eauto. - eapply nth_error_all in a; eauto. simpl in a. unfold unfold_fix in no_arg. rewrite e in no_arg. eapply (wf_fixpoint_spine wfΣ) in t0; eauto. @@ -727,7 +730,6 @@ Section classification. eapply ws_cumul_pb_Prod_l_inv in cum as (? & ? & ? & []). eapply invert_red_mkApps_tInd in c as [? [eq _]]; auto. solve_discr. - now eapply nth_error_all in a; tea. Qed. @@ -764,8 +766,8 @@ Section classification. - eapply inversion_App in typed as [na [A [B [Hf _]]]]; eauto. - eapply inversion_mkApps in typed as (? & ? & ?); eauto. eapply inversion_Fix in t as (? & ? & ? & ? & ? & ? & ?); auto. + eapply nth_error_all in a; eauto. apply isTypeRel_isType in a. eapply typing_spine_strengthen in t0; eauto. - eapply nth_error_all in a; eauto. simpl in a. rewrite /unfold_fix in e. rewrite e1 in e. noconf e. eapply (wf_fixpoint_spine wfΣ) in t0; eauto. rewrite e0 in t0. destruct t0 as [ind [u [indargs [tyarg ckind]]]]. @@ -777,7 +779,6 @@ Section classification. rewrite nth_error_map e0 in axfree. cbn in axfree. eauto. - now eapply nth_error_all in a; tea. - eapply inversion_Case in typed as (? & ? & ? & ? & [] & ?); tas; eauto. - eapply inversion_Proj in typed as (? & ? & ? & ? & ? & ? & ? & ? & ?); eauto. Qed. @@ -868,7 +869,7 @@ Section classification. intros typed unf. eapply inversion_mkApps in typed as (? & ? & ?); eauto. eapply inversion_Fix in t as (? & ? & ? & ? & ? & ? & ?); auto. - eapply nth_error_all in a; eauto. simpl in a. + eapply nth_error_all in a; eauto. apply isTypeRel_isType in a. eapply typing_spine_strengthen in t0; eauto. rewrite /unfold_fix in unf. rewrite e in unf. noconf unf. @@ -940,6 +941,7 @@ Section classification. rewrite /subst1 in t2. eapply t2. eapply (type_ws_cumul_pb (pb:=Conv)); eauto. eapply subject_reduction in IHHe2; eauto. + now eapply isTypeRel_isType. eapply ws_cumul_pb_Prod_Prod_inv in e as [eqna dom codom]; eauto. now symmetry. diff --git a/pcuic/theories/PCUICContextConversion.v b/pcuic/theories/PCUICContextConversion.v index d5efe4fbd..fc5f624c8 100644 --- a/pcuic/theories/PCUICContextConversion.v +++ b/pcuic/theories/PCUICContextConversion.v @@ -1012,7 +1012,7 @@ Section ContextConversion. End ContextConversion. -#[global] Hint Resolve isType_open PCUICClosedTyp.wf_local_closed_context : fvs. +#[global] Hint Resolve isType_is_open_term isTypeRel_is_open_term PCUICClosedTyp.wf_local_closed_context : fvs. #[global] Hint Resolve conv_ctx_refl' : pcuic. #[global] Hint Constructors All_decls_alpha_pb : pcuic. diff --git a/pcuic/theories/PCUICContexts.v b/pcuic/theories/PCUICContexts.v index 440175279..d6e71a155 100644 --- a/pcuic/theories/PCUICContexts.v +++ b/pcuic/theories/PCUICContexts.v @@ -165,12 +165,9 @@ Proof. induction Δ; simpl in *; intuition auto. { 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. + all: eapply lift_typing_fu_impl with (1 := Hj) => //= ?? Hs; auto using relevance_subst_opt. + all: eapply instantiate_minductive in Hs; eauto. + all: now rewrite subst_instance_app in Hs. Qed. Lemma sorts_local_ctx_instantiate {cf:checker_flags} Σ ind mdecl Γ Δ u s : @@ -187,7 +184,7 @@ Proof. 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 lift_typing_fu_impl with (1 := Hj) => //= ?? Hs; auto using relevance_subst_opt. all: eapply instantiate_minductive in Hs; eauto. all: now rewrite subst_instance_app in Hs. Qed. @@ -351,7 +348,9 @@ Section WfEnv. + 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 IHΓ; auto. + apply All_local_env_fold, All_local_env_impl with (1 := wfΔ). + intros Γ' j X. 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. diff --git a/pcuic/theories/PCUICCumulProp.v b/pcuic/theories/PCUICCumulProp.v index 71bd00c95..7ceea5d7e 100644 --- a/pcuic/theories/PCUICCumulProp.v +++ b/pcuic/theories/PCUICCumulProp.v @@ -1136,7 +1136,7 @@ Proof using Type. unshelve eapply declared_inductive_to_gen in isdecl, isdecl'; eauto. destruct (declared_inductive_inj isdecl isdecl'). subst. destruct data. - specialize (X7 _ _ H5 scrut_ty _ _ X10). + specialize (X7 _ _ H6 scrut_ty _ _ X10). eapply cumul_prop_sym => //. destruct e as [eqpars [eqinst [eqpctx eqpret]]]. rewrite /ptm. @@ -1208,7 +1208,7 @@ Proof using Type. destruct e as (a & _). constructor; [fvs|..]. { eapply nth_error_all in X0; tea. - now apply isType_is_open_term in X0. } + now apply isTypeRel_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. @@ -1220,7 +1220,7 @@ Proof using Type. destruct e as (a & _). constructor; [fvs|..]. { eapply nth_error_all in X0; tea. - now apply isType_is_open_term in X0. } + now apply isTypeRel_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. diff --git a/pcuic/theories/PCUICElimination.v b/pcuic/theories/PCUICElimination.v index b5341a278..e50c9449d 100644 --- a/pcuic/theories/PCUICElimination.v +++ b/pcuic/theories/PCUICElimination.v @@ -228,9 +228,9 @@ Proof. eapply (substitution0_ws_cumul_pb (t:=hd0)) in cum; auto. simpl in cum. now rewrite /subst1 subst_it_mkProd_or_LetIn Nat.add_0_r in cum. - unshelve eapply (isType_subst (Δ := [vass _ _]) [hd0]) in i0; pcuic. + unshelve eapply (isType_subst (Δ := [vass _ _]) (s := [hd0])) in i0. now rewrite subst_it_mkProd_or_LetIn in i0. - eapply subslet_ass_tip. eapply (type_ws_cumul_pb (pb:=Conv)); tea. now symmetry. + eapply subslet_ass_tip. eapply (type_ws_cumul_pb (pb:=Conv)); tea. 1: now eapply isTypeRel_isType. now symmetry. Qed. Inductive All_local_assum (P : context -> term -> Type) : context -> Type := @@ -341,7 +341,7 @@ Proof. 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. + now eapply isTypeRel_isType, declared_inductive_valid_type. * now eapply invert_cumul_ind_prod in e. @@ -399,7 +399,7 @@ Proof. eapply ws_cumul_pb_Prod_Prod_inv in e as [eqna conv cum]; auto. cbn in *. eapply isType_tProd in isty as []. have tyt : Σ ;;; Γ |- hd0 : ty. - { eapply (type_ws_cumul_pb _ (U:=ty)) in tyhd => //. now symmetry. } + { eapply (type_ws_cumul_pb _ (U:=ty)) in tyhd => //. 1: now eapply isTypeRel_isType. now symmetry. } eapply (isType_subst (Δ := [_])) in i0; revgoals. { now eapply subslet_ass_tip. } eapply typing_spine_strengthen in sp; eauto. @@ -467,7 +467,7 @@ Proof. destruct cs => //; eauto. destruct X. eapply IHΔ. intros. apply (H Γ' t0 s0). right; eauto. all:auto. destruct cs => //. destruct X. - eapply H. left; eauto. now destruct l as (_ & ? & ? & <-). + eapply H. left; eauto. now apply unlift_TypUniv in l. Qed. Lemma In_map {A B} (f : A -> B) (l : list A) x : diff --git a/pcuic/theories/PCUICExpandLetsCorrectness.v b/pcuic/theories/PCUICExpandLetsCorrectness.v index 6f86265c1..32e85f9bf 100644 --- a/pcuic/theories/PCUICExpandLetsCorrectness.v +++ b/pcuic/theories/PCUICExpandLetsCorrectness.v @@ -2437,6 +2437,21 @@ Qed. Definition TTwf_local {cf} Σ Γ := TT.All_local_env (TT.lift_typing (TT.typing (H:=cf' cf)) Σ) Γ. +Lemma trans_lift_typing_it {P Q} {tm tm' : option term} {t t' : term} {u r} : + forall tu: ST.lift_typing0 P (Judge tm t u r), + let s := tu.2.π1 in + match tm', tm with None, _ => unit | Some tm', Some tm => P tm t -> Q tm' t' | _, _ => False end -> + (P t (tSort s) -> Q t' (tSort s)) -> + TT.lift_typing0 Q (Judge tm' t' u r). +Proof. + intros (? & ? & Hs & e) s HPQc HPQs. + split. + - destruct tm, tm' => //=. now apply HPQc. + - eexists. split; [now apply HPQs|]. + destruct u => //. +Qed. + + Lemma trans_wf_local' {cf} : forall (Σ : SE.global_env_ext) Γ (wfΓ : wf_local Σ Γ), let P := @@ -2451,8 +2466,8 @@ Proof. - simpl. constructor. - simpl. econstructor. + eapply IHX. - + simpl. destruct tu. repeat (eexists; tea). - - simpl. constructor; auto. destruct tu. repeat (eexists; tea). + + eapply trans_lift_typing_it with tu => //. + - simpl. constructor; auto. eapply trans_lift_typing_it with tu => //. Qed. Lemma trans_wf_local_env {cf} Σ Γ : @@ -2471,8 +2486,8 @@ Proof. - simpl. constructor. - simpl. econstructor. + eapply IHX. - + 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. + + simpl. eapply trans_lift_typing_it with t0 => // HT. now apply HT. + - simpl. constructor; auto. red. eapply trans_lift_typing_it with t0 => // HT. all: now apply HT. Qed. Lemma trans_branches {cf} Σ Γ brs btys ps: @@ -2523,10 +2538,8 @@ Proof. } induction X;cbn;constructor;auto;cbn in *. - - destruct t0 as (_ & ? & (? & ?) & _). - repeat (eexists; tea). - - destruct t0 as ((? & ?) & ? & (? & ?) & _). - repeat (eexists; tea). + - eapply trans_lift_typing_it with t0 => // HT. now apply HT. + - eapply trans_lift_typing_it with t0 => // HT. all: now apply HT. Qed. @@ -3591,13 +3604,13 @@ 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. - unshelve eapply lift_sorting_ex_it_impl_gen with (on_ctype _) => //= Hs. + 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 []). - eexists. + eexists; split => //. 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). @@ -3638,6 +3651,7 @@ Proof. now eapply alpha_eq_trans. + rewrite <- trans_global_ext_constraints. eassumption. + + assumption. + 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. { eapply ctx_inst_impl with (1 := X5). intros ? ? [? r]; exact r. } @@ -3693,7 +3707,7 @@ Proof. now move: wfctx; rewrite on_free_vars_ctx_app /= => /andP[]. exact X. } + red. eapply Forall2_map_right, Forall2_map. - eapply Forall2_All2 in H4. + eapply Forall2_All2 in H5. eapply All2i_All2_mix_left in X8; tea. eapply All2i_nth_hyp in X8. eapply All2_Forall2. eapply All2i_All2; tea; cbv beta. @@ -3723,7 +3737,7 @@ Proof. rewrite !trans_bcontext. now eapply eq_context_upto_names_binder_annot in cd. + eapply All2i_map. eapply All2i_map_right. - eapply Forall2_All2 in H4. + eapply Forall2_All2 in H5. eapply All2i_nth_hyp in X8. eapply All2i_All2_mix_left in X8; tea. eapply All2i_impl ; tea. @@ -3817,19 +3831,19 @@ Proof. unfold map_def. reflexivity. + eapply All_map, (All_impl X1). - intros d (_ & ? & ? & _) => //. - repeat (eexists; tea). + intros d tu. + eapply trans_lift_typing_it with (tu _) => //. + fold trans. subst types. eapply All_map. 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. + eapply All_impl; tea. intros d (onb & ont); cbn in *; tas. + eapply trans_lift_typing_it with (onb _) => // HT. + all: len in HT; len; rewrite -(trans_fix_context _ _ 0 H2). + all: rewrite -trans_local_app /=. + all: rewrite -(trans_lift _ (shiftnP #|Γ| xpred0)) /=. + all: try eapply (isTypeRel_is_open_term (Σ := Σ)); tea. + all: rewrite -(fix_context_length mfix); eapply HT. + rewrite (trans_wf_fixpoint _ (shiftnP #|Γ| xpred0) #|mfix|) //. - rewrite trans_dtype. simpl. @@ -3845,19 +3859,19 @@ Proof. unfold map_def. reflexivity. + eapply All_map, (All_impl X1). - intros d (_ & ? & ? & _) => //. - repeat (eexists; tea). + intros d tu. + eapply trans_lift_typing_it with (tu _) => //. + fold trans. subst types. eapply All_map. 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. + eapply All_impl; tea. intros d (onb & ont); cbn in *; tas. + eapply trans_lift_typing_it with (onb _) => // HT. + all: len in HT; len; rewrite -(trans_fix_context _ _ 0 H2). + all: rewrite -trans_local_app /=. + all: rewrite -(trans_lift _ (shiftnP #|Γ| xpred0)) /=. + all: try eapply (isTypeRel_is_open_term (Σ := Σ)); tea. + all: rewrite -(fix_context_length mfix); eapply HT. + rewrite trans_wf_cofixpoint //. - cbn. rewrite trans_prim_ty. econstructor; eauto. rewrite prim_val_tag_map //. * now eapply trans_declared_constant in H0. @@ -4518,7 +4532,7 @@ Proof. eapply (cumulAlgo_cumulSpec _ (pb:=Cumul)). apply into_ws_cumul_pb. eapply (trans_cumul' (Σ := Σ) (Γ' := Γ' ,,, Γ'0)) => //. - eapply cumul_decorate_hetero; tea. + eapply cumul_decorate_hetero; eauto using has_sort_isType. len. now rewrite (All2_fold_length X) eqlen. now depelim ass. now depelim ass'. eapply cumulSpec_cumulAlgo_curry in eqt; eauto. @@ -4747,8 +4761,7 @@ 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 & _). - repeat (eexists; tea). + { eapply trans_lift_typing_it with on_ctype => // Hty. set (s := _.2.π1) in *; clearbody s. 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. @@ -5871,7 +5884,7 @@ Proof. 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. - apply lift_wf_term_f_impl with (u := j_univ j) (1 := X); intros t (Ht & ont). + apply lift_wf_term_f_impl with (r := j_rel j) (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. diff --git a/pcuic/theories/PCUICFirstorder.v b/pcuic/theories/PCUICFirstorder.v index e0c4d1bed..517f7ee76 100644 --- a/pcuic/theories/PCUICFirstorder.v +++ b/pcuic/theories/PCUICFirstorder.v @@ -218,10 +218,10 @@ Proof using Type. rewrite /subst1 PCUICLiftSubst.subst_it_mkProd_or_LetIn subst_mkApps //. constructor 1. reflexivity. eapply isType_tLetIn_red in i0. 2:pcuic. rewrite /subst1 PCUICLiftSubst.subst_it_mkProd_or_LetIn subst_mkApps in i0. - now eapply isType_open. + now eapply PCUICClosedTyp.isType_is_open_term. -- eapply cumul_Prod_inv in w as []. econstructor. ++ eapply type_ws_cumul_pb. 3: eapply PCUICContextConversion.ws_cumul_pb_eq_le; symmetry. all:eauto. - eapply isType_tProd in i0. eapply i0. + eapply isType_tProd in i0. eapply isTypeRel_isType, i0. ++ rewrite /subst1 PCUICLiftSubst.subst_it_mkProd_or_LetIn. autorewrite with subst. cbn. eapply X. len. eapply typing_spine_strengthen. eauto. @@ -233,10 +233,10 @@ Proof using Type. replace (it_mkProd_or_LetIn (subst_context [hd] 0 Γ0) (mkApps (tInd i u) (map (subst [hd] (#|Γ0| + 0)) pars))) with ((PCUICAst.subst10 hd (it_mkProd_or_LetIn Γ0 (mkApps (tInd i u) pars)))). 2:{ rewrite /subst1 PCUICLiftSubst.subst_it_mkProd_or_LetIn. now autorewrite with subst. } - eapply isType_subst. eapply PCUICSubstitution.subslet_ass_tip. eauto. + eapply PCUICSubstitution.isType_subst. eapply PCUICSubstitution.subslet_ass_tip. eauto. eapply isType_tProd in i0 as [_ tprod]. eapply isType_context_conversion; tea. constructor. eapply ws_cumul_ctx_pb_refl. now eapply typing_wf_local, PCUICClosedTyp.wf_local_closed_context in t. - constructor; tea. constructor. pcuic. eapply validity in t. now eauto. + constructor; tea. constructor. pcuic. now eapply isType_tProd in i1 as []. Qed. Lemma leb_spect : forall x y : nat, BoolSpecSet (x <= y) (y < x) (x <=? y). diff --git a/pcuic/theories/PCUICInductiveInversion.v b/pcuic/theories/PCUICInductiveInversion.v index ccff52183..afaa94f2f 100644 --- a/pcuic/theories/PCUICInductiveInversion.v +++ b/pcuic/theories/PCUICInductiveInversion.v @@ -160,22 +160,18 @@ Lemma type_tCoFix_inv {cf:checker_flags} (Σ : global_env_ext) Γ mfix idx T : w (Σ ;;; Γ |- subst0 (cofix_subst mfix) (dbody d) : (dtype d)) * (Σ ;;; Γ ⊢ dtype d ≤ T). Proof. - intros wfΣ H. depind H. - - exists decl. - 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. - simpl in Hty. - rewrite subst_context_nil /= in Hty. - eapply refine_type; eauto. - rewrite simpl_subst_k //. len. - apply subslet_cofix; auto. - * eapply nth_error_all in a0; tea. cbn in a0. now eapply isType_ws_cumul_pb_refl. - - destruct (IHtyping1 wfΣ) as [d [[[Hnth wfcofix] ?] ?]]. - exists d. intuition auto. - etransitivity; eauto. - now eapply PCUICConversion.cumulSpec_cumulAlgo_curry in c; tea; fvs. + intros wfΣ H. + apply typing_wf_local in H as wfΓ. + apply inversion_CoFix in H as (decl & Hguard & Hnth & Htys & Hbods & Hwf & Hty); tas. + exists decl. repeat (split; tea). + eapply All_nth_error in Hbods as Hbod; tea. + apply unlift_TermTyp in Hbod. + eapply (substitution (s := cofix_subst mfix) (Δ := [])) in Hbod. + 2: now apply subslet_cofix; auto. + simpl in Hbod. + rewrite subst_context_nil /= in Hbod. + eapply refine_type; eauto. + rewrite simpl_subst_k //. len. Qed. Lemma wf_cofixpoint_all {cf:checker_flags} (Σ : global_env_ext) mfix : @@ -314,8 +310,9 @@ Section OnConstructor. clear Hlen'. rewrite [_ ,,, _]app_context_assoc in Hinst. now exists inst. - apply lift_typing_impl with (1 := oib.(onArity)); intros ?? Hs. - now eapply weaken_ctx in Hs. + apply lift_typing_weaken_ctx with (Γ := []); tas. + eapply isTypeRel_isType. + apply oib.(onArity). Qed. End OnConstructor. @@ -544,7 +541,8 @@ Proof. (mkApps (tInd ind i) (map (subst [hd] #|Γ'|) args))). { move: wat; rewrite it_mkProd_or_LetIn_app /= /mkProd_or_LetIn /= => wat. eapply isType_tProd in wat as [isty wat]. - eapply (isType_subst (Δ := [_]) [hd]) in wat. + apply isTypeRel_isType in isty. + eapply (isType_subst (Δ := [_]) (s := [hd])) in wat. now rewrite subst_it_mkProd_or_LetIn Nat.add_0_r subst_mkApps in wat. eapply subslet_ass_tip. eapply type_ws_cumul_pb; tea. now symmetry. } rewrite subst_mkApps Nat.add_0_r in cumulB. simpl in *. @@ -574,6 +572,7 @@ Proof. constructor. constructor. rewrite subst_empty. rewrite it_mkProd_or_LetIn_app /= /mkProd_or_LetIn /= in wat. eapply isType_tProd in wat as [Hty _]; auto. + apply isTypeRel_isType in Hty. eapply type_ws_cumul_pb; eauto. now eapply ws_cumul_pb_eq_le. * pcuic. Qed. @@ -1608,11 +1607,11 @@ Proof. revert cum. induction cpos; simpl; rewrite ?subst_context_nil ?subst_context_snoc; try solve [constructor; auto]. all:rewrite ?map_length; intros cv; depelim cv; depelim wf. - assert (isType Σ + eassert (isTypeRel Σ (subst_instance u (ind_arities mdecl) ,,, subst_instance u (smash_context [] (ind_params mdecl) ,,, Γ)) - (subst_instance u t)). - { rewrite -subst_instance_app_ctx app_context_assoc //. } + (subst_instance u t) _). + { rewrite -subst_instance_app_ctx app_context_assoc //. eassumption. } depelim a. all:constructor. - eapply IHcpos. auto. now depelim ass. apply cv. @@ -3141,6 +3140,7 @@ Proof. 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. + 2: now eapply isTypeRel_isType, (declared_inductive_valid_type decli). exists (mdecl, idecl). assert (lookup_inductive Σ ind = Some (mdecl, idecl)). { destruct decli as [decli declmi]. @@ -3159,8 +3159,6 @@ Proof. eapply arity_typing_spine in Hargs; auto. destruct Hargs as [Hl Hleq ?]. rewrite Hl. len. now rewrite Nat.leb_refl. - eauto with pcuic. - now eapply (declared_inductive_valid_type decli). Qed. Lemma ctx_inst_app_weak `{checker_flags} Σ (wfΣ : wf Σ.1) ind mdecl idecl (isdecl : declared_inductive Σ.1 ind mdecl idecl)Γ (wfΓ : wf_local Σ Γ) params args u v: @@ -3247,13 +3245,6 @@ Proof. now eapply (ctx_inst_app cparams). Qed. -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. repeat (eexists; tea). -Qed. - Lemma isType_it_mkProd_or_LetIn {cf:checker_flags} {Σ Γ Δ T} : wf Σ.1 -> isType Σ (Γ ,,, Δ) T -> @@ -3294,6 +3285,7 @@ Proof. induction wfΔ using All_local_rel_ind1 in nas, ha |- *; depelim ha; cbn. constructor. specialize (IHwfΔ _ ha). apply All_local_rel_snoc; tas. + unfold set_binder_name at 2. cbn. rewrite e. apply lift_typing_impl with (1 := X) => // ?? Hs. eapply context_conversion; tea. 1: now eapply All_local_env_app. @@ -4524,3 +4516,83 @@ Proof. rewrite -eqsubst. apply: red_betas_typed; last eassumption. rewrite (ctx_inst_length instl) context_assumptions_rev //. Qed. + +Lemma ind_arity_relevant {cf : checker_flags} (Σ : global_env_ext) mdecl idecl : + wf Σ.1 -> + let ind_type := it_mkProd_or_LetIn (ind_params mdecl) (it_mkProd_or_LetIn (ind_indices idecl) (tSort (ind_sort idecl))) in + isType Σ [] ind_type -> + isTypeRel Σ [] ind_type rel_of_Type. +Proof. + intros wfΣ ind_type (_ & s & Hs & _ & _). + repeat (eexists; tea). cbn in Hs |- *. + rewrite /ind_type in Hs. + rewrite -it_mkProd_or_LetIn_app in Hs. + eapply PCUICSpine.inversion_it_mkProd_or_LetIn in Hs. + apply inversion_Sort in Hs as (_ & _ & le); tea. + apply ws_cumul_pb_Sort_inv in le. + eapply leq_relevance; tea. + eapply relevance_super. +Qed. + + +Lemma isTypeRel_cstr_type {cf : checker_flags} Σ mdecl idecl cdecl n : + wf Σ.1 -> + nth_error (ind_bodies mdecl) n = Some idecl -> + (cstr_type cdecl = + it_mkProd_or_LetIn + (mdecl.(ind_params) ,,, cdecl.(cstr_args)) + (mkApps (tRel (#|mdecl.(ind_params) ,,, cdecl.(cstr_args)| + (#|ind_bodies mdecl| - S n))) + (to_extended_list_k mdecl.(ind_params) #|cdecl.(cstr_args)| ++ + cdecl.(cstr_indices)))) -> + idecl.(ind_type) + = it_mkProd_or_LetIn mdecl.(ind_params) + (it_mkProd_or_LetIn idecl.(ind_indices) (tSort idecl.(ind_sort))) -> + isSortRel idecl.(ind_sort) idecl.(ind_relevance) -> + ctx_inst Σ (arities_context mdecl.(ind_bodies) ,,, mdecl.(ind_params) ,,, cdecl.(cstr_args)) + cdecl.(cstr_indices) + (List.rev (lift_context #|cdecl.(cstr_args)| 0 idecl.(ind_indices))) -> + isType Σ (arities_context mdecl.(ind_bodies)) (cstr_type cdecl) -> + isTypeRel Σ (arities_context mdecl.(ind_bodies)) (cstr_type cdecl) idecl.(ind_relevance). +Proof. + intros wfΣ H -> e1 <- c (_ & s & Hs & _ & _). + repeat (eexists; tea). cbn in Hs |- *. + eapply PCUICSpine.inversion_it_mkProd_or_LetIn in Hs. + eapply PCUICValidity.inversion_mkApps in Hs as [tyapp [Happ Hsp]]. + apply inversion_Rel in Happ as (decl & wfΓ & hnth & cum); tas. + eapply PCUICWeakeningTyp.All_local_env_nth_error_weaken in wfΓ as Hty; tea. + rewrite nth_error_app_ge in hnth. lia. + replace (#|ind_params mdecl,,, cstr_args cdecl| + (#|ind_bodies mdecl| - S n) - #|ind_params mdecl,,, cstr_args cdecl|) with (#|ind_bodies mdecl| - S n) in hnth by lia. + pose proof (nth_error_Some_length H) as hlen. rewrite nth_error_rev // in H. + eapply nth_error_arities_context in H. rewrite hnth in H. + eapply PCUICSpine.typing_spine_strengthen in Hsp; tea. noconf H. clear cum hnth. + cbn in Hsp. + rewrite e1 !PCUICLiftSubst.lift_it_mkProd_or_LetIn -it_mkProd_or_LetIn_app /= in Hsp. + eapply PCUICSpine.arity_typing_spine in Hsp as [_ leqs _]. now eapply leq_relevance; tea. + noconf H. cbn in Hty |- *. now apply isTypeRel_isType in Hty. +Qed. + +Lemma isType_of_constructor {cf : checker_flags} Σ mdecl idecl cdecl cstr u Γ : + wf Σ.1 -> wf_local Σ Γ -> declared_constructor Σ.1 cstr mdecl idecl cdecl -> + consistent_instance_ext Σ (ind_universes mdecl) u -> + isTypeRel Σ Γ (type_of_constructor mdecl cdecl cstr u) idecl.(ind_relevance). +Proof. + intros wfΣ wfΓ isdecl h_cuu. + pose proof (PCUICWeakeningEnvTyp.on_declared_constructor isdecl) as ((? & []) & ? & _ & []). + unshelve eapply declared_constructor_to_gen in isdecl as isdeclgen; tea. + eapply @lift_sorting_fu_it_impl with (f := fun x => _ (_ x)) (fu := fun x => x@[u]) (tu := on_ctype) => //=. + 1: apply relevance_subst. + intro Hty. + replace Γ with ([] ,,, Γ) by apply app_context_nil_l. + replace (subst0 _ _) with (lift0 #|Γ| (type_of_constructor mdecl cdecl cstr u)). + 2: { rewrite PCUICLiftSubst.lift_closed //. eapply (PCUICClosedTyp.declared_constructor_closed_type isdecl). } + change (tSort _@[u]) with (lift0 #|Γ| (tSort on_ctype.2.π1)@[u]). + eapply @weakening with (Γ := []) => //. + rewrite app_context_nil_l => //. + destruct Σ as (Σ & φ). + eapply typing_subst_instance' in Hty => //=; tea. + change ((tSort _)@[u]) with (subst0 (inds (inductive_mind cstr.1) u (ind_bodies mdecl)) (tSort on_ctype.2.π1)@[u]). + eapply @PCUICSubstitution.substitution with (Δ := []); tea. + 2: rewrite app_context_nil_l; apply Hty. + 2: { epose proof (weaken_lookup_on_global_env' _ _ (InductiveDecl mdecl) wfΣ _); eauto. now split. Unshelve. 2: apply isdeclgen.p1.p1. } + eapply subslet_inds; tea. exact isdecl. +Qed. diff --git a/pcuic/theories/PCUICInductives.v b/pcuic/theories/PCUICInductives.v index da29bb386..ee1f86ce9 100644 --- a/pcuic/theories/PCUICInductives.v +++ b/pcuic/theories/PCUICInductives.v @@ -281,35 +281,30 @@ Section OnInductives. Lemma on_inductive_inst Γ u : 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_sort u idecl.(ind_sort)))). + isTypeRel Σ Γ (it_mkProd_or_LetIn (subst_instance u (ind_params mdecl ,,, idecl.(ind_indices))) + (tSort (subst_instance_sort u idecl.(ind_sort)))) rel_of_Type. 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 & _). - eapply isType_weaken => //. + eapply lift_typing_weaken_ctx with (Γ := []) => //. rewrite -(subst_instance_it_mkProd_or_LetIn u _ (tSort _)). rewrite -it_mkProd_or_LetIn_app in ar. - eapply (typing_subst_instance_decl Σ [] _ _ _ (InductiveDecl mdecl) u wfΣ) in ar. - all:pcuic. + eapply (lift_typing_subst_instance_decl Σ [] _ _ (InductiveDecl mdecl) u wfΣ) in ar; cbn in ar; tas. unshelve eapply declared_inductive_to_gen in decli; eauto. Qed. Lemma declared_inductive_valid_type Γ u : wf_local Σ Γ -> consistent_instance_ext Σ (ind_universes mdecl) u -> - isType Σ Γ (subst_instance u idecl.(ind_type)). + isTypeRel Σ Γ (subst_instance u idecl.(ind_type)) rel_of_Type. Proof using decli wfΣ. move=> wfΓ cext. destruct (on_declared_inductive decli) as [onmind oib]. - pose proof (oib.(onArity)) as ar. - destruct ar as (_ & s & ar & _). - eapply isType_weaken => //. - eapply (typing_subst_instance_decl Σ [] _ _ _ (InductiveDecl mdecl) u wfΣ) in ar. - all:pcuic. - unshelve eapply declared_inductive_to_gen in decli; eauto. + rewrite oib.(ind_arity_eq). + rewrite -it_mkProd_or_LetIn_app subst_instance_it_mkProd_or_LetIn /=. + now apply on_inductive_inst. Qed. Local Definition oi := (on_declared_inductive decli).1. @@ -363,7 +358,7 @@ Section OnInductives. split. rewrite -subst_instance_it_mkProd_or_LetIn. rewrite -(declared_inductive_type decli). - eapply declared_inductive_valid_type; tea. + eapply isTypeRel_isType, declared_inductive_valid_type; tea. apply sp. eapply arity_spine_it_mkProd_or_LetIn_Sort. 2:reflexivity. 2:exact sp. @@ -630,7 +625,7 @@ Proof. eapply (substitution (Δ := [])) in l. eapply l. all:auto. * rewrite smash_context_acc. simpl. rewrite /map_decl /= /map_decl /=. simpl. - assert (isType Σ (smash_context [] Δ) (subst0 (extended_subst Δ 0) decl_type)). + eassert (isTypeRel Σ (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. } @@ -706,14 +701,14 @@ Proof. have wfsmpars : wf_local (Σ.1, ind_universes mdecl) (smash_context [] (ind_params mdecl)). { now apply wf_local_smash_context. } constructor => //. - eapply has_sort_isType. + eapply has_sort_isTypeRel. 1: apply oib.(ind_relevance_compat). eapply type_mkApps. econstructor; eauto. eapply consistent_instance_ext_abstract_instance; eauto with pcuic. eapply declared_inductive_wf_global_ext; eauto with pcuic. set (u := abstract_instance (ind_universes mdecl)). destruct hdecl; eauto. assert (isType (Σ.1, ind_universes mdecl) [] (ind_type idecl)). - { pose proof (onArity oib). exact X. } + { eapply isTypeRel_isType. pose proof (onArity oib). exact X. } rewrite (subst_instance_ind_type_id _ _ _ _ _ hdecl). rewrite (ind_arity_eq oib). rewrite -(app_nil_r (to_extended_list _)). @@ -1223,14 +1218,15 @@ 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. eapply has_sort_isType. + { constructor. auto. eapply has_sort_isTypeRel. + 1: apply oib.(ind_relevance_compat). eapply type_mkApps. econstructor; eauto. destruct isdecl as []; eauto. subst u. eapply consistent_instance_ext_abstract_instance; eauto with pcuic. eapply declared_inductive_wf_global_ext; eauto with pcuic. destruct decli; eauto. assert (isType (Σ.1, ind_universes mdecl) [] (ind_type idecl)@[u]). - { rewrite typeu. pose proof (onArity oib). exact X1. } + { eapply isTypeRel_isType. rewrite typeu. pose proof (onArity oib). exact X1. } rewrite (ind_arity_eq oib). rewrite subst_instance_it_mkProd_or_LetIn. rewrite -(app_nil_r (to_extended_list _)). @@ -1411,7 +1407,7 @@ Proof. 2:{ apply wf_local_smash_end; auto. } eapply wf_local_smash_end 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). + epose proof (wf_local_nth_isTypeRel (d:=arg) (n:=0) wfargs). rewrite skipn_app in X1. rewrite skipn_length in X1. len. rewrite nth_error_app_lt // in X1. rewrite skipn_length. len. len. @@ -1424,6 +1420,7 @@ Proof. rewrite H1 in X1. rewrite [skipn _ (_ ,,, _)]skipn_0_eq in X1. assert (context_assumptions (cstr_args c) > 0). lia. lia. + apply isTypeRel_isType in X1. apply X1. Qed. @@ -1532,7 +1529,8 @@ 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]. + destruct l as (_ & ss & Hs & <- & Her). + split; cbn; eexists; split; [|split; eauto using leq_relevance]. eapply type_Cumul; eauto. econstructor; pcuic. now eapply cumul_Sort. @@ -1635,7 +1633,8 @@ Proof. assert(wfsmash : wf_local Σ (smash_context [] (subst_instance u (ind_params mdecl)))). { eapply wf_local_smash_context; auto. } constructor; auto. red. - eapply @has_sort_isType with (s := subst_instance_sort u (ind_sort idecl)). + eapply has_sort_isTypeRel with (s := subst_instance_sort u (ind_sort idecl)). + 1: rewrite relevance_subst_eq; apply oib0.(ind_relevance_compat). eapply type_mkApps. econstructor; eauto. eapply decli.p1. rewrite (ind_arity_eq oib0). pose proof (on_projs_noidx _ _ _ _ _ _ onps). @@ -1644,6 +1643,7 @@ Proof. eapply typing_spine_to_extended_list; eauto. pose proof (on_inductive_inst decli _ u localenv_nil). rewrite eqind in X. simpl in X. + eapply isTypeRel_isType. now rewrite subst_instance_it_mkProd_or_LetIn. Qed. @@ -2341,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 & _ & Her). 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]. @@ -2357,7 +2357,7 @@ Proof. rewrite subst_instance_mkApps /=; cbn. erewrite subst_instance_id_mdecl; tea. eapply pre_type_mkApps_arity. econstructor; tea. - eapply declared_inductive_valid_type; tea. + eapply isTypeRel_isType, declared_inductive_valid_type; tea. pose proof (on_declared_inductive isdecl) as [onmind oib]. rewrite oib.(ind_arity_eq) !subst_instance_it_mkProd_or_LetIn subst_instance_to_extended_list subst_instance_app_ctx @@ -2415,7 +2415,10 @@ 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:{ eapply has_sort_isType. rewrite /idxctx app_context_assoc in tyass; tea. } + 2:{ eapply has_sort_isTypeRel. + 2: rewrite /idxctx app_context_assoc in tyass; tea. + rewrite relevance_subst_eq. + pose proof (on_declared_inductive isdecl) as [onmind oib]. apply oib.(ind_relevance_compat). } apply typing_wf_local in tyass. rewrite /idxctx app_context_assoc in tyass => //. Qed. @@ -2435,6 +2438,7 @@ Lemma isType_case_predicate {cf : checker_flags} {Σ : global_env_ext} {wfΣ : w (tSort ps)). Proof. move=> wfΓ isdecl cu wfps sp. + pose proof (on_declared_inductive isdecl) as [onmind oib]. rewrite /pre_case_predicate_context_gen /inst_case_context /ind_predicate_context /=. set (iass := {| decl_name := _ |}). cbn. @@ -2460,8 +2464,7 @@ Proof. erewrite subst_instance_id_mdecl; tea. rewrite subst_mkApps. cbn. eapply pre_type_mkApps_arity. econstructor; tea. - eapply declared_inductive_valid_type; tea. - pose proof (on_declared_inductive isdecl) as [onmind oib]. + eapply isTypeRel_isType, declared_inductive_valid_type; tea. rewrite oib.(ind_arity_eq) !subst_instance_it_mkProd_or_LetIn subst_instance_to_extended_list subst_instance_app_ctx subst_instance_smash. @@ -2515,11 +2518,12 @@ Proof. set (sdecl := subst_decl _ _ _). rewrite -(it_mkProd_or_LetIn_app [sdecl]). eapply type_it_mkProd_or_LetIn_sorts; tea. - constructor => //. repeat (eexists; tea). + constructor => //. repeat (eexists; tea). apply relevance_subst, oib.(ind_relevance_compat). constructor => //. simpl. constructor => //. now eapply sorts_local_ctx_wf_local; tea. repeat (eexists; tea). + apply relevance_subst, oib.(ind_relevance_compat). Qed. Lemma arity_spine_case_predicate {cf: checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} {Γ} {ci : case_info} diff --git a/pcuic/theories/PCUICInversion.v b/pcuic/theories/PCUICInversion.v index ec70dbba1..b08606ea2 100644 --- a/pcuic/theories/PCUICInversion.v +++ b/pcuic/theories/PCUICInversion.v @@ -255,6 +255,7 @@ 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) + (elim_relevance : isSortRel ps ci.(ci_relevance)) (ind_inst : ctx_inst (typing Σ) Γ (p.(pparams) ++ indices) (List.rev (subst_instance p.(puinst) (ind_params mdecl ,,, ind_indices idecl : context)))) diff --git a/pcuic/theories/PCUICPrincipality.v b/pcuic/theories/PCUICPrincipality.v index 81654c126..4580759d6 100644 --- a/pcuic/theories/PCUICPrincipality.v +++ b/pcuic/theories/PCUICPrincipality.v @@ -90,6 +90,7 @@ Section Principality. int inversion_Sort. repeat outsum. repeat outtimes. now subst. - apply inversion_Prod in hA as (dom1 & codom1 & t & t0 & w); auto. + pose proof (t.2.π2.2.p2) as Her. rewrite -t.2.π2.2.p1 /= in Her. apply unlift_TypUniv in t. specialize (IHu1 _ _ t) as [dom Hdom]. specialize (IHu2 _ _ t0) as [codom Hcodom]. @@ -114,7 +115,9 @@ Section Principality. 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. - + repeat (eexists; eauto). eapply type_reduction; eauto. eapply red. + + repeat (eexists; eauto). + 2: eapply geq_relevance, Her; tea. + eapply type_reduction; eauto. eapply red. + eapply type_reduction; eauto. eapply cored. - apply inversion_Lambda in hA => //; eauto. @@ -180,6 +183,7 @@ Section Principality. eapply (type_ws_cumul_pb (pb:=Cumul)); eauto. { eapply validity in t0; auto. eapply isType_red in t0; [|exact redA]. + eapply isTypeRel_isType. eapply isType_tProd in t0 as [? ?]; eauto. } transitivity dom' => //. transitivity A''. all:apply ws_cumul_pb_eq_le; symmetry => //. @@ -364,7 +368,7 @@ Lemma principal_type_ind {cf:checker_flags} {Σ Γ c ind ind' u u' args args'} { (∑ ui', cmp_ind_universes Σ ind #|args| ui' u * cmp_ind_universes Σ ind' #|args'| ui' u') * - ws_cumul_pb_terms Σ Γ args args' * + ws_cumul_pb_terms Σ Γ args args' * (ind = ind'). Proof. intros h h'. @@ -375,7 +379,7 @@ Proof. eapply invert_red_mkApps_tInd in redl as [args'' [-> eq0]]; auto. eapply invert_red_mkApps_tInd in redr as [args''' [eqnf eq1]]; auto. solve_discr. - repeat split; eauto. + repeat split; eauto. assert (#|args| = #|args'|). now rewrite -(All2_length eqargs) -(All2_length eqargs') (All2_length a) (All2_length a0). transitivity l'. now symmetry. @@ -484,7 +488,7 @@ Proof. 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. } + { specialize (X1 onu). rewrite e. 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. 1: now eapply lift_sorting_forget_univ. constructor; try now symmetry; now constructor. @@ -699,7 +703,7 @@ Proof. - eapply inversion_Fix in X4 as (decl' & fixguard' & Hnth & types' & bodies & wffix & cum); auto. eapply type_Cumul_alt. econstructor; eauto. - now eapply All_nth_error in X0. + eapply isTypeRel_isType; 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. @@ -708,7 +712,7 @@ Proof. - eapply inversion_CoFix in X4 as (decl' & fixguard' & Hnth & types' & bodies & wfcofix & cum); auto. eapply type_Cumul_alt. econstructor; eauto. - now eapply All_nth_error in X0. + eapply isTypeRel_isType; 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. diff --git a/pcuic/theories/PCUICSR.v b/pcuic/theories/PCUICSR.v index 4642b5c6f..8ecc8efbe 100644 --- a/pcuic/theories/PCUICSR.v +++ b/pcuic/theories/PCUICSR.v @@ -1521,7 +1521,7 @@ Proof. 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 lift_sorting_forget_rel, 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. @@ -1626,7 +1626,7 @@ Proof. - (* LetIn value *) assert (X1' : lift_typing typing Σ Γ (j_vdef na r b_ty)). - { unshelve eapply lift_judgment_SR with (1 := X1); tea. firstorder. } + { unshelve eapply lift_judgment_SR with (1 := X1); tea. split => //. right; now split. } eapply type_Cumul_alt. econstructor; eauto. unshelve eapply (closed_context_conversion _ typeb'). pcuic. @@ -1639,7 +1639,7 @@ Proof. - (* LetIn type annotation *) assert (X1' : lift_typing typing Σ Γ (j_vdef na b r)). - { unshelve eapply lift_judgment_SR with (1 := X1); tea. firstorder. } + { unshelve eapply lift_judgment_SR with (1 := X1); tea. split => //. left; now split. } eapply type_Cumul_alt. econstructor; eauto. unshelve eapply (closed_context_conversion _ typeb'). pcuic. @@ -1763,7 +1763,7 @@ Proof. set (ibrctx := (case_branch_context ci mdecl p (forget_types (bcontext br)) cdecl)) in *. set (brctx := (inst_case_context (pparams p) (puinst p) (bcontext br))). assert (wfbr : wf_branch cdecl br). - { eapply Forall2_All2, All2_nth_error in H4; tea. + { eapply Forall2_All2, All2_nth_error in H5; tea. eapply declc. } assert(eqctx : eq_context_upto_names ibrctx brctx). { rewrite /brctx /ibrctx. @@ -2343,8 +2343,8 @@ Proof. eapply type_ws_cumul_pb; tea. * eapply type_Case; eauto. constructor; eauto. constructor; eauto. epose proof (wf_case_branches_types' (p:=set_preturn p preturn') ps _ brs isdecl (validity typec) H0 - (forall_u _ X3) H4 X1). - eapply All2i_All2_mix_left in X8. 2:exact (Forall2_All2 _ _ H4). clear H4. + (forall_u _ X3) H5 X1). + eapply All2i_All2_mix_left in X8. 2:exact (Forall2_All2 _ _ H5). clear H5. eapply (All2i_All2i_mix X4) in X8. clear X4. eapply (All2i_impl X8); intuition auto; clear X8. rewrite !case_branch_type_fst in a3 a4 *. @@ -2365,7 +2365,7 @@ Proof. now apply ws_cumul_ctx_pb_refl, wf_local_closed_context. * eapply ws_cumul_pb_eq_le, ws_cumul_pb_mkApps; tea. 2:{ eapply ws_cumul_pb_terms_refl => //. - eapply isType_open in X0. + eapply isType_is_open_term in X0. rewrite on_free_vars_mkApps in X0. now move/andP: X0=> []. } eapply ws_cumul_pb_it_mkLambda_or_LetIn; tea. @@ -2395,12 +2395,12 @@ Proof. apply All2_tip. eapply (cumulAlgo_cumulSpec _ (pb:=Conv)). symmetry. now eapply closed_red1_ws_cumul_pb. + now apply wf_local_closed_context. - + eapply isType_open in X0. + + eapply isType_is_open_term in X0. rewrite on_free_vars_mkApps in X0. move/andP: X0=> []. generalize (closed_red1_open_right Hu). rewrite !forallb_app /= => clc' clptm /and3P[] -> /=. now rewrite clc'. - + eapply isType_open in X0. rewrite on_free_vars_mkApps in X0. + + eapply isType_is_open_term in X0. rewrite on_free_vars_mkApps in X0. now move/andP: X0 => []. - (* Case congruence on branches *) @@ -2411,14 +2411,14 @@ Proof. clear X5; rename Hctxi into X5. eapply type_Case; eauto. econstructor; eauto. econstructor; eauto. - * eapply Forall2_All2 in H4. - move: (All2_sym _ _ _ H4) => wfb. + * eapply Forall2_All2 in H5. + move: (All2_sym _ _ _ H5) => wfb. red. eapply All2_Forall2. apply All2_sym. eapply (OnOne2_All2_All2 X3 wfb); auto. intros [] []; simpl. intros. - destruct X0 as [_ eq]. subst bcontext0. exact H5. - * apply Forall2_All2 in H4. eapply All2i_All2_mix_left in X8; tea. + destruct X0 as [_ eq]. subst bcontext0. exact H4. + * apply Forall2_All2 in H5. eapply All2i_All2_mix_left in X8; tea. eapply (OnOne2_All2i_All2i X3 X8). intros n [] []; simpl. intros. intuition auto. intros n [ctx b] [ctx' b'] cdecl; cbn. @@ -2518,7 +2518,7 @@ Proof. eapply (declared_inductive_closed_params_inst isdecl). move/isType_subst_extended_subst. move/(isType_weaken wfΓ); rewrite app_context_assoc. - move/(isType_subst_gen _ sppars). + move/(isType_substitution sppars). rewrite -skipn_subst_instance - !skipn_subst_context. rewrite -(subst_context_smash_context _ _ []). rewrite subst_instance_smash /=. @@ -2527,7 +2527,7 @@ Proof. assert (context_assumptions (cstr_args cdecl) - (context_assumptions (cstr_args cdecl) - p.(proj_arg)) = p.(proj_arg)) by lia. rewrite H. - move/isType_open. len. + move/isType_is_open_term. len. rewrite !skipn_length; len. rewrite H //. rewrite subst_instance_subst // /indsubst. rewrite subst_instance_inds. @@ -2796,7 +2796,7 @@ Proof. rewrite !subst_instance_app_ctx -app_context_assoc. intro foo. pose (fun x y z => isType_subst_arities isdecl cu (foo x y z)). move:i. clear foo. - intro foo. pose (fun x y z => isType_open (foo x y z)). + intro foo. pose (fun x y z => isType_is_open_term _ _ (foo x y z)). move:i. clear foo. len. rewrite skipn_length; len. assert (context_assumptions (cstr_args cdecl) - @@ -2823,10 +2823,10 @@ Proof. now apply closed_red1_red. * rewrite on_free_vars_ctx_app andb_true_r on_free_vars_ctx_app H /=. cbn. rewrite shiftnP_add /=. - move/validity/isType_open: typec => -> //. + move/validity/isType_is_open_term: typec => -> //. * rewrite on_free_vars_ctx_app andb_true_r on_free_vars_ctx_app H /=. cbn. rewrite shiftnP_add /=. - move/validity/isType_open: typec => -> //. + move/validity/isType_is_open_term: typec => -> //. * have := (declared_projection_type_and_eq wf isdecl). move=> [[hctors isTy] Hdecl]. move/validity/(isType_mkApps_Ind_proj_inv _ isdecl): typec => [sppars hpars hargs cu]. @@ -2837,7 +2837,7 @@ Proof. rewrite [_@[u]](subst_instance_app_ctx _ _ [_]). rewrite app_context_assoc subst_instance_smash. move/(isType_substitution sppars) => /=. len. - move/isType_open => /=. now cbn. + move/isType_is_open_term => /=. now cbn. - (* Fix congruence *) apply mkApps_Fix_spec in H5. simpl in H5. subst args. @@ -2869,7 +2869,7 @@ Proof. { apply (OnOne2_All_All X4 X1). * intros d HT. apply lift_typing_impl with (1 := HT); now intros ?? [Hty _]. - * intros d d' [[red _] eq] HT'. + * intros d d' [[red _] eq] HT'. noconf eq. rewrite /on_def_type -H3. apply lift_sorting_it_impl_gen with HT' => //; now intros [_ IH]. } assert (wf_local Σ (Γ ,,, fix_context mfix1)). { apply All_mfix_wf; auto. } @@ -2889,9 +2889,9 @@ Proof. + move=> [na ty b rarg] [na' ty' b' rarg'] [[[red Hred] eq] HT] HB. rewrite /= in red, Hred, eq, HB, HT. noconf eq. - destruct HB as ((Hb & IHb) & s & (Ht & IHt) & e); cbn in *. + destruct HB as ((Hb & IHb) & s & (Ht & IHt) & e & er); cbn in *. eassert (Σ ;;; Γ ,,, _ |- lift0 _ ty' : tSort s). - 2: repeat (eexists; cbn). + 2: repeat (eexists; cbn). 4: eassumption. 2,3: eapply context_conversion; eauto. all: rewrite -fixl. { eapply IHt. eapply @weakening_closed_red1 with (Γ' := []); eauto. now eapply wf_local_closed_context. } @@ -2908,7 +2908,7 @@ Proof. * eapply wf_fixpoint_red1_type; eauto. eapply OnOne2_impl; tea; cbn; intuition auto. apply a2. apply a2. - * eapply All_nth_error in X0; eauto. + * eapply isTypeRel_isType. 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. @@ -2927,7 +2927,7 @@ Proof. * intros d HT. apply lift_typing_impl with (1 := HT); now intros ?? [Hty _]. * intros d d' [[red _] eq] HT'. - noconf eq. unfold on_def_type in HT' |- *. rewrite -H4. + noconf eq. unfold on_def_type in HT' |- *. rewrite -H3 -H4. apply lift_typing_impl with (1 := HT'); now intros ?? []. } assert (wf_local Σ (Γ ,,, fix_context mfix1)). { apply All_mfix_wf; auto. } @@ -2950,7 +2950,7 @@ Proof. * eapply wf_fixpoint_red1_body; eauto. eapply OnOne2_impl; tea; cbn; intuition auto. apply a2. apply a2. - * eapply All_nth_error in X0; eauto. + * eapply isTypeRel_isType. eapply All_nth_error in X0; eauto. * apply conv_cumul, conv_sym. destruct disj as [<-|[_ eq]]. reflexivity. noconf eq. rewrite H4; reflexivity. @@ -2979,7 +2979,7 @@ Proof. { apply (OnOne2_All_All X4 X1). * intros d HT. apply lift_typing_impl with (1 := HT); now intros ?? [Hty _]. - * intros d d' [[red _] eq] HT'. + * intros d d' [[red _] eq] HT'. noconf eq. rewrite /on_def_type -H3. apply lift_sorting_it_impl_gen with HT' => //. now intros [_ IH]. } assert (wf_local Σ (Γ ,,, fix_context mfix1)). { apply All_mfix_wf; auto. } @@ -2999,9 +2999,9 @@ Proof. + move=> [na ty b rarg] [na' ty' b' rarg'] [[[red Hred] eq] HT] HB. rewrite /= in red, Hred, eq, HB, HT. noconf eq. - destruct HB as ((Hb & IHb) & s & (Ht & IHt) & e); cbn in *. + destruct HB as ((Hb & IHb) & s & (Ht & IHt) & e & er); cbn in *. eassert (Σ ;;; Γ ,,, _ |- lift0 _ ty' : tSort s). - 2: repeat (eexists; cbn). + 2: repeat (eexists; cbn). 4: eassumption. 2,3: eapply context_conversion; eauto. all: rewrite -fixl. { eapply IHt. eapply @weakening_closed_red1 with (Γ' := []); eauto. now eapply wf_local_closed_context. } @@ -3018,7 +3018,7 @@ Proof. * eapply (wf_cofixpoint_red1_type _ Γ); eauto. eapply OnOne2_impl; tea; cbn; intuition auto; now eapply closed_red1_red. - * eapply All_nth_error in X0; eauto. + * eapply isTypeRel_isType. 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. @@ -3037,7 +3037,7 @@ Proof. * intros d HT. apply lift_typing_impl with (1 := HT); now intros ?? [Hty _]. * intros d d' [[red _] eq] HT'. - noconf eq. unfold on_def_type in HT' |- *. rewrite -H4. + noconf eq. unfold on_def_type in HT' |- *. rewrite -H3 -H4. apply lift_typing_impl with (1 := HT'); now intros ?? []. } assert (wf_local Σ (Γ ,,, fix_context mfix1)). { apply All_mfix_wf; auto. } @@ -3060,7 +3060,7 @@ Proof. * eapply wf_cofixpoint_red1_body; eauto. eapply OnOne2_impl; tea; cbn; intuition auto. apply a2. apply a2. - * eapply All_nth_error in X0; eauto. + * eapply isTypeRel_isType. eapply All_nth_error in X0; eauto. * apply conv_cumul, conv_sym. destruct disj as [<-|[_ eq]]. reflexivity. noconf eq. rewrite H4; reflexivity. diff --git a/pcuic/theories/PCUICSafeLemmata.v b/pcuic/theories/PCUICSafeLemmata.v index 7bc9cffb9..42ae93666 100644 --- a/pcuic/theories/PCUICSafeLemmata.v +++ b/pcuic/theories/PCUICSafeLemmata.v @@ -346,7 +346,7 @@ Section Lemmata. destruct mfix as ((?&[])&?); simpl in *. + eapply All_app in a as (_&a). depelim a. - eauto using isType_welltyped. + eauto using isTypeRel_welltyped. + eapply All_app in a0 as (_&a0). depelim a0. destruct o as (t0 & _); cbn in t0. rewrite fix_context_fix_context_alt in t0. @@ -358,7 +358,7 @@ Section Lemmata. destruct mfix as ((?&[])&?); simpl in *. + eapply All_app in a as (_&a). depelim a. - eauto using isType_welltyped. + eauto using isTypeRel_welltyped. + eapply All_app in a0 as (_&a0). depelim a0. destruct o as (t0 & _); cbn in t0. rewrite fix_context_fix_context_alt in t0. @@ -443,9 +443,9 @@ Section Lemmata. eapply closed_context_conversion; tea. now symmetry. - apply unlift_TypUniv in l. now econstructor. - - now eapply isType_welltyped. + - now eapply isTypeRel_welltyped. - apply unlift_TermTyp in l. now econstructor. - - apply lift_sorting_forget_body in l. now eapply isType_welltyped. + - apply lift_sorting_forget_body in l. now eapply isTypeRel_welltyped. - eapply inversion_Prim in typ as (?&?&[]); eauto. depelim p0. now eexists. - eapply inversion_Prim in typ as (?&?&[]); eauto. diff --git a/pcuic/theories/PCUICSpine.v b/pcuic/theories/PCUICSpine.v index 52e642bea..b36593a0e 100644 --- a/pcuic/theories/PCUICSpine.v +++ b/pcuic/theories/PCUICSpine.v @@ -79,15 +79,12 @@ Lemma wf_local_alpha {cf} {Σ} {wfΣ : wf Σ} Γ Γ' : eq_context_upto_names Γ wf_local Σ Γ -> wf_local Σ Γ'. Proof. - 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. + induction 1; intros h; depelim h; try constructor; auto. + all:depelim r; constructor; subst; auto. + all: rewrite -e. + all: eapply lift_typing_impl; tea; intros t' T Hty. + all: eapply context_conversion; eauto. + all: now apply eq_context_upto_names_conv_context. Qed. Lemma subslet_eq_context_alpha_dom {cf} {Σ} {wfΣ : wf Σ} {Γ Γ' s Δ} : @@ -388,12 +385,10 @@ Section WfEnv. rewrite subst_context_snoc /= /subst_decl /map_decl /=; simpl; intuition auto. 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. - all: simpl in Hs; rewrite Nat.add_0_r in Hs. - all: now rewrite app_context_length in Hs. + all: rewrite -app_context_assoc in b0. + all: eapply lift_typing_substitution in b0; tea. + all: rewrite subst_context_app app_context_assoc Nat.add_0_r in b0. + all: now rewrite app_context_length in b0. Qed. Lemma subst_sorts_local_ctx {Γ Γ' Δ Δ' s ctxs} : @@ -406,12 +401,10 @@ Section WfEnv. destruct a as [na [b|] ty]; rewrite subst_context_snoc /= /subst_decl /map_decl /=; intuition auto. 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. + all: rewrite -app_context_assoc in b0. + all: eapply lift_typing_substitution in b0; tea. + all: rewrite subst_context_app app_context_assoc Nat.add_0_r in b0. + all: now rewrite app_context_length in b0. Qed. Lemma wf_arity_spine_typing_spine {Γ T args T'} : @@ -502,7 +495,7 @@ Section WfEnv. rewrite context_assumptions_subst in IHn. assert (Σ ;;; Γ |- hd0 : ty). { eapply type_ws_cumul_pb; tea. - now eapply isType_tProd in typ as []. + eapply isType_tProd in typ as []; eauto with pcuic. symmetry; pcuic. } eapply typing_spine_strengthen in sp. 3:eapply cumulB. @@ -1657,6 +1650,7 @@ Section WfEnv. constructor. auto. rewrite -eql nth_error_app_lt ?app_length /=; try lia. rewrite nth_error_app_ge // ?Nat.sub_diag //. + apply lift_sorting_forget_rel in l0. apply lift_sorting_f_it_impl with (f := fun t => _ (_ t)) (tu := l0) => // Hs. set (s := l0.2.π1). change (tSort s) with @@ -1998,7 +1992,7 @@ Section WfEnv. eapply (substitution_ws_cumul_pb_vass (a:=hd0)) in cum; auto. assert (Σ ;;; Γ |- hd0 : decl_type). { eapply (type_ws_cumul_pb (pb:=Conv)); tea. 2:now symmetry. - now eapply isType_tProd in isty as []. } + now eapply isType_tProd in isty as []; eauto with pcuic. } eapply isType_apply in isty; tea. eapply typing_spine_strengthen in sp. 3:tea. 2:tas. rewrite /subst1 subst0_it_mkProd_or_LetIn in sp; auto. @@ -2092,15 +2086,15 @@ Section WfEnv. pose proof (isType_wf_local isty). eapply ws_cumul_pb_Prod_Prod_inv in e as [conv cum]; auto. eapply (type_ws_cumul_pb (pb:=Conv)); eauto. - eapply isType_tProd in isty as [dom codom]; auto. cbn in *. + eapply isType_tProd in isty as [dom codom]; eauto with pcuic. cbn in *. now symmetry. * move=> Hnth Hn'. pose proof (isType_wf_local isty). eapply isType_tProd in isty as [dom' codom']; auto. cbn in *. eapply ws_cumul_pb_Prod_Prod_inv in e as [conv cum e]; auto. simpl in codom'. assert (Σ ;;; Γ |- hd0 : ty). - { eapply (type_ws_cumul_pb (pb:=Conv)); eauto. now symmetry. } - unshelve eapply (isType_subst (Δ:=[vass na ty]) [hd0]) in codom'. + { eapply (type_ws_cumul_pb (pb:=Conv)); eauto with pcuic. now symmetry. } + unshelve eapply (isType_subst (Δ:=[vass na ty]) (s := [hd0])) in codom'. 2:{ now eapply subslet_ass_tip. } specialize (X (subst_context [hd0] 0 Γ0) ltac:(autorewrite with len; lia)). eapply substitution_ws_cumul_pb_vass in e; tea. @@ -2368,7 +2362,7 @@ Section WfEnv. specialize (IHa wf wf' subsl). constructor; auto. eapply type_ws_cumul_pb; eauto. depelim p. - eapply isType_subst. exact IHa. eauto. + eapply isType_subst. exact IHa. eauto with pcuic. depelim p. eapply (PCUICConversion.substitution_ws_cumul_pb (s:=s) (Γ' := Γ) (Γ'' := [])); eauto. Qed. @@ -3196,7 +3190,7 @@ End WfEnv. Lemma spine_subst_vass `{cf: checker_flags} Σ Γ s t σ Δ na A : wf Σ.1 -> spine_subst Σ Γ s σ Δ -> - isType Σ (Γ ,,, Δ) A -> + lift_typing typing Σ (Γ ,,, Δ) (j_vass na A) -> Σ ;;; Γ |- t : subst0 σ A -> spine_subst Σ Γ (s ++ [t]) (t :: σ) (Δ ,, vass na A). Proof. @@ -3208,20 +3202,20 @@ Proof. + rewrite /skipn /subst_context /fold_context_k /= /map_decl /=. constructor=> //=. * apply: localenv_cons_abs=> //. - apply: isType_subst; eassumption. + apply: isTypeRel_subst; eassumption. * apply: (PCUICContextSubst.context_subst_ass [] [] [] na _ t); constructor. * econstructor; first constructor. by rewrite PCUICLiftSubst.subst_empty. Qed. -Lemma wf_local_nth_isType {cf} {Σ} {Γ n d} : +Lemma wf_local_nth_isTypeRel {cf} {Σ} {Γ n d} : wf_local Σ Γ -> nth_error Γ n = Some d -> - isType Σ (skipn (S n) Γ) d.(decl_type). + isTypeRel Σ (skipn (S n) Γ) d.(decl_type) d.(decl_name).(binder_relevance). Proof. intros Hwf hnth. eapply All_local_env_nth_error in Hwf; tea. - eapply lift_sorting_it_impl_gen with Hwf => //. + now apply lift_sorting_forget_body in Hwf. Qed. @@ -3236,7 +3230,7 @@ Lemma spine_subst_vass' `{cf:checker_flags} Σ Γ s t σ Δ na A : Proof. move=> wfΣ sss Atyp /(_ sss) ttyp; apply: spine_subst_vass=> //. 2: apply: ttyp; apply: isType_subst; first (apply: inst_subslet; eassumption). - all: exact (wf_local_nth_isType (n := 0) Atyp eq_refl). + all: exact (wf_local_nth_isTypeRel (n := 0) Atyp eq_refl). Qed. Lemma mk_ctx_subst_spec' `{cf : checker_flags} {Σ Γ Δ args} (c : ctx_inst Σ Γ args (List.rev Δ)) : @@ -3268,4 +3262,3 @@ Section ClosedSpineSubst. move=> /spine_subst_smash /closed_spine_subst; by rewrite forallb_rev. Qed. End ClosedSpineSubst. - diff --git a/pcuic/theories/PCUICSubstitution.v b/pcuic/theories/PCUICSubstitution.v index 33ffb8427..a7963b00f 100644 --- a/pcuic/theories/PCUICSubstitution.v +++ b/pcuic/theories/PCUICSubstitution.v @@ -150,21 +150,6 @@ Proof. rewrite commut_lift_subst_rec. 1: lia. f_equal; lia. Qed. -Lemma All_local_env_subst {cf:checker_flags} (P Q : context -> judgment -> Type) c n k : - All_local_env Q c -> - (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 _ (Typ _)). eauto. - - simpl. eapply (Hf _ (TermTyp _ _)). eauto. -Qed. - Lemma subst_length {cf:checker_flags} Σ Γ s Γ' : subs Σ Γ s Γ' -> #|s| = #|Γ'|. Proof. induction 1; simpl; lia. @@ -579,19 +564,13 @@ 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 & _ & Her). 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. - apply X. + cbn in Her. + repeat (eexists; cbn; tea). + now eapply weaken_ctx in Hs. Qed. Lemma wf_arities_context {cf:checker_flags} {Σ : global_env} {wfΣ : wf Σ} {mind mdecl} : @@ -1803,16 +1782,56 @@ Proof. now eapply typing_wf_local in Ht. Qed. -Corollary isType_substitution {cf} {Σ} {wfΣ : wf Σ} {Γ Γ' s Δ T} : +Corollary lift_typing_substitution {cf} {Σ} {wfΣ : wf Σ} {Γ Γ' s Δ} j : subslet Σ Γ s Γ' -> - isType Σ (Γ ,,, Γ' ,,, Δ) T -> - isType Σ (Γ ,,, subst_context s 0 Δ) (subst s #|Δ| T). + lift_typing typing Σ (Γ ,,, Γ' ,,, Δ) j -> + lift_typing typing Σ (Γ ,,, subst_context s 0 Δ) (judgment_map (subst s #|Δ|) j). Proof. intros Hs HT. apply lift_typing_f_impl with (1 := HT) => // ?? Ht. eapply substitution in Ht; tea. Qed. +Corollary lift_typing_subst {cf} {Σ} {wfΣ : wf Σ} {Γ Γ' s} j : + subslet Σ Γ s Γ' -> + lift_typing typing Σ (Γ ,,, Γ') j -> + lift_typing typing Σ Γ (judgment_map (subst0 s) j). +Proof. + now apply @lift_typing_substitution with (Δ := []). +Qed. + +Corollary isType_substitution {cf} {Σ} {wfΣ : wf Σ} {Γ Γ' s Δ T} : + subslet Σ Γ s Γ' -> + isType Σ (Γ ,,, Γ' ,,, Δ) T -> + isType Σ (Γ ,,, subst_context s 0 Δ) (subst s #|Δ| T). +Proof. + apply lift_typing_substitution. +Qed. + +Corollary isType_subst {cf} {Σ} {wfΣ : wf Σ} {Γ Δ s T} : + subslet Σ Γ s Δ -> + isType Σ (Γ ,,, Δ) T -> + isType Σ Γ (subst0 s T). +Proof. + apply lift_typing_subst. +Qed. + +Corollary isTypeRel_substitution {cf} {Σ} {wfΣ : wf Σ} {Γ Γ' s Δ T r} : + subslet Σ Γ s Γ' -> + isTypeRel Σ (Γ ,,, Γ' ,,, Δ) T r -> + isTypeRel Σ (Γ ,,, subst_context s 0 Δ) (subst s #|Δ| T) r. +Proof. + apply lift_typing_substitution. +Qed. + +Corollary isTypeRel_subst {cf} {Σ} {wfΣ : wf Σ} {Γ Δ s T r} : + subslet Σ Γ s Δ -> + isTypeRel Σ (Γ ,,, Δ) T r -> + isTypeRel Σ Γ (subst0 s T) r. +Proof. + apply lift_typing_subst. +Qed. + Corollary substitution_wf_local {cf} {Σ} {wfΣ : wf Σ} {Γ Γ' s Δ} : subslet Σ Γ s Γ' -> wf_local Σ (Γ ,,, Γ' ,,, Δ) -> diff --git a/pcuic/theories/PCUICTyping.v b/pcuic/theories/PCUICTyping.v index c951c51f5..20e02aae8 100644 --- a/pcuic/theories/PCUICTyping.v +++ b/pcuic/theories/PCUICTyping.v @@ -154,6 +154,7 @@ 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) + (elim_relevance : isSortRel ps ci.(ci_relevance)) (ind_inst : ctx_inst (typing Σ) Γ (p.(pparams) ++ indices) (List.rev (subst_instance p.(puinst) (ind_params mdecl ,,, ind_indices idecl : context)))) @@ -361,6 +362,13 @@ Proof. Qed. Global Hint Resolve isType_welltyped : pcuic. +Definition isTypeRel_welltyped {cf Σ} {Γ T r} + : isTypeRel Σ Γ T r -> welltyped Σ Γ T. +Proof. + intros [_ [s []]]. now econstructor. +Qed. +Global Hint Resolve isTypeRel_welltyped : pcuic. + Definition has_sort_isType {cf Σ} {Γ T} s : Σ ;;; Γ |- T : tSort s -> isType Σ Γ T. Proof. @@ -368,6 +376,13 @@ Proof. Defined. Global Hint Resolve has_sort_isType : pcuic. +Definition has_sort_isTypeRel {cf Σ} {Γ T} s r + : isSortRel s r -> Σ ;;; Γ |- T : tSort s -> isTypeRel Σ Γ T r. +Proof. + repeat (eexists; tea). +Defined. +Global Hint Resolve has_sort_isTypeRel : pcuic. + Definition isWfArity {cf:checker_flags} Σ (Γ : context) T := (isType Σ Γ T × { ctx & { s & (destArity [] T = Some (ctx, s)) } }). @@ -610,7 +625,7 @@ Lemma wf_local_inv `{checker_flags} {Σ Γ'} (w : wf_local Σ Γ') : forall d Γ, Γ' = d :: Γ -> ∑ (w' : wf_local Σ Γ) u, - { ty : lift_typing1 (typing Σ) Γ (Judge d.(decl_body) d.(decl_type) (Some u)) | + { ty : lift_typing1 (typing Σ) Γ (j_decl_s d (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 }. @@ -734,6 +749,7 @@ Lemma typing_ind_env_app_size `{cf : checker_flags} : wf_local Σ (Γ ,,, predctx) -> PΓ Σ (Γ ,,, predctx) -> is_allowed_elimination Σ idecl.(ind_kelim) ps -> + isSortRel ps ci.(ci_relevance) -> 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) -> @@ -882,6 +898,7 @@ Proof. 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 Xg.(ind_relevance_compat). - apply (onIndices Xg). } { apply All_local_env_impl with (1 := onP); intros ?? Hj. apply Xj; tas. apply lift_typing_impl with (1 := Hj); intros ?? HT. split; tas. @@ -1189,6 +1206,7 @@ Lemma typing_ind_env `{cf : checker_flags} : P Σ (Γ ,,, predctx) p.(preturn) (tSort ps) -> PΓ Σ (Γ ,,, predctx) -> is_allowed_elimination Σ idecl.(ind_kelim) ps -> + isSortRel ps ci.(ci_relevance) -> 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) -> diff --git a/pcuic/theories/PCUICValidity.v b/pcuic/theories/PCUICValidity.v index ca255c23c..71ed906c3 100644 --- a/pcuic/theories/PCUICValidity.v +++ b/pcuic/theories/PCUICValidity.v @@ -51,18 +51,6 @@ Section Validity. now eapply inversion_Sort in Hu as [? [? ?]]. Qed. - Lemma isType_subst_instance_decl {Σ Γ T c decl u} : - wf Σ.1 -> - lookup_env Σ.1 c = Some decl -> - isType (Σ.1, universes_decl_of_decl decl) Γ T -> - consistent_instance_ext Σ (universes_decl_of_decl decl) u -> - isType Σ (subst_instance u Γ) (subst_instance u T). - Proof using Type. - destruct Σ as [Σ φ]. intros X X0 Hty X1. - 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} : wf Σ.1 -> lookup_env Σ.1 c = Some decl -> @@ -204,7 +192,7 @@ Section Validity. - (* Lambda *) eapply lift_sorting_ex_it_impl_gen with X3 => // Hs. pose proof (lift_sorting_extract X0). - eexists; constructor; eauto. + eexists; split; constructor; eauto. - (* Let *) apply lift_sorting_it_impl_gen with X3 => // Hs. @@ -228,25 +216,27 @@ Section Validity. - (* Constant *) eapply declared_constant_inv in wf as Hc; eauto. - destruct Hc as (_ & s & Hs & _); simpl in Hs. + apply lift_sorting_forget_body, lift_sorting_forget_rel in Hc. 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. + eapply isType_subst_instance_decl with (Γ:=[]); eauto. simpl. exact weaken_env_prop_typing. - (* Inductive type *) destruct (on_declared_inductive isdecl); pcuic. destruct isdecl. - apply onArity in o0. + apply onArity, isTypeRel_isType in o0. eapply isType_weakening; eauto. unshelve eapply declared_minductive_to_gen in H; eauto. - eapply (isType_subst_instance_decl (Γ:=[])); eauto. + eapply isType_subst_instance_decl with (Γ:=[]); eauto. - (* Constructor type *) destruct (on_declared_constructor isdecl) as [[oni oib] [cs [declc onc]]]. unfold type_of_constructor. - eapply lift_typing_fu_impl with (f := fun t => _ (_ t)) (1 := on_ctype onc) => //= ?? Hs. + pose proof (on_ctype onc) as Hc. + apply isTypeRel_isType in Hc. + eapply lift_typing_fu_impl with (f := fun t => _ (_ t)) (1 := Hc) => //= ?? Hs. + 2: now eapply relevance_subst_opt. eapply instantiate_minductive in Hs; eauto. 2:(destruct isdecl as [[] ?]; eauto). simpl in Hs. @@ -311,6 +301,7 @@ Section Validity. lenpars lenargs cu]]]; eauto. 2:eapply isdecl.p1. eapply lift_typing_fu_impl with (f := fun t => _ (_ t)) (1 := isdecl') => // ?? Hs. + 2: now eapply relevance_subst_opt. 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. @@ -333,9 +324,11 @@ Section Validity. assumption. - (* Fix *) + eapply isTypeRel_isType. eapply nth_error_all in X0; eauto. - (* CoFix *) + eapply isTypeRel_isType. eapply nth_error_all in X0; eauto. - (* Primitive *) @@ -345,11 +338,14 @@ Section Validity. set (s := sType (Universe.make' (array_level a))). destruct H1 as [hty' hbod huniv]. eapply has_sort_isType with s. - eapply (type_App _ _ _ _ _ (tSort s)); tea; cycle 1. + eapply (type_App _ _ _ _ (tSort s) (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)). cbn. lsets. cbn. red. destruct check_univs => //. red. red. intros v H c. csets. - + econstructor. 2: econstructor; eauto. repeat (eexists; tea). econstructor; eauto. + + econstructor. 2: econstructor; eauto. 2: constructor; tas. + all: repeat (eexists; tea; cbn). + 1,3: econstructor; eauto. + all: apply relevance_super. - (* Conv *) now eapply has_sort_isType with (s := s). diff --git a/pcuic/theories/PCUICWellScopedCumulativity.v b/pcuic/theories/PCUICWellScopedCumulativity.v index 24c749098..1f74d94b6 100644 --- a/pcuic/theories/PCUICWellScopedCumulativity.v +++ b/pcuic/theories/PCUICWellScopedCumulativity.v @@ -195,11 +195,6 @@ Arguments wt_cumul_pb_eq {cf pb Σ Γ T U}. Section EqualityLemmas. Context {cf : checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ}. - Lemma isType_open {Γ T} : isType Σ Γ T -> on_free_vars (shiftnP #|Γ| xpred0) T. - Proof using wfΣ. - move/isType_closedPT. now rewrite closedP_shiftnP. - Qed. - Lemma into_ws_cumul_pb {pb} {Γ : context} {T U} : Σ;;; Γ |- T <=[pb] U -> is_closed_context Γ -> is_open_term Γ T -> @@ -216,7 +211,7 @@ Section EqualityLemmas. Proof using wfΣ. intros H. pose proof (isType_wf_local H). - eapply (ws_cumul_pb_refl' (exist Γ (wf_local_closed_context X)) (exist T (isType_open H))). + eapply (ws_cumul_pb_refl' (exist Γ (wf_local_closed_context X)) (exist T (isType_is_open_term H))). Qed. (** From well-typed to simply well-scoped equality. *) @@ -226,7 +221,7 @@ Section EqualityLemmas. Proof using wfΣ. move=> [] dom codom equiv; cbn. generalize (wf_local_closed_context (isType_wf_local dom)). - generalize (isType_open dom) (isType_open codom). clear -wfΣ equiv. + generalize (isType_is_open_term dom) (isType_is_open_term codom). clear -wfΣ equiv. intros. apply into_ws_cumul_pb => //. Qed. @@ -441,14 +436,15 @@ Qed. Inductive wt_cumul_pb_decls {cf : checker_flags} (pb : conv_pb) (Σ : global_env_ext) (Γ Γ' : context) : context_decl -> context_decl -> Type := | wt_cumul_pb_vass {na na' : binder_annot name} {T T' : term} : - isType Σ Γ T -> isType Σ Γ' T' -> - conv_cum pb Σ Γ T T' -> + lift_typing typing Σ Γ (j_vass na T) -> + lift_typing typing Σ Γ' (j_vass na' T') -> eq_binder_annot na na' -> + conv_cum pb Σ Γ T T' -> wt_cumul_pb_decls pb Σ Γ Γ' (vass na T) (vass na' T') | wt_cumul_pb_vdef {na na' : binder_annot name} {b b' T T'} : + lift_typing typing Σ Γ (j_vdef na b T) -> + lift_typing typing Σ Γ' (j_vdef na' b' T') -> eq_binder_annot na na' -> - isType Σ Γ T -> isType Σ Γ' T' -> - Σ ;;; Γ |- b : T -> Σ ;;; Γ' |- b' : T' -> Σ ;;; Γ |- b = b' -> conv_cum pb Σ Γ T T' -> wt_cumul_pb_decls pb Σ Γ Γ' (vdef na b T) (vdef na' b' T'). @@ -498,23 +494,11 @@ Notation wt_conv_context Σ := (wt_cumul_ctx_pb Conv Σ). Section WtContextConversion. Context {cf : checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ}. - Definition wt_decl Γ d := - match d with - | {| decl_body := None; decl_type := ty |} => isType Σ Γ ty - | {| decl_body := Some b; decl_type := ty |} => isType Σ Γ ty × Σ ;;; Γ |- b : ty - end. - Lemma wf_local_All_fold Γ : wf_local Σ Γ <~> - All_fold wt_decl Γ. + All_fold (fun Γ decl => lift_typing typing Σ Γ (j_decl decl)) Γ. Proof using Type. - split. - - induction 1; constructor; 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; auto. - destruct p; split; tas. apply i. + apply All_local_env_All_fold. Qed. Lemma wt_cumul_ctx_pb_forget {pb} {Γ Γ' : context} : @@ -522,14 +506,14 @@ Section WtContextConversion. [× wf_local Σ Γ, wf_local Σ Γ' & cumul_pb_context cumulAlgo_gen pb Σ Γ Γ']. Proof using Type. move=> wteq. - eapply (All2_fold_impl (Q:=fun Γ Γ' d d' => wt_decl Γ d × wt_decl Γ' d' × cumul_pb_decls cumulAlgo_gen pb Σ Γ Γ' d d')) in wteq. + eapply (All2_fold_impl (Q:=fun Γ Γ' d d' => lift_typing typing Σ Γ (j_decl d) × lift_typing typing Σ Γ' (j_decl d') × cumul_pb_decls cumulAlgo_gen pb Σ Γ Γ' d d')) in wteq. 2:{ intros ???? []; intuition (cbn; try constructor; auto). } - eapply All2_fold_All_fold_mix_inv in wteq as [wteq [wfΓ wfΓ']]. + eapply All2_fold_All_fold_mix_inv in wteq as (wteq & wfΓ & wfΓ'). eapply wf_local_All_fold in wfΓ. eapply wf_local_All_fold in wfΓ'. split; auto. Qed. - Lemma into_wt_cumul_ctx_pb {pb} {Γ Γ' : context} {T U : term} : + Lemma into_wt_cumul_ctx_pb {pb} {Γ Γ' : context} : wf_local Σ Γ -> wf_local Σ Γ' -> cumul_pb_context cumulAlgo_gen pb Σ Γ Γ' -> wt_cumul_ctx_pb pb Σ Γ Γ'. @@ -544,30 +528,25 @@ Section WtContextConversion. destruct cum; cbn in wtd, wtd'; constructor; intuition auto. Qed. - Lemma wt_ws_ws_cumul_ctx_pb {pb} {Γ Γ' : context} {T U : term} : + Lemma wt_ws_ws_cumul_ctx_pb {pb} {Γ Γ' : context} : wt_cumul_ctx_pb pb Σ Γ Γ' -> ws_cumul_ctx_pb pb Σ Γ Γ'. Proof using wfΣ. - intros a; eapply All2_fold_impl_ind; tea. + intros a; eapply All2_fold_impl_ind; tea. clear Γ Γ' a. intros ???? wt ws eq; - pose proof (All2_fold_length wt). destruct eq. - - pose proof (isType_wf_local i). - eapply wf_local_closed_context in X. - eapply isType_open in i. apply isType_open in i0. + - apply lift_typing_wf_local in l as wfΓ. + apply wf_local_closed_context in wfΓ. + apply lift_typing_is_open_term in l as [_ clT], l0 as [_ clT']. eapply into_ws_cumul_decls with Δ; eauto with fvs. constructor; auto. rewrite (All2_fold_length ws) //. - - pose proof (isType_wf_local i). - eapply wf_local_closed_context in X. - eapply isType_open in i. apply isType_open in i0. - eapply PCUICClosedTyp.subject_closed in t. - eapply PCUICClosedTyp.subject_closed in t0. - eapply (@closedn_on_free_vars xpred0) in t. - eapply (@closedn_on_free_vars xpred0) in t0. + - apply lift_typing_wf_local in l as wfΓ. + apply wf_local_closed_context in wfΓ. + apply lift_typing_is_open_term in l as [clb clT], l0 as [clb' clT']. eapply into_ws_cumul_decls with Δ; eauto with fvs. - destruct pb; constructor; auto. - rewrite (All2_fold_length ws) //; eauto with fvs. + constructor; auto. + rewrite (All2_fold_length ws); eauto with fvs. Qed. Lemma ws_cumul_ctx_pb_inv {pb} {Γ Γ' : context} : diff --git a/pcuic/theories/PCUICWfUniverses.v b/pcuic/theories/PCUICWfUniverses.v index 63624b8ad..194cd60a9 100644 --- a/pcuic/theories/PCUICWfUniverses.v +++ b/pcuic/theories/PCUICWfUniverses.v @@ -1067,23 +1067,18 @@ Qed. Theorem wf_types : 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 Σ _ j => lift_wfbu_term (wf_universes Σ) (wf_sortb Σ) j) (fun Σ Γ => wf_ctx_universes Σ Γ). Proof using Type. - apply typing_ind_env; unfold lift_wfb_term; intros; rename_all_hyps; cbn; rewrite -!/(wf_universes _ _) ; + apply typing_ind_env; unfold lift_wfbu_term; intros; rename_all_hyps; cbn; rewrite -!/(wf_universes _ _) ; specIH; to_prop; cbn; auto. - - 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 //. + - eapply MCReflect.introT. 1: apply lift_wfu_wfbu_term. + apply lift_typing_subjtyp with (1 := X) => //= t T [] _ /andP[] //. - apply All_local_env_cst, All_forallb in X0. - apply forallb_impl with (2 := X0) => [] [na bo ty] _ //=. - rewrite andb_true_r //. + apply forallb_impl with (2 := X0) => [] [na bo ty] _ //= /andP[] //=. - rewrite wf_universes_lift. eapply forallb_nth_error with (n := n) in H0. rewrite heq_nth_error /= in H0. @@ -1131,8 +1126,8 @@ Qed. exact (weaken_lookup_on_global_env' Σ.1 _ _ wf (proj1 (proj1 isdecl))). now eapply consistent_instance_ext_wf. - - rewrite wf_universes_mkApps in H7. - move/andP: H7 => /= [] wfu; rewrite forallb_app. + - rewrite wf_universes_mkApps in H6. + move/andP: H6 => /= [] wfu; rewrite forallb_app. move/andP=> [] wfpars wfinds. cbn in wfu. rewrite wfu /= wfpars wf_universes_mkApps /= @@ -1171,7 +1166,7 @@ Qed. rewrite wfp. eapply closedu_subst_context. rewrite a. now rewrite closedu_inds. - * rewrite wf_universes_it_mkLambda_or_LetIn H6 andb_true_r. + * rewrite wf_universes_it_mkLambda_or_LetIn H7 andb_true_r. move: H4. rewrite /wf_ctx_universes forallb_app => /andP[hctx _]. apply (MCReflect.introT onctxP). diff --git a/pcuic/theories/Typing/PCUICClosedTyp.v b/pcuic/theories/Typing/PCUICClosedTyp.v index 483122f4f..eb2d8b4e6 100644 --- a/pcuic/theories/Typing/PCUICClosedTyp.v +++ b/pcuic/theories/Typing/PCUICClosedTyp.v @@ -110,7 +110,7 @@ Proof. - rewrite andb_and. split. * apply onArity in oib. now move/andP: oib => [] /=. - * pose proof (onArity oib) as X. unfold on_type in X. + * pose proof (onArity oib) as X. unfold on_type_rel in X. rewrite oib.(ind_arity_eq) in X. cbn in X. rewrite !closedn_it_mkProd_or_LetIn /= !andb_true_r in X. @@ -118,7 +118,7 @@ Proof. - pose proof (onConstructors oib). red in X. eapply All_forallb. eapply All2_All_left; eauto. intros cdecl cs X0; - move: (on_ctype X0) => /=. unfold on_type. cbn. + move: (on_ctype X0) => /=. unfold on_type_rel. cbn. simpl. unfold closed_constructor_body. intros Hty. rewrite arities_context_length in Hty. @@ -198,7 +198,7 @@ Proof. - rewrite closedn_subst_instance. eapply declared_inductive_inv in X0; eauto with extends. apply onArity in X0. - rewrite /on_type /lift_wfb_term /= in X0. + rewrite /on_type_rel /lift_wfb_term /= in X0. intuition eauto using closed_upwards with arith. - destruct isdecl as [Hidecl Hcdecl]. @@ -212,16 +212,16 @@ Proof. eapply All2_nth_error_Some in Hcdecl; eauto. destruct Hcdecl as [? [? ?]]. cbn in *. destruct o as [? ? HT _]. - rewrite /on_type /lift_wfb_term /= in HT. + rewrite /on_type_rel /lift_wfb_term /= in HT. rewrite arities_context_length in HT. eauto using closed_upwards with arith. - destruct H3 as [clret _]. - destruct H6 as [clc clty]. + destruct H7 as [clc clty]. rewrite closedn_mkApps in clty. simpl in clty. rewrite forallb_app in clty. move/andP: clty => [clpar clinds]. rewrite app_context_length in clret. - red in H8. eapply Forall2_All2 in H8. + red in H9. eapply Forall2_All2 in H9. eapply All2i_All2_mix_left in X5; eauto. eapply declared_minductive_closed_ind in X0; tea. 2:exact isdecl. pose proof (closed_ind_closed_cstrs X0 isdecl). @@ -236,16 +236,16 @@ Proof. rewrite /predctx in clret. rewrite case_predicate_context_length_indices // in clret. { now rewrite ind_predicate_context_length. } - + clear H8. solve_all. unfold test_branch_k. clear H6. solve_all. + + clear H9. solve_all. unfold test_branch_k. clear H7. solve_all. * rewrite (closedn_ctx_alpha a1). eapply closed_cstr_branch_context_gen in X0; tea. rewrite (wf_predicate_length_pars H1). now rewrite (declared_minductive_ind_npars isdecl). * rewrite (All2_length a1). - len in H8. - (*unfold case_branch_context_gen in H8. simpl in H8. + len in H9. + (*unfold case_branch_context_gen in H8. simpl in H9. rewrite case_branch_type_fst in H8. *) - rewrite case_branch_context_length_args in H8 => //. + rewrite case_branch_context_length_args in H9 => //. now rewrite cstr_branch_context_length. + rewrite closedn_mkApps; auto. rewrite closedn_it_mkLambda_or_LetIn //. @@ -374,12 +374,24 @@ Qed. #[global] Hint Extern 10 => progress unfold PCUICTypingDef.typing in * : fvs. +Lemma lift_typing_closed {cf:checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ.1} {Γ j} : + lift_typing typing Σ Γ j -> + lift_wf_term (closedn #|Γ|) j. +Proof. + intro Hj. apply lift_typing_subject with (1 := Hj) => t T. + apply subject_closed. +Qed. + +Lemma isTypeRel_closed {cf:checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ.1} {Γ T r} : isTypeRel Σ Γ T r -> closedn #|Γ| T. +Proof. move/lift_typing_closed => [] //. Qed. + Lemma isType_closed {cf:checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ.1} {Γ T} : isType Σ Γ T -> closedn #|Γ| T. -Proof. intros (_ & s & Hs & _); cbn in Hs; fvs. Qed. +Proof. move/lift_typing_closed => [] //. Qed. #[global] Hint Extern 4 (closedn #|?Γ| ?A = true) => match goal with | [ H : isType _ Γ A |- _ ] => exact (isType_closed H) + | [ H : isTypeRel _ Γ A _ |- _ ] => exact (isTypeRel_closed H) end : fvs. Lemma is_open_term_closed (Γ : context) t : @@ -419,9 +431,16 @@ Lemma lift_typing_is_open_term {cf:checker_flags} {Σ : global_env_ext} {wfΣ : 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. + move/lift_typing_closed => Hj. + apply lift_wf_term_impl with (1 := Hj) => t. + rewrite is_open_term_closed //. +Qed. + +Lemma isTypeRel_is_open_term {cf:checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ.1} {Γ T r} : + isTypeRel Σ Γ T r -> + is_open_term Γ T. +Proof. + move/isType_closed. now rewrite is_open_term_closed. Qed. Lemma isType_is_open_term {cf:checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ.1} {Γ T} : @@ -455,6 +474,11 @@ Qed. | [ H : isType _ Γ A |- _ ] => exact (isType_is_open_term H) end : fvs. +#[global] Hint Extern 4 (is_open_term ?Γ ?A = true) => + match goal with + | [ H : isTypeRel _ Γ A _ |- _ ] => exact (isTypeRel_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. diff --git a/pcuic/theories/Typing/PCUICInstTyp.v b/pcuic/theories/Typing/PCUICInstTyp.v index bc6dc5704..b14540f4a 100644 --- a/pcuic/theories/Typing/PCUICInstTyp.v +++ b/pcuic/theories/Typing/PCUICInstTyp.v @@ -437,7 +437,7 @@ Proof. rewrite inst_closed0; eauto. - intros Σ wfΣ Γ wfΓ ci p c brs indices ps mdecl idecl isdecl HΣ. intros IHΔ ci_npar eqpctx predctx wfp cup Hpctx Hpret - IHpret IHpredctx isallowed. + IHpret IHpredctx isallowed Her. intros IHctxi Hc IHc iscof ptm wfbrs Hbrs Δ f HΔ Hf. autorewrite with sigma. simpl. rewrite map_app. simpl. @@ -536,14 +536,14 @@ Proof. * now eapply fix_guard_inst. * now rewrite nth_error_map hnth. * apply All_map, (All_impl ihmfixt). - intros x t. eapply lift_typing_map with (j := TermoptTyp None _) => //. eapply t; eauto. + intros x t. eapply t; eauto. * pose proof (inst_fix_context mfix σ). setoid_rewrite <-up_Upn at 1 in H. rewrite H. 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. + 1: eapply t; eauto. + rewrite -(fix_context_length mfix). eapply well_subst_app_up => //. + eapply wf_local_app_inst; eauto. @@ -558,14 +558,14 @@ Proof. * now eapply cofix_guard_inst. * now rewrite nth_error_map hnth. * apply All_map, (All_impl ihmfixt). - intros x t. eapply lift_typing_map with (j := TermoptTyp None _) => //. eapply t; eauto. + intros x t. eapply t; eauto. * pose proof (inst_fix_context mfix σ). setoid_rewrite <-up_Upn at 1 in H. rewrite H. 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. + 1: eapply t; eauto. + rewrite -(fix_context_length mfix). eapply well_subst_app_up => //. + eapply wf_local_app_inst; eauto. diff --git a/pcuic/theories/Typing/PCUICRenameTyp.v b/pcuic/theories/Typing/PCUICRenameTyp.v index e94c5f32e..f83541888 100644 --- a/pcuic/theories/Typing/PCUICRenameTyp.v +++ b/pcuic/theories/Typing/PCUICRenameTyp.v @@ -624,7 +624,7 @@ Proof. 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 lift_typing_map with (j := j_decl _) => //. eapply X. split; auto. eapply urenaming_ext. @@ -822,8 +822,7 @@ Proof. - intros Σ wfΣ Γ wfΓ na A B s1 s2 X hA ihA hB ihB P Δ f hf. assert (lift_typing0 (typing Σ Δ) (j_vass_s na (rename f A) s1)). - + eapply lift_typing_map with (j := TypUniv _ _) => //. - eapply ihA; eauto. + + eapply ihA; eauto. + rewrite /=. econstructor; tas. eapply ihB; eauto. simpl. @@ -834,8 +833,7 @@ Proof. * 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. + + eapply ihA; eauto. + simpl. econstructor; tas. eapply iht; eauto; simpl. eapply renaming_extP. { now rewrite -(shiftnP_add 1). } @@ -882,7 +880,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 Her. intros IHctxi Hc IHc iscof ptm wfbrs Hbrs P Δ f Hf. simpl. rewrite rename_mkApps. @@ -1000,12 +998,12 @@ 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. eapply lift_typing_map with (j := TermoptTyp None _) => //. eapply t. eauto. + intros x t. eapply t. eauto. * apply All_map, (All_impl ihmfixb). 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. + 1: eapply t; eauto. 2: len; now sigma. rewrite rename_fix_context. split; auto. @@ -1026,12 +1024,12 @@ 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. eapply lift_typing_map with (j := TermoptTyp None _) => //. eapply t. eauto. + intros x t. eapply t. eauto. * apply All_map, (All_impl ihmfixb). 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. + 1: eapply t; eauto. 2: len; now sigma. rewrite rename_fix_context. split; auto. diff --git a/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v b/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v index f5a3d55c4..4bc234e91 100644 --- a/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v +++ b/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v @@ -255,8 +255,7 @@ Proof using Type. 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. + eapply lift_typing_mapu with (u := None) => //; auto using relevance_subst_opt. - intros n decl eq X u univs wfΣ' H. rewrite subst_instance_lift. rewrite map_decl_type. econstructor; aa. @@ -269,15 +268,18 @@ Proof using Type. * 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 lift_typing_mapu with (tm := None) (u := Some _) => //. apply X1; eauto. + + eapply lift_typing_mapu with (tm := None) (u := Some _) => //; + auto using relevance_subst_opt. + eapply X3; eauto. - intros n t0 b bty X X0 X1 X2 X3 u univs wfΣ' H. econstructor. - + eapply lift_typing_mapu with (tm := None) (u := None) => //. eapply X1; aa. + + eapply lift_typing_mapu with (tm := None) (u := None) => //; + auto using relevance_subst_opt. + 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 lift_typing_mapu with (tm := Some _) (u := None) => //; + auto using relevance_subst_opt. + eapply X3; 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. @@ -299,7 +301,7 @@ Proof using Type. + symmetry; apply subst_instance_two. - intros ci p c brs args u mdecl idecl isdecl hΣ hΓ indnp eqpctx wfp cup - wfpctx pty Hpty Hcpc kelim + wfpctx pty Hpty Hcpc kelim hrel IHctxi Hc IHc notCoFinite wfbrs hbrs i univs wfext cu. rewrite subst_instance_mkApps subst_instance_it_mkLambda_or_LetIn map_app. cbn. @@ -318,6 +320,7 @@ Proof using Type. + now rewrite -subst_instance_case_predicate_context -subst_instance_app_ctx. + cbn in *. eapply is_allowed_elimination_subst_instance; aa. + + now apply relevance_subst. + move: IHctxi. simpl. rewrite -subst_instance_app. rewrite -subst_instance_two_context. @@ -353,12 +356,12 @@ Proof using Type. + now eapply fix_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 lift_typing_mapu with (tm := None) (u := None); cbn; eauto using relevance_subst_opt. + 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. + eapply lift_typing_mapu with (tm := Some _) (u := None); cbn; eauto using relevance_subst_opt. + red; rewrite <- wffix. unfold wf_fixpoint, wf_fixpoint_gen. f_equal. @@ -375,12 +378,12 @@ Proof using Type. + 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 lift_typing_mapu with (tm := None) (u := None); cbn; eauto using relevance_subst_opt. + 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. + eapply lift_typing_mapu with (tm := Some _) (u := None); cbn; eauto using relevance_subst_opt. + red; rewrite <- wffix. unfold wf_cofixpoint, wf_cofixpoint_gen. rewrite map_map_compose. @@ -498,6 +501,19 @@ Proof using Type. cbn; eauto using typing_wf_local. Qed. +Lemma lift_typing_subst_instance_decl Σ Γ j c decl u : + wf Σ.1 -> + lookup_env Σ.1 c = Some decl -> + lift_typing typing (Σ.1, universes_decl_of_decl decl) Γ j -> + consistent_instance_ext Σ (universes_decl_of_decl decl) u -> + lift_typing typing Σ (subst_instance u Γ) {| j_term := option_map (subst_instance u) (j_term j); j_typ := subst_instance u (j_typ j); + j_univ := option_map (subst_instance u) (j_univ j); j_rel := j_rel j |}. +Proof. + intros wfΣ look Hj cu. + eapply lift_typing_fu_impl with (1 := Hj) => //; auto using relevance_subst_opt. + intros ?? Hs; now eapply typing_subst_instance_decl. +Qed. + Lemma isType_subst_instance_decl Σ Γ T c decl u : wf Σ.1 -> lookup_env Σ.1 c = Some decl -> @@ -505,9 +521,17 @@ Lemma isType_subst_instance_decl Σ Γ T c decl u : consistent_instance_ext Σ (universes_decl_of_decl decl) u -> isType Σ (subst_instance u Γ) (subst_instance u T). Proof using Type. - intros wfΣ look isty cu. - eapply lift_typing_fu_impl with (1 := isty) => //. - intros ?? Hs; now eapply typing_subst_instance_decl. + apply lift_typing_subst_instance_decl. +Qed. + +Lemma isTypeRel_subst_instance_decl Σ Γ T r c decl u : + wf Σ.1 -> + lookup_env Σ.1 c = Some decl -> + isTypeRel (Σ.1, universes_decl_of_decl decl) Γ T r -> + consistent_instance_ext Σ (universes_decl_of_decl decl) u -> + isTypeRel Σ (subst_instance u Γ) (subst_instance u T) r. +Proof using Type. + apply lift_typing_subst_instance_decl. Qed. Lemma isArity_subst_instance u T : @@ -525,7 +549,7 @@ Lemma wf_local_subst_instance Σ Γ ext u : Proof using Type. destruct Σ as [Σ φ]. intros X X0 X1. simpl in *. induction X1; cbn; constructor; auto. - all: eapply lift_typing_fu_impl with (1 := t0) => //= ?? Hs. + all: eapply lift_typing_fu_impl with (1 := t0) => //= ?? Hs; auto using relevance_subst_opt. all: eapply typing_subst_instance'' in Hs; eauto; apply X. Qed. @@ -538,10 +562,20 @@ Lemma wf_local_subst_instance_decl Σ Γ c decl u : Proof using Type. destruct Σ as [Σ φ]. intros X X0 X1 X2. induction X1; cbn; constructor; auto. - all: eapply lift_typing_fu_impl with (1 := t0) => // ?? Hs. + all: eapply lift_typing_fu_impl with (1 := t0) => // ?? Hs; auto using relevance_subst_opt. all: eapply typing_subst_instance_decl in Hs; eauto; apply X. Qed. + Lemma isTypeRel_subst_instance_id Σ Γ T r : + wf_ext_wk Σ -> + let u := abstract_instance Σ.2 in + isTypeRel Σ Γ T r -> 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 isType_subst_instance_id Σ Γ T : wf_ext_wk Σ -> let u := abstract_instance Σ.2 in @@ -562,7 +596,7 @@ Qed. pose proof (on_declared_inductive decli) as [onmind oib]. pose proof (onArity oib) as ona. rewrite (oib.(ind_arity_eq)) in ona. - apply isType_subst_instance_id in ona. + apply isTypeRel_subst_instance_id in ona. 2:split; simpl; auto. - rewrite !subst_instance_it_mkProd_or_LetIn in ona. eapply (f_equal (destArity [])) in ona. @@ -583,7 +617,7 @@ Qed. pose proof (on_declared_inductive decli) as [_ oib]. pose proof (onArity oib) as ona. rewrite (oib.(ind_arity_eq)) in ona |- *. - apply isType_subst_instance_id in ona; eauto. + apply isTypeRel_subst_instance_id in ona; eauto. destruct decli as [declm _]. eapply declared_inductive_wf_global_ext in declm; auto. Qed. diff --git a/pcuic/theories/Typing/PCUICWeakeningTyp.v b/pcuic/theories/Typing/PCUICWeakeningTyp.v index 71b644118..b08e8d671 100644 --- a/pcuic/theories/Typing/PCUICWeakeningTyp.v +++ b/pcuic/theories/Typing/PCUICWeakeningTyp.v @@ -165,10 +165,29 @@ Proof. eapply weakening_gen => //. Qed. +Lemma lift_typing_weakening {cf : checker_flags} Σ Γ Γ' j : + wf Σ.1 -> wf_local Σ (Γ ,,, Γ') -> + lift_typing typing Σ Γ j -> + lift_typing typing Σ (Γ ,,, Γ') (judgment_map (lift0 #|Γ'|) j). +Proof. + intros wfΣ wfΓ Hj. + apply lift_typing_f_impl with (1 := Hj) => // t T. + now apply weakening. +Qed. + +Lemma lift_typing_weaken_ctx {cf : checker_flags} Σ Γ Δ j : + wf Σ.1 -> wf_local Σ Δ -> + lift_typing typing Σ Γ j -> + lift_typing typing Σ (Δ ,,, Γ) j. +Proof. + intros wfΣ wfΔ Hj. + apply lift_typing_impl with (1 := Hj) => t T. + now apply weaken_ctx. +Qed. Corollary All_mfix_wf {cf:checker_flags} Σ Γ mfix : wf Σ.1 -> wf_local Σ Γ -> - All (fun d : def term => isType Σ Γ (dtype d)) mfix -> + All (fun d : def term => isTypeRel Σ Γ (dtype d) (dname d).(binder_relevance)) mfix -> wf_local Σ (Γ ,,, fix_context mfix). Proof. move=> wfΣ wf a; move: wf. @@ -179,21 +198,43 @@ Proof. intros Δ wfΔ. eapply All_local_env_app; auto. induction a in Δ, wfΔ |- *; simpl; auto. - + constructor. - + simpl. - 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 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. + 1: now constructor. + simpl. + eapply All_local_env_app; auto. + * constructor. 1: constructor. + now eapply lift_typing_weakening in p. + * specialize (IHa (Δ ,,, [vass (dname x) (lift0 #|Δ| (dtype x))])). + rewrite app_length in IHa. simpl in IHa. + forward IHa. + + simpl; constructor; auto. + now eapply lift_typing_weakening in p. + + eapply All_local_env_impl; eauto. + simpl; intros. + rewrite app_context_assoc. apply X. +Qed. + +Lemma lift_typing_weakening_skipn {cf:checker_flags} {Σ : global_env_ext} {n Γ j} + (isdecl : n <= #|Γ|) : + wf Σ -> wf_local Σ Γ -> + lift_typing typing Σ (skipn n Γ) j -> + lift_typing typing Σ Γ (judgment_map (lift0 n) j). +Proof. + intros wfΣ wfΓ wfty. rewrite <- (firstn_skipn n Γ) in wfΓ |- *. + assert (n = #|firstn n Γ|). + { rewrite firstn_length_le; auto with arith. } + rewrite {3}H. + now eapply lift_typing_weakening in wfty. +Qed. + +Lemma All_local_env_nth_error_weaken {cf:checker_flags} {Σ : global_env_ext} {Γ n decl} : + wf Σ -> wf_local Σ Γ -> + nth_error Γ n = Some decl -> + lift_typing typing Σ Γ (j_decl (map_decl (lift0 (S n)) decl)). +Proof using Type. + intros hΣ hΓ e. + eapply All_local_env_nth_error in e as hj; tea. + apply nth_error_Some_length in e. + now apply lift_typing_weakening_skipn in hj. Qed. Lemma isType_lift {cf:checker_flags} {Σ : global_env_ext} {n Γ ty} @@ -202,11 +243,14 @@ Lemma isType_lift {cf:checker_flags} {Σ : global_env_ext} {n Γ ty} isType Σ (skipn n Γ) ty -> isType Σ Γ (lift0 n ty). Proof. - intros wfΣ wfΓ wfty. rewrite <- (firstn_skipn n Γ) in wfΓ |- *. - assert (n = #|firstn n Γ|). - { rewrite firstn_length_le; auto with arith. } - apply lift_typing_f_impl with (1 := wfty) => // ?? Hs. - rewrite {3 4}H. - eapply (weakening_typing (Γ := skipn n Γ) (Γ' := []) (Γ'' := firstn n Γ)); - eauto with wf. + now apply @lift_typing_weakening_skipn with (j := Typ _). +Qed. + +Lemma isTypeRel_lift {cf:checker_flags} {Σ : global_env_ext} {n Γ ty r} + (isdecl : n <= #|Γ|): + wf Σ -> wf_local Σ Γ -> + isTypeRel Σ (skipn n Γ) ty r -> + isTypeRel Σ Γ (lift0 n ty) r. +Proof. + now apply @lift_typing_weakening_skipn with (j := TypRel _ _). Qed. diff --git a/quotation/theories/ToPCUIC/Common/EnvironmentTyping.v b/quotation/theories/ToPCUIC/Common/EnvironmentTyping.v index e138b1418..405ca80ee 100644 --- a/quotation/theories/ToPCUIC/Common/EnvironmentTyping.v +++ b/quotation/theories/ToPCUIC/Common/EnvironmentTyping.v @@ -127,6 +127,8 @@ Module QuoteGlobalMaps (Import T : Term) (Import E : EnvironmentSig T) (Import T #[export] Instance quote_on_type {Σ Γ T} : ground_quotable (@on_type P Σ Γ T) := ltac:(cbv [on_type]; exact _). + #[export] Instance quote_on_type_rel {Σ Γ T r} : ground_quotable (@on_type_rel P Σ Γ T r) + := ltac:(cbv [on_type_rel]; exact _). #[export] Instance quote_satisfiable_udecl {univs ϕ} : ground_quotable (@satisfiable_udecl univs ϕ) := ltac:(cbv [satisfiable_udecl]; exact _). @@ -166,6 +168,7 @@ Module QuoteGlobalMaps (Import T : Term) (Import E : EnvironmentSig T) (Import T quote_type_local_ctx quote_sorts_local_ctx quote_on_type + quote_on_type_rel quote_on_udecl quote_satisfiable_udecl quote_valid_on_mono_udecl diff --git a/quotation/theories/ToPCUIC/QuotationOf/Common/EnvironmentTyping/Sig.v b/quotation/theories/ToPCUIC/QuotationOf/Common/EnvironmentTyping/Sig.v index 9aed1cc53..b03e90c95 100644 --- a/quotation/theories/ToPCUIC/QuotationOf/Common/EnvironmentTyping/Sig.v +++ b/quotation/theories/ToPCUIC/QuotationOf/Common/EnvironmentTyping/Sig.v @@ -108,6 +108,7 @@ Module Type QuoteGlobalMapsSig (Import T: Term) (Import E: EnvironmentSig T) (Im #[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 Σ Γ j, ground_quotable (P Σ Γ j)} {Σ Γ T} : ground_quotable (@on_type P Σ Γ T). + #[export] Declare Instance quote_on_type_rel {P} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {Σ Γ T r} : ground_quotable (@on_type_rel P Σ Γ T r). #[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 ϕ). diff --git a/quotation/theories/ToPCUIC/Utils/MCOption.v b/quotation/theories/ToPCUIC/Utils/MCOption.v index ec961e3ba..ba83cb9e4 100644 --- a/quotation/theories/ToPCUIC/Utils/MCOption.v +++ b/quotation/theories/ToPCUIC/Utils/MCOption.v @@ -20,6 +20,7 @@ Proof. Defined. #[export] Polymorphic Instance quote_option_default {A P o b} {quoteP : forall x, o = Some x -> ground_quotable (P x)} {quoteP' : o = None -> ground_quotable b} : ground_quotable (@option_default A Type P o b) := ltac:(destruct o; cbv [option_default]; exact _). +#[export] Polymorphic Instance quote_option_defaultP {A o} {P : A -> Prop} {b : Prop} {quoteP : forall x, o = Some x -> ground_quotable (P x)} {quoteP' : o = None -> ground_quotable b} : ground_quotable (@option_default A Prop P o b) := ltac:(destruct o; cbv [option_default]; exact _). #[export] Instance quote_on_Some {A P o} {quoteP : forall x, o = Some x -> ground_quotable (P x:Prop)} : ground_quotable (@on_Some A P o) := ltac:(destruct o; cbv [on_Some]; exact _). #[export] Typeclasses Opaque on_Some. diff --git a/quotation/theories/ToTemplate/Common/EnvironmentTyping.v b/quotation/theories/ToTemplate/Common/EnvironmentTyping.v index bb7dc2708..b1bb199fb 100644 --- a/quotation/theories/ToTemplate/Common/EnvironmentTyping.v +++ b/quotation/theories/ToTemplate/Common/EnvironmentTyping.v @@ -127,6 +127,8 @@ Module QuoteGlobalMaps (Import T : Term) (Import E : EnvironmentSig T) (Import T #[export] Instance quote_on_type {Σ Γ T} : ground_quotable (@on_type P Σ Γ T) := ltac:(cbv [on_type]; exact _). + #[export] Instance quote_on_type_rel {Σ Γ T r} : ground_quotable (@on_type_rel P Σ Γ T r) + := ltac:(cbv [on_type_rel]; exact _). #[export] Instance quote_satisfiable_udecl {univs ϕ} : ground_quotable (@satisfiable_udecl univs ϕ) := ltac:(cbv [satisfiable_udecl]; exact _). @@ -166,6 +168,7 @@ Module QuoteGlobalMaps (Import T : Term) (Import E : EnvironmentSig T) (Import T quote_type_local_ctx quote_sorts_local_ctx quote_on_type + quote_on_type_rel quote_on_udecl quote_satisfiable_udecl quote_valid_on_mono_udecl diff --git a/quotation/theories/ToTemplate/QuotationOf/Common/EnvironmentTyping/Sig.v b/quotation/theories/ToTemplate/QuotationOf/Common/EnvironmentTyping/Sig.v index b6742d89f..de6f99187 100644 --- a/quotation/theories/ToTemplate/QuotationOf/Common/EnvironmentTyping/Sig.v +++ b/quotation/theories/ToTemplate/QuotationOf/Common/EnvironmentTyping/Sig.v @@ -108,6 +108,7 @@ Module Type QuoteGlobalMapsSig (Import T: Term) (Import E: EnvironmentSig T) (Im #[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 Σ Γ j, ground_quotable (P Σ Γ j)} {Σ Γ T} : ground_quotable (@on_type P Σ Γ T). + #[export] Declare Instance quote_on_type_rel {P} {quoteP : forall Σ Γ j, ground_quotable (P Σ Γ j)} {Σ Γ T r} : ground_quotable (@on_type_rel P Σ Γ T r). #[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 ϕ). diff --git a/quotation/theories/ToTemplate/Utils/MCOption.v b/quotation/theories/ToTemplate/Utils/MCOption.v index 6d5a37cf0..fa8924e75 100644 --- a/quotation/theories/ToTemplate/Utils/MCOption.v +++ b/quotation/theories/ToTemplate/Utils/MCOption.v @@ -20,6 +20,7 @@ Proof. Defined. #[export] Polymorphic Instance quote_option_default {A P o b} {quoteP : forall x, o = Some x -> ground_quotable (P x)} {quoteP' : o = None -> ground_quotable b} : ground_quotable (@option_default A Type P o b) := ltac:(destruct o; cbv [option_default]; exact _). +#[export] Polymorphic Instance quote_option_defaultP {A o} {P : A -> Prop} {b : Prop} {quoteP : forall x, o = Some x -> ground_quotable (P x)} {quoteP' : o = None -> ground_quotable b} : ground_quotable (@option_default A Prop P o b) := ltac:(destruct o; cbv [option_default]; exact _). #[export] Instance quote_on_Some {A P o} {quoteP : forall x, o = Some x -> ground_quotable (P x:Prop)} : ground_quotable (@on_Some A P o) := ltac:(destruct o; cbv [on_Some]; exact _). #[export] Typeclasses Opaque on_Some. diff --git a/safechecker/theories/PCUICSafeChecker.v b/safechecker/theories/PCUICSafeChecker.v index 6f17da38f..81db5cd54 100644 --- a/safechecker/theories/PCUICSafeChecker.v +++ b/safechecker/theories/PCUICSafeChecker.v @@ -438,8 +438,20 @@ Section CheckEnv. typing_result (∑ T, forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ Σ ;;; Γ |- t ▹ T ∥) := infer X_impl X_ext Γ wfΓ t. + Definition check_judgment_wf_env X_ext {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X_ext -> NormalizationIn Σ} Γ (wfΓ : forall Σ (wfΣ : abstract_env_ext_rel X_ext Σ), ∥ wf_local Σ Γ ∥) j (Hu : j_univ j = None) + : typing_result (forall Σ (wfΣ : abstract_env_ext_rel X_ext Σ), ∥ lift_typing typing Σ Γ j ∥). + eapply typing_error_forget. + pose proof (infer_judgment X_impl X_ext (infer X_impl X_ext) Γ wfΓ (j_term j) (j_typ j) (j_rel j)) as X. + rewrite -Hu in X. + apply X. + Defined. + + Definition infer_judgment_wf_env X_ext {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X_ext -> NormalizationIn Σ} Γ (wfΓ : forall Σ (wfΣ : abstract_env_ext_rel X_ext Σ), ∥ wf_local Σ Γ ∥) j (Hs : forall Σ (wfΣ : abstract_env_ext_rel X_ext Σ), option_default (fun s => ∥ wf_sort Σ s ∥) (j_univ j) True) + : typing_result (∑ s', option_default (fun s => s = s') (j_univ j) True /\ forall Σ (wfΣ : abstract_env_ext_rel X_ext Σ), ∥ lift_typing typing Σ Γ (Judge (j_term j) (j_typ j) (Some s') (j_rel j)) ∥) := + infer_judgment0 X_impl X_ext (infer X_impl X_ext) Γ wfΓ (j_term j) (j_typ j) (j_univ j) Hs (j_rel j). + Equations infer_type_wf_env X_ext {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X_ext -> NormalizationIn Σ} Γ (wfΓ : forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ wf_local Σ Γ ∥) t : - typing_result (∑ u, forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ Σ ;;; Γ |- t : tSort u∥) := + typing_result (∑ s, forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ Σ ;;; Γ |- t : tSort s∥) := infer_type_wf_env X_ext Γ wfΓ t := '(y ; H) <- typing_error_forget (infer_type X_impl X_ext (infer X_impl X_ext) Γ wfΓ t) ;; ret (y ; _). @@ -476,8 +488,8 @@ Section CheckEnv. * 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. + eapply type_Cumul'; eauto. econstructor; pcuic. + apply unlift_TypUniv in h1. pcuic. reflexivity. Qed. @@ -486,13 +498,13 @@ Section CheckEnv. typing_result (forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ type_local_ctx (lift_typing typing) Σ Γ Δ s ∥) := match Δ with | [] => 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 |} :: Δ => + | {| decl_body := None |} as decl :: Δ => checkΔ <- check_type_local_ctx X_ext Γ Δ s wfΓ ;; - checkty <- check_type_wf_env X_ext (Γ ,,, Δ) _ ty (tSort s) ;; + checkty <- infer_judgment_wf_env X_ext (Γ ,,, Δ) _ (j_decl_s decl (Some s)) _ ;; ret _ - | {| decl_body := Some b; decl_type := ty |} :: Δ => + | {| decl_body := Some b |} as decl :: Δ => checkΔ <- check_type_local_ctx X_ext Γ Δ s wfΓ ;; - checkty <- check_type_wf_env X_ext (Γ ,,, Δ) _ b ty ;; + checkty <- check_judgment_wf_env X_ext (Γ ,,, Δ) _ (j_decl decl) _ ;; ret _ end. Next Obligation. @@ -500,48 +512,50 @@ Section CheckEnv. now sq; apply/PCUICWfUniverses.wf_sort_reflect. Qed. Next Obligation. - specialize_Σ H. sq. now eapply PCUICContexts.type_local_ctx_wf_local in checkΔ. + specialize_Σ wfΣ. sq. now eapply PCUICContexts.type_local_ctx_wf_local in checkΔ. + Qed. + Next Obligation. + pose proof (abstract_env_ext_wf _ wfΣ). + specialize_Σ wfΣ. sq. + now eapply PCUICInductives.type_local_ctx_wf. Qed. 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Δ. + specialize_Σ wfΣ. sq. now eapply PCUICContexts.type_local_ctx_wf_local in checkΔ. Qed. Next Obligation. - pose proof (abstract_env_ext_wf _ H). specialize_Σ H. sq. - split; auto. split; auto. - eapply PCUICValidity.validity in checkty as (_ & ?); auto. + specialize_Σ H. sq. + split; 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 Σ Γ ∥) : typing_result (∑ s, forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ sorts_local_ctx (lift_typing typing) Σ Γ Δ s ∥) := match Δ with | [] => ret ([]; fun _ _ => sq _) - | {| decl_body := None; decl_type := ty |} :: Δ => + | {| decl_body := None |} as decl :: Δ => '(Δs; Δinfer) <- infer_sorts_local_ctx X_ext Γ Δ wfΓ ;; - '(tys; tyinfer) <- infer_type_wf_env X_ext (Γ ,,, Δ) _ ty ;; + '(tys; tyinfer) <- infer_judgment_wf_env X_ext (Γ ,,, Δ) _ (j_decl decl) _ ;; ret ((tys :: Δs); _) - | {| decl_body := Some b; decl_type := ty |} :: Δ => + | {| decl_body := Some _ |} as decl :: Δ => '(Δs; Δinfer) <- infer_sorts_local_ctx X_ext Γ Δ wfΓ ;; - checkty <- check_type_wf_env X_ext (Γ ,,, Δ) _ b ty ;; + checkty <- check_judgment_wf_env X_ext (Γ ,,, Δ) _ (j_decl decl) _ ;; ret (Δs; _) end. Next Obligation. exact tt. Qed. Next Obligation. - specialize_Σ H. sq. now eapply PCUICContexts.sorts_local_ctx_wf_local in Δinfer. + specialize_Σ wfΣ. sq. now eapply PCUICContexts.sorts_local_ctx_wf_local in Δinfer. Qed. Next Obligation. - specialize_Σ H. sq. split; auto. repeat (eexists; tea). + specialize_Σ H. sq. split; auto. Qed. Next Obligation. - specialize_Σ H. sq. now eapply PCUICContexts.sorts_local_ctx_wf_local in Δinfer. + specialize_Σ wfΣ. 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 as (_ & ?); auto. + specialize_Σ H. sq. split; auto. Qed. Definition cumul_decl Pcmp Σ Γ (d d' : context_decl) : Type := cumul_decls Pcmp Σ Γ Γ d d'. @@ -728,7 +742,7 @@ Section CheckEnv. Definition monad_mapi (l : list A) := monad_mapi_rec 0 l. End MonadMapi. - Definition check_constructor_spec Σ (ind : nat) (mdecl : mutual_inductive_body) + Definition check_constructor_spec Σ (ind : nat) (mdecl : mutual_inductive_body) (idecl : one_inductive_body) (cdecl : constructor_body) (cs : constructor_univs) := isType Σ (arities_context mdecl.(ind_bodies)) (cstr_type cdecl) * (cstr_type cdecl = @@ -778,8 +792,13 @@ Section CheckEnv. Arguments eqb : simpl never. + Definition wf_ind_type (Σ : global_env_ext) (mdecl : mutual_inductive_body) (idecl : one_inductive_body) := + isType Σ [] idecl.(ind_type) × + idecl.(ind_type) = it_mkProd_or_LetIn mdecl.(ind_params) (it_mkProd_or_LetIn idecl.(ind_indices) (tSort idecl.(ind_sort))) × + isSortRel idecl.(ind_sort) idecl.(ind_relevance). + Definition wf_ind_types (Σ : global_env_ext) (mdecl : mutual_inductive_body) := - All (fun ind => isType Σ [] ind.(ind_type)) mdecl.(ind_bodies). + All (wf_ind_type Σ mdecl) mdecl.(ind_bodies). Lemma wf_ind_types_wf_arities (Σ : global_env_ext) (wfX : wf Σ) mdecl : wf_ind_types Σ mdecl -> @@ -788,6 +807,8 @@ Section CheckEnv. rewrite /wf_ind_types. unfold arities_context. induction 1; simpl; auto. + assert (isty : isTypeRel Σ [] (ind_type x) Relevant) + by (destruct p as (a & b & c); rewrite b in a |- *; eapply PCUICInductiveInversion.ind_arity_relevant => //). rewrite rev_map_cons /=. eapply All_local_env_app. constructor; pcuic. eapply All_local_env_impl; eauto. @@ -800,11 +821,11 @@ Section CheckEnv. Program Definition check_constructor X_ext {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X_ext -> NormalizationIn Σ} (ind : nat) (mdecl : mutual_inductive_body) (wfar : forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ wf_ind_types Σ mdecl ∥) (wfpars : forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ wf_local Σ (ind_params mdecl) ∥) - (cdecl : constructor_body) : - EnvCheck X_env_ext_type (∑ cs : constructor_univs, forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ check_constructor_spec Σ ind mdecl cdecl cs ∥) := + (idecl : one_inductive_body) (cdecl : constructor_body) : + EnvCheck X_env_ext_type (∑ cs : constructor_univs, forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ check_constructor_spec Σ ind mdecl idecl cdecl cs ∥) := - '(s; Hs) <- wrap_error _ X_ext ("While checking type of constructor: " ^ cdecl.(cstr_name)) - (infer_type_wf_env X_ext (arities_context mdecl.(ind_bodies)) _ cdecl.(cstr_type)) ;; + Hs <- wrap_error _ X_ext ("While checking type of constructor: " ^ cdecl.(cstr_name)) + (check_judgment_wf_env X_ext (arities_context mdecl.(ind_bodies)) _ (Typ cdecl.(cstr_type)) _) ;; match decompose_prod_n_assum [] #|mdecl.(ind_params)| cdecl.(cstr_type) with | Some (params, concl) => eqpars <- wrap_error _ X_ext cdecl.(cstr_name) @@ -833,7 +854,7 @@ Section CheckEnv. end. Next Obligation. - pose proof (abstract_env_ext_wf _ H); specialize_Σ H; sq. + pose proof (abstract_env_ext_wf _ wfΣ); specialize_Σ wfΣ; sq. now apply wf_ind_types_wf_arities in wfar. Qed. Next Obligation. @@ -851,12 +872,11 @@ Section CheckEnv. eapply eqb_eq in eqbindices. eapply eqb_eq in eqbpars. sq; red; cbn. intuition auto. - 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. rewrite Heq_anonymous0. do 4 f_equal. len. - rewrite -(firstn_skipn (ind_npars mdecl) x1); congruence. + rewrite -(firstn_skipn (ind_npars mdecl) x0); congruence. Qed. Fixpoint All_sigma {A B} {P : A -> B -> Type} {l : list A} (a : All (fun x => ∑ y : B, P x y) l) : @@ -881,10 +901,10 @@ Section CheckEnv. Definition check_constructors_univs X_ext {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X_ext -> NormalizationIn Σ} (id : ident) (mdecl : mutual_inductive_body) (wfar : forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ wf_ind_types Σ mdecl ∥) (wfpars : forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ wf_local Σ (ind_params mdecl) ∥) - (ind : nat) + (ind : nat) (idecl : one_inductive_body) (cstrs : list constructor_body) : EnvCheck X_env_ext_type (∑ cs : list constructor_univs, - forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ All2 (fun cstr cs => check_constructor_spec Σ ind mdecl cstr cs) cstrs cs ∥) := - css <- monad_All (fun d => check_constructor X_ext ind mdecl wfar wfpars d) cstrs ;; + forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ All2 (fun cstr cs => check_constructor_spec Σ ind mdecl idecl cstr cs) cstrs cs ∥) := + css <- monad_All (fun d => check_constructor X_ext ind mdecl wfar wfpars idecl d) cstrs ;; let '(cs; all2) := All_sigma css in ret (cs ; fun Σ wfΣ => All2_sq all2 Σ wfΣ). @@ -908,25 +928,6 @@ Section CheckEnv. now exists fty, s. Qed. - Lemma nth_error_arities_context mdecl i idecl : - nth_error (List.rev mdecl.(ind_bodies)) i = Some idecl -> - nth_error (arities_context mdecl.(ind_bodies)) i = - Some {| decl_name := {| binder_name := nNamed idecl.(ind_name); binder_relevance := idecl.(ind_relevance) |}; - decl_body := None; - decl_type := idecl.(ind_type) |}. - Proof using Type. - generalize mdecl.(ind_bodies) => l. - intros hnth. - epose proof (nth_error_Some_length hnth). autorewrite with len in H. - rewrite nth_error_rev in hnth. now autorewrite with len. - rewrite List.rev_involutive in hnth. autorewrite with len in hnth. - unfold arities_context. - rewrite rev_map_spec. - rewrite nth_error_rev; autorewrite with len; auto. - rewrite List.rev_involutive nth_error_map. - rewrite hnth. simpl. reflexivity. - Qed. - Definition smash_telescope acc Γ := List.rev (smash_context acc (List.rev Γ)). @@ -1359,31 +1360,6 @@ Section CheckEnv. Defined. End PositivityCheck. - - Program Fixpoint check_wf_local X_ext {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X_ext -> NormalizationIn Σ} Γ : - typing_result (forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ wf_local Σ Γ ∥) := - match Γ with - | [] => ret (fun _ _ => sq localenv_nil) - | {| decl_body := Some b; decl_type := ty |} :: Γ => - wfΓ <- check_wf_local X_ext Γ ;; - wfty <- check_type_wf_env X_ext Γ wfΓ b ty ;; - ret _ - | {| decl_body := None; decl_type := ty |} :: Γ => - wfΓ <- check_wf_local X_ext Γ ;; - wfty <- infer_type_wf_env X_ext Γ wfΓ ty ;; - ret _ - end. - Next Obligation. - pose proof (abstract_env_ext_wf _ H); specialize_Σ H. - sq. constructor; eauto. - split; auto. - eapply validity in wfty as []; auto. - Qed. - Next Obligation. - pose proof (abstract_env_ext_wf _ H); specialize_Σ H. - sq. constructor; auto. repeat (eexists; tea). - Qed. - Definition wt_indices Σ mdecl indices cs := wf_local Σ (ind_arities mdecl,,, ind_params mdecl,,, cs.(cstr_args)) * ctx_inst Σ (ind_arities mdecl,,, ind_params mdecl,,, cs.(cstr_args)) @@ -1437,10 +1413,10 @@ Section CheckEnv. eapply All2_fold_length in H. len in H. Qed. - Lemma wt_cstrs {n mdecl cstrs cs} X_ext : + Lemma wt_cstrs {n mdecl idecl cstrs cs} X_ext : (forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ All2 (fun (cstr : constructor_body) (cs0 : constructor_univs) => - check_constructor_spec Σ n mdecl cstr cs0) cstrs cs ∥) -> + check_constructor_spec Σ n mdecl idecl cstr cs0) cstrs cs ∥) -> forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ All (fun cstr => welltyped Σ (arities_context mdecl.(ind_bodies)) (cstr_type cstr)) cstrs ∥. Proof using Type. intros. specialize_Σ H0; sq. @@ -1452,20 +1428,21 @@ Section CheckEnv. Lemma get_wt_indices {mdecl cstrs cs} X_ext (wfar : forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ wf_ind_types Σ mdecl ∥) (wfpars : forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ wf_local Σ (ind_params mdecl) ∥) - (n : nat) (idecl : one_inductive_body) (indices : context) - (hnth : nth_error mdecl.(ind_bodies) n = Some idecl) - (heq : ∥ ∑ inds, idecl.(ind_type) = it_mkProd_or_LetIn (mdecl.(ind_params) ,,, indices) (tSort inds) ∥) : + (n : nat) (idecl : one_inductive_body) + (hnth : nth_error mdecl.(ind_bodies) n = Some idecl) : (forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ All2 (fun (cstr :constructor_body) (cs0 : constructor_univs) => - check_constructor_spec Σ (S n) mdecl cstr cs0) cstrs cs ∥) -> - forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ All (fun cs => wt_indices Σ mdecl indices cs) cstrs ∥. + check_constructor_spec Σ (S n) mdecl idecl cstr cs0) cstrs cs ∥) -> + forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ All (fun cs => wt_indices Σ mdecl idecl.(ind_indices) cs) cstrs ∥. Proof using Type. intros. pose proof (abstract_env_ext_wf _ H0) as wf; specialize_Σ H0. sq. solve_all. simpl. destruct X as [[[isTy eq] sorts] eq']. simpl in *. - assert(wf_local Σ (ind_params mdecl,,, indices)). - { eapply nth_error_all in wfar; eauto. simpl in wfar. - destruct heq as [s Hs]. rewrite Hs in wfar. + pose proof (wfars := wfar). + eapply nth_error_all in wfar as (wfar & eqind); eauto. + rewrite -it_mkProd_or_LetIn_app in eqind. + assert(wf_local Σ (ind_params mdecl,,, idecl.(ind_indices))). + { rewrite eqind in wfar. eapply isType_it_mkProd_or_LetIn_wf_local in wfar. now rewrite app_context_nil_l in wfar. } red. rewrite eq in isTy. @@ -1483,7 +1460,6 @@ Section CheckEnv. noconf H1; simpl in *. eapply PCUICSpine.typing_spine_strengthen in Hsp; eauto. clear cum. - destruct heq as [inds eqind]. move: Hsp. rewrite eqind. rewrite lift_it_mkProd_or_LetIn /= => Hs. rewrite closed_ctx_lift in Hs. eapply closed_wf_local; eauto. @@ -1500,7 +1476,6 @@ Section CheckEnv. len. rewrite nth_error_rev in hnth. len. rewrite List.rev_involutive in hnth. len in hnth. - eapply nth_error_all in wfar; tea. cbn in wfar. rewrite lift_closed; [now eapply isType_closed in wfar|]. now eapply isType_weaken. Qed. @@ -1772,14 +1747,13 @@ Section CheckEnv. (wfar : forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ wf_ind_types Σ mdecl ∥) (wfpars : forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ wf_local Σ (ind_params mdecl) ∥) (mdeclvar : forall Σ0, abstract_env_rel X Σ0 -> ∥ on_variance Σ0 mdecl.(ind_universes) mdecl.(ind_variance) ∥) - (n : nat) (idecl : one_inductive_body) (indices : context) + (n : nat) (idecl : one_inductive_body) (hnth : nth_error mdecl.(ind_bodies) n = Some idecl) - (heq : ∥ ∑ inds, idecl.(ind_type) = it_mkProd_or_LetIn (mdecl.(ind_params) ,,, indices) (tSort inds) ∥) : EnvCheck X_env_ext_type (∑ cs : list constructor_univs, - forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ on_constructors cumulSpec0 (lift_typing typing) Σ mdecl n idecl indices (ind_ctors idecl) cs ∥) := + forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ on_constructors cumulSpec0 (lift_typing typing) Σ mdecl n idecl idecl.(ind_indices) (ind_ctors idecl) cs ∥) := '(cs; Hcs) <- (check_constructors_univs X_ext (string_of_kername id) mdecl wfar - wfpars (S n) idecl.(ind_ctors));; + wfpars (S n) idecl idecl.(ind_ctors));; posc <- wrap_error _ X_ext (string_of_kername id) (monad_All_All (fun x px => @@ -1788,8 +1762,8 @@ Section CheckEnv. idecl.(ind_ctors) (wt_cstrs (cs:=cs) X_ext Hcs)) ;; var <- (monad_All_All (fun cs px => @monad_lift_ext X X_ext (EnvCheck X_env_ext_type) _ _ _ - (check_cstr_variance X mdecl id indices mdeclvar cs _ _)) - idecl.(ind_ctors) (get_wt_indices X_ext wfar wfpars n idecl indices hnth heq Hcs)) ;; + (check_cstr_variance X mdecl id idecl.(ind_indices) mdeclvar cs _ _)) + idecl.(ind_ctors) (get_wt_indices X_ext wfar wfpars n idecl hnth Hcs)) ;; lets <- monad_All (P := fun x => if @lets_in_constructor_types _ as _ return Prop then true else is_assumption_context (cstr_args x)) (fun cs => if @lets_in_constructor_types _ @@ -1814,7 +1788,7 @@ Section CheckEnv. destruct lets_in_constructor_types; congruence. Qed. Next Obligation. - epose proof (get_wt_indices X_ext wfar wfpars _ _ _ hnth heq Hcs). + epose proof (get_wt_indices X_ext wfar wfpars _ _ hnth Hcs). destruct HX. specialize_Σ H. pose proof (abstract_env_ext_wf _ (H1 _ H2)) as wfΣ. @@ -1827,6 +1801,9 @@ 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. + - eapply All_nth_error in wfar; tea. + destruct wfar as (isty & e & eqrel), wtinds. + eapply PCUICInductiveInversion.isTypeRel_cstr_type; eauto. apply wfΣ. - destruct wtinds as (? & ci). apply PCUICEnvTyping.ctx_inst_impl with (1 := ci) => t T HT. split; auto. @@ -2023,8 +2000,8 @@ End monad_Alli_nth_forall. destruct a as [na [b|] ty]. - intros []. eauto. - intros []; eauto. constructor; eauto. - destruct l as (_ & ? & t0 & <-). - now eapply typing_wf_sort in t0. + apply unlift_TypUniv in l. + now eapply typing_wf_sort in l. Qed. Obligation Tactic := Program.Tactics.program_simplify. @@ -2096,7 +2073,7 @@ End monad_Alli_nth_forall. sq. red. simpl. rewrite -eqvaru. unfold variance_universes in eqvaru. - unfold check_variance in mdeclvar. + unfold on_variance in mdeclvar. rewrite -eqvar in mdeclvar. destruct (ind_universes mdecl) as [|[inst cstrs]] => //. now eapply PCUICContextConversionTyp.eq_context_cumul_Spec_rel. @@ -2111,17 +2088,54 @@ End monad_Alli_nth_forall. rewrite -eqvaru in e; discriminate. Qed. + Program Definition check_ind_type X_ext {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X_ext -> NormalizationIn Σ} (mdecl : mutual_inductive_body) (idecl : one_inductive_body) + : EnvCheck X_env_ext_type (forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ wf_ind_type Σ mdecl idecl ∥) := + + '(ctxinds; p) <- + wrap_error _ X_ext idecl.(ind_name) ((match destArity [] idecl.(ind_type) as da return da = destArity [] idecl.(ind_type) -> typing_result (∑ ctxs, idecl.(ind_type) = it_mkProd_or_LetIn ctxs.1 (tSort ctxs.2)) with + | Some (ctx, s) => fun eq => ret ((ctx, s); _) + | None => fun _ => raise (NotAnArity idecl.(ind_type)) + end eq_refl)) ;; + let '(indices, params) := split_at (#|ctxinds.1| - #|mdecl.(ind_params)|) ctxinds.1 in + eqsort <- wrap_error _ X_ext idecl.(ind_name) + (check_eq_true (eqb ctxinds.2 idecl.(ind_sort)) + (Msg "Inductive body sort does not match the sort of the inductive type"));; + check <- wrap_error _ X_ext idecl.(ind_name) (check_eq_true (relevance_of_sort (ind_sort idecl) == (ind_relevance idecl)) (Msg "Wrong relevance")) ;; + (* Can probably be deduced instead of checked *) + eqpars <- wrap_error _ X_ext idecl.(ind_name) + (check_eq_true (eqb params mdecl.(ind_params)) + (Msg "Inductive arity parameters do not match the parameters of the mutual declaration"));; + eqindices <- wrap_error _ X_ext idecl.(ind_name) + (check_eq_true (eqb indices idecl.(ind_indices)) + (Msg "Inductive arity indices do not match the indices of the mutual declaration"));; + isty <- wrap_error _ X_ext idecl.(ind_name) + (infer_type_wf_env X_ext [] (fun _ _ => sq_wfl_nil _) idecl.(ind_type)) ;; + ret _. + + Next Obligation. + symmetry in eq. + apply destArity_spec_Some in eq. now simpl in eq. + Qed. + Next Obligation. + specialize_Σ H. + sq. splits. + solve_all. + - now exists isty. + - destruct (eqb_spec params (ind_params mdecl)); [|discriminate]. subst params. + rewrite split_at_firstn_skipn in Heq_anonymous. noconf Heq_anonymous. + rewrite -it_mkProd_or_LetIn_app {1}H0. apply eqb_eq in eqindices, eqsort. + rewrite -eqindices -eqsort. now rewrite /app_context firstn_skipn. + - apply eqb_eq, check. + Qed. + Program Definition check_ind_types X_ext {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X_ext -> NormalizationIn Σ} (mdecl : mutual_inductive_body) : EnvCheck X_env_ext_type (forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ wf_ind_types Σ mdecl ∥) := - indtys <- monad_All (fun ind => wrap_error _ X_ext ind.(ind_name) - (infer_type_wf_env X_ext [] (fun _ _ => sq_wfl_nil _) ind.(ind_type))) mdecl.(ind_bodies) ;; + indtys <- monad_All (check_ind_type X_ext mdecl) mdecl.(ind_bodies) ;; ret _. - Next Obligation. - eapply All_sigma in indtys as [indus Hinds]. - eapply All2_sq in Hinds; eauto. sq. - red. - solve_all. now eapply has_sort_isType. - Qed. + Next Obligation. + eapply All_sq, All_impl with (1 := indtys) => idecl Hi. + now apply Hi. + Qed. Program Definition check_one_ind_body X X_ext {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X_ext -> NormalizationIn Σ} (mind : kername) (mdecl : mutual_inductive_body) @@ -2133,41 +2147,14 @@ End monad_Alli_nth_forall. (hnth : nth_error mdecl.(ind_bodies) i = Some idecl) : EnvCheck X_env_ext_type (forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ on_ind_body cumulSpec0 (lift_typing typing) Σ mind mdecl i idecl ∥) := let id := string_of_kername mind in - '(ctxinds; p) <- - wrap_error _ X_ext id ((match destArity [] idecl.(ind_type) as da return da = destArity [] idecl.(ind_type) -> typing_result (∑ ctxs, idecl.(ind_type) = it_mkProd_or_LetIn ctxs.1 (tSort ctxs.2)) with - | Some (ctx, s) => fun eq => ret ((ctx, s); _) - | None => fun _ => raise (NotAnArity idecl.(ind_type)) - end eq_refl)) ;; - let '(indices, params) := split_at (#|ctxinds.1| - #|mdecl.(ind_params)|) ctxinds.1 in - eqsort <- wrap_error _ X_ext id - (check_eq_true (eqb ctxinds.2 idecl.(ind_sort)) - (Msg "Inductive body sort does not match the sort of the inductive type"));; - eqpars <- wrap_error _ X_ext id - (check_eq_true (eqb params mdecl.(ind_params)) - (Msg "Inductive arity parameters do not match the parameters of the mutual declaration"));; - eqindices <- wrap_error _ X_ext id - (check_eq_true (eqb indices idecl.(ind_indices)) - (Msg "Inductive arity indices do not match the indices of the mutual declaration"));; - '(cs; oncstrs) <- (check_constructors X X_ext mind mdecl pf wfars wfpars mdeclvar i idecl idecl.(ind_indices) hnth _) ;; + '(cs; oncstrs) <- (check_constructors X X_ext mind mdecl pf wfars wfpars mdeclvar i idecl hnth) ;; onprojs <- wrap_error _ X_ext ("Checking projections of " ^ id) (check_projections X_ext mind mdecl i idecl idecl.(ind_indices) cs oncstrs) ;; onsorts <- wrap_error _ X_ext ("Checking universes of " ^ id) (do_check_ind_sorts X_ext mdecl.(ind_params) wfpars idecl.(ind_kelim) - idecl.(ind_indices) cs _ ctxinds.2 _) ;; + idecl.(ind_indices) cs _ idecl.(ind_sort) _) ;; onindices <- check_indices X mdecl mind _ mdeclvar idecl.(ind_indices) _ ;; ret _. - Next Obligation. - symmetry in eq. - apply destArity_spec_Some in eq. now simpl in eq. - Qed. - - Next Obligation. - sq. exists t0. - destruct (eqb_spec params (ind_params mdecl)); [|discriminate]. subst params. - rewrite split_at_firstn_skipn in Heq_anonymous. noconf Heq_anonymous. - rewrite {1}H. apply eqb_eq in eqindices. - rewrite -eqindices. now rewrite /app_context firstn_skipn. - Qed. Next Obligation. intros ? ?. pose proof (abstract_env_ext_wf _ H). @@ -2186,8 +2173,8 @@ 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 & _). - clear X0; rewrite p in Hs. + destruct wfars as ((_ & s & Hs & _) & e & _). + clear X0; rewrite e -it_mkProd_or_LetIn_app in Hs. eapply PCUICSpine.inversion_it_mkProd_or_LetIn in Hs; eauto. eapply inversion_Sort in Hs as [_ [? _]]; eauto. Qed. @@ -2202,36 +2189,27 @@ 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 & _). - apply eqb_eq in eqpars; apply eqb_eq in eqindices; subst. - rewrite split_at_firstn_skipn in Heq_anonymous. - noconf Heq_anonymous. - rewrite {1}H0 {1}H1 /app_context firstn_skipn. - rewrite X1 in Hs. + destruct wfars as ((_ & s & Hs & _) & e & _). + rewrite e -it_mkProd_or_LetIn_app in Hs. eapply PCUICSpine.inversion_it_mkProd_or_LetIn in Hs; eauto. eapply typing_wf_local in Hs. now rewrite app_context_nil_l in Hs. Qed. Next Obligation. - rename X0 into oncstrs. rename x into cs. destruct Σ as [Σ ext]. + rename x into cs. destruct Σ as [Σ ext]. + pose proof (abstract_env_ext_wf _ H). pose proof (abstract_env_exists X) as [[Σ0 wfΣ0]]. - destruct pf as [pf pf']; eauto. specialize_Σ H. specialize_Σ wfΣ0. - destruct (eqb_spec params (ind_params mdecl)); [|discriminate]. - subst params. sq. + destruct pf as [pf pf']; eauto. specialize_Σ H. specialize_Σ wfΣ0. sq. + eapply All_nth_error in wfars as (isty & e & eqrel); tea. refine - {| ind_arity_eq := _; onArity := _; - ind_cunivs := cs; - onConstructors := oncstrs; - onProjections := _; - onIndices := _ |}. - - cbn in eqsort; apply eqb_eq in eqindices; apply eqb_eq in eqsort; subst. - 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. - - 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. + {| ind_arity_eq := e; onArity := _; + ind_cunivs := cs; + onConstructors := oncstrs; + onProjections := onprojs; + ind_sorts := onsorts; + ind_relevance_compat := eqrel; + onIndices := _; |}. + - rewrite e in isty |- *; apply PCUICInductiveInversion.ind_arity_relevant => //. apply H0. - erewrite (abstract_env_ext_irr _ _ pf); eauto. destruct (ind_variance mdecl) => //. Unshelve. eauto. @@ -2243,11 +2221,9 @@ End monad_Alli_nth_forall. : EnvCheck X_env_ext_type (forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ on_global_decl cumulSpec0 (lift_typing typing) Σ kn d ∥) := match d with | ConstantDecl cst => - match cst.(cst_body) with - | Some term => - check_wf_judgement kn X_ext term cst.(cst_type) ;; ret _ - | None => check_wf_type kn X_ext cst.(cst_type) ;; ret _ - end + wrap_error _ X_ext (string_of_kername kn) + (check_judgment_wf_env X_ext [] (fun _ _ => sq_wfl_nil _) (TermoptTypRel (cst_body cst) (cst_type cst) (cst_relevance cst)) eq_refl) + ;; ret _ | InductiveDecl mdecl => let id := string_of_kername kn in check_var <- @check_variance X kn (ind_universes mdecl) (ind_variance mdecl) _ ;; @@ -2263,12 +2239,7 @@ End monad_Alli_nth_forall. 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. + sq. assumption. Qed. Next Obligation. edestruct pf as [? ?]; specialize_Σ H. now pose proof (abstract_env_ext_wf _ H0). diff --git a/safechecker/theories/PCUICSafeReduce.v b/safechecker/theories/PCUICSafeReduce.v index e434d90dc..fb51abed7 100644 --- a/safechecker/theories/PCUICSafeReduce.v +++ b/safechecker/theories/PCUICSafeReduce.v @@ -1373,6 +1373,7 @@ Corollary R_Acc_aux : apply inversion_mkApps in typ as (fix_ty & typ_fix & typ_args); auto. apply inversion_Fix in typ_fix as (def&?&?&?&?&?&?); auto. eapply All_nth_error in a; eauto. + apply isTypeRel_isType in a. eapply wf_fixpoint_spine in i; eauto. 2: { eapply PCUICSpine.typing_spine_strengthen; eauto. } unfold unfold_fix in uf. diff --git a/safechecker/theories/PCUICSafeRetyping.v b/safechecker/theories/PCUICSafeRetyping.v index 467e12876..f7e041047 100644 --- a/safechecker/theories/PCUICSafeRetyping.v +++ b/safechecker/theories/PCUICSafeRetyping.v @@ -402,13 +402,23 @@ Qed. sq. constructor ; tea. inversion X0. - now eapply isTypebd_isType in X1. + apply lift_sorting_forget_univ in X1. + now eapply lift_sorting_lift_typing in X1. Defined. Next Obligation. cbn ; intros. destruct s1, s2. - cbn. specialize_Σ wfΣ. sq. + cbn. specialize_Σ wfΣ. + pose (hΣ _ wfΣ). + specialize (wfΓ' _ wfΣ) as wfΓ''. + sq. constructor; eauto. - repeat (eexists; tea). + inversion wfΓ''; subst. + apply lift_typing_lift_sorting in X1; tas. + pose proof (lift_sorting_extract X1). + assert (X1.2.π1 = x) as <-; tas. + clear X2. + destruct X1 as (? & x' & X1 & ?); cbn. + eapply infering_sort_sort; tea. Defined. Next Obligation. @@ -621,7 +631,7 @@ Qed. 1: now eapply red_terms_ws_cumul_pb_terms. eapply PCUICConvCumInversion.alt_into_ws_cumul_pb_terms ; tea. * fvs. - * eapply infering_ind_typing, validity, isType_open in X1 ; auto. + * eapply infering_ind_typing, validity, isType_is_open_term in X1 ; auto. rewrite on_free_vars_mkApps in X1. move: X1 => /andP [] _ /forallb_All ?. now eapply All_forallb, All_firstn. diff --git a/safechecker/theories/PCUICTypeChecker.v b/safechecker/theories/PCUICTypeChecker.v index ad2d28790..8d3b85fd0 100644 --- a/safechecker/theories/PCUICTypeChecker.v +++ b/safechecker/theories/PCUICTypeChecker.v @@ -392,23 +392,6 @@ Section Typecheck. Unshelve. eauto. Qed. - Equations infer_isType Γ (HΓ : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥wf_local Σ Γ ∥) T : typing_result_comp (forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ isType Σ Γ T ∥) := - infer_isType Γ HΓ T := - infer_type Γ HΓ T ;; - ret _. - Next Obligation. - pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. - now eapply infering_sort_isType. - Qed. - Next Obligation. - destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. - pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. - apply absurd. - eapply isType_infering_sort in H as [u ?]. - exists u. intros. erewrite (abstract_env_ext_irr _ _ wfΣ); eauto. - Unshelve. eauto. - Qed. - Equations bdcheck Γ (HΓ : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥wf_local Σ Γ ∥) t A (hA : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ isType Σ Γ A ∥) : typing_result_comp (forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ Σ ;;; Γ |- t ◃ A ∥) := bdcheck Γ HΓ t A hA := @@ -459,63 +442,160 @@ Section Typecheck. eapply type_reduction; eauto. exact r. Qed. *) - Lemma sq_wfl_nil Σ : ∥ wf_local Σ [] ∥. - Proof using Type. - repeat constructor. - Qed. + Local Notation check_eq_true b e := + (if b as b' return (typing_result_comp (is_true b')) then ret eq_refl else raise e). - Equations check_context Γ : typing_result_comp (forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ wf_local Σ Γ ∥) - := - check_context [] := ret _ ; - check_context ({| decl_body := None; decl_type := A |} :: Γ) := - HΓ <- check_context Γ ;; - infer_type Γ HΓ A ;; - ret _ ; - check_context ({| decl_body := Some t; decl_type := A |} :: Γ) := - HΓ <- check_context Γ ;; - infer_isType Γ HΓ A ;; - bdcheck Γ HΓ t A _ ;; + Equations check_eq {A : Type} {RE : ReflectEq A} (a b : A) (e : type_error) : typing_result_comp (a = b) := + check_eq a b e := + check_eq_true (a == b) e ;; + ret _. + Next Obligation. + now eapply eqb_eq. + Qed. + Next Obligation. + apply absurd. + apply eqb_refl. + Qed. + + #[global] Transparent relevance_of_family. (* We need to compute it really now *) + + Equations infer_judgment0 Γ (HΓ : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ wf_local Σ Γ ∥) tm ty s (Hs : forall Σ (wfΣ : abstract_env_ext_rel X Σ), option_default (fun s => ∥ wf_sort Σ s ∥) s True) rel + : typing_result_comp (∑ s', option_default (fun s => s = s') s True /\ forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ lift_typing typing Σ Γ (Judge tm ty (Some s') rel) ∥) := + infer_judgment0 Γ HΓ tm ty s Hs rel := + s <- + match s return (forall Σ _, option_default _ s True) -> typing_result_comp (∑ s', option_default (fun s => s = s') s True /\ forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ Σ ;;; Γ |- ty : tSort s' ∥) + with None => fun Hs => infer_type Γ HΓ ty ;; ret _ | Some s => fun Hs => bdcheck Γ HΓ ty (tSort s) _ ;; ret _ end Hs ;; + match rel return typing_result_comp (option_default (fun r => isSortRel s.π1 r) rel True) with None => ret I | Some rel => + check_eq (relevance_of_sort s.π1) rel (Msg "Wrong relevance") + end ;; + match tm return typing_result_comp (forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ option_default (fun tm => Σ ;;; Γ |- tm : ty) tm unit ∥) + with None => ret _ | Some t => bdcheck Γ HΓ t ty _ ;; ret _ end ;; ret _. Next Obligation. + specialize_Σ wfΣ. sq. + now eapply isType_Sort. + Qed. + Next Obligation. + exists s. split; auto. intros. pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. - econstructor ; tea. - destruct s as (_ & s). - split; cbn; tas. - now eapply checking_typing. + eapply checking_typing; tea. + now eapply isType_Sort. + Defined. + Next Obligation. + eapply absurd. intros. + pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. + now eapply typing_checking. Qed. Next Obligation. - eapply absurd. intros. - pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. - inversion H ; subst. - apply unlift_TermTyp in X1. - now eapply typing_checking. + exists s0. split; auto. intros. + pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. + now eapply infering_sort_typing. + Defined. + Next Obligation. + destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. + pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. + apply absurd. + eapply typing_infering_sort in H0 as (s' & X1 & _). + exists s'. intros. unshelve erewrite (abstract_env_ext_irr _ _ wfΣ); eauto. Qed. Next Obligation. - eapply absurd. intros. - pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. - inversion H ; subst. - destruct X1 as (_ & ?); split; cbn; auto. + specialize_Σ wfΣ. sq. + now eapply has_sort_isType. Qed. Next Obligation. - eapply absurd. intros. specialize_Σ wfΣ. sq. - now inversion H. + pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. + eapply checking_typing; tea. + now eapply has_sort_isType. Qed. Next Obligation. - pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. econstructor; tas. - eapply has_sort_isType. - now eapply infering_sort_typing. + eapply absurd. intros. + pose (hΣ _ wfΣ). + specialize_Σ wfΣ. sq. + now apply typing_checking. Qed. Next Obligation. - destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. - pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. eapply absurd. - inversion H ; subst. - eapply isType_infering_sort in X1 as [] ; tea. - eexists. intros. erewrite (abstract_env_ext_irr _ _ wfΣ); eauto. - Unshelve. eauto. + exists s. split; tas. intros. + pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. + repeat (eexists; tea). + Defined. + Next Obligation. + eapply absurd. intros. + pose (hΣ _ wfΣ). + specialize_Σ wfΣ. sq. + now apply H0. Qed. Next Obligation. - eapply absurd. intros. specialize_Σ wfΣ. sq. - now inversion H. + destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. + eapply absurd. intros. + pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. + destruct rel as [rel|] => //=. + apply lift_sorting_forget_body, lift_sorting_forget_univ in H0. + eapply isTypeRel_unique; tea. + now eapply has_sort_isTypeRel. + Qed. + Next Obligation. + destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. + pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. + apply absurd. + destruct H0 as (_ & s' & Hs0 & e' & _). cbn in Hs0, e'. subst X0. + exists s'. split; tas. intros. unshelve erewrite (abstract_env_ext_irr _ _ wfΣ); eauto. + Qed. + + Equations infer_judgment Γ (HΓ : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ wf_local Σ Γ ∥) tm ty rel + : typing_result_comp (forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ lift_typing typing Σ Γ (Judge tm ty None rel) ∥) := + infer_judgment Γ HΓ tm ty rel := + H <- infer_judgment0 Γ HΓ tm ty None _ rel ;; + ret _. + Next Obligation. + specialize_Σ wfΣ. sq. + now eapply lift_sorting_forget_univ. + Qed. + Next Obligation. + destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. + pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. + pose proof (lift_sorting_extract H). + apply absurd. + eexists; split; auto. + intros. + unshelve erewrite (abstract_env_ext_irr _ _ wfΣ); eauto. + Qed. + + + Definition check_decl Γ (HΓ : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ wf_local Σ Γ ∥) decl + : typing_result_comp (forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ lift_typing typing Σ Γ (j_decl decl) ∥) := + infer_judgment Γ HΓ _ _ _. + + Definition infer_isType Γ (HΓ : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ wf_local Σ Γ ∥) T : typing_result_comp (forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ isType Σ Γ T ∥) := + infer_judgment Γ HΓ _ _ _. + + Definition infer_isTypeRel Γ (HΓ : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ wf_local Σ Γ ∥) T rel + : typing_result_comp (forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ isTypeRel Σ Γ T rel ∥) := + infer_judgment Γ HΓ _ _ _. + + Lemma sq_wfl_nil Σ : ∥ wf_local Σ [] ∥. + Proof using Type. + repeat constructor. + Qed. + + Equations check_context Γ : typing_result_comp (forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ wf_local Σ Γ ∥) := + check_context [] := ret _ ; + check_context (decl :: Γ) := + HΓ <- check_context Γ ;; + Hd <- check_decl Γ HΓ decl ;; + ret _. + Next Obligation. + specialize_Σ wfΣ. sq. + now apply All_local_env_snoc. + Qed. + Next Obligation. + eapply absurd. intros. + specialize_Σ wfΣ. sq. + now apply All_local_env_tip in H as [_ H]. + Qed. + Next Obligation. + eapply absurd. intros. + specialize_Σ wfΣ. sq. + now apply All_local_env_tip in H as [H _]. Qed. Lemma sq_wf_local_app {Γ Δ} : forall Σ (wfΣ : abstract_env_ext_rel X Σ), @@ -529,51 +609,26 @@ Section Typecheck. check_context_rel Γ wfΓ [] := ret _ ; - check_context_rel Γ wfΓ ({| decl_body := None; decl_type := A |} :: Δ) := - wfΔ <- check_context_rel Γ wfΓ Δ ;; - infer_isType (Γ ,,, Δ) (fun Σ wfΣ => sq_wf_local_app Σ wfΣ (wfΓ Σ wfΣ) (wfΔ Σ wfΣ)) A ;; - ret _ ; - - check_context_rel Γ wfΓ ({| decl_body := Some t; decl_type := A |} :: Δ) := + check_context_rel Γ wfΓ (decl :: Δ) := wfΔ <- check_context_rel Γ wfΓ Δ ;; - wfA <- infer_isType (Γ ,,, Δ) (fun Σ wfΣ => sq_wf_local_app Σ wfΣ (wfΓ Σ wfΣ) (wfΔ Σ wfΣ)) A ;; - bdcheck (Γ ,,, Δ) (fun Σ wfΣ => sq_wf_local_app Σ wfΣ (wfΓ Σ wfΣ) (wfΔ Σ wfΣ)) t A wfA ;; + wfdecl <- check_decl (Γ ,,, Δ) (fun Σ wfΣ => sq_wf_local_app Σ wfΣ (wfΓ Σ wfΣ) (wfΔ Σ wfΣ)) decl ;; ret _. Next Obligation. sq. now constructor. 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. split; cbn; auto. - Qed. - Next Obligation. - apply absurd. intros. specialize_Σ wfΣ. sq. - now inversion H. - Qed. Next Obligation. specialize_Σ wfΣ. sq. - now constructor. + now apply All_local_rel_snoc. Qed. Next Obligation. - apply absurd. intros. specialize_Σ wfΣ. sq. - now inversion H. + eapply absurd. intros. + specialize_Σ wfΣ. sq. + now apply All_local_rel_tip in H as [_ H]. Qed. Next Obligation. - apply absurd. intros ;specialize_Σ wfΣ. sq. - now inversion H. + eapply absurd. intros. + specialize_Σ wfΣ. sq. + now apply All_local_rel_tip in H as [H _]. Qed. Equations check_ws_cumul_pb_decl (le : conv_pb) Γ d d' @@ -850,6 +905,7 @@ Section Typecheck. Qed. Next Obligation. specialize_Σ wfΣ; sq. eapply All_local_env_app_inv in wfΔ as [wt _]. + eapply isTypeRel_isType. now depelim wt. Qed. Next Obligation. @@ -861,6 +917,7 @@ Section Typecheck. rewrite subst_empty. apply checking_typing ; auto. apply All_local_rel_app_inv in wfΔ as [wt _]. + eapply isTypeRel_isType. now depelim wt. Qed. Next Obligation. @@ -871,6 +928,7 @@ Section Typecheck. constructor ; tea. apply checking_typing ; auto. eapply All_local_env_app_l in wfΔ. + eapply isTypeRel_isType. now inversion wfΔ ; subst. Qed. Next Obligation. @@ -1232,21 +1290,19 @@ Section Typecheck. (Γ : context) (wfΓ : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ wf_local Σ Γ ∥). Equations check_mfix_types (mfix : mfixpoint term) - : typing_result_comp (forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ All (fun x => isType Σ Γ (dtype x)) mfix ∥) := + : typing_result_comp (forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ All (on_def_type (lift_typing1 (typing Σ)) Γ) mfix ∥) := check_mfix_types [] := Checked_comp (fun Σ wfΣ => sq All_nil) ; (* (* probably not tail recursive but needed so that next line terminates *) check_mfix_types mfix ;; infer_type infer Γ wfΓ (dtype def) ;; ret _. *) check_mfix_types (def :: mfix) := - s <- infer_type infer Γ wfΓ (dtype def) ;; + s <- infer_isTypeRel infer Γ wfΓ (dtype def) (dname def).(binder_relevance);; check_mfix_types mfix ;; ret _. Next Obligation. pose proof (heΣ _ wfΣ) as [heΣ]. specialize_Σ wfΣ; sq. constructor ; tea. - eapply has_sort_isType. - now apply infering_sort_typing. Qed. Next Obligation. apply absurd. intros; specialize_Σ wfΣ; sq. @@ -1257,19 +1313,15 @@ Section Typecheck. pose proof (heΣ _ wfΣ) as [heΣ]; specialize_Σ wfΣ. sq. depelim H. apply absurd. - apply isType_infering_sort in i as [u ?]; tea. - exists u. intros. erewrite (abstract_env_ext_irr _ _ wfΣ); eauto. - Unshelve. eauto. + intros. unshelve erewrite (abstract_env_ext_irr _ _ wfΣ); eauto. Qed. Equations check_mfix_bodies (mfix : mfixpoint term) - (wf_types : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ All (fun x => isType Σ Γ (dtype x)) mfix ∥) + (wf_types : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ All (on_def_type (lift_typing1 (typing Σ)) Γ) mfix ∥) (Δ : context) (wfΔ : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ wf_local Σ (Γ,,,Δ) ∥) - : typing_result_comp (forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ All (fun d => - Σ ;;; Γ ,,, Δ |- - dbody d ◃ (lift0 #|Δ|) (dtype d)) mfix∥) := + : typing_result_comp (forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ All (on_def_body (lift_typing1 (typing Σ)) Δ Γ) mfix ∥) := check_mfix_bodies [] _ _ _ := Checked_comp (fun _ _ => sq All_nil) ; @@ -1283,15 +1335,26 @@ Section Typecheck. apply isType_lift ; eauto. - len. - rewrite skipn_all_app. + eapply isTypeRel_isType. now depelim wf_types. Qed. Next Obligation. pose proof (heΣ _ wfΣ) as [heΣ]. specialize_Σ wfΣ; sq. now depelim wf_types. - Qed. + Qed. Next Obligation. pose proof (heΣ _ wfΣ) as [heΣ]. specialize_Σ wfΣ; sq. constructor ; tea. + depelim wf_types. + assert (o' : isTypeRel Σ (Γ,,, Δ) (lift0 #|Δ| (dtype def)) (dname def).(binder_relevance)). + - eapply isTypeRel_lift; eauto. + 1: now len. + rewrite skipn_all_app. + assumption. + - apply isTypeRel_isType in o' as o''. + apply snd in o'. + split; cbn; tas. + eapply checking_typing; eauto. Qed. Next Obligation. apply absurd. intros. pose proof (heΣ _ wfΣ) as [heΣ]. specialize_Σ wfΣ; sq. @@ -1299,7 +1362,10 @@ Section Typecheck. Qed. Next Obligation. apply absurd. intros. pose proof (heΣ _ wfΣ) as [heΣ]. specialize_Σ wfΣ; sq. - now depelim H. + depelim H. + apply lift_typing_lift_sorting in o; eauto. + apply fst in o. cbn in o. + apply o. Qed. End check_mfix. @@ -1455,16 +1521,17 @@ Section Typecheck. infer Γ HΓ (tProd na A B) := s1 <- infer_type infer Γ HΓ A ;; + check_eq_true (relevance_of_sort s1.π1 == na.(binder_relevance)) (Msg "Wrong relevance") ;; s2 <- infer_type infer (Γ,,vass na A) _ B ;; Checked_comp (tSort (Sort.sort_of_product s1.π1 s2.π1);_) ; infer Γ HΓ (tLambda na A t) := - infer_type infer Γ HΓ A ;; + infer_isTypeRel infer Γ HΓ A na.(binder_relevance) ;; B <- infer (Γ ,, vass na A) _ t ;; ret (tProd na A B.π1; _); infer Γ HΓ (tLetIn n b b_ty b') := - infer_type infer Γ HΓ b_ty ;; + infer_isTypeRel infer Γ HΓ b_ty n.(binder_relevance) ;; bdcheck infer Γ HΓ b b_ty _ ;; b'_ty <- infer (Γ ,, vdef n b b_ty) _ b' ;; ret (tLetIn n b b_ty b'_ty.π1; _) ; @@ -1515,8 +1582,6 @@ Section Typecheck. (Msg "Case on coinductives disallowed") ;; check_eq_true (eqb (ind_npars mdecl) ci.(ci_npar)) (Msg "not the right number of parameters") ;; - (* check_eq_true (eqb (ind_relevance idecl) ci.(ci_relevance)) - (Msg "invalid relevance annotation on case") ;; *) (* let '(params, indices) := chop ci.(ci_npar) args in *) let chop_args := chop ci.(ci_npar) args in let params := chop_args.1 in let indices := chop_args.2 in @@ -1531,6 +1596,7 @@ Section Typecheck. let wfp : ∥ wf_predicate mdecl idecl p ∥ := _ in ps <- infer_type infer (Γ ,,, pctx) _ p.(preturn) ;; check_is_allowed_elimination ps.π1 _ (ind_kelim idecl);; + check_eq_true (relevance_of_sort ps.π1 == ci.(ci_relevance)) (Msg "Wrong relevance") ;; let ptm := it_mkLambda_or_LetIn pctx p.(preturn) in check_brs <- check_branches infer Γ HΓ ps.π1 ci mdecl idecl p indices isdecl isty _ _ _ 0 idecl.(ind_ctors) brs _ ;; @@ -1616,11 +1682,13 @@ Section Typecheck. (* intros Γ HΓ t na A B Heq_t [s ?]; *) pose proof (heΣ _ wfΣ) as [heΣ]. specialize_Σ wfΣ. sq; econstructor ; tea. - now eapply infering_sort_isType. + apply eqb_eq in i as <-. + now eapply infering_sort_isTypeRel. Qed. Next Obligation. (* intros Γ HΓ t na A B Heq_t [s1 ?] [s2 ?]; *) cbn. specialize_Σ wfΣ. sq. econstructor; try eassumption. + apply eqb_eq in i as <-. repeat (eexists; tea). Qed. Next Obligation. @@ -1629,6 +1697,16 @@ Section Typecheck. eexists. intros. sq. erewrite (abstract_env_ext_irr _ _ wfΣ); eauto. Unshelve. all: eauto. Qed. + Next Obligation. + destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. + pose (hΣ _ wfΣ). specialize_Σ wfΣ; sq. + inversion X1; subst. apply absurd. + case: eqb_spec => i //; exfalso; apply i; clear i. + apply lift_sorting_forget_univ in X3. + apply lift_sorting_lift_typing in X3; tas. + eapply isTypeRel_unique; tea. + now eapply infering_sort_isTypeRel. + Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. specialize_Σ wfΣ; sq. inversion X1; subst. apply absurd. @@ -1640,13 +1718,12 @@ Section Typecheck. (* intros Γ HΓ t0 na A t Heq_t [s ?]; *) pose proof (heΣ _ wfΣ) as [heΣ]. specialize_Σ wfΣ. sq; econstructor; tea. - now eapply infering_sort_isType. Qed. Next Obligation. (* intros Γ HΓ t0 na A t Heq_t [s ?] [B ?]; *) cbn; pose proof (heΣ _ wfΣ) as [heΣ]. specialize_Σ wfΣ. sq; econstructor; tas. - repeat (eexists; tea). + apply lift_typing_lift_sorting; eauto. Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. @@ -1656,28 +1733,31 @@ Section Typecheck. Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. + pose proof (heΣ _ wfΣ) as [heΣ]. specialize_Σ wfΣ; sq. inversion X1; subst. apply absurd. - destruct X2 as (_ & u & Hu & _). - eexists. intros. sq. unshelve erewrite (abstract_env_ext_irr _ _ wfΣ); eauto. + intros. sq. unshelve erewrite (abstract_env_ext_irr _ _ wfΣ); eauto. + apply lift_sorting_lift_typing; eauto. Qed. (* tLetIn *) Next Obligation. - pose proof (heΣ _ wfΣ) as [heΣ]. specialize_Σ wfΣ. - sq. eapply has_sort_isType. - eapply infering_sort_typing ; eauto. + specialize_Σ wfΣ. + sq. + now eapply isTypeRel_isType. 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. - repeat (eexists; tea); cbn. - 1: apply checking_typing ; eauto. - 1: eapply has_sort_isType. - all: now eapply infering_sort_typing. + apply isTypeRel_isType in s as s'. + destruct s as [_ s]. + split; tas; cbn. + apply checking_typing ; eauto. Qed. Next Obligation. + pose proof (heΣ _ wfΣ) as [heΣ]. cbn; specialize_Σ wfΣ; sq; econstructor; tas. - repeat (eexists; tea). + eapply lift_typing_lift_sorting in s as [_ s]; auto. + split; tas. Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. @@ -1688,14 +1768,16 @@ Section Typecheck. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. specialize_Σ wfΣ; sq. inversion X1; subst. apply absurd. - destruct X3 as (Ht & u & Hu & _). - eexists. intros. sq. unshelve erewrite (abstract_env_ext_irr _ _ wfΣ); eauto. + destruct X2 as (X2 & _); cbn in X2. + intros. sq. unshelve erewrite (abstract_env_ext_irr _ _ wfΣ); eauto. Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. + pose proof (heΣ _ wfΣ) as [heΣ]. specialize_Σ wfΣ; sq. inversion X1; subst. apply absurd. - destruct X2 as (Ht & u & Hu & _). - eexists. intros. sq. unshelve erewrite (abstract_env_ext_irr _ _ wfΣ); eauto. + apply lift_sorting_lift_typing in X2; auto. + apply lift_sorting_forget_body in X2. + intros. sq. unshelve erewrite (abstract_env_ext_irr _ _ wfΣ); eauto. Qed. (* tApp *) Next Obligation. @@ -1710,10 +1792,8 @@ Section Typecheck. specialize_Σ wfΣ ; sq. eapply infering_typing, type_reduction_closed, validity in X3. 2-4: eauto. - destruct X3 as (_ & s & HH & _). - eapply inversion_Prod in HH ; auto. - destruct HH as [s1 [_ [HH _]]]. - eapply lift_sorting_forget_univ. eassumption. + apply isType_tProd in X3 as (X3 & _). + now apply lift_sorting_forget_rel in X3. Qed. Next Obligation. cbn in *; specialize_Σ wfΣ ; sq. @@ -2026,6 +2106,12 @@ Section Typecheck. cbn in *. specialize_Σ wfΣ ; now sq. Qed. + Next Obligation. + intros. cbn in *. + destruct ps ; cbn in *. + cbn in *. specialize_Σ wfΣ ; now sq. + Qed. + Next Obligation. intros. cbn in *. assumption. @@ -2065,6 +2151,7 @@ Section Typecheck. - econstructor ; tea. now apply closed_red_red. - now eapply wf_local_rel_wf_local_bd, All_local_env_app_inv, wf_case_predicate_context. + - now eapply eqb_eq. - eapply ctx_inst_typing_bd ; eauto. eapply ctx_inst_smash. now rewrite subst_instance_smash /= in wt_params. @@ -2122,6 +2209,30 @@ Section Typecheck. intuition. Qed. + Next Obligation. + intros; clearbody isty wfp. + destruct cty as [A cty]. + subst ind' u args mdecl idecl isdecl. + destruct I as [ind' [u [args s']]]. + destruct d as [mdecl [idecl isdecl]]. + destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. + destruct ps as [ps ?]. + cbn in *. apply absurd. + case: eqb_spec => _i //; exfalso; apply _i; clear _i. + pose proof (heΣ _ wfΣ) as [heΣ]. + pose proof (hΣ _ wfΣ) as [hΣ]. + cbn in *. specialize_Σ wfΣ. sq. + edestruct X0 as [? [ty]]; eauto. + inversion ty ; subst. + rewrite -H8. do 2 f_equal. + eapply infering_sort_sort. 4: tea. all: eauto. + { apply All_local_env_app, wf_local_bd_rel_typing; eauto. } + relativize predctx0; tea. + unfold pctx, predctx0. + unshelve eapply declared_inductive_to_gen in isdecl, H3; eauto. + now eapply declared_inductive_inj in isdecl as []; tea. + Qed. + Next Obligation. intros; clearbody isty wfp. destruct cty as [A cty]. @@ -2565,20 +2676,13 @@ Section Typecheck. unfold abstract_env_fixguard in guarded. erewrite <- abstract_env_guard_correct in guarded; eauto. constructor; auto. - eapply All_impl ; tea; cbn. - intros. - 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. + - eapply All_impl with (1 := wf_types); cbn. + intros. + apply lift_typing_lift_sorting; eauto. + - eapply All_impl with (1 := wf_bodies); cbn. + intros. + apply lift_typing_lift_sorting; eauto. + - erewrite abstract_wf_fixpoint in wffix; eauto. Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. @@ -2598,16 +2702,17 @@ Section Typecheck. pose proof (heΣ _ wfΣ) as [heΣ]. cbn in *. specialize_Σ wfΣ ; sq. inversion X1. - apply All_impl with (1 := X3) => d [] //. + apply All_impl with (1 := X3) => d Hj. + eapply lift_sorting_lift_typing; eauto. + eapply All_mfix_wf; eauto. Qed. Next Obligation. apply absurd; intros. pose proof (heΣ _ wfΣ) as [heΣ]. cbn in *. specialize_Σ wfΣ ; sq. - inversion X1 ; subst. - eapply All_impl with (1 := X2). - intros d (_ & ? & ? & _). - now eapply infering_sort_isType. + inversion X1. + apply All_impl with (1 := X2) => d Hj. + eapply lift_sorting_lift_typing; eauto. Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. @@ -2627,20 +2732,13 @@ Section Typecheck. unfold abstract_env_cofixguard in guarded. erewrite <- abstract_env_guard_correct in guarded; eauto. constructor; auto. - eapply All_impl ; tea; cbn. - intros. - 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. + - eapply All_impl with (1 := wf_types); cbn. + intros. + apply lift_typing_lift_sorting; eauto. + - eapply All_impl with (1 := wf_bodies); cbn. + intros. + apply lift_typing_lift_sorting; eauto. + - erewrite abstract_wf_cofixpoint in wfcofix; eauto. Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. @@ -2660,16 +2758,17 @@ Section Typecheck. pose proof (heΣ _ wfΣ) as [heΣ]. cbn in *. specialize_Σ wfΣ ; sq. inversion X1. - apply All_impl with (1 := X3) => d [] //. + apply All_impl with (1 := X3) => d Hj. + eapply lift_sorting_lift_typing; eauto. + eapply All_mfix_wf; eauto. Qed. Next Obligation. apply absurd; intros. pose proof (heΣ _ wfΣ) as [heΣ]. cbn in *. specialize_Σ wfΣ ; sq. - inversion X1 ; subst. - eapply All_impl with (1 := X2). - intros d (_ & ? & ? & _). - now eapply infering_sort_isType. + inversion X1. + apply All_impl with (1 := X2) => d Hj. + eapply lift_sorting_lift_typing; eauto. Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. @@ -2736,6 +2835,7 @@ Section Typecheck. Qed. Definition check_isType := infer_isType infer. + Definition check_isTypeRel := infer_isTypeRel infer. Equations check Γ (HΓ : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ wf_local Σ Γ ∥) t A : typing_result_comp (forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ Σ;;; Γ |- t : A ∥) := diff --git a/template-coq/theories/AstUtils.v b/template-coq/theories/AstUtils.v index 8dd36445b..7fc8d37b2 100644 --- a/template-coq/theories/AstUtils.v +++ b/template-coq/theories/AstUtils.v @@ -387,10 +387,6 @@ Definition default_sort_family (s : sort) : allowed_eliminations := | _ => IntoAny end. -Definition default_relevance (s : sort) : relevance := - if Sort.is_sprop s then Irrelevant - else Relevant. - (** Convenience functions for building constructors and inductive declarations *) (** The [indrel] argument represents the de Bruijn associated to the inductive in the mutual block. @@ -421,7 +417,7 @@ Definition make_inductive_body (id : ident) (params : context) (indices : contex ind_kelim := default_sort_family s; ind_ctors := ind_ctors; ind_projs := []; - ind_relevance := default_relevance s |}. + ind_relevance := relevance_of_sort s |}. Ltac change_Sk := repeat match goal with diff --git a/template-coq/theories/EtaExpand.v b/template-coq/theories/EtaExpand.v index 2142e448c..214b0b7f1 100644 --- a/template-coq/theories/EtaExpand.v +++ b/template-coq/theories/EtaExpand.v @@ -1179,7 +1179,7 @@ Proof. rewrite nth_error_app1. 2:now rewrite nth_error_repeat. rewrite repeat_length. lia. - cbn. econstructor; eauto. * unfold map_branches. solve_all. - clear -hrepr X1 H8. + clear -hrepr X1 H9. set (Γ'' := map _ Γ'). cbn. enough (All (expanded Σ0 Γ'') (map (eta_expand Σg Γ') (pparams p ++ indices))). now rewrite map_app in X; eapply All_app in X as []. @@ -1201,7 +1201,7 @@ Proof. unfold inst_case_context. unfold subst_context. unfold subst_instance, subst_instance_context, map_context. rewrite fold_context_k_length map_length. unfold aname. lia. - } revert H9. generalize ((case_branch_context_gen (ci_ind ci) mdecl (pparams p) + } revert H10. generalize ((case_branch_context_gen (ci_ind ci) mdecl (pparams p) (puinst p) (bcontext y) x)). clear -hrepr. induction #|bcontext y|; intros []; cbn; intros; try congruence; econstructor; eauto. - cbn. rewrite nth_error_map H0. cbn. unfold eta_fixpoint. unfold fst_ctx in *. cbn in *. diff --git a/template-coq/theories/Typing.v b/template-coq/theories/Typing.v index 2a313d667..b59e33a56 100644 --- a/template-coq/theories/Typing.v +++ b/template-coq/theories/Typing.v @@ -776,6 +776,7 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> Σ ;;; Γ |- tSort s : tSort (Sort.super s) | type_Cast c k t s : + isSortRel s rel_of_Type -> Σ ;;; Γ |- t : tSort s -> Σ ;;; Γ |- c : t -> Σ ;;; Γ |- tCast c k t : t @@ -830,6 +831,7 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> (List.rev (ind_params mdecl ,,, ind_indices idecl : context)@[p.(puinst)]) -> Σ ;;; Γ ,,, predctx |- p.(preturn) : tSort ps -> is_allowed_elimination Σ idecl.(ind_kelim) ps -> + isSortRel ps ci.(ci_relevance) -> Σ ;;; Γ |- c : mkApps (tInd ci.(ci_ind) p.(puinst)) (p.(pparams) ++ indices) -> isCoFinite mdecl.(ind_finite) = false -> let ptm := it_mkLambda_or_LetIn predctx p.(preturn) in @@ -1103,7 +1105,7 @@ Lemma wf_local_inv `{checker_flags} {Σ Γ'} (w : wf_local Σ Γ') : forall d Γ, Γ' = d :: Γ -> ∑ (w' : wf_local Σ Γ) u, - { ty : lift_typing1 (typing Σ) Γ (Judge d.(decl_body) d.(decl_type) (Some u)) | + { ty : lift_typing1 (typing Σ) Γ (j_decl_s d (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 }. @@ -1160,6 +1162,7 @@ Lemma typing_ind_env `{cf : checker_flags} : (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) (c : term) (k : cast_kind) (t : term) (s : sort), + isSortRel s rel_of_Type -> Σ ;;; Γ |- 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 Σ Γ) (na : aname) (t b : term) (s1 s2 : sort), @@ -1226,6 +1229,7 @@ Lemma typing_ind_env `{cf : checker_flags} : P Σ (Γ ,,, predctx) p.(preturn) (tSort ps) -> PΓ Σ (Γ ,,, predctx) (typing_wf_local pret) -> is_allowed_elimination Σ idecl.(ind_kelim) ps -> + isSortRel ps ci.(ci_relevance) -> Σ ;;; Γ |- 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 -> @@ -1389,6 +1393,7 @@ Proof. 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 (ind_relevance_compat Xg). -- apply (onIndices Xg). * apply All_local_env_impl with (1 := onP); intros ?? Hj. apply Xj; tas. apply lift_typing_impl with (1 := Hj); intros ?? Hs. split; tas. diff --git a/template-pcuic/theories/PCUICToTemplateCorrectness.v b/template-pcuic/theories/PCUICToTemplateCorrectness.v index d1f847baf..55b4f23cc 100644 --- a/template-pcuic/theories/PCUICToTemplateCorrectness.v +++ b/template-pcuic/theories/PCUICToTemplateCorrectness.v @@ -1526,6 +1526,21 @@ Proof. econstructor 3; tea. Qed. +Lemma trans_lift_typing_it {P Q} {tm tm' t t'} {u r} : + forall tu: ST.lift_typing0 P (Judge tm t u r), + let s := tu.2.π1 in + match tm', tm with None, _ => unit | Some tm', Some tm => P tm t -> Q tm' t' | _, _ => False end -> + (P t (PCUICAst.tSort s) -> Q t' (tSort s)) -> + TT.lift_typing0 Q (Judge tm' t' u r). +Proof. + intros (? & ? & Hs & e) s HPQc HPQs. + split. + - destruct tm, tm' => //=. now apply HPQc. + - eexists. split; [now apply HPQs|]. + destruct u => //. +Qed. + + Definition TTwf_local {cf} Σ Γ := TT.All_local_env (TT.lift_typing TT.typing Σ) Γ. Lemma trans_wf_local' {cf} : @@ -1541,9 +1556,9 @@ Proof. induction X. - simpl. constructor. - simpl. econstructor; auto. - simpl. destruct tu as (? & ? & ? & ?); cbn in *. repeat (eexists; tea). + simpl. eapply trans_lift_typing_it with tu => //. - simpl. constructor; auto. - simpl. destruct tu as (? & ? & ? & ?); cbn in *. repeat (eexists; tea); cbn. + simpl. eapply trans_lift_typing_it with tu => //. Qed. Lemma trans_wf_local_env {cf} Σ Γ : @@ -1558,9 +1573,9 @@ Proof. induction X. - simpl. constructor. - simpl. econstructor; auto. - simpl. destruct t0 as (_ & ? & (? & ?) & ?); cbn in *. repeat (eexists; tea). + simpl. eapply trans_lift_typing_it with t0 => // HT. now apply HT. - simpl. constructor; auto. - simpl. destruct t0 as ((? & ?) & ? & (? & ?) & ?); cbn in *. repeat (eexists; cbn; tea). + simpl. eapply trans_lift_typing_it with t0 => // HT. all: now apply HT. Qed. Lemma trans_branches {cf} Σ Γ brs btys ps: @@ -1610,10 +1625,8 @@ Proof. } induction X;cbn;constructor;auto;cbn in *. - - destruct t0 as (_&?&(?&?)&?). - repeat (eexists; tea); eauto. - - destruct t0 as ((?&?)&?&(?&?)&?). - repeat (eexists; tea); eauto. + - eapply trans_lift_typing_it with t0 => // HT. now apply HT. + - eapply trans_lift_typing_it with t0 => // HT. all: now apply HT. Qed. Lemma trans_mfix_All2 {cf} Σ Γ mfix xfix: @@ -1626,14 +1639,14 @@ Proof. induction 1. - constructor. - 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. + eapply trans_lift_typing_it with p => // HT. 2: set (_.2.π1) as s in *; clearbody s. + all: unfold app_context in *. + all: rewrite /trans_local map_app trans_fix_context in HT. + all: rewrite trans_lift in HT. + all: replace(#|SE.fix_context xfix|) with + (#|TT.fix_context (map(map_def trans trans) xfix)|) in HT; tas. + + all: now rewrite TT.fix_context_length map_length fix_context_length. Qed. Lemma All_over_All {cf} Σ Γ wfΓ : @@ -2297,10 +2310,11 @@ Proof. now rewrite -trans_local_app. + rewrite <- trans_global_ext_constraints. eassumption. + + assumption. + now rewrite trans_mkApps map_app in X7. + rewrite (trans_case_predicate_context Γ); tea. eapply All2i_map. eapply All2i_map_right. - eapply Forall2_All2 in H4. + eapply Forall2_All2 in H5. eapply All2i_All2_mix_left in X8; tea. eapply All2i_impl ; tea. intros i cdecl br. cbv beta. @@ -2310,10 +2324,10 @@ Proof. * rewrite /brctxty. now eapply trans_cstr_branch_context_eq. * pose proof (trans_case_branch_type ci mdecl idecl cdecl i p br isdecl H1 wf eqctx). - rewrite -/brctxty -/ptm in H5. cbn in H5. clearbody brctxty. + rewrite -/brctxty -/ptm in H6. cbn in H6. clearbody brctxty. subst brctxty. rewrite -trans_local_app. cbn. apply IHb. * pose proof (trans_case_branch_type ci mdecl idecl cdecl i p br isdecl H1 wf eqctx). - rewrite -/brctxty -/ptm in H5. cbn in H5. clearbody brctxty. + rewrite -/brctxty -/ptm in H6. cbn in H6. clearbody brctxty. subst brctxty. rewrite -trans_local_app. cbn. apply IHbty. - rewrite trans_subst trans_subst_instance /= map_rev. @@ -2333,9 +2347,7 @@ 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 X1). - intros d (_ & ? & ? & _) => //. - repeat (eexists; tea). + + eapply All_map, X1. + fold trans. subst types. eapply trans_mfix_All2. eassumption. @@ -2351,9 +2363,7 @@ 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 X1). - intros d (_ & s & IHt & _). - repeat (eexists; tea); cbn. + eapply All_map, X1. + fold trans;subst types. now apply trans_mfix_All2. + now rewrite trans_wf_cofixpoint. diff --git a/template-pcuic/theories/TemplateToPCUIC.v b/template-pcuic/theories/TemplateToPCUIC.v index 7df5f442f..7714eaa65 100644 --- a/template-pcuic/theories/TemplateToPCUIC.v +++ b/template-pcuic/theories/TemplateToPCUIC.v @@ -72,7 +72,7 @@ Section Trans. | Ast.tLambda na T M => tLambda na (trans T) (trans M) | Ast.tApp u v => mkApps (trans u) (List.map trans v) | Ast.tProd na A B => tProd na (trans A) (trans B) - | Ast.tCast c kind t => tApp (tLambda (mkBindAnn nAnon Relevant) (trans t) (tRel 0)) (trans c) + | Ast.tCast c kind t => tApp (tLambda (mkBindAnn nAnon rel_of_Type) (trans t) (tRel 0)) (trans c) | Ast.tLetIn na b t b' => tLetIn na (trans b) (trans t) (trans b') | Ast.tCase ci p c brs => let p' := Ast.map_predicate id trans trans p in diff --git a/template-pcuic/theories/TemplateToPCUICCorrectness.v b/template-pcuic/theories/TemplateToPCUICCorrectness.v index 1f7514b1c..eb162e3d4 100644 --- a/template-pcuic/theories/TemplateToPCUICCorrectness.v +++ b/template-pcuic/theories/TemplateToPCUICCorrectness.v @@ -1772,6 +1772,20 @@ Proof. eapply IHX2. auto. eapply wf_red1 in r; tea. now eapply typing_wf_sigma; auto. Qed. +Lemma trans_lift_typing_it {P Q} {tm tm'} {t t'} {u r} : + forall tu: ST.lift_typing0 P (Judge tm t u r), + let s := tu.2.π1 in + match tm', tm with None, _ => unit | Some tm', Some tm => P tm t -> Q tm' t' | _, _ => False end -> + (P t (Ast.tSort s) -> Q t' (tSort s)) -> + lift_typing0 Q (Judge tm' t' u r). +Proof. + intros (? & ? & Hs & e) s HPQc HPQs. + split. + - destruct tm, tm' => //=. now apply HPQc. + - eexists. split; [now apply HPQs|]. + destruct u => //. +Qed. + Definition TTy_wf_local {cf : checker_flags} Σ Γ := ST.All_local_env (ST.lift_typing ST.typing Σ) Γ. Lemma trans_wf_local {cf}: @@ -1791,9 +1805,9 @@ Proof. induction X0. - simpl. constructor. - simpl. econstructor; auto. - simpl. destruct tu as (? & ? & ? & ?); cbn in *. repeat (eexists; tea). eauto. + simpl. eapply trans_lift_typing_it with tu => //. eauto. - simpl. constructor; auto. - simpl. destruct tu as (? & ? & ? & ?); cbn in *. repeat (eexists; tea); cbn. all: eauto. + simpl. eapply trans_lift_typing_it with tu => //. all: eauto. Qed. Lemma trans_wf_local_env {cf} Σ Γ : @@ -1809,9 +1823,9 @@ Proof. induction X. - simpl. constructor. - simpl. econstructor; auto. - simpl. destruct t0 as (_ & ? & (? & ?) & ?); cbn in *. repeat (eexists; tea). + simpl. eapply trans_lift_typing_it with t0 => // HT. now apply HT. - simpl. constructor; auto. - simpl. destruct t0 as ((? & ?) & ? & (? & ?) & ?); cbn in *. repeat (eexists; cbn; tea). + simpl. eapply trans_lift_typing_it with t0 => // HT. all: now apply HT. Qed. #[global] @@ -2092,6 +2106,7 @@ Lemma simpl_type_Case {H : checker_flags} {Σ : global_env_ext} {Γ} {ci : case_ wf_local Σ (Γ,,, predctx) -> Σ;;; Γ,,, predctx |- preturn p : tSort ps -> is_allowed_elimination Σ (ind_kelim idecl) ps -> + isSortRel ps ci.(ci_relevance) -> Σ;;; Γ |- c : mkApps (tInd ci (puinst p)) (pparams p ++ indices) -> isCoFinite (ind_finite mdecl) = false -> let ptm := it_mkLambda_or_LetIn predctx (preturn p) in @@ -2262,17 +2277,17 @@ Proof. f_equal. - (* Casts *) - assert (lift_typing0 (typing Σ' (trans_local (trans_global_env Σ.1) Γ)) (TypUniv (trans (trans_global_env Σ.1) t) s)). + assert (lift_typing0 (typing Σ' (trans_local (trans_global_env Σ.1) Γ)) (TypUnivRel (trans (trans_global_env Σ.1) t) s rel_of_Type)). { repeat (eexists; eauto). } + apply lift_sorting_forget_univ in X4 as X4'. eapply refine_type; cbn. * eapply type_App. - 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. } + 2:{ eapply type_Lambda; eauto. eapply type_Rel. econstructor; eauto. + eapply typing_wf_local; eauto. reflexivity. } eapply type_Prod. eauto. instantiate (1 := s). simpl. eapply (weakening _ _ [_] _ (tSort _)); eauto. constructor; eauto. eapply typing_wf_local; eauto. - now eapply has_sort_isType. now eapply X2. * unfold subst1. rewrite simpl_subst; auto. now rewrite lift0_p. @@ -2305,7 +2320,7 @@ Proof. 2:apply WfAst.wf_subst; try constructor; auto. 2:now inv wfAB. specialize (p X2). specialize (p0 X2). eapply PCUICInversion.inversion_Prod in p as [s1 [s2 [HA [HB Hs]]]]; auto. - eapply (PCUICArities.isType_subst (Δ := [vass na (trans Σ' A)])); eauto. + eapply (isType_subst (Δ := [vass na (trans Σ' A)])); eauto. eapply subslet_ass_tip. eauto with pcuic. now eapply has_sort_isType. @@ -2337,7 +2352,7 @@ Proof. - cbn; rewrite trans_mkApps; auto with wf trans. pose proof (forall_decls_declared_inductive _ _ _ _ _ _ isdecl). rewrite trans_lookup_inductive. - rewrite (declared_inductive_lookup _ H4). + rewrite (declared_inductive_lookup _ H5). rewrite trans_it_mkLambda_or_LetIn. rewrite -/(trans_local Σ' (Ast.case_predicate_context _ _ _ _)). have lenpctx : #|Ast.pcontext p| = S #|Ast.Env.ind_indices idecl|. @@ -2352,9 +2367,9 @@ Proof. eapply eq_binder_annots_eq. now eapply trans_ind_predicate_context. + cbn. split. cbn. - { epose proof (declared_minductive_ind_npars (proj1 H4)). + { epose proof (declared_minductive_ind_npars (proj1 H5)). cbn in H4. len. rewrite -H0. - now rewrite context_assumptions_map in H5. } + now rewrite context_assumptions_map in H6. } cbn. rewrite map2_map2_bias_left. { rewrite PCUICCases.ind_predicate_context_length; len. } rewrite map_map2 /= map2_cst. @@ -2434,14 +2449,13 @@ Proof. eapply fix_guard_trans. assumption. now rewrite nth_error_map H0. -- 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. + intros d tu. now apply tu. + -- apply All_map. eapply (All_impl X3). + intros d tu. specialize (tu X4). unfold on_def_body, types in *; cbn in *. - rewrite H2. rewrite /trans_local !map_length map_app in Htb, Hts |- *. + rewrite H2. rewrite /trans_local !map_length map_app in tu |- *. rewrite <- trans_lift. - repeat (eexists; tea); cbn; eauto. + apply tu. -- eapply trans_wf_fixpoint => //. clear X1 X3. solve_all; destruct a as (_ & s & Hs & _), b as (Hs' & _). now eapply TypingWf.typing_wf in Hs. @@ -2459,14 +2473,13 @@ Proof. -- now eapply cofix_guard_trans. -- now rewrite nth_error_map H0. -- 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. + intros d tu. now apply tu. + -- apply All_map. eapply (All_impl X3). + intros d tu. specialize (tu X4). unfold on_def_body, types in *; cbn in *. - rewrite H2. rewrite /trans_local !map_length map_app in Htb, Hts |- *. + rewrite H2. rewrite /trans_local !map_length map_app in tu |- *. rewrite <- trans_lift. - repeat (eexists; tea); cbn; eauto. + apply tu. -- eapply trans_wf_cofixpoint => //. clear X1 X3. solve_all; destruct a as (_ & s & Hs & _), b as (Hs' & _). now eapply TypingWf.typing_wf in Hs. @@ -2911,9 +2924,9 @@ Proof. destruct a as [na [b|] ty] => //; intros [? ?]; cbn. all: split; auto. - - eapply (IH _ _ (TermTyp _ _)) in l; auto. + - eapply (IH _ _ (TermTypRel _ _ _)) in l; auto. rewrite -trans_local_app //. - - eapply (IH _ _ (TypUniv _ _)) in l => //. + - eapply (IH _ _ (TypUnivRel _ _ _)) in l => //. rewrite -trans_local_app //. Qed. @@ -3067,7 +3080,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. - * eapply (X (Σg, _) [] (TermoptTyp _ _)); auto. + * eapply (X (Σg, _) [] (TermoptTypRel _ _ _)); 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 *. @@ -3079,14 +3092,15 @@ 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. + eapply isTypeRel_isType. 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) + have onarity : on_type_rel (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) [] (Typ (Ast.Env.ind_type idecl))). } + (ind_type (trans_one_ind_body (trans_global_env Σg) idecl)) rel_of_Type. + { apply ST.onArity in oni. unfold on_type_rel in *; simpl in *. + now apply (X (Σg, Ast.Env.ind_universes m) [] (TypRel (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. @@ -3097,14 +3111,14 @@ Proof. rename b into onc. rename y into cs. rename a0 into wfctype. rename a into wfcargs. rename b1 into wfcindices. destruct onc. - have onty : on_type (PCUICEnvTyping.lift_typing typing) + have onty : on_type_rel (PCUICEnvTyping.lift_typing typing) (trans_global (Σg, Ast.Env.ind_universes m)) (arities_context (ind_bodies (trans_minductive_body (trans_global_env Σg) m))) - (cstr_type (trans_constructor_body (trans_global_env Σg) x)). + (cstr_type (trans_constructor_body (trans_global_env Σg) x)) (Ast.Env.ind_relevance idecl). { 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)) - (Typ (Ast.Env.cstr_type x))). + (TypRel (Ast.Env.cstr_type x) (Ast.Env.ind_relevance idecl))). rewrite trans_arities_context. intros H'. apply H' => //. } have ceq : cstr_type (trans_constructor_body (trans_global_env Σg) x) = @@ -3137,9 +3151,9 @@ Proof. specialize (foo (Ast.Env.arities_context (Ast.Env.ind_bodies m) ,,, Ast.Env.ind_params m ,,, c)); rewrite /trans_local !map_app in foo. - now apply (foo (TermTyp b ty)). - now apply (foo (TermTyp b ty)). - now apply (foo (TypUniv ty _)). + now apply (foo (TermTypRel b ty _)). + now apply (foo (TermTypRel b ty _)). + now apply (foo (TypUnivRel ty _ _)). + clear -X0 onu wfΣg IHond X on_cindices. revert on_cindices. rewrite -trans_lift_context /trans_local -map_rev. @@ -3255,6 +3269,8 @@ Proof. rewrite trans_subst trans_lift. f_equal. now rewrite trans_projs. --- have inds := oni.(ST.ind_sorts). eapply trans_check_ind_sorts in inds; tea. + --- have compat := oni.(ST.ind_relevance_compat). + apply compat. --- have inds := oni.(ST.onIndices). simpl. destruct Ast.Env.ind_variance => //.