From edb57eeddc0bfa481c476148035ef636c36bd8c7 Mon Sep 17 00:00:00 2001 From: Yann Leray Date: Wed, 4 May 2022 16:13:42 +0200 Subject: [PATCH 01/17] Rebase relevance Template --- template-coq/theories/AstUtils.v | 2 + template-coq/theories/BasicAst.v | 8 +- template-coq/theories/EnvironmentTyping.v | 157 +++++++++++-------- template-coq/theories/EtaExpand.v | 14 +- template-coq/theories/Typing.v | 180 ++++++++++------------ template-coq/theories/TypingWf.v | 157 ++++++------------- template-coq/theories/Universes.v | 52 +++++++ template-coq/theories/WfAst.v | 20 +-- template-coq/theories/utils/All_Forall.v | 14 ++ 9 files changed, 309 insertions(+), 295 deletions(-) diff --git a/template-coq/theories/AstUtils.v b/template-coq/theories/AstUtils.v index d971222c9..94441e478 100644 --- a/template-coq/theories/AstUtils.v +++ b/template-coq/theories/AstUtils.v @@ -439,6 +439,8 @@ Definition tFixType {A} (P P' : A -> Type) (m : mfixpoint A) := Ltac solve_all_one := try lazymatch goal with + | H: [× _, _ & _] |- _ => destruct H + | H: [× _, _, _ & _] |- _ => destruct H | H: tCasePredProp _ _ _ |- _ => destruct H end; unfold tCaseBrsProp, tFixProp, tCaseBrsType, tFixType in *; diff --git a/template-coq/theories/BasicAst.v b/template-coq/theories/BasicAst.v index 81bb51553..ce9f97c0c 100644 --- a/template-coq/theories/BasicAst.v +++ b/template-coq/theories/BasicAst.v @@ -208,19 +208,21 @@ Proof. eapply map_def_spec; eauto. Qed. -Variant typ_or_sort_ {term} := Typ (T : term) | Sort. +Variant typ_or_sort_ {term} := Typ (T : term) | Sort (relopt : option relevance). Arguments typ_or_sort_ : clear implicits. +Definition SortRel {term} rel : typ_or_sort_ term := Sort (Some rel). + Definition typ_or_sort_map {T T'} (f: T -> T') t := match t with | Typ T => Typ (f T) - | Sort => Sort + | Sort relopt => Sort relopt end. Definition typ_or_sort_default {T A} (f: T -> A) t d := match t with | Typ T => f T - | Sort => d + | Sort _ => d end. Section Contexts. diff --git a/template-coq/theories/EnvironmentTyping.v b/template-coq/theories/EnvironmentTyping.v index 336cbc254..ec79e1e88 100644 --- a/template-coq/theories/EnvironmentTyping.v +++ b/template-coq/theories/EnvironmentTyping.v @@ -273,12 +273,12 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). | localenv_cons_abs Γ na t : All_local_env Γ -> - typing Γ t Sort -> + typing Γ t (SortRel na.(binder_relevance)) -> All_local_env (Γ ,, vass na t) | localenv_cons_def Γ na b t : All_local_env Γ -> - typing Γ t Sort -> + typing Γ t (SortRel na.(binder_relevance)) -> typing Γ b (Typ t) -> All_local_env (Γ ,, vdef na b t). Derive Signature NoConfusion for All_local_env. @@ -331,13 +331,13 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). := localenv_nil. Definition All_local_rel_abs {P Γ Γ' A na} : - All_local_rel P Γ Γ' -> P (Γ ,,, Γ') A Sort + All_local_rel P Γ Γ' -> P (Γ ,,, Γ') A (SortRel na.(binder_relevance)) -> All_local_rel P Γ (Γ',, vass na A) := localenv_cons_abs. Definition All_local_rel_def {P Γ Γ' t A na} : All_local_rel P Γ Γ' -> - P (Γ ,,, Γ') A Sort -> + P (Γ ,,, Γ') A (SortRel na.(binder_relevance)) -> P (Γ ,,, Γ') t (Typ A) -> All_local_rel P Γ (Γ',, vdef na t A) := localenv_cons_def. @@ -395,29 +395,35 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). Definition on_local_decl (P : context -> term -> typ_or_sort -> Type) Γ d := match d.(decl_body) with | Some b => P Γ b (Typ d.(decl_type)) - | None => P Γ d.(decl_type) Sort + | None => P Γ d.(decl_type) (SortRel d.(decl_name).(binder_relevance)) end. + Lemma All_local_env_inv P (d : context_decl) (Γ : context) (X : All_local_env P (Γ ,, d)) : + on_local_decl P Γ d * All_local_env P Γ. + Proof. + inv X; intuition; red; simpl; eauto. + Qed. + Definition on_def_type (P : context -> term -> typ_or_sort -> Type) Γ d := - P Γ d.(dtype) Sort. + P Γ d.(dtype) (SortRel d.(dname).(binder_relevance)). Definition on_def_body (P : context -> term -> typ_or_sort -> Type) types Γ d := P (Γ ,,, types) d.(dbody) (Typ (lift0 #|types| d.(dtype))). Definition lift_judgment (check : global_env_ext -> context -> term -> term -> Type) - (infer_sort : global_env_ext -> context -> term -> Type) : + (infer_sort : global_env_ext -> context -> term -> option relevance -> Type) : (global_env_ext -> context -> term -> typ_or_sort -> Type) := fun Σ Γ t T => match T with | Typ T => check Σ Γ t T - | Sort => infer_sort Σ Γ t + | Sort relopt => infer_sort Σ Γ t relopt end. Lemma lift_judgment_impl {P Ps Q Qs Σ Σ' Γ Γ' t t' T} : lift_judgment P Ps Σ Γ t T -> (forall T, P Σ Γ t T -> Q Σ' Γ' t' T) -> - (Ps Σ Γ t -> Qs Σ' Γ' t') -> + (forall relopt, Ps Σ Γ t relopt -> Qs Σ' Γ' t' relopt) -> lift_judgment Q Qs Σ' Γ' t' T. Proof. intros HT HPQ HPsQs. @@ -428,9 +434,10 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). (* Common uses *) - Definition lift_wf_term wf_term := (lift_judgment (fun Σ Γ t T => wf_term Σ Γ t × wf_term Σ Γ T) wf_term). + Definition lift_wf_term wf_term := (lift_judgment (fun Σ Γ t T => wf_term Σ Γ t × wf_term Σ Γ T) (fun Σ Γ t relopt => wf_term Σ Γ t)). - Definition infer_sort (sorting : global_env_ext -> context -> term -> Universe.t -> Type) := (fun Σ Γ T => { s : Universe.t & sorting Σ Γ T s }). + Definition infer_sort (sorting : global_env_ext -> context -> term -> Universe.t -> Type) := + (fun Σ Γ T relopt => { s : Universe.t & isSortRelOpt s relopt × sorting Σ Γ T s }). Notation typing_sort typing := (fun Σ Γ T s => typing Σ Γ T (tSort s)). Definition lift_typing typing := lift_judgment typing (infer_sort (typing_sort typing)). @@ -440,20 +447,22 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). Definition lift_typing2 P Q := lift_typing (Prop_conj P Q). - Lemma infer_sort_impl {P Q} {Σ Σ' : global_env_ext} {Γ Γ' : context} {t t' : term} : - forall f, forall tu: infer_sort P Σ Γ t, + Lemma infer_sort_impl {P Q} {Σ Σ' : global_env_ext} {Γ Γ' : context} {t t' : term} {relopt} : + forall f, (forall s r, isSortRelOpt s r -> isSortRelOpt (f s) r) -> + forall tu: infer_sort P Σ Γ t relopt, let s := tu.π1 in (P Σ Γ t s -> Q Σ' Γ' t' (f s)) -> - infer_sort Q Σ' Γ' t'. + infer_sort Q Σ' Γ' t' relopt. Proof. - intros ? (? & Hs) s HPQ. eexists. now apply HPQ. + intros ? Hf (? & e & Hs) s HPQ. eexists. now split; [apply Hf | apply HPQ]. Qed. - Lemma infer_typing_sort_impl {P Q} {Σ Σ' : global_env_ext} {Γ Γ' : context} {t t' : term} : - forall f, forall tu: infer_sort (typing_sort P) Σ Γ t, + Lemma infer_typing_sort_impl {P Q} {Σ Σ' : global_env_ext} {Γ Γ' : context} {t t' : term} {relopt} : + forall f (Hf : forall s r, isSortRelOpt s r -> isSortRelOpt (f s) r), + forall tu: infer_sort (typing_sort P) Σ Γ t relopt, let s := tu.π1 in (P Σ Γ t (tSort s) -> Q Σ' Γ' t' (tSort (f s))) -> - infer_sort (typing_sort Q) Σ' Γ' t'. + infer_sort (typing_sort Q) Σ' Γ' t' relopt. Proof. apply (infer_sort_impl (P := typing_sort P) (Q := typing_sort Q)). Qed. @@ -465,18 +474,18 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). Proof. intros HT HPQ. apply lift_judgment_impl with (1 := HT); tas. - intro Hs; apply infer_typing_sort_impl with id Hs; apply HPQ. + intros relopt Hs; apply infer_typing_sort_impl with id Hs => //. apply HPQ. Qed. Section TypeLocalOver. Context (checking : global_env_ext -> context -> term -> term -> Type). - Context (sorting : global_env_ext -> context -> term -> Type). + Context (sorting : global_env_ext -> context -> term -> option relevance -> Type). Context (cproperty : forall (Σ : global_env_ext) (Γ : context), All_local_env (lift_judgment checking sorting Σ) Γ -> forall (t T : term), checking Σ Γ t T -> Type). Context (sproperty : forall (Σ : global_env_ext) (Γ : context), All_local_env (lift_judgment checking sorting Σ) Γ -> - forall (t : term), sorting Σ Γ t -> Type). + forall (t : term) (relopt : option relevance), sorting Σ Γ t relopt -> Type). Inductive All_local_env_over_gen (Σ : global_env_ext) : forall (Γ : context), All_local_env (lift_judgment checking sorting Σ) Γ -> Type := @@ -486,8 +495,8 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). | localenv_over_cons_abs Γ na t (all : All_local_env (lift_judgment checking sorting Σ) Γ) : All_local_env_over_gen Σ Γ all -> - forall (tu : lift_judgment checking sorting Σ Γ t Sort) - (Hs: sproperty Σ Γ all _ tu), + forall (tu : lift_judgment checking sorting Σ Γ t (SortRel na.(binder_relevance))) + (Hs: sproperty Σ Γ all _ _ tu), All_local_env_over_gen Σ (Γ ,, vass na t) (localenv_cons_abs all tu) @@ -495,8 +504,8 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). (all : All_local_env (lift_judgment checking sorting Σ) Γ) (tb : checking Σ Γ b t) : All_local_env_over_gen Σ Γ all -> forall (Hc: cproperty Σ Γ all _ _ tb), - forall (tu : lift_judgment checking sorting Σ Γ t Sort) - (Hs: sproperty Σ Γ all _ tu), + forall (tu : lift_judgment checking sorting Σ Γ t (SortRel na.(binder_relevance))) + (Hs: sproperty Σ Γ all _ _ tu), All_local_env_over_gen Σ (Γ ,, vdef na b t) (localenv_cons_def all tu tb). @@ -504,10 +513,10 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). Derive Signature for All_local_env_over_gen. Definition All_local_env_over typing property := - (All_local_env_over_gen typing (infer_sort (typing_sort typing)) property (fun Σ Γ H t tu => property _ _ H _ _ tu.π2)). + (All_local_env_over_gen typing (infer_sort (typing_sort typing)) property (fun Σ Γ H t relopt tu => property _ _ H _ _ tu.π2.2)). Definition All_local_env_over_sorting checking sorting cproperty (sproperty : forall Σ Γ _ t s, sorting Σ Γ t s -> Type) := - (All_local_env_over_gen checking (infer_sort sorting) cproperty (fun Σ Γ H t tu => sproperty Σ Γ H t tu.π1 tu.π2)). + (All_local_env_over_gen checking (infer_sort sorting) cproperty (fun Σ Γ H t relopt tu => sproperty Σ Γ H t tu.π1 tu.π2.2)). Section TypeCtxInst. Context (typing : forall (Σ : global_env_ext) (Γ : context), term -> term -> Type). @@ -574,20 +583,20 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). Section lift_judgment_size. Context {checking : global_env_ext -> context -> term -> term -> Type}. - Context {sorting : global_env_ext -> context -> term -> Type}. + Context {sorting : global_env_ext -> context -> term -> option relevance -> Type}. Context (csize : forall (Σ : global_env_ext) (Γ : context) (t T : term), checking Σ Γ t T -> size). - Context (ssize : forall (Σ : global_env_ext) (Γ : context) (t : term), sorting Σ Γ t -> size). + Context (ssize : forall (Σ : global_env_ext) (Γ : context) (t : term) (relopt : option relevance), sorting Σ Γ t relopt -> size). Definition lift_judgment_size Σ Γ t T (w : lift_judgment checking sorting Σ Γ t T) : size := match T return lift_judgment checking sorting Σ Γ t T -> size with | Typ T => csize _ _ _ _ - | Sort => ssize _ _ _ + | Sort relopt => ssize _ _ _ _ end w. End lift_judgment_size. Implicit Types (Σ : global_env_ext) (Γ : context) (t : term). - Notation infer_sort_size typing_size := (fun Σ Γ t (tu: infer_sort _ Σ Γ t) => let 'existT s d := tu in typing_size Σ Γ t s d). + Notation infer_sort_size typing_size := (fun Σ Γ t r (tu: infer_sort _ Σ Γ t r) => let '(s; (e, d)) := tu in typing_size Σ Γ t s d). Notation typing_sort_size typing_size := (fun Σ Γ t s (tu: typing_sort _ Σ Γ t s) => typing_size Σ Γ t (tSort s) tu). Section Regular. @@ -674,6 +683,7 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT Context {cf: checker_flags}. Context (Pcmp: global_env_ext -> context -> conv_pb -> term -> term -> Type). Context (P : global_env_ext -> context -> term -> typ_or_sort -> Type). + Definition on_context Σ ctx := All_local_env (P Σ) ctx. @@ -683,23 +693,26 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT Fixpoint type_local_ctx Σ (Γ Δ : context) (u : Universe.t) : Type := match Δ with | [] => wf_universe Σ u - | {| decl_body := None; decl_type := t |} :: Δ => type_local_ctx Σ Γ Δ u × P Σ (Γ ,,, Δ) t (Typ (tSort u)) - | {| decl_body := Some b; decl_type := t |} :: Δ => type_local_ctx Σ Γ Δ u × P Σ (Γ ,,, Δ) t Sort × P Σ (Γ ,,, Δ) b (Typ t) + | {| decl_name := na; decl_body := None; decl_type := t |} :: Δ => + type_local_ctx Σ Γ Δ u × isSortRel u na.(binder_relevance) × P Σ (Γ ,,, Δ) t (Typ (tSort u)) + | {| decl_name := na; decl_body := Some b; decl_type := t |} :: Δ => + type_local_ctx Σ Γ Δ u × P Σ (Γ ,,, Δ) t (SortRel na.(binder_relevance)) × P Σ (Γ ,,, Δ) b (Typ t) end. Fixpoint sorts_local_ctx Σ (Γ Δ : context) (us : list Universe.t) : Type := match Δ, us with | [], [] => unit - | {| decl_body := None; decl_type := t |} :: Δ, u :: us => - sorts_local_ctx Σ Γ Δ us × P Σ (Γ ,,, Δ) t (Typ (tSort u)) - | {| decl_body := Some b; decl_type := t |} :: Δ, us => - sorts_local_ctx Σ Γ Δ us × P Σ (Γ ,,, Δ) t Sort × P Σ (Γ ,,, Δ) b (Typ t) + | {| decl_name := na; decl_body := None; decl_type := t |} :: Δ, u :: us => + sorts_local_ctx Σ Γ Δ us × isSortRel u na.(binder_relevance) × P Σ (Γ ,,, Δ) t (Typ (tSort u)) + | {| decl_name := na; decl_body := Some b; decl_type := t |} :: Δ, us => + sorts_local_ctx Σ Γ Δ us × P Σ (Γ ,,, Δ) t (SortRel na.(binder_relevance)) × P Σ (Γ ,,, Δ) b (Typ t) | _, _ => False end. Implicit Types (mdecl : mutual_inductive_body) (idecl : one_inductive_body) (cdecl : constructor_body). - Definition on_type Σ Γ T := P Σ Γ T Sort. + Definition on_type Σ Γ T := P Σ Γ T (Sort None). + Definition on_type_rel Σ Γ T rel := P Σ Γ T (SortRel rel). Open Scope type_scope. @@ -1065,7 +1078,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) idecl.(ind_relevance); (** The sorts of the arguments contexts of each constructor *) ind_cunivs : list constructor_univs; @@ -1089,6 +1102,8 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT ind_sorts : 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. *) @@ -1189,6 +1204,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}. @@ -1218,25 +1234,17 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT destruct u; auto. intuition eauto. Qed. - Lemma on_global_env_impl {cf : checker_flags} Pcmp P Q : - (forall Σ Γ t T, - on_global_env Pcmp P Σ.1 -> - on_global_env Pcmp Q Σ.1 -> - P Σ Γ t T -> Q Σ Γ t T) -> - forall Σ, on_global_env Pcmp P Σ -> on_global_env Pcmp Q Σ. + Lemma on_global_decl_impl {cf : checker_flags} Pcmp P Q Σ kn d : + (forall Γ t T, + on_global_env Pcmp P Σ.1 -> + P Σ Γ t T -> Q Σ Γ t T) -> + on_global_env Pcmp P Σ.1 -> + on_global_decl Pcmp P Σ kn d -> on_global_decl Pcmp Q Σ kn d. Proof. - intros X Σ [cu IH]. split; auto. - revert cu IH; generalize (universes Σ) as univs, (declarations Σ). clear Σ. - induction g; intros; auto. constructor; auto. - depelim IH. specialize (IHg cu IH). constructor; auto. - pose proof (globenv_decl _ _ _ _ _ _ IH f o). - assert (X' := fun Γ t T => X ({| universes := univs; declarations := _ |}, udecl) Γ t T - (cu, IH) (cu, IHg)); clear X. - rename X' into X. - clear IH IHg. destruct d; simpl. + intros X X0. + destruct d; simpl. - destruct c; simpl. destruct cst_body0; cbn in *; now eapply X. - - red in o. simpl in *. - destruct o0 as [onI onP onNP]. + - intros [onI onP onNP]. constructor; auto. -- eapply Alli_impl; tea. intros. refine {| ind_arity_eq := X1.(ind_arity_eq); @@ -1264,11 +1272,25 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT split. * apply ind_sorts0. * destruct indices_matter; auto. - eapply type_local_ctx_impl; eauto. - eapply ind_sorts0. - --- eapply X1.(onIndices). + eapply type_local_ctx_impl. eapply ind_sorts0. eauto. + --- apply X1.(ind_relevance_compat). + --- apply X1.(onIndices). -- red in onP. red. - eapply All_local_env_impl; tea. + eapply All_local_env_impl; eauto. + Qed. + + Lemma on_global_env_impl {cf : checker_flags} Pcmp P Q : + (forall Σ Γ t T, + on_global_env Pcmp P Σ.1 -> + on_global_env Pcmp Q Σ.1 -> + P Σ Γ t T -> Q Σ Γ t T) -> + forall Σ, on_global_env Pcmp P Σ -> on_global_env Pcmp Q Σ. + Proof. + intros X [univs Σ] [cu X0]; split => /= //. cbn in *. + induction X0; constructor; auto. + eapply on_global_decl_impl; tea. + - intros; apply X => //. + - split => //. Qed. End GlobalMaps. @@ -1306,8 +1328,20 @@ Module DeclarationTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E) Import T E L TU ET CT GM CS Ty. + Notation isTypeRelOpt Σ Γ t relopt := (lift_typing typing Σ Γ t (Sort relopt)). + Definition isType `{checker_flags} (Σ : global_env_ext) (Γ : context) (t : term) := - infer_sort (typing_sort typing) Σ Γ t. + isTypeRelOpt Σ Γ t None. + (* { s : _ & True × Σ ;;; Γ |- t : tSort s }. *) + + Definition isTypeRel `{checker_flags} (Σ : global_env_ext) (Γ : context) (t : term) (rel: relevance) := + isTypeRelOpt Σ Γ t (Some rel). + (* { s : _ & isSortRel s rel × Σ ;;; Γ |- t : Sort s }. *) + + Definition isType_of_isTypeRel `{checker_flags} {Σ Γ t rel} (u: isTypeRel Σ Γ t rel) : isType Σ Γ t := + match u with + | existT s (_, σ) => existT _ s (I, σ) + end. (** This predicate enforces that there exists typing derivations for every typable term in env. *) @@ -1382,6 +1416,7 @@ Module DeclarationTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E) * destruct indices_matter; auto. eapply type_local_ctx_impl; eauto. eapply ind_sorts0. + --- apply X1.(ind_relevance_compat). --- eapply X1.(onIndices). -- red in onP. red. eapply All_local_env_impl; tea. diff --git a/template-coq/theories/EtaExpand.v b/template-coq/theories/EtaExpand.v index 04a77ca1c..c9853d590 100644 --- a/template-coq/theories/EtaExpand.v +++ b/template-coq/theories/EtaExpand.v @@ -958,8 +958,8 @@ Proof. -- eapply Forall_forall. intros x [ | (? & <- & [_ ?] % in_seq) % in_rev % in_map_iff]; subst. all: econstructor; rewrite nth_error_app1; [eapply nth_error_repeat; lia | rewrite repeat_length; lia]. + econstructor. now rewrite nth_error_map, H1. - - cbn. econstructor. eapply (H1 (up Γ')); econstructor; eauto. - - cbn. econstructor. eauto. eapply (H2 (up Γ')); econstructor; eauto. + - cbn. econstructor. eapply (H2 (up Γ')); econstructor; eauto. + - cbn. econstructor. eauto. eapply (H3 (up Γ')); econstructor; eauto. - specialize (H _ H2). assert (Forall(fun t : term => expanded Σ0 (map (fun x : option (nat × term) => @@ -1148,7 +1148,7 @@ Proof. { apply andb_and in H2. destruct H2 as [isl _]. solve_all. } solve_all. { now eapply isLambda_lift, isLambda_eta_expand. } - destruct a as (? & ? & ?). + destruct a as (s & e & Hs1 & Hs2). destruct a0 as (? & ?). rewrite !firstn_length. rewrite !Nat.min_l; try lia. eapply expanded_lift'. @@ -1170,7 +1170,7 @@ Proof. len; lia. rewrite repeat_length. len; lia. + cbn - [rev_map seq]. rewrite rev_map_spec. cbn. rewrite Nat.sub_0_r. cbn. destruct List.rev; cbn; congruence. - - cbn. econstructor; eauto. eapply All_Forall, All_map, All_impl. eapply (All_mix X X0). intros ? ((? & ? & ?) & ? & ?). cbn. + - cbn. econstructor; eauto. eapply All_Forall, All_map, All_impl. eapply (All_mix X X0). intros ? ((? & ? & ? & ?) & ? & ?). cbn. specialize (e0 (repeat None #|mfix| ++ Γ'))%list. rewrite map_app, map_repeat in e0. len. eapply e0. eapply Forall2_app; eauto. unfold types. @@ -1387,7 +1387,7 @@ Proof. destruct c as [na body ty rel]; cbn in *. destruct body. constructor => //; cbn. apply (eta_expand_expanded (Σ := g) [] [] t na wf ond). constructor. - destruct ond as [s Hs]. constructor => //. + destruct ond as (s & e & Hs). constructor => //. - destruct ond as [onI onP onN onV]. constructor. cbn. eapply eta_expand_context => //. @@ -1401,8 +1401,8 @@ Proof. pose proof onc.(on_cargs). eapply eta_expand_context_sorts in X0. now len in X0. len. len. - pose proof onc.(on_ctype). destruct X0. - epose proof (eta_expand_expanded (Σ := g) _ (repeat None #|ind_bodies m|) _ _ wf t). + pose proof onc.(on_ctype). destruct X0 as (s & e & Hs). + epose proof (eta_expand_expanded (Σ := g) _ (repeat None #|ind_bodies m|) _ _ wf Hs). forward H. rewrite -arities_context_length. clear. induction (arities_context _); constructor; auto. now rewrite map_repeat in H. diff --git a/template-coq/theories/Typing.v b/template-coq/theories/Typing.v index 0baa1ae66..2a13837a2 100644 --- a/template-coq/theories/Typing.v +++ b/template-coq/theories/Typing.v @@ -750,25 +750,29 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> Σ ;;; Γ |- tSort s : tSort (Universe.super s) | type_Cast c k t s : + isSortRel s Relevant -> Σ ;;; Γ |- t : tSort s -> Σ ;;; Γ |- c : t -> Σ ;;; Γ |- tCast c k t : t -| type_Prod n t b s1 s2 : +| type_Prod na t b s1 s2 : + isSortRel s1 na.(binder_relevance) -> Σ ;;; Γ |- t : tSort s1 -> - Σ ;;; Γ ,, vass n t |- b : tSort s2 -> - Σ ;;; Γ |- tProd n t b : tSort (Universe.sort_of_product s1 s2) + Σ ;;; Γ ,, vass na t |- b : tSort s2 -> + Σ ;;; Γ |- tProd na t b : tSort (Universe.sort_of_product s1 s2) -| type_Lambda n t b s1 bty : - Σ ;;; Γ |- t : tSort s1 -> - Σ ;;; Γ ,, vass n t |- b : bty -> - Σ ;;; Γ |- tLambda n t b : tProd n t bty +| type_Lambda na t b s bty : + isSortRel s na.(binder_relevance) -> + Σ ;;; Γ |- t : tSort s -> + Σ ;;; Γ ,, vass na t |- b : bty -> + Σ ;;; Γ |- tLambda na t b : tProd na t bty -| type_LetIn n b b_ty b' s1 b'_ty : - Σ ;;; Γ |- b_ty : tSort s1 -> +| type_LetIn na b b_ty b' s b'_ty : + isSortRel s na.(binder_relevance) -> + Σ ;;; Γ |- b_ty : tSort s -> Σ ;;; Γ |- b : b_ty -> - Σ ;;; Γ ,, vdef n b b_ty |- b' : b'_ty -> - Σ ;;; Γ |- tLetIn n b b_ty b' : tLetIn n b b_ty b'_ty + Σ ;;; Γ ,, vdef na b b_ty |- b' : b'_ty -> + Σ ;;; Γ |- tLetIn na b b_ty b' : tLetIn na b b_ty b'_ty | type_App t l t_ty t' : Σ ;;; Γ |- t : t_ty -> @@ -826,17 +830,17 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> fix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> wf_local Σ Γ -> - All (fun d => {s & Σ ;;; Γ |- d.(dtype) : tSort s}) mfix -> - All (fun d => (Σ ;;; Γ ,,, fix_context mfix |- d.(dbody) : lift0 #|fix_context mfix| d.(dtype))) mfix -> + All (on_def_type (lift_typing typing Σ) Γ) mfix -> + All (on_def_body (lift_typing typing Σ) (fix_context mfix) Γ) mfix -> wf_fixpoint Σ mfix -> - Σ ;;; Γ |- tFix mfix n : decl.(dtype) + Σ ;;; Γ |- tFix mfix n : decl.(dtype) | type_CoFix mfix n decl : cofix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> wf_local Σ Γ -> - All (fun d => {s & Σ ;;; Γ |- d.(dtype) : tSort s}) mfix -> - All (fun d => Σ ;;; Γ ,,, fix_context mfix |- d.(dbody) : lift0 #|fix_context mfix| d.(dtype)) mfix -> + All (on_def_type (lift_typing typing Σ) Γ) mfix -> + All (on_def_body (lift_typing typing Σ) (fix_context mfix) Γ) mfix -> wf_cofixpoint Σ mfix -> Σ ;;; Γ |- tCoFix mfix n : decl.(dtype) @@ -920,11 +924,10 @@ Section CtxInstSize. End CtxInstSize. -Definition typing_size `{checker_flags} {Σ Γ t T} (d : Σ ;;; Γ |- t : T) : size. +Fixpoint typing_size `{cf : checker_flags} {Σ Γ t T} (d : Σ ;;; Γ |- t : T) {struct d} : size. Proof. - revert Σ Γ t T d. - fix typing_size 5. - destruct 1 ; + specialize (typing_size cf). + destruct d ; repeat match goal with | H : typing _ _ _ _ |- _ => apply typing_size in H end; @@ -950,8 +953,8 @@ Proof. - exact (S (Nat.max d1 (Nat.max d2 (Nat.max (ctx_inst_size _ typing_size c1) (all2i_size _ (fun _ x y p => Nat.max (typing_size _ _ _ _ p.1.2) (typing_size _ _ _ _ p.2)) a0))))). - - exact (S (Nat.max (Nat.max (All_local_env_size typing_size _ _ a) (all_size _ (fun x p => infer_sort_size (typing_sort_size typing_size) Σ _ _ p) a0)) (all_size _ (fun x p => typing_size Σ _ _ _ p) a1))). - - exact (S (Nat.max (Nat.max (All_local_env_size typing_size _ _ a) (all_size _ (fun x p => infer_sort_size (typing_sort_size typing_size) Σ _ _ p) a0)) (all_size _ (fun x p => typing_size Σ _ _ _ p) a1))). + - exact (S (Nat.max (Nat.max (All_local_env_size typing_size _ _ a) (all_size _ (fun x p => lift_typing_size typing_size Σ _ _ _ p) a0)) (all_size _ (fun x p => typing_size Σ _ _ _ p) a1))). + - exact (S (Nat.max (Nat.max (All_local_env_size typing_size _ _ a) (all_size _ (fun x p => lift_typing_size typing_size Σ _ _ _ p) a0)) (all_size _ (fun x p => typing_size Σ _ _ _ p) a1))). Defined. Lemma typing_size_pos `{checker_flags} {Σ Γ t T} (d : Σ ;;; Γ |- t : T) : typing_size d > 0. @@ -1027,13 +1030,11 @@ Proof. Defined. #[global] Hint Immediate wf_local_app_l : wf. -Lemma typing_wf_local_size `{checker_flags} {Σ} {Γ t T} +Lemma typing_wf_local_size {cf : checker_flags} {Σ} {Γ t T} (d :Σ ;;; Γ |- t : T) : All_local_env_size (@typing_size _) _ _ (typing_wf_local d) < typing_size d. Proof. - induction d; simpl; - change (fun (x : global_env_ext) (x0 : context) (x1 x2 : term) - (x3 : x;;; x0 |- x1 : x2) => typing_size x3) with (@typing_size H); try lia. + induction d; simpl; try lia. Qed. Lemma wf_local_inv `{checker_flags} {Σ Γ'} (w : wf_local Σ Γ') : @@ -1061,12 +1062,12 @@ Proof. destruct w. - simpl. congruence. - intros [=]. subst d Γ0. - exists w. simpl. destruct l. exists x. exists t0. pose (typing_size_pos t0). + exists w. simpl. destruct l as (x & ? & t0). exists x. exists t0. pose (typing_size_pos t0). cbn. split. + lia. + auto with arith. - intros [=]. subst d Γ0. - exists w. simpl. simpl in l. destruct l as [u h]. + exists w. simpl. simpl in l. destruct l as (u & ? & h). simpl in l0. exists u, l0, h. cbn. pose (typing_size_pos h). @@ -1113,31 +1114,36 @@ Lemma typing_ind_env `{cf : checker_flags} : (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) (c : term) (k : cast_kind) (t : term) (s : Universe.t), - Σ ;;; Γ |- t : tSort s -> P Σ Γ t (tSort s) -> Σ ;;; Γ |- c : t -> P Σ Γ c t -> P Σ Γ (tCast c k t) t) -> + isSortRel s Relevant -> + Σ ;;; Γ |- t : tSort s -> P Σ Γ t (tSort s) -> + Σ ;;; Γ |- c : t -> P Σ Γ c t -> P Σ Γ (tCast c k t) t) -> - (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) (n : aname) (t b : term) (s1 s2 : Universe.t), + (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (t b : term) (s1 s2 : Universe.t), + isSortRel s1 na.(binder_relevance) -> PΓ Σ Γ wfΓ -> Σ ;;; Γ |- t : tSort s1 -> P Σ Γ t (tSort s1) -> - Σ ;;; Γ,, vass n t |- b : tSort s2 -> - P Σ (Γ,, vass n t) b (tSort s2) -> P Σ Γ (tProd n t b) (tSort (Universe.sort_of_product s1 s2))) -> + Σ ;;; Γ,, vass na t |- b : tSort s2 -> + P Σ (Γ,, vass na t) b (tSort s2) -> P Σ Γ (tProd na t b) (tSort (Universe.sort_of_product s1 s2))) -> - (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) (n : aname) (t b : term) - (s1 : Universe.t) (bty : term), + (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (t b : term) + (s : Universe.t) (bty : term), + isSortRel s na.(binder_relevance) -> PΓ Σ Γ wfΓ -> - Σ ;;; Γ |- t : tSort s1 -> - P Σ Γ t (tSort s1) -> - Σ ;;; Γ,, vass n t |- b : bty -> P Σ (Γ,, vass n t) b bty -> P Σ Γ (tLambda n t b) (tProd n t bty)) -> + Σ ;;; Γ |- t : tSort s -> + P Σ Γ t (tSort s) -> + Σ ;;; Γ,, vass na t |- b : bty -> P Σ (Γ,, vass na t) b bty -> P Σ Γ (tLambda na t b) (tProd na t bty)) -> - (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) (n : aname) (b b_ty b' : term) - (s1 : Universe.t) (b'_ty : term), + (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (b b_ty b' : term) + (s : Universe.t) (b'_ty : term), + isSortRel s na.(binder_relevance) -> PΓ Σ Γ wfΓ -> - Σ ;;; Γ |- b_ty : tSort s1 -> - P Σ Γ b_ty (tSort s1) -> + Σ ;;; Γ |- b_ty : tSort s -> + P Σ Γ b_ty (tSort s) -> Σ ;;; Γ |- b : b_ty -> P Σ Γ b b_ty -> - Σ ;;; Γ,, vdef n b b_ty |- b' : b'_ty -> - P Σ (Γ,, vdef n b b_ty) b' b'_ty -> P Σ Γ (tLetIn n b b_ty b') (tLetIn n b b_ty b'_ty)) -> + Σ ;;; Γ,, vdef na b b_ty |- b' : b'_ty -> + P Σ (Γ,, vdef na b b_ty) b' b'_ty -> P Σ Γ (tLetIn na b b_ty b') (tLetIn na b b_ty b'_ty)) -> (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) (t : term) (l : list term) (t_ty t' : term), Σ ;;; Γ |- t : t_ty -> P Σ Γ t t_ty -> @@ -1314,6 +1320,7 @@ Proof. eapply type_local_ctx_impl. eapply ind_sorts0. intros ??? HT. apply lift_typing_impl with (1 := HT); intros ? Hs. apply (IH (_; _; _; Hs)). + -- apply Xg.(ind_relevance_compat). -- apply (onIndices Xg). * red in onP |- *. apply All_local_env_impl with (1 := onP); intros ??? HT. @@ -1343,7 +1350,7 @@ Proof. clearbody foo. assert (All_local_env_size (@typing_size cf) Σ Γ' foo < typing_size H) by lia. clear H1 H0 Hty. revert foo H2. - induction foo; simpl in *; try destruct t3 as [u Hu]; cbn in *; constructor. + induction foo; simpl in *; try destruct t3 as (u & e & Hu); cbn in *; constructor. - simpl in *. apply IHfoo. lia. - red. apply (X14 _ _ _ Hu). lia. - simpl in *. apply IHfoo. lia. @@ -1371,16 +1378,10 @@ Proof. eapply X4 with t_ty t0; eauto. clear X4. unshelve eapply X14; simpl; auto with arith. simpl in X14. - assert(X: forall Γ0 : context, - wf_local Σ Γ0 -> - forall (t1 T : term) (Hty : Σ;;; Γ0 |- t1 : T), - typing_size Hty < - S - ((typing_spine_size - (fun x (x0 : context) (x1 x2 : term) (x3 : x;;; x0 |- x1 : x2) => - typing_size x3) Σ Γ t_ty l t' t0)) -> - on_global_env cumul_gen (lift_typing P) Σ.1 * P Σ Γ0 t1 T). { - intros. unshelve eapply X14; eauto. lia. } + assert (X : forall Γ', wf_local Σ Γ' -> forall t T (Hty : Σ;;; Γ' |- t : T), + typing_size Hty < S (typing_spine_size (@typing_size _) Σ Γ t_ty l t' t0) -> + on_global_env cumul_gen (lift_typing P) Σ.1 * P Σ Γ' t T). + { intros. unshelve eapply X14; eauto. lia. } clear X14. simpl in pΓ. clear n e H pΓ. induction t0; constructor. unshelve eapply X; clear X; simpl; auto with arith. @@ -1404,12 +1405,8 @@ Proof. ++ eapply (X14 _ _ _ H0); eauto. simpl; auto with arith. lia. ++ clear -c1 X14. assert (forall (Γ' : context) (t T : term) (Hty : Σ;;; Γ' |- t : T), - typing_size Hty <= ctx_inst_size _ (@typing_size _) c1 -> - P Σ Γ' t T). - { intros. eapply (X14 _ _ _ Hty). simpl. - change (fun (x : global_env_ext) (x0 : context) (x1 x2 : term) - (x3 : x;;; x0 |- x1 : x2) => typing_size x3) with (@typing_size cf). - lia. } + typing_size Hty <= ctx_inst_size _ (@typing_size _) c1 -> P Σ Γ' t T). + { intros. eapply (X14 _ _ _ Hty). simpl. lia. } clear -X c1. revert c1 X. generalize (List.rev (subst_instance (puinst p) (ind_params mdecl ,,, ind_indices idecl))). @@ -1445,32 +1442,25 @@ Proof. unshelve eapply X14; eauto. -- clear X X0 Xcast X1 X2 X3 X4 X5 X6 X7 X8 X9 X11 X12. - eapply X10; eauto; clear X10. simpl in *. - * assert(forall (t T : term) (Hty : Σ;;; Γ |- t : T), - typing_size Hty < - S (all_size (fun x : def term => - ∑ s : Universe.t, Σ;;; Γ |- dtype x : tSort s) - (fun (x : def term) '(existT s d) => typing_size d) a0) -> + eapply X10; eauto; clear X10. + * assert (forall t T (Hty : Σ;;; Γ |- t : T), + typing_size Hty < S (all_size _ (fun x p => lift_typing_size (@typing_size _) Σ _ _ _ p) a0) -> on_global_env cumul_gen (lift_typing P) Σ.1 * P Σ Γ t T). - intros; eauto. eapply (X14 _ _ _ Hty); eauto. unfold infer_sort. lia. + { intros; eauto. eapply (X14 _ _ _ Hty); eauto. simpl in *. lia. } clear X13 X14 a pΓ. clear -a0 X. induction a0; constructor. - destruct p as [s Hs]. exists s; split; auto. + destruct p as (s & e & Hs). exists s; repeat split; auto. apply (X (dtype x) (tSort s) Hs). simpl. lia. apply IHa0. intros. eapply (X _ _ Hty); eauto. - simpl. lia. + cbn [all_size]. lia. * simpl in X14. - assert(forall Γ0 : context, - wf_local Σ Γ0 -> - forall (t T : term) (Hty : Σ;;; Γ0 |- t : T), - typing_size Hty < - S - (all_size (fun x : def term => (Σ;;; Γ ,,, fix_context mfix |- dbody x : lift0 #|fix_context mfix| (dtype x))%type - )%type - (fun (x : def term) p => typing_size p) a1) -> - on_global_env cumul_gen (lift_typing P) Σ.1 * P Σ Γ0 t T). - {intros. eapply (X14 _ _ _ Hty); eauto. lia. } + assert (forall Γ' : context, + wf_local Σ Γ' -> + forall (t T : term) (Hty : Σ;;; Γ' |- t : T), + typing_size Hty < S (all_size _ (fun x p => typing_size p) a1) -> + on_global_env cumul_gen (lift_typing P) Σ.1 * P Σ Γ' t T). + { intros. eapply (X14 _ _ _ Hty); eauto. lia. } clear X14 X13. clear e decl i a0 i0 pΓ. remember (fix_context mfix) as mfixcontext. clear Heqmfixcontext. @@ -1482,30 +1472,24 @@ Proof. eapply (X _ X0 _ _ Hty). simpl; lia. -- clear X X0 Xcast X1 X2 X3 X4 X5 X6 X7 X8 X9 X10 X12. - eapply X11; eauto; clear X11. simpl in *. - * assert(forall (t T : term) (Hty : Σ;;; Γ |- t : T), - typing_size Hty < - S (all_size (fun x : def term => - ∑ s : Universe.t, Σ;;; Γ |- dtype x : tSort s) - (fun (x : def term) '(existT s d) => typing_size d) a0) -> - on_global_env cumul_gen (lift_typing P) Σ.1 * P Σ Γ t T). - intros; eauto. eapply (X14 _ _ _ Hty); eauto. unfold infer_sort. lia. + eapply X11; eauto; clear X11. + * assert (forall t T (Hty : Σ;;; Γ |- t : T), + typing_size Hty < S (all_size _ (fun x p => lift_typing_size (@typing_size _) Σ _ _ _ p) a0) -> + on_global_env cumul_gen (lift_typing P) Σ.1 * P Σ Γ t T). + { intros; eauto. eapply (X14 _ _ _ Hty); eauto. simpl in *. lia. } clear X13 X14 a pΓ. clear -a0 X. induction a0; constructor. - destruct p as [s Hs]. exists s; split; auto. + destruct p as (s & e & Hs). exists s; repeat split; auto. apply (X (dtype x) (tSort s) Hs). simpl. lia. apply IHa0. intros. eapply (X _ _ Hty); eauto. - simpl. lia. + cbn [all_size]. lia. * simpl in X14. - assert(forall Γ0 : context, - wf_local Σ Γ0 -> - forall (t T : term) (Hty : Σ;;; Γ0 |- t : T), - typing_size Hty < - S - (all_size (fun x : def term => (Σ;;; Γ ,,, fix_context mfix |- dbody x : lift0 #|fix_context mfix| (dtype x))%type) - (fun (x : def term) p => typing_size p) a1) -> - on_global_env cumul_gen (lift_typing P) Σ.1 * P Σ Γ0 t T). + assert (forall Γ' : context, + wf_local Σ Γ' -> + forall (t T : term) (Hty : Σ;;; Γ' |- t : T), + typing_size Hty < S (all_size _ (fun x p => typing_size p) a1) -> + on_global_env cumul_gen (lift_typing P) Σ.1 * P Σ Γ' t T). { intros. eapply (X14 _ _ _ Hty); eauto. lia. } clear X14 X13. clear e decl a0 i i0 pΓ. @@ -1642,7 +1626,7 @@ Definition on_wf_local_decl `{checker_flags} {Σ Γ} | {| decl_name := na; decl_body := Some b; decl_type := ty |} => fun H => P Σ Γ wfΓ b ty H | {| decl_name := na; decl_body := None; decl_type := ty |} - => fun H => P Σ Γ wfΓ _ _ (projT2 H) + => fun H => P Σ Γ wfΓ _ _ (H.π2.2) end H. diff --git a/template-coq/theories/TypingWf.v b/template-coq/theories/TypingWf.v index 039211c4e..5900dc6d8 100644 --- a/template-coq/theories/TypingWf.v +++ b/template-coq/theories/TypingWf.v @@ -41,65 +41,17 @@ Lemma on_global_decl_impl `{checker_flags} Σ P Q kn d : (forall Σ Γ t T, on_global_env cumul_gen P Σ.1 -> P Σ Γ t T -> Q Σ Γ t T) -> on_global_env cumul_gen P Σ.1 -> on_global_decl cumul_gen P Σ kn d -> on_global_decl cumul_gen Q Σ kn d. -Proof. - unfold on_global_env. - intros X X0 o. - destruct d; simpl. - - destruct c; simpl. destruct cst_body0; simpl in *. - red in o |- *. simpl in *. now eapply X. - red in o |- *. simpl in *. now eapply X. - - simpl in *. - destruct o as [onI onP onNP]. - constructor; auto. - -- eapply Alli_impl. exact onI. eauto. intros. - - refine {| ind_arity_eq := X1.(ind_arity_eq); - ind_cunivs := X1.(ind_cunivs) |}. - --- apply onArity in X1. unfold on_type in *; simpl in *. - now eapply X. - --- pose proof X1.(onConstructors) as X11. red in X11. - eapply All2_impl; eauto. - simpl. intros. destruct X2 as [? ? ? ?]; unshelve econstructor; eauto. - * apply X; eauto. - * clear -X0 X on_cargs. revert on_cargs. - generalize (cstr_args x0), y. - induction c; destruct y0; simpl; auto; - destruct a as [na [b|] ty]; simpl in *; auto; - split; intuition eauto. - * clear -X0 X on_cindices. - revert on_cindices. - generalize (List.rev (lift_context #|cstr_args x0| 0 (ind_indices x))). - generalize (cstr_indices x0). - induction 1; simpl; constructor; auto. - --- simpl; intros. apply (onProjections X1 H0). - --- destruct X1. simpl. unfold check_ind_sorts in *. - destruct Universe.is_prop; auto. - destruct Universe.is_sprop; auto. - split. apply ind_sorts. destruct indices_matter; auto. - eapply type_local_ctx_impl. eapply ind_sorts. auto. - --- apply (onIndices X1). - -- red in onP. red. - eapply All_local_env_impl. eauto. - intros. now apply X. -Qed. +Proof. intros. eapply on_global_decl_impl; tea. apply (X Σ). Qed. Lemma on_global_env_impl `{checker_flags} Σ P Q : (forall Σ Γ t T, on_global_env cumul_gen P Σ.1 -> P Σ Γ t T -> Q Σ Γ t T) -> on_global_env cumul_gen P Σ -> on_global_env cumul_gen Q Σ. -Proof. - destruct Σ as [univs Σ]; cbn. - intros X [cu X0]; split => /= //. cbn in *. - induction X0; constructor; auto. - clear IHX0. - eapply on_global_decl_impl; tea. split => //. -Qed. +Proof. intros. eapply on_global_env_impl; tea. auto. Qed. Lemma All_local_env_wf_decl_inv Σ (a : context_decl) (Γ : list context_decl) (X : All_local_env (wf_decl_pred Σ) (a :: Γ)) : on_local_decl (wf_decl_pred Σ) Γ a * All_local_env (wf_decl_pred Σ) Γ. -Proof. - inv X; intuition; red; simpl; eauto. -Qed. +Proof. apply All_local_env_inv, X. Qed. Lemma unfold_fix_wf: forall Σ (mfix : mfixpoint term) (idx : nat) (narg : nat) (fn : term), @@ -138,31 +90,11 @@ Proof. induction 1 using red1_ind_all; simpl; try discriminate; auto. Qed. -Lemma OnOne2_All_All {A} {P Q} {l l' : list A} : +(* Lemma OnOne2_All_All {A} {P Q} {l l' : list A} : OnOne2 P l l' -> (forall x y, P x y -> Q x -> Q y) -> All Q l -> All Q l'. -Proof. intros Hl H. induction Hl; intros H'; inv H'; constructor; eauto. Qed. - -Lemma All_mapi {A B} (P : B -> Type) (l : list A) (f : nat -> A -> B) : - Alli (fun i x => P (f i x)) 0 l -> All P (mapi f l). -Proof. - unfold mapi. generalize 0. - induction 1; constructor; auto. -Qed. - -Lemma Alli_id {A} (P : nat -> A -> Type) n (l : list A) : - (forall n x, P n x) -> Alli P n l. -Proof. - intros H. induction l in n |- *; constructor; auto. -Qed. - - -Lemma All_Alli {A} {P : A -> Type} {Q : nat -> A -> Type} {l n} : - All P l -> - (forall n x, P x -> Q n x) -> - Alli Q n l. -Proof. intro H. revert n. induction H; constructor; eauto. Qed. +Proof. intros Hl H. induction Hl; intros H'; inv H'; constructor; eauto. Admitted. *) Ltac wf := intuition try (eauto with wf || congruence || solve [constructor]). @@ -188,6 +120,7 @@ Proof. depelim f. cbn in *. subst. contradiction. Qed. +(* TODO : Put in a shared theory *) Lemma lookup_env_extends {cf : checker_flags} (Σ : global_env) k d (Σ' : global_env) P : on_global_env cumul_gen P Σ' -> lookup_env Σ k = Some d -> @@ -228,7 +161,7 @@ Qed. Lemma wf_decl_extends {cf} {Σ : global_env} T {Σ' : global_env} P : on_global_env cumul_gen P Σ' -> wf_decl Σ T -> extends_decls Σ Σ' -> wf_decl Σ' T. Proof. - intros wf [] ext. red. destruct decl_body; split; eauto using wf_extends. + intros wf [] ext. red. destruct decl_body; split; cbn; eauto using wf_extends. Qed. Arguments lookup_on_global_env {H} {Pcmp P Σ c decl}. @@ -275,6 +208,20 @@ Proof. solve_all. now eapply wf_decl_extends. Qed. +Lemma sorts_local_ctx_All_wf_decl {cf:checker_flags} {Σ} {mdecl} {u: list Universe.t} {args} : + sorts_local_ctx (fun Σ : global_env_ext => wf_decl_pred Σ) Σ + (arities_context (ind_bodies mdecl),,, ind_params mdecl) + args u -> All (wf_decl Σ) args. +Proof. + induction args as [|[na [b|] ty] args] in u |- * ; + constructor. + - constructor; now destruct X as (?&(?&?)&(?&?)). + - eapply IHargs; now apply X. + - destruct u => //; constructor; cbnr; now destruct X as (?&?&?&?). + - destruct u => //; eapply IHargs; now apply X. +Qed. + + Lemma declared_inductive_wf_ctors {cf:checker_flags} {Σ} {ind} {mdecl idecl} : on_global_env cumul_gen wf_decl_pred Σ -> declared_inductive Σ ind mdecl idecl -> @@ -286,15 +233,10 @@ Proof. apply onInductives in prf. eapply nth_error_alli in Hidecl; eauto. pose proof (onConstructors Hidecl). red in X0. - solve_all. destruct X0. - clear -X ext on_cargs. - induction (cstr_args x) as [|[na [b|] ty] args] in on_cargs, y |- * ; - try destruct on_cargs; - constructor; unfold wf_decl in *; cbn in *; intuition eauto using wf_extends; simpl in *. - destruct b0. intuition eauto using wf_extends. - destruct a. intuition eauto using wf_extends. - destruct y => //. destruct on_cargs. destruct w; eauto using wf_extends. - destruct y => //. eapply IHargs; intuition eauto. + solve_all. + apply on_cargs in X0. + eapply sorts_local_ctx_All_wf_decl in X0. + eapply All_impl; eauto using wf_decl_extends. Qed. Lemma All_local_env_wf_decls Σ ctx : @@ -302,7 +244,7 @@ Lemma All_local_env_wf_decls Σ ctx : All (wf_decl Σ) ctx. Proof. induction 1; constructor; auto. - destruct t0 as [s Hs]. split; simpl; intuition auto. + destruct t0 as [Hw Hs]. now split. Qed. Lemma on_global_inductive_wf_params {cf:checker_flags} {Σ : global_env_ext} {kn mdecl} : @@ -310,9 +252,8 @@ Lemma on_global_inductive_wf_params {cf:checker_flags} {Σ : global_env_ext} {kn All (wf_decl Σ) (ind_params mdecl). Proof. intros prf. - apply onParams in prf. red in prf. - apply All_local_env_wf_decls in prf. - solve_all. + apply onParams in prf. + now apply All_local_env_wf_decls in prf. Qed. Lemma destArity_spec ctx T : @@ -500,7 +441,7 @@ Section WfAst. on_global_env cumul_gen wf_decl_pred Σ. Proof using Type. apply on_global_env_impl => Σ' Γ t []; simpl; unfold wf_decl_pred; - intros; auto. destruct X0 as [s []]; intuition auto. + intros; auto. destruct X0 as [s ?]; now split. Qed. (* Hint Resolve on_global_wf_Forall_decls : wf. *) @@ -519,8 +460,8 @@ Section WfAst. Proof using Type. intros oib. apply onParams in oib. red in oib. - induction (ind_params mdecl) as [|[? [] ?] ?]; simpl in oib; inv oib; constructor; - try red in X0; try red in X1; try red; simpl; intuition auto. + induction (ind_params mdecl) as [|[? [] ?] ?]; + inv oib; constructor; auto. now destruct X0. Qed. Lemma declared_inductive_wf_params {ind mdecl idecl} : @@ -641,14 +582,8 @@ Section WfLookup. { unfold on_constructors in onConstructors. clear -onConstructors. induction onConstructors; constructor; auto. - destruct r. - clear -on_cargs. - revert on_cargs. revert y. generalize (cstr_args x). - induction c as [|[? [] ?] ?]; simpl; - destruct y; intuition auto; - constructor; - try red; simpl; try red in a0, b0; intuition eauto. - now red in b. } + apply on_cargs in r. + eapply sorts_local_ctx_All_wf_decl; tea. } split => //. - now destruct onArity. - rewrite ind_arity_eq in onArity . @@ -783,8 +718,7 @@ Section WfRed. - constructor; auto. induction X; inv X0; constructor; intuition auto; congruence. - constructor; auto. solve_all. - pose proof X0 as H'. revert X0. - apply (OnOne2_All_All X). clear X. + eapply OnOne2_All_All; eauto. clear X. intros [na bo ty ra] [nb bb tb rb] [[r ih] e] [? ?]. simpl in *. inversion e. subst. clear e. @@ -793,15 +727,14 @@ Section WfRed. solve_all. apply All_app_inv. 2: assumption. unfold fix_context. apply All_rev. eapply All_mapi. - eapply All_Alli. 1: exact H'. + eapply All_Alli. 1: exact X0. cbn. unfold wf_decl. simpl. intros ? [? ? ? ?] ?. simpl in *. intuition eauto with wf. - constructor; auto. induction X; inv X0; constructor; intuition auto; congruence. - constructor; auto. solve_all. - pose proof X0 as H'. revert X0. - apply (OnOne2_All_All X). clear X. + eapply OnOne2_All_All; eauto. clear X. intros [na bo ty ra] [nb bb tb rb] [[r ih] e] [? ?]. simpl in *. inversion e. subst. clear e. @@ -809,7 +742,7 @@ Section WfRed. eapply ih. 2: assumption. solve_all. apply All_app_inv. 2: assumption. unfold fix_context. apply All_rev. eapply All_mapi. - eapply All_Alli. 1: exact H'. + eapply All_Alli. 1: exact X0. cbn. unfold wf_decl. simpl. intros ? [? ? ? ?] ?. simpl in *. intuition eauto with wf. @@ -906,14 +839,6 @@ End WfRed. #[global] Hint Resolve wf_extends : wf. -Lemma All2i_All2 {A B} {P : nat -> A -> B -> Type} {Q : A -> B -> Type} n l l' : - All2i P n l l' -> - (forall i x y, P i x y -> Q x y) -> - All2 Q l l'. -Proof. - induction 1; constructor; eauto. -Qed. - Lemma cstr_branch_context_length ind mdecl cdecl : #|cstr_branch_context ind mdecl cdecl| = #|cdecl.(cstr_args)|. Proof. rewrite /cstr_branch_context. now len. Qed. @@ -941,7 +866,9 @@ Section TypingWf. try solve [split; try constructor; intuition auto with wf]. - eapply All_local_env_wf_decls. - induction X; constructor; auto; red; intuition auto. + induction X; constructor; auto; red; + unfold wf_decl_pred, typ_or_sort_default, SortRel; + intuition auto. - split; wf. apply wf_lift. apply (nth_error_all H X). - split. constructor; auto. wf. @@ -1014,7 +941,9 @@ Section TypingWf. pose proof (env_prop_sigma typing_wf_gen _ wfΣ). red in X. do 2 red in wfΣ. eapply on_global_env_impl; eauto; simpl; intros. - destruct T. red. apply X1. red. destruct X1 as [x [a wfs]]. split; auto. + unfold wf_decl_pred. + destruct T. apply X1. + now destruct X1 as (x & a & wf1 & wf2). Qed. Lemma typing_wf Σ (wfΣ : wf Σ.1) Γ t T : diff --git a/template-coq/theories/Universes.v b/template-coq/theories/Universes.v index 37c510990..b17f136d8 100644 --- a/template-coq/theories/Universes.v +++ b/template-coq/theories/Universes.v @@ -1153,6 +1153,16 @@ End Universe. Definition is_propositional u := Universe.is_prop u || Universe.is_sprop u. +Definition relevance_of_sort (s: Universe.t) := + match s with + | Universe.lSProp => Irrelevant + | _ => Relevant + end. + +Notation isSortRel s rel := (relevance_of_sort s = rel). +Definition isSortRelOpt s relopt := + match relopt with None => True | Some rel => isSortRel s rel end. + Notation "⟦ u ⟧_ v" := (Universe.to_cuniv v u) (at level 0, format "⟦ u ⟧_ v", v name) : univ_scope. @@ -2073,6 +2083,29 @@ Tactic Notation "unfold_univ_rel" "eqn" ":"ident(H) := Ltac cong := intuition congruence. +Lemma leq_relevance_eq {cf φ} {s s'} : + leq_universe φ 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_universe φ s s' -> isSortRelOpt s rel -> isSortRelOpt s' rel. +Proof. + now destruct s, s'. +Qed. + +Lemma leq_relevance {cf φ} {s s' rel} : + leq_universe φ s s' -> isSortRel s rel -> isSortRel s' rel. +Proof. + now destruct s, s'. +Qed. + +Lemma geq_relevance {cf φ} {s s' rel} : + leq_universe φ s' s -> isSortRel s rel -> isSortRel s' rel. +Proof. + now destruct s, s'. +Qed. Lemma leq_universe_product_mon {cf} ϕ s1 s1' s2 s2' : leq_universe ϕ s1 s1' -> @@ -2367,6 +2400,25 @@ Notation "x @[ u ]" := (subst_instance u x) (at level 3, #[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_univ u s) = relevance_of_sort s. +Proof. + now destruct s. +Qed. + +Theorem relevance_subst_opt {cf} u s rel : + isSortRelOpt s rel -> isSortRelOpt (subst_instance_univ u s) rel. +Proof. + now destruct s. +Qed. + +Theorem relevance_subst {cf} u s rel : + isSortRel s rel -> isSortRel (subst_instance_univ 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/template-coq/theories/WfAst.v b/template-coq/theories/WfAst.v index c7f0994c8..076119bb4 100644 --- a/template-coq/theories/WfAst.v +++ b/template-coq/theories/WfAst.v @@ -60,6 +60,8 @@ Inductive wf {Σ} : term -> Type := Arguments wf : clear implicits. Derive Signature for wf. +Definition wf_def Σ def := wf Σ def.(dtype) × wf Σ def.(dbody). + (** * Inversion lemmas for the well-formedness judgement *) Definition wf_Inv Σ (t : term) : Type := @@ -88,8 +90,8 @@ Definition wf_Inv Σ (t : term) : Type := wf_nactx br.(bcontext) (cstr_branch_context ci.(ci_ind) mdecl cdecl) × wf Σ (bbody br)) idecl.(ind_ctors) brs] | tProj p t => wf Σ t - | tFix mfix k => All (fun def => wf Σ def.(dtype) * wf Σ def.(dbody)) mfix - | tCoFix mfix k => All (fun def => wf Σ def.(dtype) * wf Σ def.(dbody)) mfix + | tFix mfix k => All (wf_def Σ) mfix + | tCoFix mfix k => All (wf_def Σ) mfix end. Lemma wf_inv {Σ t} (w : wf Σ t) : wf_Inv Σ t. @@ -182,16 +184,10 @@ Proof. Qed. Definition wf_decl Σ d := - match decl_body d with - | Some b => wf Σ b - | None => True - end * wf Σ (decl_type d). - -Definition wf_decl_pred Σ : context -> term -> typ_or_sort -> Type := - (fun _ t T => wf Σ t * match T with - | Typ T => wf Σ T - | Sort => True - end). + option_default (wf Σ) (decl_body d) True * wf Σ (decl_type d). + +Definition wf_decl_pred Σ (Γ : context) t T : Type := + wf Σ t × typ_or_sort_default (wf Σ) T True. Lemma wf_mkApp Σ u a : wf Σ u -> wf Σ a -> wf Σ (mkApp u a). Proof. diff --git a/template-coq/theories/utils/All_Forall.v b/template-coq/theories/utils/All_Forall.v index 4adaadac6..26b097dd3 100644 --- a/template-coq/theories/utils/All_Forall.v +++ b/template-coq/theories/utils/All_Forall.v @@ -727,6 +727,12 @@ Proof. induction 1; simpl; constructor; auto. Qed. +Lemma Alli_refl {A} (P : nat -> A -> Type) n (l : list A) : + (forall n x, P n x) -> Alli P n l. +Proof. + intros H. induction l in n |- *; constructor; auto. +Qed. + Lemma Alli_rev {A} {P : nat -> A -> Type} k l : Alli P k l -> Alli (fun k' => P (Nat.pred #|l| - k' + k)) 0 (List.rev l). @@ -3002,6 +3008,14 @@ Proof. congruence. Qed. +Lemma All2i_All2 {A B} {P : nat -> A -> B -> Type} {Q : A -> B -> Type} n l l' : + All2i P n l l' -> + (forall i x y, P i x y -> Q x y) -> + All2 Q l l'. +Proof. + induction 1; constructor; eauto. +Qed. + Lemma All2i_All2_All2i_All2i {A B C n} {P : nat -> A -> B -> Type} {Q : B -> C -> Type} {R : nat -> A -> C -> Type} {S : nat -> A -> C -> Type} {l l' l''} : All2i P n l l' -> From ba7ef6c8a97753b1dd660197d2116567a81a5396 Mon Sep 17 00:00:00 2001 From: Yann Leray Date: Thu, 5 May 2022 21:28:54 +0200 Subject: [PATCH 02/17] Rebase relevance pcuic --- pcuic/theories/Bidirectional/BDFromPCUIC.v | 87 ++++---- .../theories/Bidirectional/BDStrengthening.v | 28 ++- pcuic/theories/Bidirectional/BDToPCUIC.v | 78 ++++--- pcuic/theories/Bidirectional/BDTyping.v | 116 +++++----- pcuic/theories/Bidirectional/BDUnique.v | 18 +- pcuic/theories/Conversion/PCUICClosedConv.v | 6 +- pcuic/theories/Conversion/PCUICInstConv.v | 6 +- .../Conversion/PCUICUnivSubstitutionConv.v | 11 +- .../theories/Conversion/PCUICWeakeningConv.v | 4 +- pcuic/theories/PCUICAlpha.v | 103 +++++---- pcuic/theories/PCUICArities.v | 77 ++++--- pcuic/theories/PCUICCanonicity.v | 32 +-- pcuic/theories/PCUICContextConversion.v | 4 +- pcuic/theories/PCUICContexts.v | 126 ++++++----- pcuic/theories/PCUICConversion.v | 6 +- pcuic/theories/PCUICCumulProp.v | 28 +-- pcuic/theories/PCUICElimination.v | 21 +- pcuic/theories/PCUICExpandLetsCorrectness.v | 130 ++++++----- pcuic/theories/PCUICGeneration.v | 10 +- pcuic/theories/PCUICInductiveInversion.v | 81 +++---- pcuic/theories/PCUICInductives.v | 66 +++--- pcuic/theories/PCUICInversion.v | 28 +-- pcuic/theories/PCUICParallelReduction.v | 4 +- pcuic/theories/PCUICPrincipality.v | 53 ++--- pcuic/theories/PCUICProgress.v | 24 ++- pcuic/theories/PCUICSR.v | 180 ++++++++-------- pcuic/theories/PCUICSafeLemmata.v | 28 +-- pcuic/theories/PCUICSpine.v | 103 +++++---- pcuic/theories/PCUICSubstitution.v | 40 ++-- pcuic/theories/PCUICToTemplateCorrectness.v | 60 +++--- pcuic/theories/PCUICTyping.v | 201 ++++++++---------- pcuic/theories/PCUICValidity.v | 57 ++--- pcuic/theories/PCUICWellScopedCumulativity.v | 27 ++- pcuic/theories/PCUICWfUniverses.v | 37 ++-- pcuic/theories/Syntax/PCUICDepth.v | 14 +- pcuic/theories/Syntax/PCUICInduction.v | 17 +- pcuic/theories/TemplateToPCUICCorrectness.v | 145 +++++++------ pcuic/theories/TemplateToPCUICExpanded.v | 12 +- pcuic/theories/Typing/PCUICClosedTyp.v | 25 ++- .../Typing/PCUICContextConversionTyp.v | 19 +- pcuic/theories/Typing/PCUICInstTyp.v | 31 +-- pcuic/theories/Typing/PCUICNamelessTyp.v | 5 +- pcuic/theories/Typing/PCUICRenameTyp.v | 25 ++- .../Typing/PCUICUnivSubstitutionTyp.v | 46 ++-- pcuic/theories/Typing/PCUICWeakeningEnvTyp.v | 12 +- pcuic/theories/Typing/PCUICWeakeningTyp.v | 11 +- safechecker/theories/PCUICEqualityDec.v | 2 +- template-coq/theories/BasicAst.v | 2 +- template-coq/theories/EnvironmentTyping.v | 162 +++++++------- template-coq/theories/Typing.v | 6 +- template-coq/theories/TypingWf.v | 2 +- template-coq/theories/Universes.v | 2 + 52 files changed, 1288 insertions(+), 1130 deletions(-) diff --git a/pcuic/theories/Bidirectional/BDFromPCUIC.v b/pcuic/theories/Bidirectional/BDFromPCUIC.v index abd5f9a09..51e865279 100644 --- a/pcuic/theories/Bidirectional/BDFromPCUIC.v +++ b/pcuic/theories/Bidirectional/BDFromPCUIC.v @@ -23,8 +23,8 @@ Proof. lia. Qed. -Lemma ctx_inst_length {ty Σ Γ args Δ} : -PCUICTyping.ctx_inst ty Σ Γ args Δ -> +Lemma ctx_inst_length {ty Γ args Δ} : +PCUICTyping.ctx_inst ty Γ args Δ -> #|args| = context_assumptions Δ. Proof. induction 1; simpl; auto. @@ -34,9 +34,9 @@ rewrite context_assumptions_mapi in IHX. congruence. Qed. -Lemma ctx_inst_app_impl {P Q Σ Γ} {Δ : context} {Δ' args} (c : PCUICTyping.ctx_inst P Σ Γ args (Δ ++ Δ')) : - (forall Γ' t T, P Σ Γ' t T -> Q Σ Γ' t T) -> - PCUICTyping.ctx_inst Q Σ Γ (firstn (context_assumptions Δ) args) Δ. +Lemma ctx_inst_app_impl {P Q Γ} {Δ : context} {Δ' args} (c : PCUICTyping.ctx_inst P Γ args (Δ ++ Δ')) : + (forall Γ' t T, P Γ' t T -> Q Γ' t T) -> + PCUICTyping.ctx_inst Q Γ (firstn (context_assumptions Δ) args) Δ. Proof. revert args Δ' c. induction Δ using ctx_length_ind; intros. @@ -120,17 +120,15 @@ Proof. - intros bdwfΓ. induction bdwfΓ. all: constructor ; auto. - + apply conv_infer_sort in Hs ; auto. - 2: by destruct tu. - destruct Hs as (?&?&?). + + destruct tu as (?&?&?). + apply conv_infer_sort in Hs as (? & ? & ?); auto. eexists. - eassumption. - + apply conv_check in Hc ; auto. - apply conv_infer_sort in Hs ; auto. - 2: by destruct tu. - destruct Hs as (?&?&?). - 1: eexists. - all: eassumption. + split; [eapply geq_relevance|]; tea. + + destruct tu as (?&?&?). + apply conv_check in Hc ; auto. + apply conv_infer_sort in Hs as (? & ? & ?); auto. + eexists. + split; [eapply geq_relevance|]; tea. + by apply conv_check in Hc. - intros. @@ -147,38 +145,37 @@ Proof. constructor. assumption. - - intros n A B ? ? ? ? CumA ? CumB. + - intros n A B ? ? e ? ? CumA ? CumB. apply conv_infer_sort in CumA ; auto. destruct CumA as (?&?&?). apply conv_infer_sort in CumB ; auto. destruct CumB as (?&?&?). eexists. split. - + constructor ; eassumption. + + constructor; [apply (geq_relevance l)| |]; tea. + constructor ; cbn ; auto. 1: by apply wf_local_closed_context. constructor. apply leq_universe_product_mon. all: assumption. - - intros n A t ? ? ? ? CumA ? (?&?&?). + - intros n A t ? ? e ? ? CumA ? (?&?&?). apply conv_infer_sort in CumA ; auto. destruct CumA as (?&?&?). eexists. split. - + econstructor. all: eassumption. + + econstructor; [apply (geq_relevance l)| |]; tea. + apply ws_cumul_pb_Prod ; auto. eapply isType_ws_cumul_pb_refl. - by eexists ; eauto. + by eexists ; split; eauto. - - intros n t A u ? ? ? ? CumA ? Cumt ? (?&?&?). + - intros n t A u ? ? e ? ? CumA ? Cumt ? (?&?&?). apply conv_infer_sort in CumA ; auto. destruct CumA as (?&?&?). apply conv_check in Cumt ; auto. eexists. split. - + econstructor. - all: eassumption. + + econstructor; [apply (geq_relevance l)| | |]; tea. + apply ws_cumul_pb_LetIn_bo. eassumption. @@ -305,12 +302,12 @@ Proof. by (apply infering_ind_typing in i ; auto). assert (consistent_instance_ext Σ (ind_universes mdecl) u). { destruct isdecl. - apply validity in X1 as []. + apply validity in X1 as (s & e & Hs). eapply invert_type_mkApps_ind ; eauto. } assert (consistent_instance_ext Σ (ind_universes mdecl) ui'). { destruct isdecl. - apply validity in X2 as [] ; auto. + apply validity in X2 as (s & e & Hs). eapply invert_type_mkApps_ind ; eauto. } unshelve epose proof (wf_projection_context _ _ _ _) ; eauto. @@ -343,14 +340,14 @@ Proof. split. 2:{ apply isType_ws_cumul_pb_refl. - eapply nth_error_all in Alltypes as [? []] ; tea. - eexists ; tea. + eapply nth_error_all in Alltypes as (s & e & Hs & Hs'); tea. + eexists; split; cbnr; tea. } constructor ; eauto. + apply (All_impl Alltypes). - intros ? [? [? s]]. - apply conv_infer_sort in s as [? []] ; auto. - eexists ; eauto. + intros ? (s & e & Hs & Hs'). + apply conv_infer_sort in Hs' as (s' & e' & Hleq) ; auto. + eexists; split; [apply (geq_relevance Hleq)|]; tea. + apply (All_impl Allbodies). intros ? [? s]. by apply conv_check in s ; auto. @@ -360,14 +357,14 @@ Proof. split. 2:{ apply isType_ws_cumul_pb_refl. - eapply nth_error_all in Alltypes as [? []] ; tea. - eexists ; tea. + eapply nth_error_all in Alltypes as (? & ? & ? & ?) ; tea. + eexists; split; cbnr; tea. } constructor ; eauto. + apply (All_impl Alltypes). - intros ? [? [? s]]. - apply conv_infer_sort in s as [? []] ; auto. - eexists ; eauto. + intros ? (s & e & Hs & Hs'). + apply conv_infer_sort in Hs' as (s' & e' & Hleq) ; auto. + eexists; split; [apply (geq_relevance Hleq)|]; tea. + apply (All_impl Allbodies). intros ? [? s]. by apply conv_check in s ; auto. @@ -410,11 +407,19 @@ Qed. Lemma isType_infering_sort `{checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} {Γ t} : isType Σ Γ t -> ∑ u', Σ ;;; Γ |- t ▹□ u'. Proof. - intros [? ty]. - eapply typing_infering_sort in ty as [? []]; tea. + intros (? & e & ty). + eapply typing_infering_sort in ty as (? & ? & ?); tea. now eexists. Qed. +Lemma isTypeRel_infering_sort `{checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} {Γ t r} : + isTypeRel Σ Γ t r -> ∑ u', relevance_of_sort u' = r × Σ ;;; Γ |- t ▹□ u'. +Proof. + intros (s & e & ty). + eapply typing_infering_sort in ty as (s' & e' & Hleq); tea. + exists s'; split; [apply (geq_relevance Hleq)|]; 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'), @@ -453,14 +458,14 @@ Proof. induction Γ' as [|[? [] ?]] in wfΓ' |- *. all: constructor ; inversion wfΓ' ; subst ; cbn in *. 1,4: now apply IHΓ'. - - now apply isType_infering_sort. + - now apply isTypeRel_infering_sort. - now apply typing_checking. - - now apply isType_infering_sort. + - now apply isTypeRel_infering_sort. Qed. Theorem ctx_inst_typing_bd `{checker_flags} (Σ : global_env_ext) Γ l Δ (wfΣ : wf Σ) : - PCUICTyping.ctx_inst typing Σ Γ l Δ -> - PCUICTyping.ctx_inst checking Σ Γ l Δ. + PCUICTyping.ctx_inst (typing Σ) Γ l Δ -> + PCUICTyping.ctx_inst (checking Σ) Γ l Δ. Proof. intros inl. induction inl. diff --git a/pcuic/theories/Bidirectional/BDStrengthening.v b/pcuic/theories/Bidirectional/BDStrengthening.v index 005c0fd6e..0aa7f2475 100644 --- a/pcuic/theories/Bidirectional/BDStrengthening.v +++ b/pcuic/theories/Bidirectional/BDStrengthening.v @@ -390,14 +390,14 @@ Section OnFreeVars. - easy. - easy. - intros until bty. - move => _ _ _ Hbty ? ? /= /andP [] ? ?. + move => _ _ _ _ Hbty ? ? /= /andP [] ? ?. apply /andP ; split ; tea. apply Hbty ; tea. rewrite on_ctx_free_vars_snoc. apply /andP ; split ; tea. - intros until A. - move => _ _ _ _ _ Ht ? ? /= /andP [] ? /andP [] ? ?. + move => _ _ _ _ _ _ Ht ? ? /= /andP [] ? /andP [] ? ?. repeat (apply /andP ; split ; tea). apply Ht ; tea. rewrite on_ctx_free_vars_snoc. @@ -580,8 +580,8 @@ Lemma rename_telescope P f Γ Δ tel tys: on_ctx_free_vars P Γ -> forallb (on_free_vars P) tys -> on_free_vars_ctx P (List.rev tel) -> - PCUICTyping.ctx_inst (fun _ => Pcheck) Σ Γ tys tel -> - PCUICTyping.ctx_inst checking Σ Δ (map (rename f) tys) (rename_telescope f tel). + PCUICTyping.ctx_inst Pcheck Γ tys tel -> + PCUICTyping.ctx_inst (checking Σ) Δ (map (rename f) tys) (rename_telescope f tel). Proof using Type. intros ur hΓ htys htel ins. induction ins in Δ, ur, hΓ, htys, htel |- *. @@ -618,8 +618,10 @@ Proof using wfΣ. induction hΓ. + constructor. + constructor ; tea. + destruct tu as (s' & e & H). eexists ; eauto. + constructor ; tea. + destruct tu as (s' & e & H). eexists ; eauto. - intros Γ Γ' wfΓ' allΓ'. red. move => P Δ f hf hΓ hΓ'. @@ -630,7 +632,9 @@ Proof using wfΣ. move: hΓ' => /andP [] ? ?. constructor ; eauto. 1: by eapply IHallΓ' ; eauto. + destruct tu as (s' & e & H). eexists. + split; [apply e|]. eapply Hs. * eapply urenaming_context ; tea. * rewrite on_ctx_free_vars_concat. @@ -642,7 +646,9 @@ Proof using wfΣ. move: hΓ' => /andP [] ? /andP /= [] ? ?. constructor ; eauto. * by eapply IHallΓ' ; eauto. - * eexists. + * destruct tu as (s' & e & H). + eexists. + split; [apply e|]. eapply Hs. 1: eapply urenaming_context ; tea. 2: eauto. @@ -816,14 +822,14 @@ Proof using wfΣ. + by rewrite nth_error_map H0 /=. + eapply All_mix in X ; tea. eapply All_map, All_impl ; tea. - move => ? [] /andP [] ? ? [] ? [] ? p. - rewrite -map_dtype. + move => ? [] /andP [] ? ? [] ? [] e [] ? p. eexists. + split; [apply e|]. eapply p ; tea. + eapply All_mix in X0 ; tea. eapply All_map, All_impl ; tea. move => ? [] /andP [] ? ? [] ? p. - rewrite -map_dbody -map_dtype rename_fix_context rename_context_length -(fix_context_length mfix) -rename_shiftn. + rewrite /on_def_body -map_dbody -map_dtype rename_fix_context rename_context_length -(fix_context_length mfix) -rename_shiftn. eapply p ; tea. * rewrite -(fix_context_length mfix). eapply urenaming_context ; tea. @@ -840,14 +846,14 @@ Proof using wfΣ. + by rewrite nth_error_map H0 /=. + eapply All_mix in X ; tea. eapply All_map, All_impl ; tea. - move => ? [] /andP [] ? ? [] ? [] ? p. - rewrite -map_dtype. + move => ? [] /andP [] ? ? [] ? [] e [] ? p. eexists. + split; [apply e|]. eapply p ; tea. + eapply All_mix in X0 ; tea. eapply All_map, All_impl ; tea. move => ? [] /andP [] ? ? [] ? p. - rewrite -map_dbody -map_dtype rename_fix_context rename_context_length -(fix_context_length mfix) -rename_shiftn. + rewrite /on_def_body -map_dbody -map_dtype rename_fix_context rename_context_length -(fix_context_length mfix) -rename_shiftn. eapply p ; tea. * rewrite -(fix_context_length mfix). eapply urenaming_context ; tea. diff --git a/pcuic/theories/Bidirectional/BDToPCUIC.v b/pcuic/theories/Bidirectional/BDToPCUIC.v index 52f57b9ab..e38a82697 100644 --- a/pcuic/theories/Bidirectional/BDToPCUIC.v +++ b/pcuic/theories/Bidirectional/BDToPCUIC.v @@ -37,7 +37,7 @@ Proof. change (d :: Γ') with (Γ' ,, d). destruct d as [na' [bd|] ty]; rewrite !app_context_cons; intro HH. - rewrite subst_context_snoc0. simpl. - inversion HH; subst; cbn in *. destruct X0 as [s X0]. + inversion HH; subst. change (Γ,, vdef na b t ,,, Γ') with (Γ ,,, [vdef na b t] ,,, Γ') in *. assert (subslet Σ (Δ ,,, Γ) [b] [vdef na b t]). { pose proof (cons_let_def Σ (Δ ,,, Γ) [] [] na b t) as XX. @@ -46,14 +46,14 @@ Proof. } constructor; cbn; auto. 1: apply IHΓ' ; exact X. - 1: exists s. + 1: destruct X0 as (s & e & Hs); exists s; split; [apply e|]. 1: change (tSort s) with (subst [b] #|Γ'| (tSort s)). all: rewrite app_context_assoc. all: eapply substitution; tea. - 1: rewrite !app_context_assoc in X0 ; assumption. - rewrite !app_context_assoc in X1 ; assumption. + 1: rewrite !app_context_assoc in Hs; assumption. + red in X1; rewrite !app_context_assoc in X1 ; assumption. - rewrite subst_context_snoc0. simpl. - inversion HH; subst; cbn in *. destruct X0 as [s X0]. + inversion HH; subst; cbn in *. change (Γ,, vdef na b t ,,, Γ') with (Γ ,,, [vdef na b t] ,,, Γ') in *. assert (subslet Σ (Δ ,,, Γ) [b] [vdef na b t]). { pose proof (cons_let_def Σ (Δ ,,, Γ) [] [] na b t) as XX. @@ -61,11 +61,11 @@ Proof. apply All_local_env_app_l in X. inversion X; subst; cbn in *; assumption. } constructor; cbn; auto. 1: apply IHΓ' ; exact X. - exists s. + destruct X0 as (s & e & Hs); exists s; split; [apply e|]. change (tSort s) with (subst [b] #|Γ'| (tSort s)). rewrite app_context_assoc. all: eapply substitution; tea. - rewrite !app_context_assoc in X0. eassumption. + rewrite !app_context_assoc in Hs. eassumption. Qed. Lemma wf_rel_weak `{checker_flags} (Σ : global_env_ext) (wfΣ : wf Σ) Γ (wfΓ : wf_local Σ Γ) Γ' : @@ -76,13 +76,13 @@ Proof. - intros wfl. inversion_clear wfl. constructor. + apply IHΓ'. assumption. - + apply infer_typing_sort_impl with id X0; intros Hs. + + apply infer_typing_sort_impl with id X0 => //; intros Hs. apply weaken_ctx ; eauto. + apply weaken_ctx ; auto. - intros wfl. inversion_clear wfl. constructor. + apply IHΓ'. assumption. - + apply infer_typing_sort_impl with id X0; intros Hs. + + apply infer_typing_sort_impl with id X0 => //; intros Hs. apply weaken_ctx ; eauto. Qed. @@ -143,7 +143,8 @@ Section BDToPCUICTyping. 1,3: assumption. all: do 2 red. 3: apply Hc; auto. - all: apply infer_sort_impl with id tu; now intros Ht. + 3: eapply isType_of_isTypeRel. + all: apply infer_sort_impl with id tu => //; now intros Ht. Qed. Lemma bd_wf_local_rel Γ (wfΓ : wf_local Σ Γ) Γ' (all: wf_local_bd_rel Σ Γ Γ') : @@ -159,13 +160,13 @@ Section BDToPCUICTyping. all: constructor. 1,3: assumption. all: red. - 3: apply Hc; [by apply wf_local_app|]. - all: apply infer_sort_impl with id tu; intros Ht. + 3: apply Hc; [by apply wf_local_app|eapply isType_of_isTypeRel]. + all: apply infer_sort_impl with id tu => //; intros Ht. all: now apply Hs, wf_local_app. Qed. Lemma ctx_inst_impl Γ (wfΓ : wf_local Σ Γ) (Δ : context) (wfΔ : wf_local_rel Σ Γ (List.rev Δ)) : - forall args, PCUICTyping.ctx_inst (fun _ => Pcheck) Σ Γ args Δ -> ctx_inst Σ Γ args Δ. + forall args, PCUICTyping.ctx_inst Pcheck Γ args Δ -> ctx_inst Σ Γ args Δ. Proof using wfΣ. revert wfΔ. induction Δ using ctx_length_ind. @@ -173,12 +174,13 @@ Section BDToPCUICTyping. intros wfΔ args ctxi ; inversion ctxi. - subst d. subst. - assert (isType Σ Γ t). + assert (isTypeRel Σ Γ t na.(binder_relevance)). { eapply wf_local_rel_app_inv in wfΔ as [wfd _]. inversion_clear wfd. eassumption. } + pose proof (isType_of_isTypeRel X2). constructor ; auto. apply X ; auto. 1: by rewrite subst_telescope_length ; reflexivity. @@ -194,12 +196,13 @@ Section BDToPCUICTyping. eassumption. - subst d. subst. - assert (isType Σ Γ t). + assert (isTypeRel Σ Γ t na.(binder_relevance)). { eapply wf_local_rel_app_inv in wfΔ as [wfd _]. inversion_clear wfd. eassumption. } + pose proof (isType_of_isTypeRel X1). constructor ; auto. apply X ; auto. 1: by rewrite subst_telescope_length ; reflexivity. @@ -235,21 +238,21 @@ Section BDToPCUICTyping. - red ; intros ; econstructor ; eauto. + apply X2 ; auto. - eexists. eauto. + eexists. split; [cbnr|]. eauto. + apply X4. constructor ; auto. * eexists. eauto. - * apply X2 ; auto. eexists. eauto. + * apply X2 ; auto. eexists. split; [cbnr|]. eauto. - red ; intros. eapply type_App' ; auto. apply X2 ; auto. specialize (X0 X3). - apply validity in X0 as [? X0]. - apply inversion_Prod in X0 as (? & ? & ? & _). + apply validity in X0 as (? & ? & X0). + apply inversion_Prod in X0 as (? & ? & ? & ? & _). 2: done. - eexists. eassumption. + eexists. split; [cbnr|]. eassumption. - red ; intros ; econstructor ; eauto. @@ -306,6 +309,7 @@ Section BDToPCUICTyping. (pparams p ++ skipn (ci_npar ci) args))) as [? tyapp]. { eexists. + split; [cbnr|]. eapply type_mkApps_arity. 1: econstructor ; eauto. erewrite PCUICGlobalMaps.ind_arity_eq. @@ -342,7 +346,7 @@ Section BDToPCUICTyping. fold ptm in brctxty. repeat split ; auto. apply Hbody ; auto. - eexists. + eexists. split; [cbnr|]. eassumption. - red ; intros ; econstructor ; eauto. @@ -352,17 +356,19 @@ Section BDToPCUICTyping. + clear H H0 H1 X0. induction X. all: constructor ; auto. - destruct p as (? & ? & ?). + destruct p as (? & e & ? & ?). eexists. + split; [apply e|]. apply p. auto. - + have Htypes : All (fun d => isType Σ Γ (dtype d)) mfix. + + have Htypes : All (on_def_type (lift_typing typing Σ) Γ) mfix. { clear H H0 H1 X0. induction X. all: constructor ; auto. - destruct p as (? & ? & ?). + destruct p as (? & e & ? & ?). eexists. + split; [apply e|]. apply p. auto. } @@ -375,8 +381,8 @@ Section BDToPCUICTyping. 2:{ inversion_clear X. apply IHX0 ; auto. inversion_clear Htypes. auto. } destruct p. apply p ; auto. - inversion_clear Htypes as [| ? ? [u]]. - exists u. + inversion_clear Htypes as [| ? ? (u & e & ?)]. + exists u. split; [cbnr|]. change (tSort u) with (lift0 #|Γ'| (tSort u)). apply weakening. all: auto. @@ -385,17 +391,19 @@ Section BDToPCUICTyping. + clear H H0 H1 X0. induction X. all: constructor ; auto. - destruct p as (? & ? & ?). + destruct p as (? & e & ? & ?). eexists. + split; [apply e|]. apply p. auto. - + have Htypes : All (fun d => isType Σ Γ (dtype d)) mfix. + + have Htypes : All (on_def_type (lift_typing typing Σ) Γ) mfix. { clear H H0 H1 X0. induction X. all: constructor ; auto. - destruct p as (? & ? & ?). + destruct p as (? & e & ? & ?). eexists. + split; [apply e|]. apply p. auto. } @@ -407,8 +415,8 @@ Section BDToPCUICTyping. 2:{ inversion_clear X. apply IHX0 ; auto. inversion_clear Htypes. auto. } destruct p. apply p ; auto. - inversion_clear Htypes as [| ? ? [u]]. - exists u. + inversion_clear Htypes as [| ? ? (u & e & ?)]. + exists u. split; [cbnr|]. change (tSort u) with (lift0 #|Γ'| (tSort u)). apply weakening. all: auto. @@ -423,7 +431,7 @@ Section BDToPCUICTyping. now eapply type_reduction. - red ; intros. - destruct X3. + destruct X3 as (s & e & Hs). econstructor ; eauto. eapply (cumulAlgo_cumulSpec _ (pb := Cumul)), into_ws_cumul_pb ; tea. + fvs. @@ -464,6 +472,7 @@ Theorem infering_sort_isType `{checker_flags} (Σ : global_env_ext) Γ t u (wfΣ Proof. intros wfΓ Ht. exists u. + split; [cbnr|]. now apply infering_sort_typing. Qed. @@ -506,8 +515,8 @@ Qed. Theorem ctx_inst_bd_typing `{checker_flags} (Σ : global_env_ext) Γ l Δ (wfΣ : wf Σ) : wf_local Σ (Γ,,,Δ) -> - PCUICTyping.ctx_inst checking Σ Γ l (List.rev Δ) -> - PCUICTyping.ctx_inst typing Σ Γ l (List.rev Δ). + PCUICTyping.ctx_inst (checking Σ) Γ l (List.rev Δ) -> + PCUICTyping.ctx_inst (typing Σ) Γ l (List.rev Δ). Proof. intros wfΓ inl. rewrite -(List.rev_involutive Δ) in wfΓ. @@ -520,6 +529,7 @@ Proof. rewrite /= app_context_assoc in wfΓ. eapply wf_local_app_inv in wfΓ as [wfΓ _]. inversion wfΓ ; subst. + apply isType_of_isTypeRel in X0. now apply checking_typing. } constructor ; tea. diff --git a/pcuic/theories/Bidirectional/BDTyping.v b/pcuic/theories/Bidirectional/BDTyping.v index 527ec59d4..f99a9d3c9 100644 --- a/pcuic/theories/Bidirectional/BDTyping.v +++ b/pcuic/theories/Bidirectional/BDTyping.v @@ -34,16 +34,19 @@ Inductive infering `{checker_flags} (Σ : global_env_ext) (Γ : context) : term Σ ;;; Γ |- tSort s ▹ tSort (Universe.super s) | infer_Prod na A B s1 s2 : + isSortRel s1 na.(binder_relevance) -> Σ ;;; Γ |- A ▹□ s1 -> Σ ;;; Γ ,, vass na A |- B ▹□ s2 -> Σ ;;; Γ |- tProd na A B ▹ tSort (Universe.sort_of_product s1 s2) | infer_Lambda na A t s B : + isSortRel s na.(binder_relevance) -> Σ ;;; Γ |- A ▹□ s -> Σ ;;; Γ ,, vass na A |- t ▹ B -> Σ ;;; Γ |- tLambda na A t ▹ tProd na A B | infer_LetIn na b B t s A : + isSortRel s na.(binder_relevance) -> Σ ;;; Γ |- B ▹□ s -> Σ ;;; Γ |- b ◃ B -> Σ ;;; Γ ,, vdef na b B |- t ▹ A -> @@ -82,7 +85,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 -> - ctx_inst checking Σ Γ (pparams p) + ctx_inst (checking Σ) Γ (pparams p) (List.rev mdecl.(ind_params)@[p.(puinst)]) -> isCoFinite mdecl.(ind_finite) = false -> R_global_instance Σ (eq_universe Σ) (leq_universe Σ) @@ -107,16 +110,16 @@ Inductive infering `{checker_flags} (Σ : global_env_ext) (Γ : context) : term | infer_Fix mfix n decl : fix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> - All (fun d => {s & Σ ;;; Γ |- d.(dtype) ▹□ s}) mfix -> - All (fun d => Σ ;;; Γ ,,, fix_context mfix |- d.(dbody) ◃ lift0 #|fix_context mfix| d.(dtype)) mfix -> + All (on_def_type (lift_sorting checking infering_sort Σ) Γ) mfix -> + All (on_def_body (lift_sorting checking infering_sort Σ) (fix_context mfix) Γ) mfix -> wf_fixpoint Σ mfix -> Σ ;;; Γ |- tFix mfix n ▹ dtype decl | infer_CoFix mfix n decl : cofix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> - All (fun d => {s & Σ ;;; Γ |- d.(dtype) ▹□ s}) mfix -> - All (fun d => Σ ;;; Γ ,,, fix_context mfix |- d.(dbody) ◃ lift0 #|fix_context mfix| d.(dtype)) mfix -> + All (on_def_type (lift_sorting checking infering_sort Σ) Γ) mfix -> + All (on_def_body (lift_sorting checking infering_sort Σ) (fix_context mfix) Γ) mfix -> wf_cofixpoint Σ mfix -> Σ ;;; Γ |- tCoFix mfix n ▹ dtype decl @@ -199,9 +202,9 @@ Proof. | |- _ => exact 1 end. - exact (S (Nat.max a0 (Nat.max i (Nat.max i0 (Nat.max (ctx_inst_size _ (checking_size _) c1) (branches_size (checking_size _) (infering_sort_size _) a2)))))). - - exact (S (Nat.max (all_size _ (fun d p => infering_sort_size _ _ _ _ _ p.π2) a) + - exact (S (Nat.max (all_size _ (fun d p => infering_sort_size _ _ _ _ _ p.π2.2) a) (all_size _ (fun x p => checking_size _ _ _ _ _ p) a0))). - - exact (S (Nat.max (all_size _ (fun d p => infering_sort_size _ _ _ _ _ p.π2) a) + - exact (S (Nat.max (all_size _ (fun d p => infering_sort_size _ _ _ _ _ p.π2.2) a) (all_size _ (fun x => checking_size _ _ _ _ _) a0))). Defined. @@ -313,6 +316,8 @@ Section BidirectionalInduction. let Pdecl_sort := fun Σ Γ wfΓ t u tyT => Psort Γ t u in let Pdecl_check_rel Γ := fun _ Δ _ t T _ => Pcheck (Γ,,,Δ) t T in let Pdecl_sort_rel Γ := fun _ Δ _ t u _ => Psort (Γ,,,Δ) t u in + let PcheckΣ := fun Σ => Pcheck in + let PsortΣ := fun Σ => Psort in (forall (Γ : context) (wfΓ : wf_local_bd Σ Γ), All_local_env_over_sorting checking infering_sort Pdecl_check Pdecl_sort Σ Γ wfΓ -> PΓ Γ) -> @@ -328,28 +333,31 @@ Section BidirectionalInduction. wf_universe Σ s-> Pinfer Γ (tSort s) (tSort (Universe.super s))) -> - (forall (Γ : context) (n : aname) (t b : term) (s1 s2 : Universe.t), + (forall (Γ : context) (na : aname) (t b : term) (s1 s2 : Universe.t), + isSortRel s1 na.(binder_relevance) -> Σ ;;; Γ |- t ▹□ s1 -> Psort Γ t s1 -> - Σ ;;; Γ,, vass n t |- b ▹□ s2 -> - Psort (Γ,, vass n t) b s2 -> - Pinfer Γ (tProd n t b) (tSort (Universe.sort_of_product s1 s2))) -> + Σ ;;; Γ,, vass na t |- b ▹□ s2 -> + Psort (Γ,, vass na t) b s2 -> + Pinfer Γ (tProd na t b) (tSort (Universe.sort_of_product s1 s2))) -> - (forall (Γ : context) (n : aname) (t b : term) (s : Universe.t) (bty : term), + (forall (Γ : context) (na : aname) (t b : term) (s : Universe.t) (bty : term), + isSortRel s na.(binder_relevance) -> Σ ;;; Γ |- t ▹□ s -> Psort Γ t s -> - Σ ;;; Γ,, vass n t |- b ▹ bty -> - Pinfer (Γ,, vass n t) b bty -> - Pinfer Γ (tLambda n t b) (tProd n t bty)) -> + Σ ;;; Γ,, vass na t |- b ▹ bty -> + Pinfer (Γ,, vass na t) b bty -> + Pinfer Γ (tLambda na t b) (tProd na t bty)) -> - (forall (Γ : context) (n : aname) (b B t : term) (s : Universe.t) (A : term), + (forall (Γ : context) (na : aname) (b B t : term) (s : Universe.t) (A : term), + isSortRel s na.(binder_relevance) -> Σ ;;; Γ |- B ▹□ s -> Psort Γ B s -> Σ ;;; Γ |- b ◃ B -> Pcheck Γ b B -> - Σ ;;; Γ,, vdef n b B |- t ▹ A -> - Pinfer (Γ,, vdef n b B) t A -> - Pinfer Γ (tLetIn n b B t) (tLetIn n b B A)) -> + Σ ;;; Γ,, vdef na b B |- t ▹ A -> + Pinfer (Γ,, vdef na b B) t A -> + Pinfer Γ (tLetIn na b B t) (tLetIn na b B A)) -> (forall (Γ : context) (t : term) na A B u, Σ ;;; Γ |- t ▹Π (na, A, B) -> @@ -392,9 +400,9 @@ Section BidirectionalInduction. wf_local_bd_rel Σ Γ predctx -> PΓ_rel Γ predctx -> is_allowed_elimination Σ (ind_kelim idecl) ps -> - ctx_inst checking Σ Γ (pparams p) + ctx_inst (checking Σ) Γ (pparams p) (List.rev mdecl.(ind_params)@[p.(puinst)]) -> - ctx_inst (fun _ => Pcheck) Σ Γ p.(pparams) + ctx_inst Pcheck Γ p.(pparams) (List.rev (subst_instance p.(puinst) mdecl.(ind_params))) -> isCoFinite mdecl.(ind_finite) = false -> R_global_instance Σ (eq_universe Σ) (leq_universe Σ) @@ -423,18 +431,16 @@ Section BidirectionalInduction. (forall (Γ : context) (mfix : mfixpoint term) (n : nat) decl, fix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> - All (fun d => {s & (Σ ;;; Γ |- d.(dtype) ▹□ s) × Psort Γ d.(dtype) s}) mfix -> - All (fun d => (Σ ;;; Γ ,,, fix_context mfix |- d.(dbody) ◃ lift0 #|fix_context mfix| d.(dtype)) × - Pcheck (Γ ,,, fix_context mfix) d.(dbody) (lift0 #|fix_context mfix| d.(dtype))) mfix -> + All (on_def_type (lift_sorting (Prop_conj checking PcheckΣ) (Prop_conj infering_sort PsortΣ) Σ) Γ) mfix -> + All (on_def_body (lift_sorting (Prop_conj checking PcheckΣ) (Prop_conj infering_sort PsortΣ) Σ) (fix_context mfix) Γ) mfix -> wf_fixpoint Σ mfix -> Pinfer Γ (tFix mfix n) (dtype decl)) -> (forall (Γ : context) (mfix : mfixpoint term) (n : nat) decl, cofix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> - All (fun d => {s & (Σ ;;; Γ |- d.(dtype) ▹□ s) × Psort Γ d.(dtype) s}) mfix -> - All (fun d => (Σ ;;; Γ ,,, fix_context mfix |- d.(dbody) ◃ lift0 #|fix_context mfix| d.(dtype)) × - Pcheck (Γ ,,, fix_context mfix) d.(dbody) (lift0 #|fix_context mfix| d.(dtype))) mfix -> + All (on_def_type (lift_sorting (Prop_conj checking PcheckΣ) (Prop_conj infering_sort PsortΣ) Σ) Γ) mfix -> + All (on_def_body (lift_sorting (Prop_conj checking PcheckΣ) (Prop_conj infering_sort PsortΣ) Σ) (fix_context mfix) Γ) mfix -> wf_cofixpoint Σ mfix -> Pinfer Γ (tCoFix mfix n) (dtype decl)) -> @@ -465,7 +471,7 @@ Section BidirectionalInduction. env_prop_bd. Proof using Type. - intros Pdecl_check Pdecl_sort Pdecl_check_rel Pdecl_sort_rel HΓ HΓRel HRel HSort HProd HLambda HLetIn HApp HConst HInd HConstruct HCase + intros Pdecl_check Pdecl_sort Pdecl_check_rel Pdecl_sort_rel PcheckΣ PsortΣ HΓ HΓRel HRel HSort HProd HLambda HLetIn HApp HConst HInd HConstruct HCase HProj HFix HCoFix HiSort HiProd HiInd HCheck ; unfold env_prop_bd. pose (@Fix_F typing_sum (precompose lt typing_sum_size) Ptyping_sum) as p. forward p. @@ -490,7 +496,7 @@ Section BidirectionalInduction. - eapply HΓ. dependent induction wfΓ. + constructor. - + destruct t0 as (u & d). + + destruct t0 as (s & e & Hs). constructor. * apply IHwfΓ. intros ; apply IH. @@ -500,18 +506,18 @@ Section BidirectionalInduction. cbn. assert (0 < wfl_size wfΓ) by apply All_local_env_size_pos. lia. - + destruct t0 as [u h]. + + destruct t0 as (s & e & Hs). constructor. * apply IHwfΓ. intros ; apply IH. - cbn in H |- *. pose proof (infering_sort_size_pos h). lia. - * red. applyIH. pose proof (infering_sort_size_pos h). cbn in H |- *. lia. + cbn in H |- *. pose proof (infering_sort_size_pos Hs). lia. + * red. applyIH. pose proof (infering_sort_size_pos Hs). cbn in H |- *. lia. * red. applyIH. pose proof (checking_size_pos t1). cbn in H |- *. lia. - eapply HΓRel. dependent induction wfΓ'. + constructor. - + destruct t0 as (u & d). + + destruct t0 as (s & e & Hs). constructor. * apply IHwfΓ'. intros ; apply IH. @@ -524,14 +530,14 @@ Section BidirectionalInduction. assert (0 < wfl_size_rel wfΓ') by apply All_local_env_size_pos. unfold wfl_size_rel in H. lia. - + destruct t0 as [u h]. + + destruct t0 as (s & e & Hs). constructor. * apply IHwfΓ'. intros ; apply IH. cbn in H |- *. fold (wfl_size_rel wfΓ'). - pose proof (infering_sort_size_pos h). lia. - * red. applyIH. pose proof (infering_sort_size_pos h). cbn in H |- *. lia. + pose proof (infering_sort_size_pos Hs). lia. + * red. applyIH. pose proof (infering_sort_size_pos Hs). cbn in H |- *. lia. * red. applyIH. pose proof (checking_size_pos t1). cbn in H |- *. lia. @@ -609,23 +615,24 @@ Section BidirectionalInduction. - unshelve eapply HFix ; eauto. - all: simpl in IH. + all: cbn [typing_sum_size infering_size] in IH. all: remember (fix_context mfix) as mfixcontext. - - + remember (all_size _ _ a0) as s. + all: unfold on_def_body, PcheckΣ in a0 |- *; cbn in a0 |- *. + + remember (all_size _ _ a0) as size. clear -IH. dependent induction a. 1: by constructor. - constructor ; cbn. - * destruct p ; eexists ; split. - 1: eassumption. + constructor. + * destruct p as (s & e & Hs). + exists s; + repeat split; [apply e|apply Hs|hnf]. applyIH. - * apply (IHa s). + * apply (IHa size). intros. apply IH. - cbn. lia. + cbn [all_size]. lia. - + remember (all_size _ _ a) as s. + + remember (all_size _ (fun d p => _ p.π2.2) a) as size. clear -IH. induction a0. 1: by constructor. @@ -639,23 +646,24 @@ Section BidirectionalInduction. - unshelve eapply HCoFix ; eauto. - all: simpl in IH. + all: cbn [typing_sum_size infering_size] in IH. all: remember (fix_context mfix) as mfixcontext. - - + remember (all_size _ _ a0) as s. + all: unfold on_def_body, PcheckΣ in a0 |- *; cbn in a0 |- *. + + remember (all_size _ _ a0) as size. clear -IH. dependent induction a. 1: by constructor. - constructor ; cbn. - * destruct p ; eexists ; split. - 1: eassumption. + constructor. + * destruct p as (s & e & Hs). + exists s; + repeat split; [apply e|apply Hs|hnf]. applyIH. - * apply (IHa s). + * apply (IHa size). intros. apply IH. - cbn. lia. + cbn [all_size]. lia. - + remember (all_size _ _ a) as s. + + remember (all_size _ (fun d p => _ p.π2.2) a) as size. clear -IH. induction a0 as [| ? ? ?]. 1: by constructor. diff --git a/pcuic/theories/Bidirectional/BDUnique.v b/pcuic/theories/Bidirectional/BDUnique.v index 1cf71de61..857760357 100644 --- a/pcuic/theories/Bidirectional/BDUnique.v +++ b/pcuic/theories/Bidirectional/BDUnique.v @@ -68,21 +68,21 @@ Proof using wfΣ. - eexists ; split. all: eapply closed_red_refl ; fvs. - - apply H in X2 => //. - apply H0 in X3. - 2:{ constructor ; auto. now eapply infering_sort_isType. } + - apply H0 in X2 => //. + apply H1 in X3. + 2:{ constructor ; auto. exists s1; split; [auto|]. now eapply infering_sort_typing. } subst. eexists ; split. all: eapply closed_red_refl ; fvs. - apply X1 in X4 as [bty' []]. - 2:{ constructor ; auto. now eapply infering_sort_isType. } - exists (tProd n t bty') ; split. + 2:{ constructor ; auto. eexists; split; eauto; now eapply infering_sort_typing. } + exists (tProd na t bty') ; split. all: now eapply closed_red_prod_codom. - apply X2 in X6 as [A' []]. - 2:{ constructor ; auto. 2: eapply checking_typing ; tea. all: now eapply infering_sort_isType. } - exists (tLetIn n b B A'). + 2:{ constructor ; auto. 2: eapply checking_typing ; tea; now eapply infering_sort_isType. eexists; split; eauto; now apply infering_sort_typing. } + exists (tLetIn na b B A'). assert (Σ ;;; Γ |- b : B) by (eapply checking_typing ; tea ; now eapply infering_sort_isType). split. @@ -101,12 +101,12 @@ Proof using wfΣ. 1: constructor. rewrite subst_empty. eapply checking_typing ; tea. - now eapply isType_tProd, validity, infering_prod_typing. + now eapply isType_of_isTypeRel, isType_tProd, validity, infering_prod_typing. + constructor. 1: constructor. rewrite subst_empty. eapply checking_typing ; tea. - now eapply isType_tProd, validity, infering_prod_typing. + now eapply isType_of_isTypeRel, isType_tProd, validity, infering_prod_typing. - replace decl0 with decl by (eapply declared_constant_inj ; eassumption). eexists ; split. diff --git a/pcuic/theories/Conversion/PCUICClosedConv.v b/pcuic/theories/Conversion/PCUICClosedConv.v index 3ecc6004c..c63c2b0ee 100644 --- a/pcuic/theories/Conversion/PCUICClosedConv.v +++ b/pcuic/theories/Conversion/PCUICClosedConv.v @@ -46,9 +46,9 @@ Proof. rewrite app_context_length in b0. now rewrite Nat.add_comm. - apply Alli_app_inv; auto. constructor. simpl. rewrite List.rev_length. 2:constructor. - unfold closed_decl. unfold Pclosed in b. simpl. - rewrite app_context_length in b. rewrite Nat.add_comm. - now rewrite andb_true_r in b. + unfold closed_decl. unfold Pclosed in b0. simpl. + rewrite app_context_length in b0. rewrite Nat.add_comm. + now rewrite andb_true_r in b0. Qed. Lemma sorts_local_ctx_Pclosed Σ Γ Δ s : diff --git a/pcuic/theories/Conversion/PCUICInstConv.v b/pcuic/theories/Conversion/PCUICInstConv.v index 7463430b0..3562b7037 100644 --- a/pcuic/theories/Conversion/PCUICInstConv.v +++ b/pcuic/theories/Conversion/PCUICInstConv.v @@ -295,7 +295,7 @@ Proof. - unfold inst_context, snoc. rewrite fold_context_k_snoc0. unfold snoc. f_equal. all: auto. unfold map_decl. simpl. unfold vass. f_equal. - destruct t0 as [s ht]. eapply typed_inst. all: eauto. + destruct t0 as (s & e & Hs). eapply typed_inst. all: eauto. - unfold inst_context, snoc. rewrite fold_context_k_snoc0. unfold snoc. f_equal. all: auto. unfold map_decl. simpl. unfold vdef. f_equal. @@ -1616,11 +1616,11 @@ Proof using Type. induction X. - now simpl. - rewrite inst_context_snoc /=. constructor; auto. - apply infer_typing_sort_impl with id t0; intros Hs. + apply infer_typing_sort_impl with id t0 => //; intros Hs. eapply (Hs (Δ' ,,, inst_context σ Γ0) (⇑^#|Γ0| σ)) => //. eapply well_subst_app; auto. - rewrite inst_context_snoc /=. constructor; auto. - * apply infer_typing_sort_impl with id t0; intros Hs. + * apply infer_typing_sort_impl with id t0 => //; intros Hs. eapply (Hs (Δ' ,,, inst_context σ Γ0) (⇑^#|Γ0| σ)) => //. eapply well_subst_app; auto. * simpl. apply t1 => //. diff --git a/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v b/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v index 0445e23c3..fa3bff8c1 100644 --- a/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v +++ b/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v @@ -1613,7 +1613,8 @@ 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 infer_typing_sort_impl with _ tu; cbn in *; eauto. + all: eapply infer_typing_sort_impl with _ tu; [apply relevance_subst_opt|]. + all: cbn in *; eauto. Qed. #[global] Hint Resolve All_local_env_over_subst_instance : univ_subst. @@ -1997,10 +1998,10 @@ Section SubstIdentity. rewrite /subst_instance_list. now rewrite map_rev Hpars. * rewrite [subst_instance_constr _ _]subst_instance_two. noconf Hi. now rewrite [subst_instance _ u]H. - - solve_all. destruct a as [s [? ?]]; solve_all. - - clear X0. eapply nth_error_all in X as [s [Hs [IHs _]]]; eauto. - - solve_all. destruct a as [s [? ?]]. solve_all. - - clear X0. eapply nth_error_all in X as [s [Hs [IHs _]]]; eauto. + - solve_all. now destruct a as (s & ? & ? & ?). + - clear X0. now eapply nth_error_all in X as (s & e & Hs & IHs & _). + - solve_all. now destruct a as (s & ? & ? & ?). + - clear X0. now eapply nth_error_all in X as (s & e & Hs & IHs & _). Qed. Lemma typed_subst_abstract_instance Σ Γ t T : diff --git a/pcuic/theories/Conversion/PCUICWeakeningConv.v b/pcuic/theories/Conversion/PCUICWeakeningConv.v index 927d7c43f..57889ec69 100644 --- a/pcuic/theories/Conversion/PCUICWeakeningConv.v +++ b/pcuic/theories/Conversion/PCUICWeakeningConv.v @@ -339,7 +339,7 @@ Qed. Lemma isType_on_free_vars {cf} {Σ : global_env_ext} {wfΣ : wf Σ} {Γ T} : isType Σ Γ T -> on_free_vars xpredT T. Proof. - intros [s Hs]. + intros (s & e & Hs). eapply subject_closed in Hs. rewrite closedP_on_free_vars in Hs. eapply on_free_vars_impl; tea => //. @@ -348,7 +348,7 @@ Qed. Lemma isType_on_ctx_free_vars {cf} {Σ : global_env_ext} {wfΣ : wf Σ} {Γ T} : isType Σ Γ T -> on_ctx_free_vars xpredT Γ. Proof. - intros [s Hs]. + intros (s & e & Hs). eapply typing_wf_local in Hs. eapply closed_wf_local in Hs; tea. eapply (closed_ctx_on_free_vars xpredT) in Hs. diff --git a/pcuic/theories/PCUICAlpha.v b/pcuic/theories/PCUICAlpha.v index bc86190c2..b4739323d 100644 --- a/pcuic/theories/PCUICAlpha.v +++ b/pcuic/theories/PCUICAlpha.v @@ -42,14 +42,15 @@ Section Alpha. wf Σ.1 -> wf_local Σ Γ -> nth_error Γ i = Some (vass na ty) -> - lift_typing typing Σ Γ (lift0 (S i) ty) Sort. + lift_typing typing Σ Γ (lift0 (S i) ty) (Sort None). Proof using Type. intros Σ Γ i na ty hΣ hΓ e. simpl. induction i in Γ, hΓ, e |- *. - destruct Γ. 1: discriminate. simpl in e. apply some_inj in e. subst. inversion hΓ. subst. - apply infer_typing_sort_impl with id X0; intros Hs. + apply isType_of_isTypeRel in X0. + apply infer_typing_sort_impl with id X0 => //; intros Hs. change (tSort _) with (lift0 1 (tSort X0.π1)). eapply weakening with (Γ' := [ vass na ty ]). all: assumption. @@ -57,7 +58,7 @@ Section Alpha. simpl in e. specialize IHi with (2 := e). forward IHi by inversion hΓ; tas. - apply infer_typing_sort_impl with id IHi; intros Hs. + apply infer_typing_sort_impl with id IHi => //; intros Hs. change (tSort _) with (lift0 1 (lift0 (S i) (tSort IHi.π1))). rewrite simpl_lift0. eapply weakening with (Γ' := [ c ]). @@ -465,18 +466,18 @@ Section Alpha. now eapply eq_context_gen_eq_univ_subst_preserved. Qed. - Lemma ctx_inst_impl {Σ P Q Γ s Δ}: - (forall Σ Γ u U, P Σ Γ u U -> Q Σ Γ u U) -> - PCUICTyping.ctx_inst P Σ Γ s Δ -> PCUICTyping.ctx_inst Q Σ Γ s Δ. + Lemma ctx_inst_impl {P Q Γ s Δ}: + (forall Γ u U, P Γ u U -> Q Γ u U) -> + PCUICTyping.ctx_inst P Γ s Δ -> PCUICTyping.ctx_inst Q Γ s Δ. Proof using Type. intros HP. induction 1; constructor; auto. Qed. - Lemma ctx_inst_eq_context {Σ P Q Γ Γ' s Δ}: - (forall Σ Γ Γ' u U, (Γ ≡Γ Γ') -> P Σ Γ u U -> Q Σ Γ' u U) -> + Lemma ctx_inst_eq_context {P Q Γ Γ' s Δ}: + (forall Γ Γ' u U, (Γ ≡Γ Γ') -> P Γ u U -> Q Γ' u U) -> Γ ≡Γ Γ' -> - PCUICTyping.ctx_inst P Σ Γ s Δ -> PCUICTyping.ctx_inst Q Σ Γ' s Δ. + PCUICTyping.ctx_inst P Γ s Δ -> PCUICTyping.ctx_inst Q Γ' s Δ. Proof using Type. intros HP e. induction 1; constructor; eauto. @@ -491,7 +492,7 @@ Section Alpha. eapply All2_fold_All2 in eq. induction eq; depelim a; cbn; try solve [constructor; auto]; depelim r; subst; constructor; auto. - 1,2: apply infer_typing_sort_impl with id l; intros Hs. + 1,2: rewrite e; apply infer_typing_sort_impl with id l => //; intros Hs. 3: rename l0 into Hs. all: eapply (closed_context_cumulativity _ (pb:=Conv)); tea; [apply IHeq; pcuic|]. all: eapply ws_cumul_ctx_pb_rel_app. @@ -552,7 +553,7 @@ Section Alpha. wf_local Σ Δ -> isType Σ Δ T. Proof using Type. - intros Hty eq wfΔ; apply infer_typing_sort_impl with id Hty; intros Hs. + intros Hty eq wfΔ; apply infer_typing_sort_impl with id Hty => //; intros Hs. now eapply eq_context_conversion. Qed. @@ -570,10 +571,10 @@ Section Alpha. * intros Δ eq; destruct Δ; depelim eq; constructor. * intros Δ eq; depelim eq. depelim c. constructor; auto. - apply infer_typing_sort_impl with id tu; intros _. eapply Hs; auto. + rewrite -e. apply infer_typing_sort_impl with id tu => //; intros _. eapply Hs; auto. * intros Δ eq; depelim eq. depelim c. constructor; auto. - apply infer_typing_sort_impl with id tu; intros _. eapply Hs; auto. + rewrite -e. apply infer_typing_sort_impl with id tu => //; intros _. eapply Hs; auto. red. specialize (Hc _ _ e0 eq). specialize (Hs _ _ e1 eq). @@ -598,42 +599,47 @@ Section Alpha. specialize (ih _ eqctx). 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. + - intros na A B s1 s2 e' ih hA ihA hB ihB Δ v e eqctx; invs e. econstructor. + + rewrite <- H3. apply e'. + eapply ihA. assumption. eauto. + eapply context_conversion. * eapply ihB. assumption. reflexivity. * constructor; eauto. - simpl. eexists. eapply ihA; tea. + eexists. split. rewrite <- H3; apply e'. eapply ihA; tea. * constructor. -- now eapply eq_context_upto_empty_conv_context. -- constructor. assumption. constructor. eapply upto_names_impl_eq_term. assumption. - - intros na A t s1 B ih hA ihA hB ihB Δ v e e'; invs e. + - intros na A t s1 B e2 ih hA ihA hB ihB Δ v e e'; invs e. eapply type_Cumul'. + econstructor. + * rewrite <- H3; apply e2. * eapply ihA; tea. * eapply eq_context_conversion. -- eapply ihB. assumption. reflexivity. -- constructor. 1: assumption. simpl. constructor; auto. -- constructor; eauto. exists s1. + split. rewrite <- H3; apply e2. now eapply ihA. + eapply validity in hB;tea. eapply isType_tProd; eauto. split; eauto with pcuic. + exists s1; split => //. apply ihA => //. reflexivity. eapply validity. eapply ihB; eauto. constructor; auto. constructor ; auto. reflexivity. + apply eq_term_upto_univ_cumulSpec, eq_term_leq_term. symmetry. constructor; auto. all: try (eapply upto_names_impl_eq_term ; assumption). all: reflexivity. - - intros na b B t s1 A ih hB ihB hb ihb hA ihA Δ v e e'; invs e. + - intros na b B t s1 A e2 ih hB ihB hb ihb hA ihA Δ v e e'; invs e. specialize (ihB _ _ X0 e'). specialize (ihb _ _ X e'). assert (isType Σ Γ (tLetIn na b B A)). { eapply validity. econstructor; eauto. } eapply type_Cumul'. + econstructor. + * rewrite <- H4; apply e2. * eapply ihB; trea. * econstructor. -- eapply ihb; trea. @@ -646,12 +652,12 @@ Section Alpha. ++ assumption. ++ constructor; auto. -- constructor; auto. - ++ exists s1; eapply ihB; eauto. + ++ exists s1; split. rewrite <- H4; apply e2. eapply ihB; eauto. ++ eapply type_Cumul'; tea. - exists s1. eapply ihB; eauto. + exists s1. split; [cbnr|]. eapply ihB; eauto. eapply eq_term_upto_univ_cumulSpec, eq_term_leq_term. now apply upto_names_impl_eq_term. - + apply infer_typing_sort_impl with id X2; intros Hs. + + apply infer_typing_sort_impl with id X2 => //; intros Hs. eapply eq_context_conversion; tea. eauto. + eapply eq_term_upto_univ_cumulSpec, eq_term_leq_term. symmetry; constructor. assumption. @@ -667,7 +673,8 @@ Section Alpha. all:typeclasses eauto. * eapply ihu; trea. * eapply ihty. reflexivity. auto. - + apply infer_typing_sort_impl with id X1; intros Hs. eapply eq_context_conversion; tea. eauto. + + apply infer_typing_sort_impl with id X1 => //; intros Hs. + eapply eq_context_conversion; tea. eauto. + eapply eq_term_upto_univ_cumulSpec, eq_term_leq_term. symmetry. eapply upto_names_impl_eq_term. @@ -686,8 +693,8 @@ Section Alpha. wfcpc kelim HIHctxi Hc IHc iscof ptm wfbrs Hbrs Δ v e e'; invs e. have eqp := X1. eassert (ctx_inst _ _ _ _) as Hctxi by now eapply ctx_inst_impl with (2 := HIHctxi). - eassert (PCUICEnvTyping.ctx_inst _ _ _ _ _) as IHctxi. - { eapply ctx_inst_impl with (2 := HIHctxi). intros ? ? ? ? [? r]; exact r. } + eassert (PCUICEnvTyping.ctx_inst _ _ _ _) as IHctxi. + { eapply ctx_inst_impl with (2 := HIHctxi). intros ? ? ? [? r]; exact r. } destruct X1 as [eqpars [eqinst [eqctx eqret]]]. assert (wf_predicate mdecl idecl p'). { destruct wfp. split; auto. @@ -756,7 +763,7 @@ Section Alpha. econstructor; tea; eauto. 2,3: constructor; tea ; eauto. * eapply (type_ws_cumul_pb (pb:=Cumul)). eapply IHc; eauto. - eexists; eapply isType_mkApps_Ind; tea. + eexists; split; [cbnr|]; eapply isType_mkApps_Ind; tea. unshelve eapply (ctx_inst_spine_subst _ ctxinst'). eapply weaken_wf_local; tea. now eapply (on_minductive_wf_params_indices_inst isdecl). @@ -782,7 +789,7 @@ Section Alpha. 2:now eapply Forall2_All2 in wfbrs. epose proof (wf_case_branches_types' (p:=p') ps args brs' isdecl) as a. forward a. - { eexists; eapply isType_mkApps_Ind; tea. + { eexists; split; [cbnr|]; eapply isType_mkApps_Ind; tea. unshelve eapply (ctx_inst_spine_subst _ ctxinst'). eapply weaken_wf_local; tea. eapply (on_minductive_wf_params_indices_inst isdecl _ cu'). } @@ -837,14 +844,15 @@ Section Alpha. - intros mfix n decl types hguard hnth hwf ihmfix ihmfixb wffix Δ v e e'; invs e. eapply All2_nth_error_Some in hnth as hnth' ; eauto. destruct hnth' as [decl' [hnth' hh]]. - destruct hh as [[[ety eqann] ebo] era]. + destruct hh as [[[ety ebo] era] eqann]. assert (hwf' : wf_local Σ (Γ ,,, fix_context mfix')). { apply PCUICWeakeningTyp.All_mfix_wf; auto. eapply (All2_All_mix_left ihmfix) in X. clear -X. induction X; constructor; simpl; auto. - destruct r as [Hty [[[eqty eqann] eqbod] eqrarg]]. - apply infer_typing_sort_impl with id Hty; intros [Hs IH]. + destruct r as [Hty [[[eqty eqbod] eqrarg] eqann]]. + rewrite /on_def_type -eqann. + apply infer_typing_sort_impl with id Hty => //; intros [Hs IH]. apply IH; eauto. reflexivity. } assert (convctx : conv_context cumulAlgo_gen Σ (Γ ,,, fix_context mfix) (Γ ,,, fix_context mfix')). { eapply eq_context_upto_univ_conv_context. @@ -883,8 +891,9 @@ Section Alpha. * eapply (All2_All_mix_left ihmfix) in X. clear -X. induction X; constructor; simpl; auto. - destruct r as [Hty [[[eqty eqann] eqbod] eqrarg]]. - apply infer_typing_sort_impl with id Hty; intros [Hs IH]. + destruct r as [Hty [[[eqty eqbod] eqrarg] eqann]]. + rewrite /on_def_type -eqann. + apply infer_typing_sort_impl with id Hty => //; intros [Hs IH]. apply IH; eauto. reflexivity. * solve_all. destruct b0 as [Hb IHb]. @@ -892,7 +901,8 @@ Section Alpha. forward IHb. { apply eq_context_upto_cat; reflexivity. } eapply context_conversion; eauto. eapply (type_Cumul' (lift0 #|fix_context mfix| (dtype x))); auto. - apply infer_typing_sort_impl with id a1; intros [Hs IH]. + eapply isType_of_isTypeRel. + apply infer_typing_sort_impl with id a1 => //; intros [Hs IH]. specialize (IH _ _ a0 e'). rewrite <-H. eapply (weakening _ _ _ _ (tSort _)); eauto. @@ -914,22 +924,24 @@ Section Alpha. move: ho; enough (map check_one_fix mfix = map check_one_fix mfix') as ->; auto. apply upto_names_check_fix. solve_all. + eapply All_nth_error in ihmfix; tea. - now apply infer_typing_sort_impl with id ihmfix; intros []. + eapply isType_of_isTypeRel. + now apply infer_typing_sort_impl with id ihmfix => //; intros []. + apply eq_term_upto_univ_cumulSpec, eq_term_leq_term, upto_names_impl_eq_term. now symmetry. - intros mfix n decl types hguard hnth hwf ihmfix ihmfixb wffix Δ v e e'; invs e. eapply All2_nth_error_Some in hnth as hnth' ; eauto. destruct hnth' as [decl' [hnth' hh]]. - destruct hh as [[[ety eqann] ebo] era]. + destruct hh as [[[ety ebo] era] eqann]. assert (hwf' : wf_local Σ (Γ ,,, fix_context mfix')). { apply PCUICWeakeningTyp.All_mfix_wf; auto. eapply (All2_All_mix_left ihmfix) in X. clear -X. induction X; constructor; simpl; auto. - destruct r as [Hty [[[eqty eqann] eqbod] eqrarg]]. - apply infer_typing_sort_impl with id Hty; intros [Hs IH]. - apply IH; eauto. reflexivity. } + destruct r as [Hty [[[eqty eqbod] eqrarg] eqann]]. + rewrite /on_def_type -eqann. + apply infer_typing_sort_impl with id Hty => //; intros [Hs IH]. + apply IH; eauto. reflexivity. } assert (convctx : conv_context cumulAlgo_gen Σ (Γ ,,, fix_context mfix) (Γ ,,, fix_context mfix')). { eapply eq_context_upto_univ_conv_context. eapply (eq_context_impl _ eq). intros x y eqx. subst. reflexivity. @@ -967,16 +979,18 @@ Section Alpha. * eapply (All2_All_mix_left ihmfix) in X. clear -X. induction X; constructor; simpl; auto. - destruct r as [Hty [[[eqty eqann] eqbod] eqrarg]]. - apply infer_typing_sort_impl with id Hty; intros [Hs IH]. - apply IH; eauto. reflexivity. + destruct r as [Hty [[[eqty eqbod] eqrarg] eqann]]. + rewrite /on_def_type -eqann. + apply infer_typing_sort_impl with id Hty => //; intros [Hs IH]. + apply IH; eauto. reflexivity. * solve_all. destruct b0 as [Hb IHb]. specialize (IHb (Γ ,,, types) _ b2). forward IHb. { apply eq_context_upto_cat; reflexivity. } eapply context_conversion; eauto. eapply (type_Cumul' (lift0 #|fix_context mfix| (dtype x))); auto. - apply infer_typing_sort_impl with id a1; intros [Hs IH]. + eapply isType_of_isTypeRel. + apply infer_typing_sort_impl with id a1 => //; intros [Hs IH]. specialize (IH _ _ a0 e'). rewrite <-H. eapply (weakening _ _ _ _ (tSort _)); eauto. @@ -991,14 +1005,15 @@ Section Alpha. enough (map check_one_cofix mfix = map check_one_cofix mfix') as ->; auto. apply upto_names_check_cofix. solve_all. + eapply All_nth_error in ihmfix; tea. - now apply infer_typing_sort_impl with id ihmfix; intros []. + eapply isType_of_isTypeRel. + now apply infer_typing_sort_impl with id ihmfix => //; intros []. + apply eq_term_upto_univ_cumulSpec, eq_term_leq_term, upto_names_impl_eq_term. now symmetry. - intros t A B X wf ht iht har ihar hcu Δ v e e'. eapply (type_ws_cumul_pb (pb:=Cumul)). + eapply iht; tea. - + exists X; eauto. + + exists X; split; [cbnr|]; eauto. specialize (wf _ e'). now eapply eq_context_conversion. + eapply (ws_cumul_pb_ws_cumul_ctx (pb':=Conv)); tea. 2:eapply PCUICInversion.into_ws_cumul; tea. @@ -1105,7 +1120,7 @@ Section Alpha. isType Σ Γ v. Proof using Type. intros Hty eq. - apply infer_typing_sort_impl with id Hty; intros Hs. + apply infer_typing_sort_impl with id Hty => //; intros Hs. eapply typing_alpha; eauto. Qed. @@ -1116,7 +1131,7 @@ Section Alpha. isType Σ Δ v. Proof using Type. intros Hty eqctx eqtm. - apply infer_typing_sort_impl with id Hty; intros Hs. + apply infer_typing_sort_impl with id Hty => //; intros Hs. eapply typing_alpha_prop; eauto. Qed. diff --git a/pcuic/theories/PCUICArities.v b/pcuic/theories/PCUICArities.v index c78025c9c..00aa92213 100644 --- a/pcuic/theories/PCUICArities.v +++ b/pcuic/theories/PCUICArities.v @@ -25,7 +25,8 @@ Lemma isType_Sort {cf:checker_flags} {Σ Γ s} : isType Σ Γ (tSort s). Proof. intros wfs wfΓ. - eexists; econstructor; eauto. + eexists; split; cbnr. + econstructor; eauto. Qed. #[global] Hint Resolve isType_Sort : pcuic. @@ -77,7 +78,7 @@ Proof. (Σ, ind_universes mdecl) ;;; [] |- tInd {| inductive_mind := inductive_mind ind; inductive_ind := i |} u : (ind_type x)) 0 (ind_bodies mdecl)). { apply forall_nth_error_Alli. intros. eapply Alli_nth_error in oind; eauto. simpl in oind. - destruct oind. destruct onArity as [s Hs]. + apply onArity in oind as (s & e & Hs). eapply type_Cumul; eauto. econstructor; eauto. split; eauto with pcuic. eapply consistent_instance_ext_abstract_instance; eauto. @@ -205,17 +206,19 @@ 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]. - apply inversion_Prod in H; tas. destruct H as [s1 [s2 [HA [HB Hs]]]]. + - destruct HH as (s & e & H). + apply inversion_Prod in H; tas. destruct H as (s1 & s2 & e' & HA & HB & Hs). split. - * eexists; tea. - * eexists; tea. + * eexists; split; tea. + * eexists; split; [|eassumption]; auto. - destruct HH as [HA HB]. - destruct HA as [sA HA], HB as [sB HB]. - eexists. econstructor; eassumption. + destruct HA as (sA & e & HA), HB as (sB & e' & HB). + exists (Universe.sort_of_product sA sB). + split; cbnr. + econstructor; eassumption. Defined. Lemma isType_subst {Γ Δ A} s : @@ -224,7 +227,7 @@ Section WfEnv. isType Σ Γ (subst0 s A). Proof using wfΣ. intros sub HT. - apply infer_typing_sort_impl with id HT; intros Hs. + apply infer_typing_sort_impl with id HT => //; intros Hs. have wf := typing_wf_local Hs. now eapply (substitution (Δ := []) (T := tSort _)). Qed. @@ -235,7 +238,7 @@ Section WfEnv. isType Σ (Γ ,,, subst_context s 0 Δ') (subst s #|Δ'| A). Proof using wfΣ. intros sub HT. - apply infer_typing_sort_impl with id HT; intros Hs. + apply infer_typing_sort_impl with id HT => //; intros Hs. now eapply (substitution (T:=tSort _)). Qed. @@ -256,13 +259,14 @@ Section WfEnv. Lemma isType_tLetIn_red {Γ} (HΓ : wf_local Σ Γ) {na t A B} : isType Σ Γ (tLetIn na t A B) -> isType Σ Γ (B {0:=t}). Proof using wfΣ. - intro HH. - apply infer_typing_sort_impl with id HH; intros H. - assert (Hs := typing_wf_universe _ H). - apply inversion_LetIn in H; tas. destruct H as (s1 & A' & HA & Ht & HB & H). + intro HH. + apply infer_typing_sort_impl with id HH => //; intros Hs. + assert (wf := typing_wf_universe _ Hs). + apply inversion_LetIn in Hs; tas. + destruct Hs as (s1 & A' & e' & HA & Ht & HB & H). eapply (type_ws_cumul_pb (pb:=Cumul)) with (A' {0 := t}). eapply substitution_let in HB; eauto. - * econstructor; eauto with pcuic. econstructor; eauto. - * eapply ws_cumul_pb_Sort_r_inv in H as [s' [H H']]. + * now apply isType_Sort. + * eapply ws_cumul_pb_Sort_r_inv in H as (s' & H & H'). transitivity (tSort s'); eauto. eapply red_ws_cumul_pb. apply invert_red_letin in H as [H|H] => //. @@ -274,13 +278,14 @@ Section WfEnv. Lemma isType_tLetIn_dom {Γ} (HΓ : wf_local Σ Γ) {na t A B} : isType Σ Γ (tLetIn na t A B) -> Σ ;;; Γ |- t : A. Proof using wfΣ. - intros (s & H). - apply inversion_LetIn in H; tas. now destruct H as (s1 & A' & HA & Ht & HB & H). + intros (? & _ & H). + apply inversion_LetIn in H; tas. + now destruct H as (s1 & A' & e' & HA & Ht & HB & H). Qed. 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. @@ -288,7 +293,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. @@ -383,7 +388,8 @@ Section WfEnv. induction 1; auto. Qed. - Lemma type_mkProd_or_LetIn {Γ} d {u t s} : + Lemma type_mkProd_or_LetIn {Γ} d {u t s} : + isSortRel u d.(decl_name).(binder_relevance) -> Σ ;;; Γ |- decl_type d : tSort u -> Σ ;;; Γ ,, d |- t : tSort s -> match decl_body d return Type with @@ -391,10 +397,11 @@ Section WfEnv. | None => Σ ;;; Γ |- mkProd_or_LetIn d t : tSort (Universe.sort_of_product u s) end. Proof using wfΣ. - destruct d as [na [b|] dty] => [Hd Ht|Hd Ht]; rewrite /mkProd_or_LetIn /=. + destruct d as [na [b|] dty] => [e Hd Ht|e Hd Ht]; rewrite /mkProd_or_LetIn /=. - have wf := typing_wf_local Ht. depelim wf. clear l. - eapply type_Cumul. econstructor; eauto. + eapply type_Cumul. + econstructor; eauto. econstructor; eauto. now eapply typing_wf_universe in Ht; pcuic. eapply convSpec_cumulSpec, red1_cumulSpec. constructor. - have wf := typing_wf_local Ht. @@ -422,7 +429,8 @@ Section WfEnv. destruct a as [na [b|] ty]; intuition auto. { 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. + + exact X0.π2.1. + + exact X0.π2.2. + eapply type_Cumul; eauto. econstructor; eauto with pcuic. eapply cumul_Sort. eapply leq_universe_product. } @@ -458,11 +466,12 @@ Section WfEnv. induction Γ'; simpl; auto; move=> Γ us s t equ Ht. - destruct us => //. - destruct a as [na [b|] ty]; intuition auto. - * destruct a0 as [s' Hs]. + * destruct a0 as (s' & e & Hs). eapply IHΓ'; eauto. eapply (type_mkProd_or_LetIn {| decl_body := Some b |}); auto. + apply e. simpl. exact Hs. - * destruct us => //. destruct equ. + * destruct us => //. destruct equ as (s0 & e & p). simpl. eapply IHΓ'; eauto. apply (type_mkProd_or_LetIn {| decl_body := None |}) => /=; eauto. @@ -586,7 +595,7 @@ Section WfEnv. isType Σ Γ (subst0 s T). Proof using wfΣ. intros sub HT. - apply infer_typing_sort_impl with id HT; intros Hs. + apply infer_typing_sort_impl with id HT => //; intros Hs. destruct HT as (s' & t); cbn in Hs |- *; clear t. revert Γ s sub Hs. generalize (le_n #|Δ|). @@ -606,7 +615,7 @@ Section WfEnv. intros Hs. assert (wfs' := typing_wf_universe wfΣ Hs). - eapply inversion_LetIn in Hs as (? & ? & ? & ? & ? & ?); auto. + eapply inversion_LetIn in Hs as (? & ? & ? & ? & ? & ? & ?); auto. eapply substitution_let in t1; auto. eapply ws_cumul_pb_LetIn_l_inv in w; auto. pose proof (subslet_app_inv sub) as [subl subr]. @@ -622,7 +631,7 @@ Section WfEnv. intros Hs. assert (wfs' := typing_wf_universe wfΣ Hs). - eapply inversion_Prod in Hs as (? & ? & ? & ? & ?); auto. + eapply inversion_Prod in Hs as (? & ? & ? & ? & ? & ?); auto. pose proof (subslet_app_inv sub) as [subl subr]. depelim subl; depelim subl. rewrite subst_empty in t2. rewrite H0 in subr. epose proof (substitution0 t0 t2). @@ -658,14 +667,14 @@ Section WfEnv. + simpl. intros. now eapply typing_wf_local in X. + rewrite it_mkProd_or_LetIn_app. destruct x as [na [b|] ty]; cbn; move=> H. - * apply inversion_LetIn in H as (s1 & A & H0 & H1 & H2 & H3); auto. + * apply inversion_LetIn in H as (s1 & A & e & H0 & H1 & H2 & H3); auto. eapply All_local_env_app; split; pcuic. eapply All_local_env_app. split. repeat constructor. now exists s1. auto. apply IHΔ in H2. eapply All_local_env_app_inv in H2. intuition auto. eapply All_local_env_impl; eauto. simpl. intros. now rewrite app_context_assoc. - * apply inversion_Prod in H as (s1 & A & H0 & H1 & H2); auto. + * apply inversion_Prod in H as (s1 & A & e & H0 & H1 & H2); auto. eapply All_local_env_app; split; pcuic. eapply All_local_env_app. split. repeat constructor. now exists s1. apply IHΔ in H1. @@ -677,7 +686,7 @@ Section WfEnv. Lemma isType_it_mkProd_or_LetIn_wf_local {Γ Δ T} : isType Σ Γ (it_mkProd_or_LetIn Δ T) -> wf_local Σ (Γ ,,, Δ). Proof using wfΣ. - move=> [s Hs]. + intros (? & _ & Hs). now eapply it_mkProd_or_LetIn_wf_local in Hs. Qed. @@ -687,7 +696,7 @@ Section WfEnv. isType Σ Γ T. Proof using wfΣ. intros wfΓ HT. - apply infer_typing_sort_impl with id HT; intros hs. + apply infer_typing_sort_impl with id HT => //; intros hs. unshelve epose proof (subject_closed hs); eauto. eapply (weakening _ _ Γ) in hs => //. rewrite lift_closed in hs => //. diff --git a/pcuic/theories/PCUICCanonicity.v b/pcuic/theories/PCUICCanonicity.v index 56b1e9ace..99fada1fd 100644 --- a/pcuic/theories/PCUICCanonicity.v +++ b/pcuic/theories/PCUICCanonicity.v @@ -309,6 +309,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 isType_of_isTypeRel. now eapply isType_tProd in i as []. Qed. @@ -338,6 +339,7 @@ Qed. now len. eapply isType_apply in i; tea. eapply (type_ws_cumul_pb (pb:=Conv)); tea. 2:now symmetry. + eapply isType_of_isTypeRel. now eapply isType_tProd in i as []. Qed. @@ -389,6 +391,7 @@ Qed. 2:{ eapply ws_cumul_pb_eq_le. symmetry. now eapply red_ws_cumul_pb. } eapply ws_cumul_pb_Prod; eauto. + eapply ws_cumul_pb_Prod_Prod_inv in w as [eqna dom codom]. + apply isType_of_isTypeRel in l. assert (Σ ;;; Γ |- hd : ty). { eapply type_ws_cumul_pb; pcuic. eapply ws_cumul_pb_eq_le. now symmetry. } eapply (substitution0_ws_cumul_pb (t:=hd)) in codom; eauto. @@ -481,7 +484,7 @@ Qed. - eapply typing_spine_nth_error in sp; eauto; cycle 1. * eapply smash_context_assumption_context; constructor. * eapply wf_local_smash_end; eauto. - destruct isty as [s Hs]. + destruct isty as (s & e & Hs). eapply inversion_it_mkProd_or_LetIn in Hs; eauto. now eapply typing_wf_local. * destruct (decompose_app (decl_type c)) as [hd tl] eqn:da. @@ -494,7 +497,7 @@ Qed. - eapply typing_spine_nth_error_None_prod in sp; eauto. * eapply smash_context_assumption_context; constructor. * eapply wf_local_smash_end; eauto. - destruct isty as [s Hs]. + destruct isty as (s & e & Hs). eapply inversion_it_mkProd_or_LetIn in Hs; eauto. now eapply typing_wf_local. Qed. @@ -542,7 +545,7 @@ Qed. * len in sp. * eapply smash_context_assumption_context; constructor. * eapply wf_local_smash_end; eauto. - destruct isty as [s Hs]. + destruct isty as (s & e & Hs). eapply inversion_it_mkProd_or_LetIn in Hs; eauto. now eapply typing_wf_local. * len; simpl. eapply nth_error_None in hargs => //. @@ -812,7 +815,7 @@ Section WeakNormalization. eapply inversion_mkApps in tyarg as [A [Hcof sp]]; auto. eapply inversion_CoFix in Hcof as (? & ? & ? & ? & ? & ? & ?); auto. eapply nth_error_all in a; eauto. - simpl in a. + pose proof (isType_of_isTypeRel 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. @@ -882,8 +885,9 @@ Section WeakNormalization. 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. + pose proof (isType_of_isTypeRel 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. @@ -891,7 +895,6 @@ Section WeakNormalization. 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. Fixpoint axiom_free_value Σ args t := @@ -944,8 +947,9 @@ Section WeakNormalization. - 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. + pose proof (isType_of_isTypeRel 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]]]]. @@ -957,7 +961,6 @@ Section WeakNormalization. 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. @@ -981,9 +984,9 @@ Section WeakNormalization. - elimtype False. eauto using wh_neutral_empty. - eapply inversion_Sort in typed as (? & ? & e); auto. now eapply invert_cumul_sort_ind in e. - - eapply inversion_Prod in typed as (? & ? & ? & ? & e); auto. + - eapply inversion_Prod in typed as (? & ? & ? & ? & ? & e); auto. now eapply invert_cumul_sort_ind in e. - - eapply inversion_Lambda in typed as (? & ? & ? & ? & e); auto. + - eapply inversion_Lambda in typed as (? & ? & ? & ? & ? & e); auto. now eapply invert_cumul_prod_ind in e. - now rewrite head_mkApps /= /head /=. - exfalso; eapply invert_ind_ind; eauto. @@ -1021,7 +1024,8 @@ Section WeakNormalization. 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. + pose proof (isType_of_isTypeRel a). eapply typing_spine_strengthen in t0; eauto. rewrite /unfold_fix in unf. rewrite e in unf. noconf unf. @@ -1084,17 +1088,17 @@ Section WeakNormalization. pose proof (subject_closed t1); auto. eapply eval_closed in H; eauto. eapply subject_reduction in IHHe1. 2-3:eauto. - eapply inversion_Lambda in IHHe1 as (? & ? & ? & ? & e); eauto. + eapply inversion_Lambda in IHHe1 as (? & ? & ? & ? & ? & e); eauto. eapply (substitution0 (Γ := [])) in t3; eauto. eapply IHHe3. rewrite (closed_subst a' 0 b); auto. rewrite /subst1 in t3. eapply t3. eapply (type_ws_cumul_pb (pb:=Conv)); eauto. - eapply subject_reduction in IHHe2; eauto. now eexists. + eapply subject_reduction in IHHe2; eauto. now eexists; split. eapply ws_cumul_pb_Prod_Prod_inv in e as [eqna dom codom]; eauto. now symmetry. - - eapply inversion_LetIn in Ht as (? & ? & ? & ? & ? & ?); auto. + - eapply inversion_LetIn in Ht as (? & ? & ? & ? & ? & ? & ?); auto. redt (tLetIn na b0' t b1); eauto. eapply red_letin; eauto. redt (b1 {0 := b0'}); auto. diff --git a/pcuic/theories/PCUICContextConversion.v b/pcuic/theories/PCUICContextConversion.v index 509afdbf4..f0835e127 100644 --- a/pcuic/theories/PCUICContextConversion.v +++ b/pcuic/theories/PCUICContextConversion.v @@ -1161,8 +1161,8 @@ Qed. end. *) Definition on_local_decl (P : context -> term -> typ_or_sort -> Type) (Γ : context) (d : context_decl) := match decl_body d with - | Some b => P Γ b (Typ (decl_type d)) * P Γ (decl_type d) Sort - | None => P Γ (decl_type d) Sort + | Some b => P Γ b (Typ (decl_type d)) × P Γ (decl_type d) (SortRel d.(decl_name).(binder_relevance)) + | None => P Γ (decl_type d) (SortRel d.(decl_name).(binder_relevance)) end. Lemma nth_error_All_local_env {P Γ n} (isdecl : n < #|Γ|) : diff --git a/pcuic/theories/PCUICContexts.v b/pcuic/theories/PCUICContexts.v index 62639a82a..747c76a1e 100644 --- a/pcuic/theories/PCUICContexts.v +++ b/pcuic/theories/PCUICContexts.v @@ -161,18 +161,19 @@ Lemma type_local_ctx_instantiate {cf:checker_flags} Σ ind mdecl Γ Δ u s : type_local_ctx (lift_typing typing) Σ (subst_instance u Γ) (subst_instance u Δ) (subst_instance_univ u s). Proof. intros Hctx Hu. - induction Δ; simpl in *; intuition auto. + induction Δ; intuition auto. { destruct Σ as [Σ univs]. eapply (wf_universe_subst_instance (Σ, ind_universes mdecl)); eauto. } - destruct a as [na [b|] ty]; simpl; [destruct X as (Hwfctx & Ht & Hb) | destruct X as (Hwfctx & Ht)]; repeat split. + destruct a as [na [b|] ty]; destruct X as (Hwfctx & Ht & Hb); repeat split. - now apply IHΔ. - - eapply infer_typing_sort_impl with _ Ht; intros Hs. + - eapply infer_typing_sort_impl with _ Ht; [apply relevance_subst_opt|]; intros Hs. eapply instantiate_minductive in Hs; eauto. now rewrite subst_instance_app in Hs. - eapply instantiate_minductive in Hb; eauto. now rewrite subst_instance_app in Hb. - now apply IHΔ. - - eapply instantiate_minductive in Ht; eauto. - now rewrite subst_instance_app in Ht. + - now apply relevance_subst. + - eapply instantiate_minductive in Hb; eauto. + now rewrite subst_instance_app in Hb. Qed. Lemma sorts_local_ctx_instantiate {cf:checker_flags} Σ ind mdecl Γ Δ u s : @@ -184,17 +185,20 @@ Lemma sorts_local_ctx_instantiate {cf:checker_flags} Σ ind mdecl Γ Δ u s : (List.map (subst_instance_univ u) s). Proof. intros Hctx Hu. - induction Δ in s |- *; simpl in *; intuition auto. - destruct s; simpl; intuition eauto. - destruct a as [na [b|] ty]; simpl. intuition eauto. - - eapply infer_typing_sort_impl with _ a0; intros Hs. + induction Δ in s |- *; intros. + simpl in *. now destruct s. + destruct a as [na [b|] ty]; [|destruct s; simpl; intuition eauto]; + destruct X as (Hwfctx & Ht & Hb); repeat split. + - now apply IHΔ. + -eapply infer_typing_sort_impl with _ Ht; [apply relevance_subst_opt|]; intros Hs. eapply instantiate_minductive in Hs; eauto. now rewrite subst_instance_app in Hs. - - eapply instantiate_minductive in b1; eauto. - now rewrite subst_instance_app in b1. - - destruct s; simpl; intuition eauto. - eapply instantiate_minductive in b; eauto. - now rewrite subst_instance_app in b. + - eapply instantiate_minductive in Hb; eauto. + now rewrite subst_instance_app in Hb. + - now apply IHΔ. + - now apply relevance_subst. + - eapply instantiate_minductive in Hb; eauto. + now rewrite subst_instance_app in Hb. Qed. Lemma subst_type_local_ctx {cf:checker_flags} Σ Γ Δ Δ' s ctxs : @@ -204,16 +208,17 @@ Lemma subst_type_local_ctx {cf:checker_flags} Σ Γ Δ Δ' s ctxs : subslet Σ Γ s Δ -> type_local_ctx (lift_typing typing) Σ Γ (subst_context s 0 Δ') ctxs. Proof. - induction Δ'; simpl; auto. - destruct a as [na [b|] ty]; simpl; intuition auto. - + simpl; rewrite subst_context_snoc /= /subst_decl /map_decl /= Nat.add_0_r; - repeat split; tas. - - apply infer_typing_sort_impl with id a0; intros Hs. - now eapply substitution in Hs. - - now eapply substitution in b1. - + rewrite subst_context_snoc /= /subst_decl /map_decl /= Nat.add_0_r. - intuition auto. - eapply substitution in b; eauto. + induction Δ'; auto. + intros wfΣ wfΓΔ X X0. + destruct a as [na [b|] ty]; destruct X as (Hwfctx & Ht & Hb); + rewrite subst_context_snoc /subst_decl /map_decl Nat.add_0_r; repeat split. + - now apply IHΔ'. + - eapply infer_typing_sort_impl with id Ht => //; intros Hs. + eapply substitution in Hs; eauto. + - eapply substitution in Hb; eauto. + - now apply IHΔ'. + - assumption. + - eapply substitution in Hb; eauto. Qed. Lemma subst_sorts_local_ctx {cf:checker_flags} Σ Γ Δ Δ' s ctxs : @@ -223,19 +228,21 @@ Lemma subst_sorts_local_ctx {cf:checker_flags} Σ Γ Δ Δ' s ctxs : subslet Σ Γ s Δ -> sorts_local_ctx (lift_typing typing) Σ Γ (subst_context s 0 Δ') ctxs. Proof. - induction Δ' in ctxs |- *; simpl; auto. - destruct a as [na [b|] ty]; simpl; intuition auto. - + simpl; rewrite subst_context_snoc /= /subst_decl /map_decl /= Nat.add_0_r. + induction Δ' in ctxs |- *; auto. + intros wfΣ wfΓΔ X X0. + destruct a as [na [b|] ty]; + rewrite subst_context_snoc /subst_decl /map_decl Nat.add_0_r; + [|destruct ctxs; simpl; intuition eauto]; + destruct X as (Hwfctx & Ht & Hb); repeat split. - - now apply IHΔ'. - - apply infer_typing_sort_impl with id a0; intros Hs. - now eapply substitution in Hs. - - now eapply substitution in b1. - + rewrite subst_context_snoc /= /subst_decl /map_decl /= Nat.add_0_r. - destruct ctxs; auto. - intuition auto. - eapply substitution in b; eauto. -Qed. + - now apply IHΔ'. + - apply infer_typing_sort_impl with id Ht => //; intros Hs. + eapply substitution in Hs; eauto. + - eapply substitution in Hb; eauto. + - now apply IHΔ'. + - assumption. + - eapply substitution in Hb; eauto. + Qed. Lemma weaken_type_local_ctx {cf:checker_flags} Σ Γ Γ' Δ ctxs : wf Σ.1 -> @@ -243,11 +250,19 @@ Lemma weaken_type_local_ctx {cf:checker_flags} Σ Γ Γ' Δ ctxs : type_local_ctx (lift_typing typing) Σ Γ' Δ ctxs -> type_local_ctx (lift_typing typing) Σ (Γ ,,, Γ') Δ ctxs. Proof. - induction Δ; simpl; auto. - destruct a as [na [b|] ty]; simpl; intuition auto. - 1: apply infer_typing_sort_impl with id a0; intros Hs. - all: rewrite -app_context_assoc. - all: eapply (weaken_ctx Γ); auto. + induction Δ; auto. + intros wfΣ wfΓΔ X. + destruct a as [na [b|] ty]; destruct X as (Hwfctx & Ht & Hb); repeat split. + - now apply IHΔ. + - apply infer_typing_sort_impl with id Ht => //; intros Hs. + rewrite -app_context_assoc. + eapply (weaken_ctx Γ); auto. + - rewrite -app_context_assoc. + eapply (weaken_ctx Γ); auto. + - now apply IHΔ. + - assumption. + - rewrite -app_context_assoc. + eapply (weaken_ctx Γ); auto. Qed. Lemma weaken_sorts_local_ctx {cf:checker_flags} Σ Γ Γ' Δ ctxs : @@ -256,13 +271,20 @@ Lemma weaken_sorts_local_ctx {cf:checker_flags} Σ Γ Γ' Δ ctxs : sorts_local_ctx (lift_typing typing) Σ Γ' Δ ctxs -> sorts_local_ctx (lift_typing typing) Σ (Γ ,,, Γ') Δ ctxs. Proof. - induction Δ in ctxs |- *; simpl; auto. - destruct a as [na [b|] ty]; simpl; intuition auto. - 1: apply infer_typing_sort_impl with id a0; intros Hs. - all: rewrite -app_context_assoc. - 3: destruct ctxs; auto. - 3: destruct X1; split; auto. - all: now eapply (weaken_ctx Γ). + induction Δ in ctxs |- *; auto. + intros wfΣ wfΓΔ X. + destruct a as [na [b|] ty]; [|destruct ctxs; simpl; intuition eauto]; + destruct X as (Hwfctx & Ht & Hb); repeat split. + - now apply IHΔ. + - apply infer_typing_sort_impl with id Ht => //; intros Hs. + rewrite -app_context_assoc. + eapply (weaken_ctx Γ); auto. + - rewrite -app_context_assoc. + eapply (weaken_ctx Γ); auto. + - now apply IHΔ. + - assumption. + - rewrite -app_context_assoc. + eapply (weaken_ctx Γ); auto. Qed. Lemma reln_app acc Γ Δ k : reln acc k (Γ ++ Δ) = @@ -364,7 +386,7 @@ Section WfEnv. now rewrite app_context_assoc. - apply IHΓ. auto. eapply All_local_env_subst; eauto. simpl; intros. destruct T; simpl in *; pcuicfo auto. - 2: apply infer_typing_sort_impl with id X; intros Hs. + 2: apply infer_typing_sort_impl with id X => //; intros Hs. 1: rename X into Hs. all: rewrite Nat.add_0_r. all: eapply (substitution (Γ':=[vdef na b t]) (s := [b])) in Hs; eauto. @@ -674,8 +696,8 @@ Proof. rewrite -(lift_context_lift_context 1 _). eapply (subslet_lift _ [_]); eauto. constructor. - { pose proof l.π2. eapply wf_local_smash_end; pcuic. } - apply infer_typing_sort_impl with id l; intros Hs. + { pose proof l.π2.2. eapply wf_local_smash_end; pcuic. } + apply infer_typing_sort_impl with id l => //; intros Hs. eapply (weakening_typing (Γ'' := smash_context [] Δ)) in Hs. len in Hs. simpl in Hs. simpl. 2:{ eapply wf_local_smash_end; pcuic. } @@ -683,7 +705,7 @@ Proof. - eapply meta_conv. econstructor. constructor. apply wf_local_smash_end; auto. eapply wf_local_app; eauto. - apply infer_typing_sort_impl with id l; intros Hs. + apply infer_typing_sort_impl with id l => //; intros Hs. eapply (weakening_typing (Γ'' := smash_context [] Δ)) in Hs. len in Hs. simpl in Hs. simpl. 2:{ eapply wf_local_smash_end; pcuic. } diff --git a/pcuic/theories/PCUICConversion.v b/pcuic/theories/PCUICConversion.v index ba2da71e6..e286bf69e 100644 --- a/pcuic/theories/PCUICConversion.v +++ b/pcuic/theories/PCUICConversion.v @@ -3134,7 +3134,7 @@ Proof. * cbn; constructor. + apply IHΔ => //. + destruct d as [na [b|] ty]; constructor; cbn in *; auto; try congruence. - destruct l as [s Hs]. + destruct l as (s & e & Hs). constructor. 1:eauto with fvs. { now eapply subject_closed in Hs; rewrite is_open_term_closed in Hs. } { erewrite on_free_vars_subst_instance. @@ -3144,14 +3144,14 @@ Proof. * cbn; constructor. + apply IHΔ => //. + destruct d as [na [b'|] ty]; constructor; cbn in *; auto; try congruence; noconf H. - { destruct l as [s Hs]. + { destruct l as (s & e & Hs). constructor. 1:eauto with fvs. { now eapply subject_closed in l0; rewrite is_open_term_closed in l0. } { erewrite on_free_vars_subst_instance. eapply subject_closed in l0; rewrite is_open_term_closed in l0. now rewrite on_free_vars_subst_instance in l0. } apply eq_term_upto_univ_subst_instance; try typeclasses eauto. auto. } - { destruct l as [s Hs]. + { destruct l as (s & e & Hs). constructor. 1:eauto with fvs. { now eapply subject_closed in Hs; rewrite is_open_term_closed in Hs. } { erewrite on_free_vars_subst_instance. diff --git a/pcuic/theories/PCUICCumulProp.v b/pcuic/theories/PCUICCumulProp.v index e2ceb2c6b..b02adeff5 100644 --- a/pcuic/theories/PCUICCumulProp.v +++ b/pcuic/theories/PCUICCumulProp.v @@ -1132,14 +1132,14 @@ Proof using Hcf Hcf'. split; split; intros H'. 1,2:now eapply is_prop_superE in H'. 1,2:now eapply is_sprop_superE in H'. - - eapply inversion_Prod in X4 as [s1' [s2' [Ha [Hb Hs]]]]; auto. - specialize (X1 _ _ H Ha). + - eapply inversion_Prod in X4 as (s1' & s2' & e' & Ha & Hb & Hs); auto. + specialize (X1 _ _ H0 Ha). specialize (X1 _ (eq_term_upto_univ_napp_leq X5_1)). eapply context_conversion in Hb. 3:{ constructor. apply conv_ctx_refl. constructor. eassumption. constructor. eauto. } 2:{ constructor; eauto. now exists s1. } - specialize (X3 _ _ H Hb _ X5_2). + specialize (X3 _ _ H0 Hb _ X5_2). eapply cumul_cumul_prop in Hs => //. eapply cumul_prop_trans; eauto. constructor; fvs. constructor. @@ -1149,22 +1149,22 @@ Proof using Hcf Hcf'. * split; intros Hs'; apply is_sprop_sort_prod in Hs'; eapply is_sprop_prod; eapply cumul_sprop_props; eauto. now eapply cumul_prop_sym; eauto. - - eapply inversion_Lambda in X4 as (s & B & dom & bod & cum). - specialize (X1 _ _ H dom _ (eq_term_upto_univ_napp_leq X5_1)). - specialize (X3 t0 B H). - assert(conv_context cumulAlgo_gen Σ (Γ ,, vass na ty) (Γ ,, vass n t)). + - eapply inversion_Lambda in X4 as (s & B & e' & dom & bod & cum). + specialize (X1 _ _ H0 dom _ (eq_term_upto_univ_napp_leq X5_1)). + specialize (X3 t0 B H0). + assert(conv_context cumulAlgo_gen Σ (Γ ,, vass na0 ty) (Γ ,, vass na t)). { repeat constructor; pcuic. } forward X3 by eapply context_conversion; eauto; pcuic. specialize (X3 _ X5_2). eapply cumul_cumul_prop in cum; eauto. eapply cumul_prop_trans; eauto. eapply cumul_prop_tProd; eauto. now symmetry. now symmetry. auto. - - eapply inversion_LetIn in X6 as (s1' & A & dom & bod & codom & cum); auto. - specialize (X1 _ _ H dom _ (eq_term_upto_univ_napp_leq X7_2)). - specialize (X3 _ _ H bod _ (eq_term_upto_univ_napp_leq X7_1)). - assert(conv_context cumulAlgo_gen Σ (Γ ,, vdef na t ty) (Γ ,, vdef n b b_ty)). + - eapply inversion_LetIn in X6 as (s1' & A & e' & dom & bod & codom & cum); auto. + specialize (X1 _ _ H0 dom _ (eq_term_upto_univ_napp_leq X7_2)). + specialize (X3 _ _ H0 bod _ (eq_term_upto_univ_napp_leq X7_1)). + assert(conv_context cumulAlgo_gen Σ (Γ ,, vdef na0 t ty) (Γ ,, vdef na b b_ty)). { repeat constructor; pcuic. } - specialize (X5 u A H). + specialize (X5 u A H0). forward X5 by eapply context_conversion; eauto; pcuic. specialize (X5 _ X7_3). eapply cumul_cumul_prop in cum; eauto. @@ -1298,7 +1298,7 @@ Proof using Hcf Hcf'. eapply All2_nth_error in a; eauto. destruct a as [[[a _] _] _]. constructor; [fvs|..]. - { eapply nth_error_all in X0 as [? [dty ?]]; tea. + { eapply nth_error_all in X0 as (? & e & dty & ?); tea. now apply subject_is_open_term in dty. } { now eapply cumul_prop_is_open in cum as []. } eapply eq_term_eq_term_prop_impl; eauto. @@ -1310,7 +1310,7 @@ Proof using Hcf Hcf'. eapply All2_nth_error in a; eauto. destruct a as [[[a _] _] _]. constructor; [fvs|..]. - { eapply nth_error_all in X0 as [? [dty ?]]; tea. + { eapply nth_error_all in X0 as (? & ? & dty & ?); tea. now apply subject_is_open_term in dty. } { now eapply cumul_prop_is_open in cum as []. } eapply eq_term_eq_term_prop_impl; eauto. diff --git a/pcuic/theories/PCUICElimination.v b/pcuic/theories/PCUICElimination.v index aac7e085b..598f1b593 100644 --- a/pcuic/theories/PCUICElimination.v +++ b/pcuic/theories/PCUICElimination.v @@ -231,7 +231,7 @@ Proof. now rewrite /subst1 subst_it_mkProd_or_LetIn Nat.add_0_r in cum. unshelve eapply (isType_subst (Δ := [vass _ _]) [hd0]) in i0; pcuic. 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)); apply isType_of_isTypeRel in i; tea. now symmetry. Qed. Inductive All_local_assum (P : context -> term -> Type) : context -> Type := @@ -303,7 +303,7 @@ Lemma isType_ws_cumul_ctx_pb {cf Σ Γ Δ T} {wfΣ : wf Σ}: isType Σ Δ T. Proof. intros HT wf eq. - apply infer_sort_impl with id HT; intros Hs. + apply infer_sort_impl with id HT => //; intros Hs. eapply closed_context_conversion; tea. Qed. @@ -393,7 +393,7 @@ Proof. eapply subject_reduction_closed in Ht; eauto. intros. pose proof (PCUICWfUniverses.typing_wf_universe wfΣ Ht). - eapply inversion_Prod in Ht as [s1 [s2 [dom [codom cum']]]]; auto. + eapply inversion_Prod in Ht as (s1 & s2 & e & dom & codom & cum'); auto. specialize (H Γ0 ltac:(reflexivity) (Γ ,, vass na' dom') args' []). eapply (type_Cumul _ _ _ _ (tSort s)) in codom; cycle 1; eauto. { econstructor; pcuic. } @@ -414,6 +414,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. + apply isType_of_isTypeRel in i. { eapply (type_ws_cumul_pb _ (U:=ty)) in tyhd => //. now symmetry. } eapply (isType_subst (Δ := [_])) in i0; revgoals. { now eapply subslet_ass_tip. } @@ -427,7 +428,7 @@ Proof. intros prs;eapply All_local_assum_app in prs as [prd prs]. depelim prd. eapply (type_ws_cumul_pb (pb:=Conv) _ (U:=ty)) in tyhd. - 2:{ destruct s0 as [s' [Hs' _]]. exists s'; auto. } + 2:{ destruct s0 as (s' & Hs' & _). exists s'; split; [cbnr|apply Hs']. } 2:now symmetry. destruct H as [H _]. forward H. { @@ -482,9 +483,9 @@ Proof. induction Δ in cs, H |- *; simpl; intros. constructor; intuition auto. destruct a as [na [b|] ty]; constructor; intuition auto. destruct cs => //; eauto. - destruct cs => //; eauto. destruct X. + destruct cs => //; eauto. destruct X as (? & ? & ?). eapply IHΔ. intros. apply (H Γ' t1 s0). right; eauto. all:auto. - destruct cs => //. destruct X. + destruct cs => //. destruct X as (? & ? & ?). eapply H. left; eauto. eauto. Qed. @@ -921,7 +922,7 @@ Lemma cumul_prop1 (Σ : global_env_ext) Γ A B u : Σ ;;; Γ |- A : tSort u. Proof using Hcf Hcf'. intros. - destruct X0 as [s Hs]. + destruct X0 as (s & e & Hs). eapply cumul_prop_inv in H. 4:eauto. pcuicfo. auto. right; eauto. Qed. @@ -935,7 +936,7 @@ Lemma cumul_prop2 (Σ : global_env_ext) Γ A B u : Σ ;;; Γ |- B : tSort u. Proof using Hcf Hcf'. intros. - destruct X0 as [s Hs]. + destruct X0 as (s & e & Hs). eapply cumul_prop_inv in H. 4:eauto. pcuicfo. auto. left; eauto. Qed. @@ -949,7 +950,7 @@ Lemma cumul_sprop1 (Σ : global_env_ext) Γ A B u : Σ ;;; Γ |- A : tSort u. Proof using Hcf Hcf'. intros. - destruct X0 as [s Hs]. + destruct X0 as (s & e & Hs). eapply cumul_sprop_inv in H. 4:eauto. pcuicfo. auto. right; eauto. Qed. @@ -963,7 +964,7 @@ Lemma cumul_sprop2 (Σ : global_env_ext) Γ A B u : Σ ;;; Γ |- B : tSort u. Proof using Hcf Hcf'. intros. - destruct X0 as [s Hs]. + destruct X0 as (s & e & Hs). eapply cumul_sprop_inv in H. 4:eauto. pcuicfo. auto. left; eauto. Qed. diff --git a/pcuic/theories/PCUICExpandLetsCorrectness.v b/pcuic/theories/PCUICExpandLetsCorrectness.v index 0d5ac6eb7..849f5a1b7 100644 --- a/pcuic/theories/PCUICExpandLetsCorrectness.v +++ b/pcuic/theories/PCUICExpandLetsCorrectness.v @@ -941,11 +941,11 @@ Section wtsub. destruct t; simpl; intros [T h]; try exact tt. - now eapply typing_wf_local in h. - now eapply inversion_Evar in h. - - eapply inversion_Prod in h as (?&?&?&?&?); tea. + - eapply inversion_Prod in h as (?&?&?&?&?&?); tea. split; eexists; eauto. - - eapply inversion_Lambda in h as (?&?&?&?&?); tea. + - eapply inversion_Lambda in h as (?&?&?&?&?&?); tea. split; eexists; eauto. - - eapply inversion_LetIn in h as (?&?&?&?&?&?); tea. + - eapply inversion_LetIn in h as (?&?&?&?&?&?&?); tea. repeat split; eexists; eauto. - eapply inversion_App in h as (?&?&?&?&?&?); tea. split; eexists; eauto. @@ -986,11 +986,11 @@ Section wtsub. eexists; eauto. - eapply inversion_Fix in h as (?&?&?&?&?&?&?); tea. eapply All_prod. - eapply (All_impl a). intros ? h; exact h. + eapply (All_impl a). intros ? h; apply isType_of_isTypeRel in h; exact h. eapply (All_impl a0). intros ? h; eexists; tea. - eapply inversion_CoFix in h as (?&?&?&?&?&?&?); tea. eapply All_prod. - eapply (All_impl a). intros ? h; exact h. + eapply (All_impl a). intros ? h; apply isType_of_isTypeRel in h; exact h. eapply (All_impl a0). intros ? h; eexists; tea. Qed. End wtsub. @@ -2466,8 +2466,8 @@ Proof. - simpl. constructor. - simpl. econstructor. + eapply IHX. - + simpl. destruct tu. exists x. eapply Hs. - - simpl. constructor; auto. red. destruct tu. exists x. auto. + + simpl. destruct tu as (s & e & Hty). exists s. split; [apply e|]. eapply Hs. + - simpl. constructor; auto. red. destruct tu as (s & e & Hty). exists s. split; auto. Qed. Lemma trans_wf_local_env {cf} Σ Γ : @@ -2486,8 +2486,9 @@ Proof. - simpl. constructor. - simpl. econstructor. + eapply IHX. - + simpl. destruct t0. exists x. now eapply p. - - simpl. constructor; auto. red. destruct t0. exists x. intuition auto. + + simpl. destruct t0 as (s & e & Hs). exists s. split; [apply e|]. eapply Hs; auto. + - simpl. constructor; auto. red. destruct t0 as (s & e & Hs). exists s. split; [apply e|]. + intuition auto. red. red in t1. destruct t1 => //. Qed. @@ -2539,11 +2540,12 @@ Proof. } induction X;cbn;constructor;auto;cbn in *. - - destruct t0 as (?&?&?). + - destruct t0 as (?&e&?&?). exists x. + split; [apply e|]. apply t1. - - destruct t0 as (?&?&?). - eexists;eassumption. + - destruct t0 as (?&e&?&?). + eexists;split;eassumption. - destruct t1. assumption. Qed. @@ -3350,7 +3352,7 @@ Proof. assert (wfΣ := extends_decls_wf _ _ Hw X). eapply (weakening_env (trans_global (Σ, u))); eauto. tc. - - intros [s Hs]. exists s. intros Hw. + - intros (s & e & Hs). exists s. split; [apply e|]. intros Hw. pose proof (extends_decls_trans ext). pose proof (extends_decls_wf _ _ Hw X). specialize (Hs X0). @@ -3510,13 +3512,15 @@ Proof. - constructor. + apply IHAll_local_env_over_gen. + cbn in *. - destruct tu. - eexists;split;auto;try assumption. + destruct tu as (s & e & Hty). + exists s. split; [apply e|]. + intros; split; eauto. - constructor. + apply IHAll_local_env_over_gen. + cbn in *. - destruct tu. - eexists;split;auto;eassumption. + destruct tu as (s & e & Hty). + exists s. split; [apply e|]. + intros; split; eauto. + cbn in *. split;auto;eassumption. Qed. @@ -3571,9 +3575,9 @@ Proof. + now apply trans_consistent_instance_ext. + red in X. epose proof (declared_constructor_inv_decls weaken_prop _ X isdecl) as [cs [hnth onc]]. destruct onc. red in on_ctype. - destruct on_ctype as [s Hs]. + destruct on_ctype as (s & e & Hs). rewrite /type_of_constructor. forward Hs. eauto. - exists (s@[u]). + exists (s@[u]); split; [apply e|]. 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. @@ -3617,7 +3621,7 @@ Proof. + rewrite <- trans_global_ext_constraints. eassumption. + eassert (ctx_inst _ _ _ _) as Hctxi by (eapply ctx_inst_impl with (1 := X5); now intros ? []). - eassert (PCUICEnvTyping.ctx_inst (fun Σ _ _ _ => wf_trans Σ -> @typing (cf' _) _ _ _ _) _ _ _ _) as IHctxi. + eassert (PCUICEnvTyping.ctx_inst (fun _ _ _ => wf_trans Σ -> @typing (cf' _) _ _ _ _) _ _ _) as IHctxi. { eapply ctx_inst_impl with (1 := X5). intros ? ? [? r]; exact r. } move: Hctxi IHctxi. cbn. have wfctx : wf_local Σ (Γ ,,, (ind_params mdecl,,, ind_indices idecl)@[puinst p]). @@ -3800,21 +3804,22 @@ Proof. intuition auto. destruct X as [s [? ?]]. exists s; intuition auto. + fold trans. + unfold TT.on_def_body, on_def_body in *. subst types. eapply All_map. eapply All_prod in X0; tea. clear X1. - eapply All_impl; tea. intros d [[Hdb IHdb] [hs [hdty ihdty]]]. + eapply All_impl; tea. intros d [[Hdb IHdb] (hs & e & hdty & ihdty)]. len. rewrite -(trans_fix_context _ _ _ H2). rewrite -trans_local_app. rewrite (trans_lift _ (shiftnP #|Γ| xpred0)) in IHdb. eapply (subject_is_open_term (Σ := Σ)); tea. - len in IHdb. eauto. + len in IHdb. cbn; eauto. + rewrite (trans_wf_fixpoint _ (shiftnP #|Γ| xpred0) #|mfix|) //. - rewrite trans_dtype. simpl. assert (is_open_term Γ (tCoFix mfix n)). - { eapply (subject_is_open_term (Σ := Σ)). econstructor; tea. solve_all. - destruct a as [s Hs]. exists s; intuition eauto. + { eapply (subject_is_open_term (Σ := Σ)). econstructor; tea. unfold on_def_body in *. solve_all. + destruct a as (s & e & Hs). exists s; intuition eauto. solve_all. now destruct b. } eapply TT.type_CoFix; auto. + rewrite /trans_local map_app in X. @@ -3828,24 +3833,25 @@ Proof. + fold trans. eapply All_map, (All_impl X0). intros x [s ?]; exists s; intuition auto. - + fold trans;subst types. + + fold trans; subst types. + unfold TT.on_def_body, on_def_body in *. eapply All_map. eapply All_prod in X0; tea. clear X1. - eapply All_impl; tea. intros d [[Hdb IHdb] [hs [hdty ihdty]]]. + eapply All_impl; tea. intros d [[Hdb IHdb] (hs & e & hdty & ihdty)]. len. rewrite -(trans_fix_context _ _ _ H2). exact 0. rewrite -trans_local_app. rewrite (trans_lift _ (shiftnP #|Γ| xpred0)) in IHdb. eapply (subject_is_open_term (Σ := Σ)); tea. - len in IHdb. eauto. + len in IHdb. cbn; eauto. + rewrite trans_wf_cofixpoint //. - eapply (type_ws_cumul_pb (pb:=Cumul)). + eauto. - + now exists s. + + now exists s; split. + eapply cumulSpec_cumulAlgo_curry in X4; fvs. eapply ws_cumul_pb_forget in X4. eapply cumul_decorate in X4; tea. 2:eapply validity; tea. - 2:now exists s. + 2:now exists s; split. eapply into_ws_cumul_pb. eapply (trans_cumul (Σ := Σ)); eauto. eapply (trans_on_free_vars_ctx 0). now eapply wf_local_closed_context in wfΓ. @@ -3915,9 +3921,9 @@ Proof. cbn. now rewrite -distr_lift_subst_context. Qed. -Lemma ctx_inst_expand_lets {cf} {Σ : global_env_ext} {wfΣ : wf Σ} {univs} {Γ Δ} {s ctx} : - PCUICTyping.ctx_inst (λ (Σ : global_env_ext) (Γ : context) (t T : term), Σ;;; Γ |- t : T) (Σ, univs) (Γ ,,, Δ) s ctx -> - PCUICTyping.ctx_inst (λ (Σ : global_env_ext) (Γ : context) (t T : term), Σ;;; Γ |- t : T) (Σ, univs) (Γ ,,, smash_context [] Δ) (map (expand_lets Δ) s) (List.rev (expand_lets_ctx Δ (List.rev ctx))). +Lemma ctx_inst_expand_lets {cf} {Σ : global_env_ext} {wfΣ : wf Σ} {Γ Δ} {s ctx} : + PCUICTyping.ctx_inst (λ (Γ : context) (t T : term), Σ;;; Γ |- t : T) (Γ ,,, Δ) s ctx -> + PCUICTyping.ctx_inst (λ (Γ : context) (t T : term), Σ;;; Γ |- t : T) (Γ ,,, smash_context [] Δ) (map (expand_lets Δ) s) (List.rev (expand_lets_ctx Δ (List.rev ctx))). Proof. induction 1. - cbn. constructor. @@ -3958,9 +3964,9 @@ Qed. Lemma trans_ctx_inst_expand_lets {cf} {Σ : global_env_ext} {wfΣ : wf Σ} {Γ Δ} {s} : wf_trans Σ -> wf_local Σ (Γ ,,, List.rev Δ) -> - PCUICTyping.ctx_inst (λ (Σ : global_env_ext) (Γ : context) (t T : term), Σ;;; Γ |- t : T) Σ Γ s Δ -> - PCUICTyping.ctx_inst (λ (Σ : global_env_ext) (Γ : context) (t T : term), - typing (H:=cf' cf) Σ Γ t T) (trans_global Σ) + PCUICTyping.ctx_inst (λ (Γ : context) (t T : term), Σ;;; Γ |- t : T) Γ s Δ -> + PCUICTyping.ctx_inst (λ (Γ : context) (t T : term), + typing (H:=cf' cf) (trans_global Σ) Γ t T) (trans_local Γ) (map trans s) (trans_local Δ). Proof. intros wf wfctx i. @@ -4184,6 +4190,8 @@ Proof. depelim wfΓ; depelim wfΓ'; depelim p; constructor; auto; auto. - now apply IHX. - constructor; auto. + pose proof (isType_of_isTypeRel l). + pose proof (isType_of_isTypeRel l0). eapply cumulSpec_cumulAlgo_curry in eqt; tea. 2-4:fvs. red in l. now eapply ws_cumul_pb_forget in eqt. rewrite (All2_fold_length X). fvs. @@ -4223,10 +4231,10 @@ Proof. pose proof (trans_cumul wfΣ' cum); tea. eapply (cumulAlgo_cumulSpec _ (pb:=Cumul)). eapply into_ws_cumul_pb; tea. - destruct wtT. apply trans_is_closed_context. fvs. + destruct wtT as (s & e & Hs). apply trans_is_closed_context. fvs. apply trans_on_free_vars; len; fvs. apply trans_on_free_vars; len; fvs. - destruct wtT. fvs. + destruct wtT as (s & e & Hs). fvs. Qed. Lemma trans_convSpec {cf} {Σ : PCUICEnvironment.global_env_ext} {Γ T U} {wfΣ : PCUICTyping.wf Σ} : @@ -4263,10 +4271,10 @@ Proof. destruct p; constructor; cbn in *; auto. - rewrite -trans_local_app. depelim wfl; depelim wfr. red in l, l0. - destruct l0 as [s Hs]. destruct l as [s' Hs']. + destruct l0 as (s & e & Hs), l as (s' & e' & Hs'). eapply trans_cumulSpec in eqt; tea. { now exists s'. } - { exists s. eapply context_cumulativity_spec; tea. + { exists s. split; [cbnr|]. eapply context_cumulativity_spec; tea. eapply All2_fold_app. reflexivity. apply X. } - rewrite -trans_local_app. depelim wfl; depelim wfr. red in l, l0. eapply (trans_convSpec (Σ := Σ)) => //. @@ -4276,7 +4284,9 @@ Proof. - rewrite -trans_local_app. depelim wfl; depelim wfr. red in l, l0. eapply (trans_cumulSpec (Σ := Σ)) => //. - { destruct l1 as [s Hs]. exists s. eapply context_cumulativity_spec; tea. + eapply isType_of_isTypeRel; eapply l. + { destruct l1 as (s & e & Hs). exists s. split; [cbnr|]. + eapply context_cumulativity_spec; tea. eapply All2_fold_app. reflexivity. apply X. } Qed. @@ -4484,15 +4494,15 @@ Proof. eapply IHX. now depelim wfl. now depelim wfr. now depelim ass. now depelim ass'. destruct p; constructor; cbn in *; auto. - rewrite -trans_local_app. - depelim wfl; depelim wfr. red in l, l0. destruct l, l0. + depelim wfl; depelim wfr. red in l, l0. destruct l as (s & e & Hs), l0 as (s' & e' & Hs'). eapply (cumulAlgo_cumulSpec _ (pb:=Cumul)). apply into_ws_cumul_pb. eapply (trans_cumul' (Σ := Σ) (Γ' := Γ' ,,, Γ'0)) => //. - eapply cumul_decorate_hetero; tea. now eexists. now eexists. + eapply cumul_decorate_hetero; tea. now eexists; split. now eexists; split. len. now rewrite (All2_fold_length X) eqlen. now depelim ass. now depelim ass'. eapply cumulSpec_cumulAlgo_curry in eqt; eauto. - now apply ws_cumul_pb_forget in eqt. fvs. eapply (subject_is_open_term t0). + now apply ws_cumul_pb_forget in eqt. fvs. eapply (subject_is_open_term Hs). len. rewrite (All2_fold_length X) eqlen. rewrite -app_length; fvs. len. eapply All2_fold_length in X. lia. @@ -4537,7 +4547,7 @@ Proof. unfold PCUICLookup.wf_universe, wf_universe. destruct s => //. destruct a as [? [?|] ?] => /= //; intuition auto. - destruct a0 as [s' Hs]. exists s'. + destruct a0 as (s' & e & Hs). exists s'. split; [apply e|]. all:rewrite -trans_local_app. now eapply (pcuic_expand_lets _ _ _ (tSort _)). now eapply (pcuic_expand_lets _ _ _ _). @@ -4551,8 +4561,8 @@ Lemma trans_on_context {cf} {Σ Γ} : Proof. intros wf wf'. induction 1; cbn; constructor; auto. - destruct t0 as [s Hs]. exists s. now eapply (pcuic_expand_lets _ _ _ (tSort _)). - destruct t0 as [s Hs]. exists s. now eapply (pcuic_expand_lets _ _ _ (tSort _)). + destruct t0 as (s & e & Hs). exists s. split; [apply e|]. now eapply (pcuic_expand_lets _ _ _ (tSort _)). + destruct t0 as (s & e & Hs). exists s. split; [apply e|]. now eapply (pcuic_expand_lets _ _ _ (tSort _)). now eapply (pcuic_expand_lets _ _ _ _). Qed. @@ -4619,7 +4629,7 @@ Lemma wf_ind_indices {cf:checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} {mi Proof. intros []. rewrite ind_arity_eq in onArity. - destruct onArity as [s Hs]. + destruct onArity as (s & e & Hs). eapply type_it_mkProd_or_LetIn_inv in Hs as [Δs [ts [_ Hs]]]. eapply type_it_mkProd_or_LetIn_inv in Hs as [Δs' [ts' []]]. eapply typing_wf_local in t. now rewrite app_context_nil_l in t. @@ -4689,14 +4699,14 @@ Proof. * cbn. red. move: ond; rewrite /on_constant_decl. destruct c as [type [body|] univs] => /=. intros Hty; eapply (pcuic_expand_lets (Σ0, univs) [] _ _ X Hty IHX). - intros [s Hty]. exists s. + intros (s & e & Hty). exists s. split; [apply e|]. exact (pcuic_expand_lets (Σ0, univs) [] _ _ X Hty IHX). * generalize ond. intros []; econstructor; eauto. + cbn. eapply (Alli_mapi _ _ _).1, Alli_impl; tea. intros n idecl onind; generalize onind; intros []; econstructor; tea. { cbn. rewrite ind_arity_eq !trans_it_mkProd_or_LetIn //. } - { cbn. destruct onArity as [s Hty]. exists s. + { cbn. destruct onArity as (s & e & Hty). exists s. split; [apply e|]. exact (pcuic_expand_lets (Σ0, ind_universes m) [] _ _ X Hty IHX). } { instantiate (1 := ind_cunivs). red in onConstructors. @@ -4709,11 +4719,11 @@ Proof. { eapply sorts_local_ctx_wf_local in on_cargs. eapply wf_local_closed_context in on_cargs. now move/onfvs_app: on_cargs; len. - rewrite cstr_eq in on_ctype. destruct on_ctype as [s Hs]. + rewrite cstr_eq in on_ctype. destruct on_ctype as (s & e & Hs). eapply typing_wf_local in Hs. eapply weaken_wf_local; tea. } have oncindices: on_free_vars_terms (shiftnP (#|cstr_args cdecl| + #|ind_params m| + #|ind_bodies m|) xpred0) (cstr_indices cdecl). - { rewrite cstr_eq in on_ctype. destruct on_ctype as [s Hs]. + { rewrite cstr_eq in on_ctype. destruct on_ctype as (s & e & Hs). eapply subject_is_open_term in Hs. move: Hs. rewrite !on_free_vars_it_mkProd_or_LetIn. move/and3P => [] _ _. rewrite /cstr_concl. @@ -4723,8 +4733,8 @@ Proof. { cbn; rewrite -cstr_args_length. rewrite context_assumptions_smash_context context_assumptions_map //. } { cbn; rewrite /trans_cstr_concl /trans_cstr_concl_head /cstr_concl_head. now len. } - { destruct on_ctype as [s Hty]. - exists s. + { destruct on_ctype as (s & e & Hty). + exists s. split; [apply e|]. 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. @@ -4751,7 +4761,9 @@ Proof. rewrite expand_lets_ctx_tip /=. destruct univs => //. split. cbn in IHc. apply IHc, on_cargs. - destruct on_cargs as [hs ht]. red in ht. + destruct on_cargs as (hs & e & ht). + split; [apply e|]. + red in ht. have fvsc : on_free_vars_ctx (shiftnP (#|ind_params m| + #|ind_bodies m|) xpred0) c. { eapply typing_wf_local, wf_local_closed_context in ht. move/onfvs_app: ht. now len. } @@ -4777,7 +4789,7 @@ Proof. rewrite shiftnP_add Nat.add_assoc //. rewrite -(trans_lift_context (shiftnP #|ind_params m| xpred0)). { rewrite ind_arity_eq in onArity. - destruct onArity as [s Hs]. + destruct onArity as (s & e & Hs). eapply subject_is_open_term in Hs. move: Hs; rewrite !on_free_vars_it_mkProd_or_LetIn /=. move/and3P => [] _. now rewrite shiftnP0. } @@ -4790,7 +4802,7 @@ Proof. eapply weaken_wf_local; tea. eapply wf_arities_context'; tea. rewrite ind_arity_eq in onArity. - destruct onArity as [s Hs]. + destruct onArity as (s & e & Hs). rewrite -it_mkProd_or_LetIn_app in Hs. eapply type_it_mkProd_or_LetIn_inv in Hs as [? [? [Hs _]]]. eapply PCUICClosedConv.sorts_local_ctx_All_local_env in Hs; eauto. @@ -4808,7 +4820,7 @@ Proof. rewrite -(trans_expand_lets (shiftnP (#|ind_params m| + #|ind_bodies m|) xpred0)) //. rewrite shiftnP_add Nat.add_assoc //. { rewrite cstr_eq in on_ctype. - destruct on_ctype as [s Hs]. + destruct on_ctype as (s & e & Hs). eapply subject_is_open_term in Hs. len in Hs. move: Hs; rewrite !on_free_vars_it_mkProd_or_LetIn. move/and3P => [] onpars onargs onconcl. @@ -4819,7 +4831,7 @@ Proof. eapply (positive_cstr_trans _ _ _ (shiftnP #|ind_bodies m| xpred0)) in on_ctype_positive. exact on_ctype_positive. rewrite cstr_eq in on_ctype. - destruct on_ctype as [s Hs]. + destruct on_ctype as (s & e & Hs). eapply subject_is_open_term in Hs. len in Hs. move: Hs; rewrite !on_free_vars_it_mkProd_or_LetIn. @@ -4949,7 +4961,7 @@ Proof. depelim onConstructors. depelim onConstructors. have wfargs : wf_local (Σ0, ind_universes m) (arities_context (ind_bodies m),,, ind_params m,,, cstr_args c). { destruct o. eapply sorts_local_ctx_wf_local in on_cargs => //. - rewrite cstr_eq in on_ctype. destruct on_ctype as [s Hs]. + rewrite cstr_eq in on_ctype. destruct on_ctype as (s & e & Hs). eapply typing_wf_local in Hs. eapply weaken_wf_local; tea. } have oncargs: on_free_vars_ctx (shiftnP (#|ind_params m| + #|ind_bodies m|) xpred0) (cstr_args c). diff --git a/pcuic/theories/PCUICGeneration.v b/pcuic/theories/PCUICGeneration.v index 3eeea8b70..4da6b95dd 100644 --- a/pcuic/theories/PCUICGeneration.v +++ b/pcuic/theories/PCUICGeneration.v @@ -33,7 +33,7 @@ Section Generation. intros. specialize (IHHsp (tApp t0 hd)). apply IHHsp. - destruct i as [s Hs]. + destruct i as (s & e & Hs). eapply type_App; eauto. eapply type_Cumul; eauto. Qed. @@ -49,11 +49,13 @@ Section Generation. - simpl. cbn. eapply ih. simpl in h. pose proof (typing_wf_local h) as hc. dependent induction hc. - econstructor; try eassumption. exact t0.π2. + destruct t0 as (s & e & Hs). + econstructor; try eassumption. - simpl. cbn. eapply ih. pose proof (typing_wf_local h) as hc. cbn in hc. - dependent induction hc; - econstructor; try eassumption. exact t0.π2. + dependent induction hc. + destruct t0 as (s & e & Hs). + econstructor; try eassumption. Qed. End Generation. diff --git a/pcuic/theories/PCUICInductiveInversion.v b/pcuic/theories/PCUICInductiveInversion.v index 4716b5b1d..e4acc8256 100644 --- a/pcuic/theories/PCUICInductiveInversion.v +++ b/pcuic/theories/PCUICInductiveInversion.v @@ -60,7 +60,7 @@ Lemma type_tFix_inv {cf:checker_flags} (Σ : global_env_ext) Γ mfix idx T : wf wf_fixpoint Σ mfix * (Σ ;;; Γ |- f : T') * (Σ ;;; Γ ⊢ T' ≤ T) }}}%type. Proof. - intros wfΣ H. depind H. + intros wfΣ H. depind H. unfold on_def_type, on_def_body in *. - unfold unfold_fix. rewrite e. specialize (nth_error_all e a0) as [s Hs]. specialize (nth_error_all e a1) as Hty. @@ -115,13 +115,11 @@ Qed. Lemma subslet_cofix {cf:checker_flags} (Σ : global_env_ext) Γ mfix : wf_local Σ Γ -> cofix_guard Σ Γ mfix -> - All (fun d : def term => ∑ s : Universe.t, Σ;;; Γ |- dtype d : tSort s) mfix -> - All - (fun d : def term => - Σ;;; Γ ,,, fix_context mfix |- dbody d - : lift0 #|fix_context mfix| (dtype d)) mfix -> + All (on_def_type (lift_typing typing Σ) Γ) mfix -> + All (on_def_body (lift_typing typing Σ) (fix_context mfix) Γ) mfix -> wf_cofixpoint Σ mfix -> subslet Σ Γ (cofix_subst mfix) (fix_context mfix). Proof. + unfold on_def_body, on_def_type. intros wfΓ hguard types bodies wfcofix. pose proof bodies as X1. apply All_rev in X1. unfold cofix_subst, fix_context. simpl. @@ -173,7 +171,7 @@ Proof. 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. + * eapply nth_error_all in a0; tea. apply isType_of_isTypeRel in a0. now eapply isType_ws_cumul_pb_refl. - destruct (IHtyping1 wfΣ) as [d [[[Hnth wfcofix] ?] ?]]. exists d. intuition auto. etransitivity; eauto. @@ -282,12 +280,12 @@ Section OnConstructor. eapply weaken_wf_local => //. eapply wf_arities_context; eauto; eapply declc. now eapply onParams. - destruct (on_ctype onc). - rewrite onc.(cstr_eq) in t. - rewrite -it_mkProd_or_LetIn_app in t. - eapply inversion_it_mkProd_or_LetIn in t => //. - unfold cstr_concl_head in t. simpl in t. - eapply inversion_mkApps in t as [A [ta sp]]. + destruct (on_ctype onc) as (s & e & Hs). + rewrite onc.(cstr_eq) in Hs. + rewrite -it_mkProd_or_LetIn_app in Hs. + eapply inversion_it_mkProd_or_LetIn in Hs => //. + unfold cstr_concl_head in Hs. simpl in Hs. + eapply inversion_mkApps in Hs as [A [ta sp]]. eapply inversion_Rel in ta as [decl [wfΓ [nth cum']]]. rewrite nth_error_app_ge in nth. len. len in nth. @@ -302,9 +300,9 @@ Section OnConstructor. rewrite Hdecl in cum'; clear Hdecl. assert(closed (ind_type idecl)). { pose proof (oib.(onArity)). rewrite (oib.(ind_arity_eq)) in X0 |- *. - destruct X0 as [s Hs]. now apply subject_closed in Hs. } + destruct X0 as (s' & e' & Hs'). now apply subject_closed in Hs'. } rewrite lift_closed in cum' => //. - eapply typing_spine_strengthen in sp; simpl. + eapply typing_spine_strengthen in sp. 3:tea. move: sp. rewrite (oib.(ind_arity_eq)). @@ -314,7 +312,8 @@ Section OnConstructor. clear Hlen'. rewrite [_ ,,, _]app_context_assoc in Hinst. now exists inst. - apply infer_typing_sort_impl with id oib.(onArity); intros Hs. + eapply isType_of_isTypeRel. + apply infer_typing_sort_impl with id oib.(onArity) => //; intros Hs. now eapply weaken_ctx in Hs. Qed. End OnConstructor. @@ -544,6 +543,7 @@ 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]. + apply isType_of_isTypeRel in isty. eapply (isType_subst (Δ := [_]) [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. } @@ -574,6 +574,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 isType_of_isTypeRel in Hty. eapply type_ws_cumul_pb; eauto. now eapply ws_cumul_pb_eq_le. * pcuic. Qed. @@ -686,7 +687,7 @@ Proof. now rewrite closedn_subst_instance_context. eapply subslet_app_inv in Hs. move: Hs. len. intuition auto. - rewrite closed_ctx_subst in a0 => //. + rewrite closed_ctx_subst in a1 => //. now rewrite closedn_subst_instance_context. (*rewrite -Heqp in spars sargs. simpl in *. clear Heqp. *) @@ -701,7 +702,7 @@ Proof. split; auto; try split; auto. - apply weaken_wf_local => //. - - pose proof (subslet_length a0). rewrite subst_instance_length in H1. + - pose proof (subslet_length a1). rewrite subst_instance_length in H1. rewrite -H1 -subst_app_context. eapply (substitution_wf_local (Γ' := subst_instance u (arities_context (ind_bodies mdecl) ,,, ind_params mdecl))); eauto. rewrite subst_instance_app; eapply subslet_app; eauto. @@ -711,13 +712,13 @@ Proof. rewrite -app_context_assoc. eapply weaken_wf_local => //. rewrite -subst_instance_app_ctx. - apply a. + apply a0. - exists (map (subst_instance_univ u') x). split. * move/onParams: onmind. rewrite /on_context. pose proof (@wf_local_instantiate _ Σ (InductiveDecl mdecl) (ind_params mdecl) u'). move=> H'. eapply X in H'; eauto. 2:destruct decli; eauto. - clear -wfar wfpars wfΣ hΓ cons decli t cargs sargs H0 H' a spars a0. + clear -wfar wfpars wfΣ hΓ cons decli cargs sargs H0 H' a spars a0. eapply (subst_sorts_local_ctx (Γ' := []) (Δ := subst_context (inds (inductive_mind i) u' (ind_bodies mdecl)) 0 (subst_instance u' (ind_params mdecl)))) => //. @@ -1011,7 +1012,7 @@ Lemma isType_mkApps_Ind_smash {cf:checker_flags} {Σ Γ ind u params args} {wfΣ Proof. move=> isdecl isty hpars. pose proof (isType_wf_local isty). - destruct isty as [s Hs]. + destruct isty as (s & e & Hs). eapply invert_type_mkApps_ind in Hs as [sp cu]; tea. move=> parctx argctx. erewrite ind_arity_eq in sp. @@ -3153,7 +3154,7 @@ Lemma wt_ind_app_variance {cf} {Σ} {wfΣ : wf Σ} {Γ ind u l}: ∑ mdecl, (lookup_inductive Σ ind = Some mdecl) * (global_variance Σ (IndRef ind) #|l| = ind_variance (fst mdecl)). Proof. - move => [s wat]. + intros (s & e & wat). eapply inversion_mkApps in wat as [ty [Hind Hargs]]; auto. eapply inversion_Ind in Hind as [mdecl [idecl [wfΓ [decli [cu cum]]]]]; auto. eapply typing_spine_strengthen in Hargs; eauto. clear cum. @@ -3185,10 +3186,11 @@ Lemma ctx_inst_app_weak `{checker_flags} Σ (wfΣ : wf Σ.1) ind mdecl idecl (is ctx_inst Σ Γ (params ++ skipn (ind_npars mdecl) args) (List.rev (subst_instance v (ind_params mdecl ,,, ind_indices idecl))). Proof. - intros [? ty_args] ? cparams cum. - pose proof (wt_ind_app_variance (x; ty_args)) as [mdecl' [idecl' gv]]. + intros isty ? cparams cum. + pose proof (wt_ind_app_variance isty) as [mdecl' [idecl' gv]]. + destruct isty as (s & e & Hs). rewrite (declared_inductive_lookup isdecl) in idecl'. noconf idecl'. - eapply invert_type_mkApps_ind in ty_args as [ty_args ?] ; eauto. + eapply invert_type_mkApps_ind in Hs as [ty_args ?] ; eauto. erewrite ind_arity_eq in ty_args. 2: eapply PCUICInductives.oib ; eauto. @@ -3258,10 +3260,10 @@ Proof. Qed. Lemma wf_local_vass {cf:checker_flags} Σ {Γ na A} s : - Σ ;;; Γ |- A : tSort s -> wf_local Σ (Γ ,, vass na A). + isSortRel s na.(binder_relevance) -> Σ ;;; Γ |- A : tSort s -> wf_local Σ (Γ ,, vass na A). Proof. - intro X; apply typing_wf_local in X as Y. - constructor; tea. eexists; eassumption. + intros e X; apply typing_wf_local in X as Y. + constructor; tea. eexists; split; eassumption. Qed. Lemma isType_it_mkProd_or_LetIn {cf:checker_flags} {Σ Γ Δ T} : @@ -3275,23 +3277,22 @@ Proof. - intros T Hty. rewrite /= /mkProd_or_LetIn /=. eapply IHΔ. - apply infer_sort_impl with id Hty; intros Hs. + apply infer_sort_impl with id Hty => //; intros Hs. have wf := typing_wf_local Hs. depelim wf. - destruct l as [s1 Hs1]. red in l0. + destruct l as (s1 & e1 & Hs1). red in l0. eapply type_Cumul'. econstructor; eauto. eapply isType_Sort; eauto. now eapply PCUICWfUniverses.typing_wf_universe in Hs. eapply convSpec_cumulSpec, red1_cumulSpec. repeat constructor. - - intros T [s Hs]. + - intros T (s & e & Hs). apply IHΔ. - red. - unfold PCUICTypingDef.typing in *. have wf := typing_wf_local Hs. depelim wf. - destruct l as [s1 Hs1]. + destruct l as (s1 & e1 & Hs1). exists (Universe.sort_of_product s1 s). + split; [apply e|]. econstructor; eauto. Qed. @@ -3305,7 +3306,7 @@ Proof. eapply wf_local_app => //. induction w in nas, ha |- *; depelim ha; cbn. constructor. all: constructor; eauto; try apply IHw; auto. - 1,2: apply infer_typing_sort_impl with id t0; intros Hs. + 1,2: rewrite e; apply infer_typing_sort_impl with id t0 => //; intros Hs. all: eapply context_conversion; tea. 1,3,5: eapply wf_local_app, IHw; eauto. all: eapply eq_binder_annots_eq_ctx in ha. @@ -3328,7 +3329,7 @@ Proof. split. 2:{ eexists _, _. rewrite destArity_it_mkProd_or_LetIn. reflexivity. } rewrite /case_predicate_context /case_predicate_context_gen. - have wfΓ := typing_wf_local X.π2. + have wfΓ := typing_wf_local X.π2.2. eapply isType_mkApps_Ind_inv in X; tea. destruct X as [parsubst [argsubst [sppars spargs parslen argslen cu]]]. epose proof (isType_case_predicate (puinst p) (pparams p) ps wfΓ isdecl cu wfps). @@ -3428,7 +3429,7 @@ Qed. Lemma isType_is_open_term {cf} {Σ} {wfΣ : wf Σ} Γ T : isType Σ Γ T -> is_open_term Γ T. Proof. - intros [s hs]. now apply subject_is_open_term in hs. + intros (s & e & hs). now apply subject_is_open_term in hs. Qed. #[global] Hint Immediate isType_is_open_term : pcuic. @@ -3456,7 +3457,7 @@ Lemma isType_subst_all_rels {cf} {Σ} {wfΣ : wf Σ} {Γ Δ} {T} : isType Σ (Γ ,,, Δ) (subst0 (all_rels Δ 0 #|Δ|) (lift #|Δ| #|Δ| T)). Proof. intros Hty. - apply infer_typing_sort_impl with id Hty; intros Hs. + apply infer_typing_sort_impl with id Hty => //; intros Hs. pose proof (typing_wf_local Hs). eapply weakening_typing in Hs; tea. rewrite -(app_nil_l (lift_context _ _ _)) -/(app_context _ _) app_context_assoc in Hs. @@ -3537,7 +3538,7 @@ Proof. eapply X. 2:pcuic. eauto with fvs. clear X. - apply infer_typing_sort_impl with id isty; intros Hs. + apply infer_typing_sort_impl with id isty => //; intros Hs. eapply (weakening_typing) in Hs; tea. eapply (substitution (Δ := [])) in Hs. 2:{ epose proof (spine_subst_to_extended_list_k a). apply X. } @@ -3967,6 +3968,7 @@ Proof. PCUICLiftSubst.map_subst_instance_to_extended_list_k //. } rewrite /pre_case_branch_context /subst_let_expand_k. eexists (subst_instance p.(puinst) (ind_sort idecl)). + split; [cbnr|]. relativize #|cstr_args cdecl|. eapply (substitution (s := List.rev (pparams p)) (T := tSort _)); tea. eapply sppars. @@ -4327,6 +4329,7 @@ Proof. PCUICLiftSubst.map_subst_instance_to_extended_list_k //. } rewrite /pre_case_branch_context /subst_let_expand_k. eexists (subst_instance p.(puinst) (ind_sort idecl)). + split; [cbnr|]. relativize #|cstr_args cdecl|. eapply (substitution (s := List.rev (pparams p)) (T := tSort _)); tea. eapply sppars. diff --git a/pcuic/theories/PCUICInductives.v b/pcuic/theories/PCUICInductives.v index c9a42cd14..3803b814a 100644 --- a/pcuic/theories/PCUICInductives.v +++ b/pcuic/theories/PCUICInductives.v @@ -236,7 +236,7 @@ Section OnInductives. eapply on_declared_inductive in decli as [onmind oib]. pose proof (oib.(onArity)). rewrite oib.(ind_arity_eq) in X. - destruct X as [s Hs]. + destruct X as (s & e & Hs). rewrite -it_mkProd_or_LetIn_app in Hs. eapply it_mkProd_or_LetIn_wf_local in Hs. now rewrite app_context_nil_l in Hs. @@ -282,7 +282,7 @@ Section OnInductives. destruct (on_declared_inductive decli) as [onmind oib]. pose proof (oib.(onArity)) as ar. rewrite oib.(ind_arity_eq) in ar. - destruct ar as [s ar]. + destruct ar as (s & e & ar). eapply isType_weaken => //. rewrite -(subst_instance_it_mkProd_or_LetIn u _ (tSort _)). rewrite -it_mkProd_or_LetIn_app in ar. @@ -298,7 +298,7 @@ Section OnInductives. move=> wfΓ cext. destruct (on_declared_inductive decli) as [onmind oib]. pose proof (oib.(onArity)) as ar. - destruct ar as [s ar]. + destruct ar as (s & e & ar). eapply isType_weaken => //. eapply (typing_subst_instance_decl Σ [] _ _ _ (InductiveDecl mdecl) u wfΣ) in ar. all:pcuic. eapply decli. @@ -311,7 +311,7 @@ Section OnInductives. Proof using decli wfΣ. pose proof (oib.(onArity)) as ar. rewrite oib.(ind_arity_eq) in ar. - destruct ar as [s ar]. + destruct ar as (s & e & ar). eapply typing_wf_universes in ar; auto. move/andP: ar => []. rewrite wf_universes_it_mkProd_or_LetIn => /andP [] _. @@ -568,7 +568,7 @@ Lemma isType_it_mkProd_or_LetIn_inv {cf:checker_flags} {Σ : global_env_ext} {wf isType Σ Γ (it_mkProd_or_LetIn Δ T) -> isType Σ (Γ ,,, Δ) T. Proof. - intros HT; apply infer_typing_sort_impl with id HT; intros Hs. + intros HT; apply infer_typing_sort_impl with id HT => //; intros Hs. now eapply inversion_it_mkProd_or_LetIn in Hs. Qed. @@ -623,21 +623,21 @@ Proof. eapply (substitution (Δ := [])) in l0. eapply l0. all:auto. * rewrite smash_context_acc. simpl. rewrite /map_decl /= /map_decl /=. simpl. - pose proof (l.π2). + pose proof (l.π2.2). rewrite lift_closed. fvs. rewrite (lift_extended_subst _ 1). rewrite -{4}(closed_ctx_lift 1 0 Δ); fvs. constructor. { eapply (subslet_lift _ [_]); eauto. constructor. eapply wf_local_smash_context; auto. - apply infer_typing_sort_impl with id l; intros Hs. + apply infer_typing_sort_impl with id l => //; intros Hs. eapply weaken_ctx in Hs. 3:eapply wf_local_smash_context; eauto. 2:auto. eapply (substitution (Δ := [])) in Hs. eapply Hs. all:auto. } { eapply refine_type. econstructor. rewrite smash_context_acc /=. constructor. apply wf_local_smash_context; auto. - apply infer_typing_sort_impl with id l; intros Hs. + apply infer_typing_sort_impl with id l => //; intros Hs. eapply weaken_ctx in Hs. 3:eapply wf_local_smash_context; eauto. 2:auto. eapply (substitution (Δ := [])) in Hs. eapply Hs. all:auto. reflexivity. @@ -703,12 +703,13 @@ Proof. { now apply wf_local_smash_context. } constructor => //. red. exists (ind_sort idecl). + split; [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)). assert (isType (Σ.1, ind_universes mdecl) [] (ind_type idecl)). - { pose proof (onArity oib). exact X. } + { pose proof (onArity oib). eapply isType_of_isTypeRel. exact X. } rewrite (subst_instance_ind_type_id _ _ _ _ _ hdecl). rewrite (ind_arity_eq oib). rewrite -(app_nil_r (to_extended_list _)). @@ -1071,7 +1072,7 @@ Proof. rewrite H //. split => //. apply Hnth'. eapply meta_conv. econstructor. - destruct IH as [[s isTy] _]. + destruct IH as ((s & e & isTy) & _). now eapply typing_wf_local in isTy. simpl. reflexivity. simpl. rewrite lift_mkApps. simpl. destruct ind; simpl. @@ -1121,7 +1122,9 @@ Proof. intros Hwf hnth. epose proof (nth_error_All_local_env (nth_error_Some_length hnth) Hwf). rewrite hnth /= in X. unfold on_local_decl in X. + eapply isType_of_isTypeRel. destruct decl_body => //. destruct X => //. + all: eauto. Qed. Lemma map_expand_lets_to_extended_list_k Γ : @@ -1227,12 +1230,13 @@ Proof. vass {| binder_name := nNamed (ind_name idecl); binder_relevance := idecl.(ind_relevance) |} (mkApps (tInd ind u) (to_extended_list (ind_params mdecl))))). { constructor. auto. red. exists (ind_sort idecl). + split; [apply (ind_relevance_compat oib)|]. 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. assert (isType (Σ.1, ind_universes mdecl) [] (ind_type idecl)@[u]). - { rewrite typeu. pose proof (onArity oib). exact X1. } + { rewrite typeu. pose proof (onArity oib). eapply isType_of_isTypeRel. exact X1. } rewrite (ind_arity_eq oib). rewrite subst_instance_it_mkProd_or_LetIn. rewrite -(app_nil_r (to_extended_list _)). @@ -1255,7 +1259,7 @@ Proof. constructor. reflexivity. rewrite -subst_instance_it_mkProd_or_LetIn. pose proof (onArity oib). rewrite -(oib.(ind_arity_eq)). - apply infer_typing_sort_impl with id X1. + apply infer_typing_sort_impl with id X1 => //. eapply (weaken_ctx (Γ:=[])); auto. } intros wf. generalize (weakening_wf_local (Γ'':=[_]) wf X1). @@ -1310,7 +1314,7 @@ Proof. rewrite /subst_context /lift_context. apply assumption_context_fold, smash_context_assumption_context. constructor. } rewrite H in onnth. 2:{ simpl in onnth; contradiction. } - destruct onnth as [s Hs]. + destruct onnth as (s & e & Hs). Ltac havefst H := match goal with |- ?T × _ => have H : T; [|split => //] @@ -1405,7 +1409,7 @@ Proof. now rewrite /indsubst inds_length. } split. unfold projection_type in H0. - rewrite H0. exists s; auto. + rewrite H0. exists s; split; auto. rewrite -/indb in Hs. rewrite /projection_type' -/indb -/indsubst -/projsubst. rewrite Nat.add_1_r in Hs. exact Hs. @@ -1533,7 +1537,8 @@ Lemma type_local_ctx_cum {cf:checker_flags} {Σ Γ Δ s s'} : Proof. intros wfΣ wfs leqs. induction Δ as [|[na [b|] ty] Δ]; simpl; auto; - intros []; split; auto. + intros (? & e & p); split; auto. + split; [eapply (leq_relevance leqs)|]; tea. eapply type_Cumul; eauto. econstructor; pcuic. now eapply cumul_Sort. @@ -1568,24 +1573,24 @@ Lemma isType_it_mkProd_or_LetIn_smash {cf:checker_flags} Σ Γ Δ s : isType Σ Γ (it_mkProd_or_LetIn Δ (tSort s)) -> isType Σ Γ (it_mkProd_or_LetIn (smash_context [] Δ) (tSort s)). Proof. - move=> wfΣ [u Hu]. - exists u. unfold PCUICTypingDef.typing in *. revert Γ u Hu. + intros wfΣ (u & _ & Hu); exists u; split; [cbnr|]. + revert Γ u Hu. induction Δ as [|[na [b|] ty] Δ] using ctx_length_rev_ind; simpl in *; auto; rewrite !it_mkProd_or_LetIn_app /= /mkProd_or_LetIn /=; rewrite smash_context_app; intros Γ u Hu. - simpl. assert (wfu := typing_wf_universe wfΣ Hu). - eapply inversion_LetIn in Hu as [? [? [? [? [? e]]]]]; auto. + eapply inversion_LetIn in Hu as (? & ? & ? & ? & ? & ? & e); auto. eapply substitution_let in t1; auto. rewrite /subst1 subst_it_mkProd_or_LetIn /= in t1. specialize (X (subst_context [b] 0 Δ) ltac:(autorewrite with len; lia)). eapply ws_cumul_pb_LetIn_l_inv in e; eauto. eapply type_ws_cumul_pb in t1. 3:eauto. - 2:{ pcuic. eexists; econstructor; pcuic. } + 2:{ apply isType_Sort; pcuic. } specialize (X _ _ t1). now rewrite -smash_context_subst /= subst_context_nil. - simpl. rewrite it_mkProd_or_LetIn_app. simpl. assert (wfu := typing_wf_universe wfΣ Hu). - eapply inversion_Prod in Hu as [? [? [? [? ?]]]]; auto. + eapply inversion_Prod in Hu as (? & ? & ? & ? & ? & ?); auto. specialize (X Δ ltac:(reflexivity)). specialize (X _ _ t0). eapply type_ws_cumul_pb; eauto. @@ -1636,6 +1641,7 @@ Proof. { eapply wf_local_smash_context; auto. } constructor; auto. red. exists (subst_instance_univ u (ind_sort idecl)). + split; [apply relevance_subst; apply oib0.(ind_relevance_compat)|]. eapply type_mkApps. econstructor; eauto. eapply decli.p1. rewrite (ind_arity_eq oib0). pose proof (on_projs_noidx _ _ _ _ _ _ onps). @@ -1686,8 +1692,7 @@ Lemma isType_mkApps_Ind_inv {cf:checker_flags} {Σ Γ ind u args} (wfΣ : wf Σ. #|skipn (ind_npars mdecl) args| = context_assumptions idecl.(ind_indices) & consistent_instance_ext Σ (ind_universes mdecl) u]. Proof. - move=> wfΓ isType. - destruct isType as [s Hs]. + intros wfΓ (s & e & Hs). eapply invert_type_mkApps_ind in Hs as [tyargs cu]; eauto. set (decli' := on_declared_inductive declm). rename declm into decli. @@ -1730,8 +1735,7 @@ Lemma isType_mkApps_Ind_inv' {cf:checker_flags} {Σ Γ ind u args} (wfΣ : wf Σ #|indices| = context_assumptions idecl.(ind_indices) & consistent_instance_ext Σ (ind_universes mdecl) u]. Proof. - move=> wfΓ isType. - destruct isType as [s Hs]. + intros wfΓ (s & e & Hs). eapply invert_type_mkApps_ind in Hs as [tyargs cu]; eauto. set (decli' := on_declared_inductive declm). rename declm into decli. @@ -2338,7 +2342,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 & e & Hs). eapply type_it_mkProd_or_LetIn_inv in Hs as (idxs & inds & [sortsidx sortind leq]). set (idxctx := smash_context [] (ind_params mdecl)@[u] ,,, expand_lets_ctx (ind_params mdecl)@[u] (ind_indices idecl)@[u]) in *. have tyass : Σ ;;; Γ ,,, idxctx |- (decl_type iass)@[u] : tSort (ind_sort idecl)@[u]. @@ -2412,7 +2416,11 @@ Proof. rewrite subst_extended_lift //; len. now rewrite closedn_subst_instance_context. rewrite (lift_extended_subst _ #|ind_indices _|) subst_map_lift_lift_context; len. } constructor => //. - 2:{ eexists; tea. rewrite /idxctx app_context_assoc in tyass; tea. } + 2:{ + eexists. + split; [apply relevance_subst; apply (on_declared_inductive isdecl).2.(ind_relevance_compat)|]. + rewrite /idxctx app_context_assoc in tyass; tea. + } apply typing_wf_local in tyass. rewrite /idxctx app_context_assoc in tyass => //. Qed. @@ -2440,13 +2448,14 @@ 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 & e & Hs). eapply isType_it_mkProd_or_LetIn_app in Hs. 2:eapply sp. rewrite subst_let_expand_it_mkProd_or_LetIn in Hs. eapply type_it_mkProd_or_LetIn_inv in Hs as (idxs & inds & [sortsidx sortind leq]). eexists (sort_of_products (subst_instance u (ind_sort idecl) :: idxs) (Universe.super ps)). + split; [apply e|]. set (idxctx := subst_context_let_expand _ _ _) in *. have tyass : Σ ;;; Γ ,,, idxctx |- subst (List.rev params) #|ind_indices idecl| (decl_type iass)@[u] : tSort (ind_sort idecl)@[u]. @@ -2513,10 +2522,11 @@ Proof. rewrite -(it_mkProd_or_LetIn_app [sdecl]). eapply type_it_mkProd_or_LetIn_sorts; tea. constructor => //. + split; [ apply relevance_subst; apply (on_declared_inductive isdecl).2.(ind_relevance_compat) | apply tyass]. constructor => //. simpl. constructor => //. now eapply sorts_local_ctx_wf_local; tea. red. - eexists; tea. + eexists; split; [apply relevance_subst; apply (on_declared_inductive isdecl).2.(ind_relevance_compat) | idtac ]; tea. 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 23d7c369d..ebaab7f8e 100644 --- a/pcuic/theories/PCUICInversion.v +++ b/pcuic/theories/PCUICInversion.v @@ -144,6 +144,7 @@ Section Inversion. forall {Γ na A B T}, Σ ;;; Γ |- tProd na A B : T -> ∑ s1 s2, + isSortRel s1 na.(binder_relevance) × Σ ;;; Γ |- A : tSort s1 × Σ ;;; Γ ,, vass na A |- B : tSort s2 × Σ ;;; Γ ⊢ tSort (Universe.sort_of_product s1 s2) ≤ T. @@ -166,6 +167,7 @@ Section Inversion. forall {Γ na A t T}, Σ ;;; Γ |- tLambda na A t : T -> ∑ s B, + isSortRel s na.(binder_relevance) × Σ ;;; Γ |- A : tSort s × Σ ;;; Γ ,, vass na A |- t : B × Σ ;;; Γ ⊢ tProd na A B ≤ T. @@ -176,8 +178,9 @@ Section Inversion. Lemma inversion_LetIn : forall {Γ na b B t T}, Σ ;;; Γ |- tLetIn na b B t : T -> - ∑ s1 A, - Σ ;;; Γ |- B : tSort s1 × + ∑ s A, + isSortRel s na.(binder_relevance) × + Σ ;;; Γ |- B : tSort s × Σ ;;; Γ |- b : B × Σ ;;; Γ ,, vdef na b B |- t : A × Σ ;;; Γ ⊢ tLetIn na b B A ≤ T. @@ -255,7 +258,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) - (ind_inst : ctx_inst typing Σ Γ (p.(pparams) ++ indices) + (ind_inst : ctx_inst (typing Σ) Γ (p.(pparams) ++ indices) (List.rev (subst_instance p.(puinst) (ind_params mdecl ,,, ind_indices idecl)))) (scrut_ty : Σ ;;; Γ |- c : mkApps (tInd ci.(ci_ind) p.(puinst)) (p.(pparams) ++ indices)) @@ -309,9 +312,8 @@ Section Inversion. let types := fix_context mfix in fix_guard Σ Γ mfix × nth_error mfix n = Some decl × - All (fun d => isType Σ Γ (dtype d)) mfix × - All (fun d => - Σ ;;; Γ ,,, types |- dbody d : (lift0 #|types|) (dtype d)) mfix × + All (on_def_type (lift_typing typing Σ) Γ) mfix × + All (on_def_body (lift_typing typing Σ) types Γ) mfix × wf_fixpoint Σ mfix × Σ ;;; Γ ⊢ dtype decl ≤ T. Proof using wfΣ. @@ -325,10 +327,8 @@ Section Inversion. cofix_guard Σ Γ mfix × let types := fix_context mfix in nth_error mfix idx = Some decl × - All (fun d => isType Σ Γ (dtype d)) mfix × - All (fun d => - Σ ;;; Γ ,,, types |- d.(dbody) : lift0 #|types| d.(dtype) - ) mfix × + All (on_def_type (lift_typing typing Σ) Γ) mfix × + All (on_def_body (lift_typing typing Σ) types Γ) mfix × wf_cofixpoint Σ mfix × Σ ;;; Γ ⊢ decl.(dtype) ≤ T. Proof using wfΣ. @@ -362,16 +362,16 @@ Section Inversion. - simpl. apply ih in h. cbn in h. destruct h as [B [h c]]. apply inversion_LetIn in h as hh. - destruct hh as [s1 [A' [? [? [? ?]]]]]. - exists A'. split ; eauto. + destruct hh as (s1 & A' & ? & ? & ? & ? & ?). + exists A'. split; eauto. cbn. etransitivity; tea. eapply ws_cumul_pb_it_mkProd_or_LetIn_codom. assumption. - simpl. apply ih in h. cbn in h. destruct h as [B [h c]]. apply inversion_Lambda in h as hh. - pose proof hh as [s1 [B' [? [? ?]]]]. - exists B'. split ; eauto. + pose proof hh as (s1 & B' & ? & ? & ? & ?). + exists B'. split; eauto. cbn. etransitivity; tea. eapply ws_cumul_pb_it_mkProd_or_LetIn_codom. assumption. diff --git a/pcuic/theories/PCUICParallelReduction.v b/pcuic/theories/PCUICParallelReduction.v index d86563555..37a92f5cf 100644 --- a/pcuic/theories/PCUICParallelReduction.v +++ b/pcuic/theories/PCUICParallelReduction.v @@ -775,7 +775,7 @@ Section ParallelReduction. constructor. eapply t0. apply on_contexts_app => //. - destruct t1. + destruct t1 as [p p0]. constructor; [eapply p|eapply p0]; apply on_contexts_app => //. Qed. @@ -793,7 +793,7 @@ Section ParallelReduction. constructor. eapply t0. apply on_contexts_app => //. - destruct t1. + destruct t1 as [p p0]. constructor; [eapply p|eapply p0]; apply on_contexts_app => //. Qed. diff --git a/pcuic/theories/PCUICPrincipality.v b/pcuic/theories/PCUICPrincipality.v index 7d0723bc4..02ad31605 100644 --- a/pcuic/theories/PCUICPrincipality.v +++ b/pcuic/theories/PCUICPrincipality.v @@ -92,18 +92,18 @@ Section Principality. repeat outtimes. specialize (IHu1 _ _ t) as [dom Hdom]. specialize (IHu2 _ _ t0) as [codom Hcodom]. - destruct (Hdom _ t) as [e e']. - eapply ws_cumul_pb_Sort_r_inv in e as [domu [red leq]]. - destruct (Hcodom _ t0) as [e e'']. - eapply ws_cumul_pb_Sort_r_inv in e as [codomu [cored coleq]]. + destruct (Hdom _ t) as [le le']. + eapply ws_cumul_pb_Sort_r_inv in le as [domu [red leq]]. + destruct (Hcodom _ t0) as [le le'']. + eapply ws_cumul_pb_Sort_r_inv in le as [codomu [cored coleq]]. exists (tSort (Universe.sort_of_product domu codomu)). int inversion_Prod. repeat outsum; repeat outtimes. + etransitivity. 1: auto. 2:eapply w0. - destruct (Hdom _ t1) as [le' u1']. - eapply ws_cumul_pb_Sort_r_inv in le' as [u' [redu' leu']]. - destruct (Hcodom _ t2) as [le'' u2']. - eapply ws_cumul_pb_Sort_r_inv in le'' as [u'' [redu'' leu'']]. + destruct (Hdom _ t1) as [lle' u1']. + eapply ws_cumul_pb_Sort_r_inv in lle' as [u' [redu' leu']]. + destruct (Hcodom _ t2) as [lle'' u2']. + eapply ws_cumul_pb_Sort_r_inv in lle'' as [u'' [redu'' leu'']]. constructor => //. fvs. constructor. apply leq_universe_product_mon; auto. pose proof (closed_red_confluence red redu') as [v' [redl redr]]. @@ -112,6 +112,7 @@ 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. + + eapply (geq_relevance leq); eauto. + eapply type_reduction; eauto. eapply red. + eapply type_reduction; eauto. eapply cored. @@ -130,7 +131,7 @@ Section Principality. eapply ws_cumul_pb_Prod => //; auto. transitivity A' => //. now symmetry. - - eapply inversion_LetIn in hA as (s1 & bty & Hu2 & Hu1 & Hu3 & Hcum); auto. + - eapply inversion_LetIn in hA as (s1 & bty & e & Hu2 & Hu1 & Hu3 & Hcum); auto. destruct (IHu1 _ _ Hu1) as [? p]. destruct (p _ Hu1). destruct (IHu2 _ _ Hu2) as [? p']. @@ -139,7 +140,7 @@ Section Principality. destruct (p'' _ Hu3). exists (tLetIn n u1 u2 x1). int inversion_LetIn. - destruct hB as (s1' & bty' & Hu2' & Hu1' & Hu3' & Hcum'); eauto. + destruct hB as (s1' & bty' & e' & Hu2' & Hu1' & Hu3' & Hcum'); eauto. etransitivity; eauto. eapply ws_cumul_pb_LetIn; eauto using wt_cumul_pb_refl. now specialize (p'' _ Hu3') as [? ?]. @@ -180,7 +181,8 @@ Section Principality. eapply type_App'. tea. eapply type_reduction; eauto. eapply redA. eapply (type_ws_cumul_pb (pb:=Cumul)); eauto. - { eapply validity in t0; auto. + { eapply isType_of_isTypeRel. + eapply validity in t0; auto. eapply isType_red in t0; [|exact redA]. eapply isType_tProd in t0 as [? ?]; eauto. } transitivity dom' => //. transitivity A''. @@ -272,11 +274,11 @@ Section Principality. assert (consistent_instance_ext Σ (ind_universes x6) u'). { eapply type_reduction in t2. 2:eapply redr. eapply validity in t2; eauto. - destruct t2 as [s Hs]. + destruct t2 as (s & e' & Hs). eapply invert_type_mkApps_ind in Hs. intuition eauto. all:auto. eapply d. } assert (consistent_instance_ext Σ (ind_universes x6) x5). { eapply validity in t1; eauto. - destruct t1 as [s Hs]. + destruct t1 as (s & e' & Hs). eapply invert_type_mkApps_ind in Hs. intuition eauto. all:auto. eapply d. } set (p := {| proj_ind := ind; proj_npars := k; proj_arg := pars |}). transitivity (subst0 (u :: List.rev x0') (subst_instance x5 (proj_type x3))); cycle 1. @@ -479,7 +481,7 @@ Proof. constructor; auto. eapply PCUICArities.isType_Sort; pcuic. apply cumul_Sort. now apply leq_universe_super. - - eapply inversion_Prod in X4 as [s1' [s2' [Ha [Hb Hs]]]]; auto. + - eapply inversion_Prod in X4 as (s1' & s2' & e' & Ha & Hb & Hs); auto. specialize (X1 onu _ _ Ha). specialize (X1 (eq_term_empty_leq_term X5_1)). apply eq_term_empty_eq_term in X5_1. @@ -489,6 +491,7 @@ Proof. 2:{ constructor; eauto. now exists s1. } specialize (X3 onu _ _ Hb X5_2). econstructor; eauto. + rewrite e; auto. apply leq_term_empty_leq_term in X5_2. eapply context_conversion; eauto. constructor; pcuic. constructor; try now symmetry; now constructor. @@ -496,18 +499,18 @@ Proof. constructor; pcuic. constructor. now symmetry. - - eapply inversion_Lambda in X4 as (s & B & dom & codom & cum); auto. + - eapply inversion_Lambda in X4 as (s & B & e' & dom & codom & cum); auto. specialize (X1 onu _ _ dom (eq_term_empty_leq_term X5_1)). apply eq_term_empty_eq_term in X5_1. - assert(conv_context cumulAlgo_gen Σ (Γ ,, vass na ty) (Γ ,, vass n t)). + assert(conv_context cumulAlgo_gen Σ (Γ ,, vass na0 ty) (Γ ,, vass na t)). { repeat constructor; pcuic. } specialize (X3 onu t0 B). forward X3 by eapply context_conversion; eauto; pcuic. eapply (type_ws_cumul_pb (pb:=Conv)). - * econstructor. eauto. instantiate (1 := bty). + * econstructor. 1,2: eauto. instantiate (1 := bty). eapply context_conversion; eauto; pcuic. constructor; pcuic. constructor; pcuic. symmetry; constructor; auto. - * have tyl := type_Lambda _ _ _ _ _ _ _ X0 X2. + * have tyl := type_Lambda _ _ _ _ _ _ _ H X0 X2. now eapply PCUICValidity.validity in tyl. * eapply ws_cumul_pb_Prod; eauto. constructor; auto; fvs. @@ -515,23 +518,23 @@ Proof. eapply type_closed, closedn_on_free_vars in X2. now len in X2; len. - - eapply inversion_LetIn in X6 as (s1' & A & dom & bod & codom & cum); auto. + - eapply inversion_LetIn in X6 as (s1' & A & e' & dom & bod & codom & cum); auto. specialize (X1 onu _ _ dom (eq_term_empty_leq_term X7_2)). specialize (X3 onu _ _ bod (eq_term_empty_leq_term X7_1)). apply eq_term_empty_eq_term in X7_1. apply eq_term_empty_eq_term in X7_2. - assert(Σ ⊢ Γ ,, vdef na t ty = Γ ,, vdef n b b_ty). + assert(Σ ⊢ Γ ,, vdef na0 t ty = Γ ,, vdef na b b_ty). { constructor. eapply ws_cumul_ctx_pb_refl. fvs. constructor => //. constructor; fvs. constructor; fvs. } specialize (X5 onu u A). forward X5 by eapply closed_context_conversion; eauto; pcuic. specialize (X5 X7_3). eapply leq_term_empty_leq_term in X7_3. - have uty : Σ ;;; Γ ,, vdef na t ty |- u : b'_ty. + have uty : Σ ;;; Γ ,, vdef na0 t ty |- u : b'_ty. { eapply closed_context_conversion; eauto. pcuic. now symmetry. } eapply type_ws_cumul_pb. - * econstructor. eauto. eauto. + * econstructor. eauto. eauto. eauto. now instantiate (1 := b'_ty). * eapply PCUICValidity.validity; eauto. econstructor; eauto. @@ -616,7 +619,7 @@ Proof. eapply PCUICEquality.subst_eq_term. eapply PCUICUnivSubstitutionConv.eq_term_upto_univ_subst_instance; eauto; typeclasses eauto. - - eassert (ctx_inst _ _ _ _ _) as Hctxi by now eapply ctx_inst_impl with (1 := X5). + - eassert (ctx_inst _ _ _ _) as Hctxi by now eapply ctx_inst_impl with (1 := X5). assert (isType Σ Γ (mkApps ptm (indices ++ [c]))). { eapply validity. econstructor; eauto. all:split; eauto. solve_all. } @@ -698,7 +701,7 @@ Proof. eapply PCUICValidity.validity; eauto. econstructor. 3:eapply H0. all:eauto. eapply (All_impl X0); pcuicfo. - apply infer_typing_sort_impl with id X2; now intros []. + apply infer_typing_sort_impl with id X2 => //; now intros []. eapply (All_impl X1); pcuicfo; now destruct X2. eapply All2_nth_error in a; eauto. destruct a as [[[eqty _] _] _]. @@ -711,7 +714,7 @@ Proof. eapply PCUICValidity.validity; eauto. eapply type_CoFix. 3:eapply H0. all:eauto. eapply (All_impl X0); pcuicfo. - apply infer_typing_sort_impl with id X2; now intros []. + apply infer_typing_sort_impl with id X2 => //; now intros []. eapply (All_impl X1); pcuicfo; now destruct X2. eapply All2_nth_error in a; eauto. destruct a as [[[eqty _] _] _]. diff --git a/pcuic/theories/PCUICProgress.v b/pcuic/theories/PCUICProgress.v index 0393eef9c..c31e32f6f 100644 --- a/pcuic/theories/PCUICProgress.v +++ b/pcuic/theories/PCUICProgress.v @@ -79,7 +79,7 @@ Section WfEnv. - specialize IHX with (T' := (B {0 := hd})). assert (isType Σ Γ (B {0 := hd})) as HH. { clear p. - eapply inversion_Prod in H' as (? & ? & ? & ? & ?); tea. + eapply inversion_Prod in H' as (? & ? & ? & ? & ? & ?); tea. eapply isType_subst. econstructor. econstructor. rewrite subst_empty; eauto. econstructor; cbn; eauto. } @@ -101,7 +101,9 @@ Proof. revert f T. induction u; intros f T. simpl. intros. { exists T, H, s, HT. intuition pcuic. - econstructor. eexists; eauto. eexists; eauto. eapply isType_ws_cumul_pb_refl. eexists; eauto. } + econstructor. + 3: eapply isType_ws_cumul_pb_refl. + all: eexists; split; [cbnr|]; eauto. } intros Hf Ht. simpl in Hf. specialize (IHu (tApp f a) T). epose proof (IHu Hf) as (T' & H' & s' & H1 & H2 & H3 & H4); tea. @@ -112,14 +114,14 @@ Proof. unshelve econstructor. 5: eauto. 1: eauto. - 3:eapply isType_ws_cumul_pb_refl; eexists; eauto. - 1: eexists; eauto. + 3:eapply isType_ws_cumul_pb_refl; eexists; split; [cbnr|]; eauto. + 1: eexists; split; [cbnr|]; eauto. 1, 2: rewrite <- H2; lia. eapply typing_spine_pred_strengthen; tea. - eexists; eauto. clear Hs3. - eapply inversion_Prod in HA as (? & ? & ? & ? & ?); tea. + eexists; split; [cbnr|]; eauto. clear Hs3. + eapply inversion_Prod in HA as (? & ? & ? & ? & ? & ?); tea. eapply isType_subst. econstructor. econstructor. rewrite subst_empty; eauto. - econstructor; cbn; eauto. + econstructor; cbn; eauto. Unshelve. eauto. Qed. @@ -212,7 +214,7 @@ forall (P : global_env_ext -> context -> term -> term -> Type) wf_local Σ (Γ ,,, predctx) -> PΓ Σ (Γ ,,, predctx) -> is_allowed_elimination Σ idecl.(ind_kelim) ps -> - PCUICTyping.ctx_inst (Prop_conj typing P) Σ Γ (p.(pparams) ++ indices) + PCUICTyping.ctx_inst (Prop_conj typing P Σ) Γ (p.(pparams) ++ indices) (List.rev (subst_instance p.(puinst) (mdecl.(ind_params) ,,, idecl.(ind_indices)))) -> Σ ;;; Γ |- c : mkApps (tInd ci.(ci_ind) p.(puinst)) (p.(pparams) ++ indices) -> P Σ Γ c (mkApps (tInd ci.(ci_ind) p.(puinst)) (p.(pparams) ++ indices)) -> @@ -288,7 +290,7 @@ Proof. eapply typing_closed_context; eauto. eapply type_is_open_term. eapply type_App; eauto. + cbn. inversion HL. subst. clear HL. - eapply inversion_Prod in H' as Hx; eauto. destruct Hx as (? & ? & ? & ? & ?). + eapply inversion_Prod in H' as Hx; eauto. destruct Hx as (? & ? & ? & ? & ? & ?). econstructor. 7: unshelve eapply IHL. now eauto. now eauto. split. now eauto. unshelve eapply IH. eauto. lia. @@ -386,7 +388,7 @@ Lemma typing_ind_env `{cf : checker_flags} : wf_local Σ (Γ ,,, predctx) -> PΓ Σ (Γ ,,, predctx) -> is_allowed_elimination Σ idecl.(ind_kelim) ps -> - PCUICTyping.ctx_inst (Prop_conj typing P) Σ Γ (p.(pparams) ++ indices) + PCUICTyping.ctx_inst (Prop_conj typing P Σ) Γ (p.(pparams) ++ indices) (List.rev (subst_instance p.(puinst) (mdecl.(ind_params) ,,, idecl.(ind_indices)))) -> Σ ;;; Γ |- c : mkApps (tInd ci.(ci_ind) p.(puinst)) (p.(pparams) ++ indices) -> P Σ Γ c (mkApps (tInd ci.(ci_ind) p.(puinst)) (p.(pparams) ++ indices)) -> @@ -557,7 +559,7 @@ Proof. * eapply inversion_Sort in hfn as [? [? cu]]; tea. eapply typing_spine_strengthen in hcum. 3:tea. 2:{ eapply validity; econstructor; eauto. } now eapply typing_spine_sort, app_tip_nil in hcum. - * eapply inversion_Prod in hfn as [? [? [? [? cu]]]]; tea. + * eapply inversion_Prod in hfn as (? & ? & ? & ? & ? & cu); tea. eapply typing_spine_strengthen in hcum. 3:tea. 2:{ eapply validity. econstructor; eauto. } now eapply typing_spine_sort, app_tip_nil in hcum. * (* Lambda *) left. destruct args. diff --git a/pcuic/theories/PCUICSR.v b/pcuic/theories/PCUICSR.v index a94045c04..b01ff0533 100644 --- a/pcuic/theories/PCUICSR.v +++ b/pcuic/theories/PCUICSR.v @@ -416,7 +416,7 @@ Lemma isType_weaken {cf} {Σ Γ Δ T} {wfΣ : wf Σ} : isType Σ (Δ ,,, Γ) T. Proof. intros wfΔ HT. - apply infer_typing_sort_impl with id HT; intros Hs. + apply infer_typing_sort_impl with id HT => //; intros Hs. eapply weaken_ctx => //. Qed. @@ -782,8 +782,8 @@ Qed. *) Lemma ctx_inst_merge {cf} {Σ} {wfΣ : wf Σ} Γ inst inst' Δ : wf_local Σ (Γ ,,, (List.rev Δ)) -> PCUICTyping.ctx_inst - (fun (Σ : global_env_ext) (Γ : context) (t T : term) => - forall u : term, closed_red1 Σ Γ t u -> Σ;;; Γ |- u : T) Σ Γ inst Δ -> + (fun (Γ : context) (t T : term) => + forall u : term, closed_red1 Σ Γ t u -> Σ;;; Γ |- u : T) Γ inst Δ -> ctx_inst Σ Γ inst Δ -> OnOne2 (closed_red1 Σ Γ) inst inst' -> ctx_inst Σ Γ inst' Δ. @@ -829,8 +829,8 @@ Qed. Lemma ctx_inst_merge' {cf} {Σ} {wfΣ : wf Σ} Γ inst inst' Δ : wf_local Σ (Γ ,,, Δ) -> PCUICTyping.ctx_inst - (fun (Σ : global_env_ext) (Γ : context) (t T : term) => - forall u : term, closed_red1 Σ Γ t u → Σ;;; Γ |- u : T) Σ Γ inst (List.rev Δ) -> + (fun (Γ : context) (t T : term) => + forall u : term, closed_red1 Σ Γ t u → Σ;;; Γ |- u : T) Γ inst (List.rev Δ) -> ctx_inst Σ Γ inst (List.rev Δ) -> OnOne2 (closed_red1 Σ Γ) inst inst' -> ctx_inst Σ Γ inst' (List.rev Δ). @@ -1456,7 +1456,7 @@ Lemma isType_expand_lets {cf} {Σ} {wfΣ : wf Σ} {Γ Δ T} : isType Σ (Γ ,,, Δ) T -> isType Σ (smash_context [] Γ ,,, expand_lets_ctx Γ Δ) (expand_lets_k Γ #|Δ| T). Proof. - move=> [] s. + move=> [] s [] e. rewrite -[_ ,,, _]app_context_nil_l app_context_assoc. move/typing_expand_lets_gen; rewrite app_context_nil_l => hs. now exists s. @@ -1467,13 +1467,12 @@ Lemma isType_subst_extended_subst {cf} {Σ} {wfΣ : wf Σ} {Γ Δ T} : isType Σ (smash_context [] Γ ,,, subst_context (extended_subst Γ 0) 0 Δ) (subst (extended_subst Γ 0) #|Δ| T). Proof. - move=> [] s. - intros Hs. + move=> [] s [] e Hs. have wfctx := typing_wf_local Hs. generalize Hs. rewrite -[_ ,,, _]app_context_nil_l app_context_assoc. move/typing_expand_lets_gen; rewrite app_context_nil_l => hs. - exists s. + exists s. split; [apply e|]. rewrite /expand_lets_ctx /expand_lets_k_ctx in hs. rewrite closed_ctx_lift /= in hs. move/closed_wf_local: wfctx. rewrite closedn_ctx_app => /andP[] //. @@ -1546,18 +1545,18 @@ Proof. induction 1. * depelim p. subst. depelim X. constructor. now eapply wf_local_app_inv. - apply infer_typing_sort_impl with id tu; intros _. + apply infer_typing_sort_impl with id tu => //; intros _. now eapply Hs. * depelim X. constructor. now eapply wf_local_app_inv. depelim p. destruct s as [[red <-]|[red <-]]; subst. - apply infer_typing_sort_impl with id tu; intros _. + apply infer_typing_sort_impl with id tu => //; intros _. now eapply Hs. exact tu. red. depelim p. destruct s as [[red <-]|[red <-]]; subst. specialize (Hs _ red). eapply type_ws_cumul_pb; tea. - apply infer_typing_sort_impl with id tu; intros _. - apply Hs. eapply (red_ws_cumul_pb (pb:=Cumul)). + eapply isType_of_isTypeRel. apply infer_typing_sort_impl with id tu => //. + eapply (red_ws_cumul_pb (pb:=Cumul)). now eapply closed_red1_red. now eapply Hc. @@ -1566,7 +1565,7 @@ Proof. pose proof (wf_local_closed_context all). eapply wf_local_app_inv in all as []. eapply wf_local_app in IHX0; tea. - apply infer_typing_sort_impl with id tu; intros Hty. + apply infer_typing_sort_impl with id tu => //; intros Hty. eapply closed_context_conversion; tea. eapply red_one_decl_ws_cumul_ctx_pb => //. eapply OnOne2_local_env_impl; tea. @@ -1574,7 +1573,7 @@ Proof. + constructor; auto. clear X. { eapply wf_local_app_inv in all as []. eapply wf_local_app in IHX0; tea. - apply infer_typing_sort_impl with id tu; intros Hty. + apply infer_typing_sort_impl with id tu => //; intros Hty. eapply closed_context_conversion; tea. eapply red_one_decl_ws_cumul_ctx_pb => //. eapply OnOne2_local_env_impl; tea. @@ -1599,7 +1598,7 @@ Proof. - (* Prod *) constructor; eauto. intuition auto. - unshelve eapply (closed_context_conversion _ typeb); pcuics. + unshelve eapply (closed_context_conversion _ typeb). econstructor; [| eexists; split]; pcuics. constructor. now eapply ws_cumul_ctx_pb_refl, wf_local_closed_context. constructor; auto. eapply red_conv. now eapply closed_red1_red. @@ -1607,19 +1606,19 @@ Proof. - (* Lambda *) eapply type_Cumul_alt. eapply type_Lambda; eauto. intuition auto. - unshelve eapply (closed_context_conversion _ typeb); pcuics. + unshelve eapply (closed_context_conversion _ typeb). econstructor; [| eexists; split]; pcuics. constructor. now eapply ws_cumul_ctx_pb_refl, wf_local_closed_context. constructor; auto. eapply red_conv. now eapply closed_red1_red. - assert (Σ ;;; Γ |- tLambda n t b : tProd n t bty). econstructor; pcuics. + assert (Σ ;;; Γ |- tLambda na t b : tProd na t bty). econstructor; pcuics. now eapply validity in X0. econstructor 3. eapply cumul_refl'. constructor. apply Hu. - (* LetIn body *) eapply type_Cumul_alt. - apply (@substitution_let _ _ _ Γ n b b_ty b' b'_ty typeb'). + apply (@substitution_let _ _ _ Γ na b b_ty b' b'_ty typeb'). specialize (typing_wf_local typeb') as wfd. - assert (Σ ;;; Γ |- tLetIn n b b_ty b' : tLetIn n b b_ty b'_ty). + assert (Σ ;;; Γ |- tLetIn na b b_ty b' : tLetIn na b b_ty b'_ty). econstructor; eauto. eapply (validity X0). eapply cumul_red_r. @@ -1628,10 +1627,10 @@ Proof. - (* LetIn value *) eapply type_Cumul_alt. econstructor; eauto. - unshelve eapply (closed_context_conversion _ typeb'); pcuics. + unshelve eapply (closed_context_conversion _ typeb'). econstructor; [| eexists; split |]; pcuics. constructor. eapply ws_cumul_ctx_pb_refl; eauto. constructor; auto. now eapply red_conv, closed_red1_red. apply ws_cumul_pb_refl; eauto. now repeat inv_on_free_vars. - assert (Σ ;;; Γ |- tLetIn n b b_ty b' : tLetIn n b b_ty b'_ty). econstructor; eauto. + assert (Σ ;;; Γ |- tLetIn na b b_ty b' : tLetIn na b b_ty b'_ty). econstructor; eauto. eapply (validity X0). eapply cumul_red_r. apply cumul_refl'. constructor. apply Hu. @@ -1640,7 +1639,7 @@ Proof. specialize (forall_u _ Hu). eapply type_Cumul_alt. econstructor; eauto. - eapply type_Cumul_alt. eauto. exists s1; auto. + eapply type_Cumul_alt. eauto. exists s1; split; cbnr. auto. apply: red_cumul Hu. unshelve eapply (closed_context_conversion _ typeb'). constructor; pcuic. @@ -1648,7 +1647,7 @@ Proof. constructor; auto. eapply ws_cumul_ctx_pb_refl; eauto. constructor; auto. eapply ws_cumul_pb_refl; eauto; repeat inv_on_free_vars =>//. apply: red_conv Hu. - assert (Σ ;;; Γ |- tLetIn n b b_ty b' : tLetIn n b b_ty b'_ty). econstructor; eauto. + assert (Σ ;;; Γ |- tLetIn na b b_ty b' : tLetIn na b b_ty b'_ty). econstructor; eauto. eapply (validity X0). eapply cumul_red_r. apply cumul_refl'. constructor. apply Hu. @@ -1656,12 +1655,12 @@ Proof. - (* Application *) eapply substitution0; eauto. pose proof typet as typet'. - eapply inversion_Lambda in typet' as [s1 [B' [Ht [Hb HU]]]]=>//. + eapply inversion_Lambda in typet' as (s1 & B' & e & Ht & Hb & HU) => //. apply cumul_Prod_inv in HU as [eqann eqA leqB] => //. pose proof (validity typet) as i. eapply isType_tProd in i as [Hdom Hcodom]; auto. eapply type_ws_cumul_pb; eauto. - unshelve eapply (closed_context_conversion _ Hb); pcuics. + unshelve eapply (closed_context_conversion _ Hb). econstructor. auto. apply Hdom. constructor. now apply ws_cumul_ctx_pb_refl. constructor; auto. - (* Fixpoint unfolding *) @@ -1672,7 +1671,7 @@ Proof. clear e. specialize (inversion_mkApps typet) as [T' [appty spty]]. have vT' := (validity appty). - eapply type_tFix_inv in appty as [T [arg [fn' [[[Hnth wffix] Hty]]]]]; auto. + eapply type_tFix_inv in appty as (T & arg & fn' & [[[Hnth wffix] Hty]]); auto. rewrite H in Hnth. noconf Hnth. eapply type_App; eauto. eapply type_mkApps; [|tea]. @@ -1712,8 +1711,8 @@ Proof. clear forall_u forall_u0 X X0. destruct X4 as [wfcpc IHcpc]. eassert (ctx_inst _ _ _ _) as Hctxi by (eapply ctx_inst_impl with (2 := X5); now intros ? []). - eassert (PCUICEnvTyping.ctx_inst _ _ _ _ _) as IHctxi. - { eapply ctx_inst_impl with (2 := X5). intros ? ? ? ? [? r]; exact r. } + eassert (PCUICEnvTyping.ctx_inst _ _ _ _) as IHctxi. + { eapply ctx_inst_impl with (2 := X5). intros ? ? ? [? r]; exact r. } hide X8. pose proof typec as typec''. unfold iota_red. @@ -2090,8 +2089,8 @@ Proof. - (* Case congruence on a parameter *) eassert (ctx_inst _ _ _ _) as Hctxi by (eapply ctx_inst_impl with (2 := X5); now intros ? []). - eassert (PCUICEnvTyping.ctx_inst _ _ _ _ _) as X6. - { eapply ctx_inst_impl with (2 := X5). intros ? ? ? ? [? r]; exact r. } + eassert (PCUICEnvTyping.ctx_inst _ _ _ _) as X6. + { eapply ctx_inst_impl with (2 := X5). intros ? ? ? [? r]; exact r. } clear X5; rename Hctxi into X5. destruct X0, X4. assert (isType Σ Γ (mkApps (it_mkLambda_or_LetIn (case_predicate_context ci mdecl idecl p) (preturn p)) (indices ++ [c]))). @@ -2171,7 +2170,7 @@ Proof. eapply closedn_smash_context. now eapply declared_inductive_closed_params_inst. } have isty' : isType Σ Γ (mkApps (tInd ci (puinst p)) (params' ++ indices)). - { eexists; eapply isType_mkApps_Ind; tea. } + { eexists; split; [cbnr|]. eapply isType_mkApps_Ind; tea. } have wfcpc' : wf_local Σ (Γ ,,, case_predicate_context ci mdecl idecl (set_pparams p params')). { eapply wf_case_predicate_context; tea. } have eqindices : ws_cumul_pb_terms Σ Γ indices indices. @@ -2245,7 +2244,7 @@ Proof. apply weaken_wf_local => //. eapply on_minductive_wf_params; tea. eapply isdecl. } eapply (type_ws_cumul_pb (pb:=Conv)); tea. + eapply closed_context_conversion; tea. - + exists ps. exact wfcbcty'. + + exists ps. split; [cbnr|]. exact wfcbcty'. + eapply ws_cumul_pb_mkApps; tea. { rewrite /ptm. cbn [preturn set_pparams]. rewrite !lift_it_mkLambda_or_LetIn. @@ -2324,8 +2323,8 @@ Proof. - (* Case congruence on the return clause context *) clear IHHu. destruct X0, X4 as []. eassert (ctx_inst _ _ _ _) as Hctxi by (eapply ctx_inst_impl with (2 := X5); now intros ? []). - eassert (PCUICEnvTyping.ctx_inst _ _ _ _ _) as X6. - { eapply ctx_inst_impl with (2 := X5). intros ? ? ? ? [? r]; exact r. } + eassert (PCUICEnvTyping.ctx_inst _ _ _ _) as X6. + { eapply ctx_inst_impl with (2 := X5). intros ? ? ? [? r]; exact r. } clear X5; rename Hctxi into X5. set (ptm := it_mkLambda_or_LetIn _ _). assert (isType Σ Γ (mkApps ptm (indices ++ [c]))). @@ -2373,8 +2372,8 @@ Proof. - (* Case congruence on discriminee *) destruct X0, X4. eassert (ctx_inst _ _ _ _) as Hctxi by (eapply ctx_inst_impl with (2 := X5); now intros ? []). - eassert (PCUICEnvTyping.ctx_inst _ _ _ _ _) as X6. - { eapply ctx_inst_impl with (2 := X5). intros ? ? ? ? [? r]; exact r. } + eassert (PCUICEnvTyping.ctx_inst _ _ _ _) as X6. + { eapply ctx_inst_impl with (2 := X5). intros ? ? ? [? r]; exact r. } clear X5; rename Hctxi into X5. set (ptm := it_mkLambda_or_LetIn _ _). assert (isType Σ Γ (mkApps ptm (indices ++ [c]))). @@ -2403,8 +2402,8 @@ Proof. - (* Case congruence on branches *) destruct X0, X4. eassert (ctx_inst _ _ _ _) as Hctxi by (eapply ctx_inst_impl with (2 := X5); now intros ? []). - eassert (PCUICEnvTyping.ctx_inst _ _ _ _ _) as X6. - { eapply ctx_inst_impl with (2 := X5). intros ? ? ? ? [? r]; exact r. } + eassert (PCUICEnvTyping.ctx_inst _ _ _ _) as X6. + { eapply ctx_inst_impl with (2 := X5). intros ? ? ? [? r]; exact r. } clear X5; rename Hctxi into X5. eapply type_Case; eauto. econstructor; eauto. econstructor; eauto. @@ -2593,7 +2592,7 @@ Proof. eapply (type_ws_cumul_pb (pb:=Cumul)); eauto. { rewrite firstn_skipn. eapply (isType_subst_instance_decl _ _ _ _ _ u wf isdecl.p1.p1.p1) in projty; eauto. - apply infer_typing_sort_impl with id projty; intros Hs. + apply infer_typing_sort_impl with id projty => //; intros Hs. rewrite /= /map_decl /= in Hs. eapply (weaken_ctx Γ) in Hs; auto. rewrite (subst_app_simpl [_]). @@ -2855,13 +2854,14 @@ Proof. * cbn. intros ???? []; constructor; eauto; now apply ws_cumul_pb_forget_conv. } assert (All (on_def_type (lift_typing typing Σ) Γ) mfix). { apply (All_impl X0); intros d HT. - apply infer_typing_sort_impl with id HT; now intros [Hty _]. } + apply infer_typing_sort_impl with id HT => //; now intros [Hty _]. } assert (All (on_def_type (lift_typing typing Σ) Γ) mfix1). { apply (OnOne2_All_All X2 X0). * intros d HT. - apply infer_typing_sort_impl with id HT; now intros [Hty _]. + apply infer_typing_sort_impl with id HT => //; now intros [Hty _]. * intros d d' [[red _] eq] HT'. - apply infer_typing_sort_impl with id HT'; now intros [_ IH]. } + noconf eq. rewrite /on_def_type -H3. + apply infer_typing_sort_impl with id HT' => //; now intros [_ IH]. } assert (wf_local Σ (Γ ,,, fix_context mfix1)). { apply All_mfix_wf; auto. } destruct (OnOne2_nth_error _ _ _ decl _ X2 heq_nth_error) as [decl' [eqnth disj]]. @@ -2882,14 +2882,14 @@ Proof. eapply context_conversion; eauto. rewrite -fixl. eapply type_Cumul_alt. eapply Hb. - red. simple apply infer_typing_sort_impl with id HT; move=> /= [Hs IH]. + eapply isType_of_isTypeRel; apply infer_typing_sort_impl with id HT => //; move=> /= [Hs IH]. specialize (IH _ red). eapply (weakening _ _ _ _ (tSort _)); auto. apply All_mfix_wf; auto. apply red_cumul, red1_red. eapply (weakening_red1 _ []); auto. 2:eapply red. - pose proof (Hs := HT.π2.1). + pose proof (Hs := HT.π2.2.1). eapply subject_closed in Hs. eapply (closedn_on_free_vars (P:=xpredT)) in Hs. now eapply on_free_vars_any_xpredT. @@ -2897,7 +2897,8 @@ Proof. * eapply wf_fixpoint_red1_type; eauto. eapply OnOne2_impl; tea; cbn; intuition auto. apply a2. apply a2. - * eapply All_nth_error in X3; eauto. + * eapply All_nth_error in X3; tea. + now eapply isType_of_isTypeRel in X3. * apply conv_cumul, conv_sym. destruct disj as [<-|[[red Hred] eq]] => //. reflexivity. eapply PCUICCumulativity.red_conv. apply red1_red, red. @@ -2913,14 +2914,14 @@ Proof. simpl. intros n'; f_equal. apply IHX2. } assert (All (on_def_type (lift_typing typing Σ) Γ) mfix). { apply (All_impl X0); intros d HT. - apply infer_typing_sort_impl with id HT; now intros [Hty _]. } + apply infer_typing_sort_impl with id HT => //; now intros [Hty _]. } assert (All (on_def_type (lift_typing typing Σ) Γ) mfix1). { apply (OnOne2_All_All X2 X0). * intros d HT. - apply infer_typing_sort_impl with id HT; now intros [Hty _]. + apply infer_typing_sort_impl with id HT => //; now intros [Hty _]. * intros d d' [[red _] eq] HT'. - noconf eq. - apply infer_typing_sort_impl with id HT'; intros [Hs IH]. now rewrite -H4. } + noconf eq. rewrite /on_def_type -H3. + apply infer_typing_sort_impl with id HT' => //; intros [Hs IH]. now rewrite -H4. } assert (wf_local Σ (Γ ,,, fix_context mfix1)). { apply All_mfix_wf; auto. } destruct (OnOne2_nth_error _ _ _ decl _ X2 heq_nth_error) as [decl' [eqnth disj]]. @@ -2942,7 +2943,8 @@ Proof. * eapply wf_fixpoint_red1_body; eauto. eapply OnOne2_impl; tea; cbn; intuition auto. apply a2. apply a2. - * eapply All_nth_error in X3; eauto. + * eapply All_nth_error in X3; tea. + now eapply isType_of_isTypeRel in X3. * apply conv_cumul, conv_sym. destruct disj as [<-|[_ eq]]. reflexivity. noconf eq. rewrite H4; reflexivity. @@ -2969,13 +2971,14 @@ Proof. * cbn. intros ???? []; constructor; eauto; now apply ws_cumul_pb_forget_conv. } assert (All (on_def_type (lift_typing typing Σ) Γ) mfix). { apply (All_impl X0); intros d HT. - apply infer_typing_sort_impl with id HT; now intros [Hty _]. } + apply infer_typing_sort_impl with id HT => //; now intros [Hty _]. } assert (All (on_def_type (lift_typing typing Σ) Γ) mfix1). { apply (OnOne2_All_All X2 X0). * intros d HT. - apply infer_typing_sort_impl with id HT; now intros [Hty _]. + apply infer_typing_sort_impl with id HT => //; now intros [Hty _]. * intros d d' [[red _] eq] HT'. - apply infer_typing_sort_impl with id HT'; now intros [_ IH]. } + noconf eq; rewrite /on_def_type -H3. + apply infer_typing_sort_impl with id HT' => //; now intros [_ IH]. } assert (wf_local Σ (Γ ,,, fix_context mfix1)). { apply All_mfix_wf; auto. } destruct (OnOne2_nth_error _ _ _ decl _ X2 heq_nth_error) as [decl' [eqnth disj]]. @@ -2996,14 +2999,14 @@ Proof. eapply context_conversion; eauto. rewrite -fixl. eapply type_Cumul_alt. eapply Hb. - red. simple apply infer_typing_sort_impl with id HT; move=> /= [Hs IH]. + eapply isType_of_isTypeRel; apply infer_typing_sort_impl with id HT => //; move=> /= [Hs IH]. specialize (IH _ red). eapply (weakening _ _ _ _ (tSort _)); auto. apply All_mfix_wf; auto. apply red_cumul, red1_red. eapply (weakening_red1 _ []); auto. 2:eapply red. - pose proof (Hs := HT.π2.1). + pose proof (Hs := HT.π2.2.1). eapply subject_closed in Hs. eapply (closedn_on_free_vars (P:=xpredT)) in Hs. now eapply on_free_vars_any_xpredT. @@ -3011,7 +3014,8 @@ Proof. * eapply (wf_cofixpoint_red1_type _ Γ); eauto. eapply OnOne2_impl; tea; cbn; intuition auto; now eapply closed_red1_red. - * eapply All_nth_error in X3; eauto. + * eapply All_nth_error in X3; tea. + now eapply isType_of_isTypeRel in X3. * apply conv_cumul, conv_sym. destruct disj as [<-|[[red Hred] eq]] => //. reflexivity. eapply PCUICCumulativity.red_conv. apply red1_red, red. @@ -3027,14 +3031,14 @@ Proof. simpl. intros n'; f_equal. apply IHX2. } assert (All (on_def_type (lift_typing typing Σ) Γ) mfix). { apply (All_impl X0); intros d HT. - apply infer_typing_sort_impl with id HT; now intros [Hty _]. } + apply infer_typing_sort_impl with id HT => //; now intros [Hty _]. } assert (All (on_def_type (lift_typing typing Σ) Γ) mfix1). { apply (OnOne2_All_All X2 X0). * intros d HT. - apply infer_typing_sort_impl with id HT; now intros [Hty _]. + apply infer_typing_sort_impl with id HT => //; now intros [Hty _]. * intros d d' [[red _] eq] HT'. - noconf eq. - apply infer_typing_sort_impl with id HT'; intros [Hs IH]. now rewrite -H4. } + noconf eq. rewrite /on_def_type -H3. + apply infer_typing_sort_impl with id HT' => //; intros [Hs IH]. now rewrite -H4. } assert (wf_local Σ (Γ ,,, fix_context mfix1)). { apply All_mfix_wf; auto. } destruct (OnOne2_nth_error _ _ _ decl _ X2 heq_nth_error) as [decl' [eqnth disj]]. @@ -3056,7 +3060,8 @@ Proof. * eapply wf_cofixpoint_red1_body; eauto. eapply OnOne2_impl; tea; cbn; intuition auto. apply a2. apply a2. - * eapply All_nth_error in X3; eauto. + * eapply All_nth_error in X3; tea. + now eapply isType_of_isTypeRel in X3. * apply conv_cumul, conv_sym. destruct disj as [<-|[_ eq]]. reflexivity. noconf eq. rewrite H4; reflexivity. @@ -3125,7 +3130,7 @@ Lemma type_reduction {cf} {Σ} {wfΣ : wf Σ} {Γ t A B} : Proof. intros Ht Hr. eapply type_Cumul_alt. eassumption. - apply infer_typing_sort_impl with id (validity Ht); intros HA. + apply infer_typing_sort_impl with id (validity Ht) => //; intros HA. eapply subject_reduction; eassumption. eapply conv_cumul. now apply PCUICCumulativity.red_conv. Defined. @@ -3173,14 +3178,14 @@ Section SRContext. Lemma wf_local_isType_nth {Σ} {wfΣ : wf Σ} Γ n decl : wf_local Σ Γ -> nth_error Γ n = Some decl -> - ∑ s, Σ ;;; Γ |- lift0 (S n) (decl_type decl) : tSort s. + isType Σ Γ (lift0 (S n) (decl_type decl)). Proof using Type. induction n in Γ, decl |- *; intros hΓ e; destruct Γ; cbn; inversion e; inversion hΓ; subst. - 1,2: rename X0 into H0. + 1,2: apply isType_of_isTypeRel in X0; rename X0 into H0. 3,4: eapply IHn in H0; tas. 3,4: rewrite simpl_lift0. - all: eapply infer_typing_sort_impl with id H0; intros Hs. + all: eapply infer_typing_sort_impl with id H0 => //; intros Hs. all: eapply (weakening _ _ [_] _ (tSort _)); tas. Qed. @@ -3214,18 +3219,18 @@ Section SRContext. induction H in Γ', r, X |- *; depelim r. - constructor; auto. cbn in o. destruct o as [<- r]. - apply infer_typing_sort_impl with id t1; intros Hs. + apply infer_typing_sort_impl with id t1 => //; intros Hs. eapply subject_reduction1 in Hs; eauto. - depelim X. constructor; auto. - apply infer_typing_sort_impl with id t1; intros Hs. + apply infer_typing_sort_impl with id t1 => //; intros Hs. eapply context_conversion; eauto. - depelim X. red in o. simpl in t2. destruct o as [<- [[r ->]|[r <-]]]. constructor; auto. - apply infer_typing_sort_impl with id t1; intros Hs. + apply infer_typing_sort_impl with id t1 => //; intros Hs. eapply subject_reduction1; eauto. eapply type_reduction; tea. pcuic. constructor; auto. @@ -3233,7 +3238,7 @@ Section SRContext. - depelim X. simpl in t2. constructor; auto. - apply infer_typing_sort_impl with id t1; intros Hs. + apply infer_typing_sort_impl with id t1 => //; intros Hs. eapply context_conversion; eauto. red; eapply context_conversion; eauto. - eapply context_conversion; eauto. @@ -3245,18 +3250,19 @@ Section SRContext. induction 1; cbn in *. - destruct p as [-> r]. intro e. inversion e; subst; cbn in *. constructor; tas. - apply infer_typing_sort_impl with id X0; intros Hs. + apply infer_typing_sort_impl with id X0 => //; intros Hs. eapply subject_reduction1; tea. - intro e. inversion e; subst; cbn in *. destruct p as [-> [[? []]|[? []]]]; constructor; cbn; tas. - + apply infer_typing_sort_impl with id X0; intros Hs; eapply subject_reduction1; tea. + + apply infer_typing_sort_impl with id X0 => //; intros Hs; eapply subject_reduction1; tea. + eapply type_Cumul_alt; tea. - apply infer_typing_sort_impl with id X0; intros Hs. + eapply isType_of_isTypeRel. + apply infer_typing_sort_impl with id X0 => //; intros Hs. eapply subject_reduction1; tea. econstructor 2. eassumption. eapply cumul_refl'. + eapply subject_reduction1; tea. - intro H; inversion H; subst; constructor; cbn in *; auto. - 1,2: apply infer_typing_sort_impl with id X1; intros Hs. + 1,2: apply infer_typing_sort_impl with id X1 => //; intros Hs. all: eapply subject_reduction_ctx; tea. Qed. @@ -3314,7 +3320,7 @@ Section SRContext. apply wf_local_app_l in X. inversion X; subst; cbn in *; assumption. } constructor; cbn; auto. - 1: apply infer_typing_sort_impl with id X0; intros Hs. + 1: apply infer_typing_sort_impl with id X0 => //; intros Hs. 1: change (tSort _) with (subst [b] #|Γ'| (tSort X0.π1)). all: eapply PCUICSubstitution.substitution; tea. - rewrite subst_context_snoc0. simpl. @@ -3325,7 +3331,7 @@ Section SRContext. rewrite !subst_empty in XX. apply XX. constructor. apply wf_local_app_l in X. inversion X; subst; cbn in *; assumption. } constructor; cbn; auto. - apply infer_typing_sort_impl with id X0; intros Hs. + apply infer_typing_sort_impl with id X0 => //; intros Hs. change (tSort _) with (subst [b] #|Γ'| (tSort X0.π1)). all: eapply PCUICSubstitution.substitution; tea. Qed. @@ -3344,7 +3350,7 @@ Section SRContext. isType Σ Γ B. Proof using Type. intros HT red. - apply infer_typing_sort_impl with id HT; intros Hs. + apply infer_typing_sort_impl with id HT => //; intros Hs. eapply subject_reduction1; eauto. Qed. @@ -3403,7 +3409,7 @@ Section SRContext. Lemma isType_red {Σ} {wfΣ : wf Σ} {Γ T U} : isType Σ Γ T -> red Σ Γ T U -> isType Σ Γ U. Proof using Type. - intros HT red; apply infer_typing_sort_impl with id HT; intros Hs. + intros HT red; apply infer_typing_sort_impl with id HT => //; intros Hs. eapply subject_reduction; eauto. Qed. @@ -3411,17 +3417,17 @@ End SRContext. Lemma isType_tLetIn {cf} {Σ} {HΣ' : wf Σ} {Γ} {na t A B} : isType Σ Γ (tLetIn na t A B) - <~> (isType Σ Γ A × (Σ ;;; Γ |- t : A) × isType Σ (Γ,, vdef na t A) B). + <~> (isTypeRel Σ Γ A na.(binder_relevance) × (Σ ;;; Γ |- t : A) × isType Σ (Γ,, vdef na t A) B). Proof. split; intro HH. - - destruct HH as [s H]. - apply inversion_LetIn in H; tas. destruct H as [s1 [A' [HA [Ht [HB H]]]]]. - repeat split; tas. 1: eexists; eassumption. + - destruct HH as (s & e & H). + apply inversion_LetIn in H; tas. destruct H as (s1 & A' & e' & HA & Ht & HB & H). + repeat split; tas. 1: eexists; split; eassumption. apply ws_cumul_pb_Sort_r_inv in H. - destruct H as [s' [H H']]. - exists s'. eapply type_reduction; tea. + destruct H as (s' & H & H'). + exists s'. split; [cbnr|]. eapply type_reduction; tea. apply invert_red_letin in H; tas. - destruct H as [[? [? [? [H ? ? ?]]]]|H]. + destruct H as [(? & ? & ? & [H ? ? ?]) | H]. * discriminate. * etransitivity. + exact (@red_rel_all _ (Γ ,, vdef na t A) 0 t A' eq_refl). @@ -3429,10 +3435,10 @@ Proof. erewrite -> on_free_vars_ctx_on_ctx_free_vars. apply H. apply H. apply H. - destruct HH as (HA & Ht & HB). - apply infer_typing_sort_impl with id HB; intros Hs. + apply infer_typing_sort_impl with id HB => //; intros Hs. eapply type_reduction; tas. * econstructor; tea. - apply HA.π2. + all: apply HA.π2. * apply red1_red. apply red_zeta with (b':=tSort _). Defined. diff --git a/pcuic/theories/PCUICSafeLemmata.v b/pcuic/theories/PCUICSafeLemmata.v index 92f2d72b7..a83c24ebc 100644 --- a/pcuic/theories/PCUICSafeLemmata.v +++ b/pcuic/theories/PCUICSafeLemmata.v @@ -344,7 +344,7 @@ Section Lemmata. apply wf_local_app_inv in wfl as (_&wf). apply wf_local_rel_app_inv in wf as (wf&_). destruct h; depelim wf; simpl in *. - all: destruct l; econstructor; eauto. + all: destruct l as (s & e & Hs); econstructor; eauto. Qed. (* todo: rename alpha_eq *) Lemma compare_decls_conv Γ Γ' : @@ -384,38 +384,38 @@ Section Lemmata. all: apply IHπ in h as (?&typ). all: try apply inversion_App in typ as (?&?&?&?&?&?); auto. all: try apply inversion_Proj in typ as (?&?&?&?&?&?&?&?&?); auto. - all: try apply inversion_Prod in typ as (?&?&?&?&?); auto. - all: try apply inversion_Lambda in typ as (?&?&?&?&?); auto. - all: try apply inversion_LetIn in typ as (?&?&?&?&?&?); auto. + all: try apply inversion_Prod in typ as (?&?&?&?&?&?); auto. + all: try apply inversion_Lambda in typ as (?&?&?&?&?&?); auto. + all: try apply inversion_LetIn in typ as (?&?&?&?&?&?&?); auto. all: try solve [econstructor; eauto]. - apply inversion_Fix in typ as (?&?&?&?&?&?&?); eauto. destruct mfix as ((?&[])&?); simpl in *. + eapply All_app in a as (_&a). depelim a. + apply isType_of_isTypeRel in o. eauto using isType_welltyped. + eapply All_app in a0 as (_&a0). depelim a0. - rewrite fix_context_fix_context_alt in t0. - rewrite map_app in t0. - simpl in t0. + rewrite fix_context_fix_context_alt in o. + rewrite map_app in o. rewrite app_context_assoc. econstructor; eauto. - apply inversion_CoFix in typ as (?&?&?&?&?&?&?); eauto. destruct mfix as ((?&[])&?); simpl in *. + eapply All_app in a as (_&a). depelim a. + apply isType_of_isTypeRel in o. eauto using isType_welltyped. + eapply All_app in a0 as (_&a0). depelim a0. - rewrite fix_context_fix_context_alt in t0. - rewrite map_app in t0. - simpl in t0. + rewrite fix_context_fix_context_alt in o. + rewrite map_app in o. rewrite app_context_assoc. econstructor; eauto. - apply inversion_Case in typ as (?&?&?&?&[]&?); auto. rewrite app_context_assoc. destruct p. - + apply validity in scrut_ty as (?&typ). + + apply validity in scrut_ty as (?&?&typ). clear brs_ty. apply inversion_mkApps in typ as (?&_&spine); auto; simpl in *. clear -spine. @@ -598,12 +598,12 @@ Section Lemmata. - simpl. apply ih in h. cbn in h. destruct h as [T h]. apply inversion_LetIn in h as hh ; auto. - destruct hh as [s1 [A' [? [? [? ?]]]]]. + destruct hh as (s1 & A' & ? & ? & ? & ? & ?). exists A'. assumption. - simpl. apply ih in h. cbn in h. destruct h as [T h]. apply inversion_Lambda in h as hh ; auto. - pose proof hh as [s1 [B [? [? ?]]]]. + pose proof hh as (s1 & B & ? & ? & ? & ?). exists B. assumption. Qed. @@ -759,7 +759,7 @@ Section Lemmata. destruct hw as [T hw']. apply inversion_App in hw' as ihw' ; auto. destruct ihw' as [na' [A' [B' [hP [? ?]]]]]. - apply inversion_Prod in hP as [s1 [s2 [? [? bot]]]] ; auto. + apply inversion_Prod in hP as (s1 & s2 & ? & ? & ? & bot) ; auto. apply ws_cumul_pb_Sort_Prod_inv in bot ; auto. Qed. diff --git a/pcuic/theories/PCUICSpine.v b/pcuic/theories/PCUICSpine.v index 9b2df3b68..450d3c12d 100644 --- a/pcuic/theories/PCUICSpine.v +++ b/pcuic/theories/PCUICSpine.v @@ -24,7 +24,7 @@ Implicit Types (cf : checker_flags) (Σ : global_env_ext). Derive Signature for ctx_inst. -Notation ctx_inst := (ctx_inst typing). +Notation ctx_inst Σ := (ctx_inst (typing Σ)). Ltac splits := repeat split. @@ -90,6 +90,7 @@ Lemma wf_local_alpha {cf} {Σ} {wfΣ : wf Σ} Γ Γ' : eq_context_upto_names Γ Proof. induction 1; intros h; depelim h; try constructor; auto. all:depelim r; constructor; subst; auto. + 1,2: rewrite -e. all: eapply lift_typing_impl; tea; intros T Hty. all: eapply context_conversion; eauto. all: now apply eq_context_alpha_conv. @@ -360,11 +361,11 @@ Section WfEnv. destruct u. simpl in Hf, IHu. - edestruct @inversion_App_size with (H0 := Hf) as (na' & A' & B' & s & Hf' & Ha & HA & _ & _ & _ & HA'''); tea. eexists _, _; intuition eauto. - econstructor; eauto with pcuic. + econstructor; eauto with pcuic. all: pcuic. eapply isType_ws_cumul_pb_refl; eexists; eauto. econstructor. all:eauto with pcuic. - eapply inversion_Prod in HA as (? & ? & ? & ? & ?); tea. + eapply inversion_Prod in HA as (? & ? & ? & ? & ? & ?); tea. eapply isType_subst. econstructor. econstructor. rewrite subst_empty; eauto. econstructor; cbn; eauto. - specialize (IHu (tApp f a) T). @@ -372,12 +373,12 @@ Section WfEnv. edestruct @inversion_App_size with (H0 := H') as (na' & A' & B' & s & Hf' & Ha & HA & _ & _ & _ & HA'''); tea. exists (tProd na' A' B'). exists s. intuition; eauto. econstructor; eauto with wf. - 1,2: eexists; eauto. 1:eapply isType_ws_cumul_pb_refl; eexists; eauto. - + 3:eapply isType_ws_cumul_pb_refl. + 1,2,3: eexists; split; [cbnr|eauto]. eapply typing_spine_strengthen; tea. - eapply inversion_Prod in HA as (? & ? & ? & ? & ?); tea. + eapply inversion_Prod in HA as (? & ? & ? & ? & ? & ?); tea. eapply isType_subst. econstructor. econstructor. rewrite subst_empty; eauto. econstructor; cbn; eauto. Unshelve. eauto. @@ -393,10 +394,10 @@ Section WfEnv. destruct a as [na [b|] ty]; rewrite subst_context_snoc /= /subst_decl /map_decl /=; simpl; intuition auto. - 1: apply infer_typing_sort_impl with id a0; intros Hs. + 1: apply infer_typing_sort_impl with id a0 => //; intros Hs. 1: destruct a0 as (? & t); cbn in Hs |- *; clear t. 2: rename b1 into Hs. - 3: rename b into Hs. + 3: rename b0 into Hs. all: rewrite -app_context_assoc in Hs. all: eapply substitution in Hs; eauto. all: rewrite subst_context_app app_context_assoc in Hs. @@ -410,27 +411,29 @@ Section WfEnv. subslet Σ Γ s Δ -> sorts_local_ctx (lift_typing typing) Σ (Γ ,,, subst_context s 0 Γ') (subst_context s #|Γ'| Δ') ctxs. Proof using wfΣ. - induction Δ' in Γ', ctxs |- *; simpl; auto. - destruct a as [na [b|] ty]; rewrite subst_context_snoc /= /subst_decl /map_decl /=; - intuition auto. - + eapply infer_typing_sort_impl with id a0; intros Hs. - destruct a0 as (? & t); cbn in Hs |- *; clear t. + induction Δ' in Γ', ctxs |- *; auto. + intros wfΓΔΓ' X X0. + destruct a as [na [b|] ty]; rewrite subst_context_snoc /subst_decl /map_decl; + [|destruct ctxs; simpl; intuition eauto]; + destruct X as (wgctxs & Ht & Hb); repeat split. + - now apply IHΔ'. + - apply infer_typing_sort_impl with id Ht => //; intros Hs. + destruct Ht as (s' & p); cbn in Hs |- *; clear p. rewrite -app_context_assoc in Hs. eapply substitution in Hs; eauto. rewrite subst_context_app app_context_assoc in Hs. simpl in Hs. rewrite Nat.add_0_r in Hs. now rewrite app_context_length in Hs. - + rewrite -app_context_assoc in b1. - eapply substitution in b1; eauto. - rewrite subst_context_app app_context_assoc Nat.add_0_r in b1. - now rewrite app_context_length in b1. - + destruct ctxs => //. - intuition auto. - rewrite -app_context_assoc in b. - eapply substitution in b; eauto. - rewrite subst_context_app app_context_assoc in b. - rewrite Nat.add_0_r in b. - now rewrite app_context_length in b. + - rewrite -app_context_assoc in Hb. + eapply substitution in Hb; eauto. + rewrite subst_context_app app_context_assoc Nat.add_0_r in Hb. + now rewrite app_context_length in Hb. + - now apply IHΔ'. + - assumption. + - rewrite -app_context_assoc in Hb. + eapply substitution in Hb; eauto. + rewrite subst_context_app app_context_assoc Nat.add_0_r in Hb. + now rewrite app_context_length in Hb. Qed. Lemma wf_arity_spine_typing_spine {Γ T args T'} : @@ -521,7 +524,8 @@ 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 [i _]. + now apply isType_of_isTypeRel in i. symmetry; pcuic. } eapply typing_spine_strengthen in sp. 3:eapply cumulB. @@ -1680,12 +1684,12 @@ Section WfEnv. constructor. auto. rewrite -eql nth_error_app_lt ?app_length /=; try lia. rewrite nth_error_app_ge // ?Nat.sub_diag //. - destruct l0. - exists x. - change (tSort x) with - (subst0 (all_rels c (S #|l|) #|Δ|) (lift #|Δ| #|c| (tSort x))). + destruct l0 as (s & e & Hs). + exists s; split; [cbnr|]. + change (tSort s) with + (subst0 (all_rels c (S #|l|) #|Δ|) (lift #|Δ| #|c| (tSort s))). { eapply (substitution (Γ' := lift_context #|Δ| 0 c) (Δ := [])); cbn; auto. - change (tSort x) with (lift #|Δ| #|c| (tSort x)). + change (tSort s) with (lift #|Δ| #|c| (tSort s)). eapply (weakening_typing); eauto. } eapply ws_cumul_pb_eq_le. simpl. rewrite -{1}eql. simpl. @@ -1702,6 +1706,7 @@ Section WfEnv. rewrite !app_tip_assoc /= in X. rewrite -app_assoc. forward X by auto. + destruct l0 as (s & e & Hs); unfold ",,," in Hs. apply X; auto. all:eauto with fvs. rewrite -app_tip_assoc app_assoc -[(l ++ _) ++ _]app_assoc eql. eapply wf_local_app_inv in Hwf as []. eauto with fvs. @@ -1736,9 +1741,9 @@ Section WfEnv. * specialize (IHΔ _ _ _ h) as (Δs & ts & [sorts IHΔ leq]). exists Δs, ts. pose proof (PCUICWfUniverses.typing_wf_universe _ IHΔ) as wfts. - eapply inversion_LetIn in IHΔ as [s' [? [? [? [? e]]]]]; auto. + eapply inversion_LetIn in IHΔ as (s' & ? & ? & ? & ? & ? & e); auto. splits; eauto. now eexists. - eapply (type_ws_cumul_pb (pb:=Cumul)). eapply t2. apply isType_Sort; pcuic. + eapply (type_ws_cumul_pb (pb:=Cumul)). eapply t2. apply isType_Sort; now pcuic. eapply ws_cumul_pb_LetIn_l_inv in e; auto. eapply ws_cumul_pb_Sort_r_inv in e as [u' [redu' cumu']]. transitivity (tSort u'). @@ -1758,7 +1763,7 @@ Section WfEnv. eapply closed_red_refl; eauto with fvs. * specialize (IHΔ _ _ _ h) as (Δs & ts & [sorts IHΔ leq]). - eapply inversion_Prod in IHΔ as [? [? [? [? e]]]]; tea. + eapply inversion_Prod in IHΔ as (? & ? & ? & ? & ? & e); tea. exists (x :: Δs), x0. splits; tea. eapply ws_cumul_pb_Sort_inv in e. transitivity (sort_of_products Δs ts); auto using leq_universe_product. @@ -2021,6 +2026,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. + eapply isType_of_isTypeRel. now eapply isType_tProd in isty as []. } eapply isType_apply in isty; tea. eapply typing_spine_strengthen in sp. 3:tea. 2:tas. @@ -2043,9 +2049,9 @@ Section WfEnv. now autorewrite with len => <-. Qed. - Arguments ctx_inst_nil {typing} {Σ} {Γ}. - Arguments PCUICTyping.ctx_inst_ass {typing} {Σ} {Γ} {na t i inst Δ}. - Arguments PCUICTyping.ctx_inst_def {typing} {Σ} {Γ} {na b t inst Δ}. + Arguments ctx_inst_nil {typing} {Γ}. + Arguments PCUICTyping.ctx_inst_ass {typing} {Γ} {na t i inst Δ}. + Arguments PCUICTyping.ctx_inst_def {typing} {Γ} {na b t inst Δ}. Lemma spine_subst_ctx_inst_sub {Γ args argsub Δ} (sp : spine_subst Σ Γ args argsub Δ) : ctx_inst_sub (spine_subst_ctx_inst sp) = argsub. @@ -2115,14 +2121,16 @@ 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_of_isTypeRel. + eapply isType_tProd in isty as [dom codom]; cbn in dom. + apply dom. 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. } + { eapply (type_ws_cumul_pb (pb:=Conv)); eauto. eapply isType_of_isTypeRel. eauto. now symmetry. } unshelve eapply (isType_subst (Δ:=[vass na ty]) [hd0]) in codom'. 2:{ now eapply subslet_ass_tip. } specialize (X (subst_context [hd0] 0 Γ0) ltac:(autorewrite with len; lia)). @@ -2396,7 +2404,8 @@ 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. + eapply isType_of_isTypeRel. eauto. depelim p. eapply (PCUICConversion.substitution_ws_cumul_pb (s:=s) (Γ' := Γ) (Γ'' := [])); eauto. Qed. @@ -2866,7 +2875,7 @@ Section WfEnv. rewrite -[context_assumptions Γ0](smash_context_length []); cbn. relativize #|Γ0|. eapply is_open_term_lift. - destruct l0 as [s Hs]. eapply subject_closed in Hs. + destruct l0 as (s & e & Hs). eapply subject_closed in Hs. rewrite is_open_term_closed in Hs. move: Hs. now rewrite !app_length -(All2_fold_length cum). reflexivity. * split; auto. @@ -2994,7 +3003,7 @@ Section WfEnv. rewrite skipn_all_app_eq // in H. noconf H. intros HΔ; depelim HΔ. intros HΔ'; depelim HΔ'. - destruct l0 as [s Hs]. simpl. + destruct l0 as (s & e & Hs). simpl. rewrite (ctx_inst_sub_subst dom) in t1. rewrite firstn_app_left // in dom. specialize (IHa _ dom HΔ HΔ'). @@ -3120,8 +3129,8 @@ Section WfEnv. (Γ ,,, subst_context (List.rev s') 0 Δ)) -> wf_local Σ (Γ ,,, (List.rev Δ)) -> PCUICTyping.ctx_inst - (fun (Σ : global_env_ext) (Γ : context) (t T : term) => - forall u : term, P Σ Γ t u -> Σ;;; Γ |- u : T) Σ Γ inst Δ -> + (fun (Γ : context) (t T : term) => + forall u : term, P Σ Γ t u -> Σ;;; Γ |- u : T) Γ inst Δ -> ctx_inst Σ Γ inst Δ -> OnOne2 (P Σ Γ) inst inst' -> ctx_inst Σ Γ inst' Δ. @@ -3167,8 +3176,8 @@ Section WfEnv. (Γ ,,, subst_context (List.rev s') 0 Δ)) -> wf_local Σ (Γ ,,, (List.rev Δ)) -> PCUICTyping.ctx_inst - (fun (Σ : global_env_ext) (Γ : context) (t T : term) => - forall u : term, P Σ Γ t u -> Σ;;; Γ |- u : T) Σ Γ inst Δ -> + (fun (Γ : context) (t T : term) => + forall u : term, P Σ Γ t u -> Σ;;; Γ |- u : T) Γ inst Δ -> ctx_inst Σ Γ inst Δ -> All2 (P Σ Γ) inst inst' -> ctx_inst Σ Γ inst' Δ. @@ -3220,8 +3229,8 @@ Section WfEnv. Lemma ctx_inst_eq_context {Γ Δ : context} {args args'} : wf_local Σ (Γ ,,, List.rev Δ) -> PCUICTyping.ctx_inst - (fun (Σ : global_env_ext) (Γ : context) (u A : term) => - forall v : term, upto_names' u v -> Σ;;; Γ |- v : A) Σ Γ args Δ -> + (fun (Γ : context) (u A : term) => + forall v : term, upto_names' u v -> Σ;;; Γ |- v : A) Γ args Δ -> ctx_inst Σ Γ args Δ -> All2 upto_names' args args' -> ctx_inst Σ Γ args' Δ. diff --git a/pcuic/theories/PCUICSubstitution.v b/pcuic/theories/PCUICSubstitution.v index 644d7960a..9cf3ae74a 100644 --- a/pcuic/theories/PCUICSubstitution.v +++ b/pcuic/theories/PCUICSubstitution.v @@ -161,8 +161,8 @@ Proof. intros Hq Hf. induction Hq in |- *; try econstructor; eauto; simpl; unfold snoc; rewrite subst_context_snoc; econstructor; eauto. - - simpl. eapply (Hf _ _ Sort). eauto. - - simpl. eapply (Hf _ _ Sort). eauto. + - simpl. eapply (Hf _ _ (SortRel _)). eauto. + - simpl. eapply (Hf _ _ (SortRel _)). eauto. - simpl. eapply (Hf _ _ (Typ t)). eauto. Qed. @@ -347,7 +347,7 @@ Proof. induction 1; auto; unfold subst_context, snoc; rewrite fold_context_k_snoc0; auto; unfold snoc; f_equal; auto; unfold map_decl; simpl. - - destruct t0 as [s Hs]. unfold vass. simpl. f_equal. + - destruct t0 as (s & e & Hs). unfold vass. simpl. f_equal. eapply typed_subst; eauto. lia. - unfold vdef. f_equal. @@ -584,18 +584,17 @@ 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]. specialize (IHl Hl). econstructor; eauto. fold (arities_context l) in *. + apply infer_typing_sort_impl with id (onArity X) => //; intros Hs. unshelve epose proof (weakening Σ [] (arities_context l) _ _ wfΣ _ Hs). 1: now rewrite app_context_nil_l. - simpl in X. eapply (env_prop_typing typecheck_closed) in Hs; eauto. rewrite -> andb_and in Hs. destruct Hs as [Hs Ht]. simpl in Hs. apply (lift_closed #|arities_context l|) in Hs. - rewrite -> Hs, app_context_nil_l in X. simpl. exists s. - apply X. + rewrite -> Hs, app_context_nil_l in X0. + apply X0. Qed. Lemma wf_arities_context {cf:checker_flags} {Σ : global_env} {wfΣ : wf Σ} {mind mdecl} : @@ -610,7 +609,7 @@ Lemma on_constructor_closed_arities {cf:checker_flags} {Σ : global_env} {wfΣ : on_constructor cumulSpec0 (lift_typing typing) (Σ, ind_universes mdecl) mdecl (inductive_ind mind) indices idecl cdecl cs -> closedn #|arities_context (ind_bodies mdecl)| cdecl.(cstr_type). Proof. - intros [? ? [s Hs] _ _ _ _]. + intros [? ? (s & e & Hs) _ _ _ _]. pose proof (typing_wf_local Hs). destruct cdecl as [id cty car]. apply subject_closed in Hs; eauto. @@ -622,7 +621,7 @@ Lemma on_constructor_closed {cf:checker_flags} {Σ : global_env} {wfΣ : wf Σ} (subst_instance u cdecl.(cstr_type)) in closed cty. Proof. - intros [? ? [s Hs] _ _ _ _]. + intros [? ? (s & e & Hs) _ _ _ _]. pose proof (typing_wf_local Hs). destruct cdecl as [id cty car]. apply subject_closed in Hs; eauto. @@ -1343,12 +1342,18 @@ Qed. Section SubstitutionLemmas. Context {cf} {Σ} {wfΣ : wf Σ}. - Lemma isType_closedPT Γ T : isType Σ Γ T -> on_free_vars (closedP #|Γ| xpredT) T. + Lemma isTypeRelOpt_closedPT Γ T relopt : isTypeRelOpt Σ Γ T relopt -> on_free_vars (closedP #|Γ| xpredT) T. Proof using wfΣ. - move/isType_closed/(closedn_on_free_vars (P:=xpred0)). + move/isTypeRelOpt_closed/(closedn_on_free_vars (P:=xpred0)). now rewrite shiftnPF_closedPT. Qed. + + Lemma isType_closedPT Γ T : isType Σ Γ T -> on_free_vars (closedP #|Γ| xpredT) T. + Proof using wfΣ. + apply isTypeRelOpt_closedPT. + Qed. + Lemma wt_cumul_pb_equalityP {le} {Γ : context} {T U} : wt_cumul_pb le Σ Γ T U -> cumulP le Σ (closedP #|Γ| xpredT) Γ T U. Proof using wfΣ. move=> [] dom. @@ -1813,8 +1818,17 @@ Corollary isType_substitution {cf} {Σ} {wfΣ : wf Σ} {Γ Γ' s Δ T} : isType Σ (Γ ,,, Γ' ,,, Δ) T -> isType Σ (Γ ,,, subst_context s 0 Δ) (subst s #|Δ| T). Proof. - intros Hs [s' Ht]. - eapply substitution in Ht; tea. now eexists. + intros Hs (s' & e & Hs'). + eapply substitution in Hs'; tea. now eexists. +Qed. + +Corollary isTypeRel_substitution {cf} {Σ} {wfΣ : wf Σ} {Γ Γ' s rel Δ T} : + subslet Σ Γ s Γ' -> + isTypeRel Σ (Γ ,,, Γ' ,,, Δ) T rel -> + isTypeRel Σ (Γ ,,, subst_context s 0 Δ) (subst s #|Δ| T) rel. +Proof. + intros Hs (s' & e & Hs'). + eapply substitution in Hs'; tea. now eexists. Qed. Corollary substitution_wf_local {cf} {Σ} {wfΣ : wf Σ} {Γ Γ' s Δ} : diff --git a/pcuic/theories/PCUICToTemplateCorrectness.v b/pcuic/theories/PCUICToTemplateCorrectness.v index 3734232bb..1cf4342ff 100644 --- a/pcuic/theories/PCUICToTemplateCorrectness.v +++ b/pcuic/theories/PCUICToTemplateCorrectness.v @@ -690,11 +690,11 @@ Section wtsub. Proof. destruct t; simpl; intros [T h]; try exact tt. - now eapply inversion_Evar in h. - - eapply inversion_Prod in h as (?&?&?&?&?); tea. + - eapply inversion_Prod in h as (?&?&?&?&?&?); tea. split; eexists; eauto. - - eapply inversion_Lambda in h as (?&?&?&?&?); tea. + - eapply inversion_Lambda in h as (?&?&?&?&?&?); tea. split; eexists; eauto. - - eapply inversion_LetIn in h as (?&?&?&?&?&?); tea. + - eapply inversion_LetIn in h as (?&?&?&?&?&?&?); tea. repeat split; eexists; eauto. - eapply inversion_App in h as (?&?&?&?&?&?); tea. split; eexists; eauto. @@ -734,11 +734,11 @@ Section wtsub. eexists; eauto. - eapply inversion_Fix in h as (?&?&?&?&?&?&?); tea. eapply All_prod. - eapply (All_impl a). intros ? h; exact h. + eapply (All_impl a). intros ? h; apply isType_of_isTypeRel in h; exact h. eapply (All_impl a0). intros ? h; eexists; tea. - eapply inversion_CoFix in h as (?&?&?&?&?&?&?); tea. eapply All_prod. - eapply (All_impl a). intros ? h; exact h. + eapply (All_impl a). intros ? h; apply isType_of_isTypeRel in h; exact h. eapply (All_impl a0). intros ? h; eexists; tea. Qed. End wtsub. @@ -1510,8 +1510,8 @@ Proof. - simpl. constructor. - simpl. econstructor. + eapply IHX. - + simpl. destruct tu. exists x. eapply Hs. - - simpl. constructor; auto. red. destruct tu. exists x. auto. + + simpl. destruct tu as (s & e & Hty). exists s. split; [apply e|]; eapply Hs. + - simpl. constructor; auto. red. destruct tu as (s & e & Hty). exists s. split; auto. Qed. Lemma trans_wf_local_env {cf} Σ Γ : @@ -1527,7 +1527,7 @@ Proof. - simpl. constructor. - simpl. econstructor. + eapply IHX. - + simpl. destruct t0. exists x. eapply p. + + simpl. destruct t0 as (s & e & Hs & p). exists s. split; [apply e|]; eapply p. - simpl. constructor; auto. red. destruct t0. exists x. intuition auto. red. red in t1. destruct t1. eapply t2. Qed. @@ -1579,11 +1579,12 @@ Proof. } induction X;cbn;constructor;auto;cbn in *. - - destruct t0 as (?&?&?). + - destruct t0 as (?&e&?&?). exists x. + split;[apply e|]. apply t1. - - destruct t0 as (?&?&?). - eexists;eassumption. + - destruct t0 as (?&?&?&?). + eexists;split;eassumption. - destruct t1. assumption. Qed. @@ -1641,13 +1642,13 @@ Proof. - constructor. + apply IHAll_local_env_over_gen. + cbn in *. - destruct tu. - eexists;split;eassumption. + destruct tu as (s & e & Hty). + eexists;split;[|split]; eassumption. - constructor. + apply IHAll_local_env_over_gen. + cbn in *. - destruct tu. - eexists;split;eassumption. + destruct tu as (s & e & Hty). + eexists;split;[|split];eassumption. + cbn in *. split;eassumption. Qed. @@ -1870,9 +1871,9 @@ Section Typing_Spine_size. Fixpoint typing_spine_size {t T U} (s : typing_spine Σ Γ t T U) : size := match s with | type_spine_eq ty => 0 - | type_spine_nil ty ty' ist cum => typing_size ist.π2 + | type_spine_nil ty ty' ist cum => typing_size ist.π2.2 | type_spine_cons hd tl na A B T B' typrod cumul ty s' => - (max (typing_size typrod.π2) (max (typing_size ty) (typing_spine_size s')))%nat + (max (typing_size typrod.π2.2) (max (typing_size ty) (typing_spine_size s')))%nat end. End Typing_Spine_size. @@ -1901,11 +1902,12 @@ Lemma typing_spine_weaken_concl_size {cf:checker_flags} {Σ Γ T args S S'} (Hs : Σ ;;; Γ ⊢ S ≤ S') (ist : isType Σ Γ S') : typing_spine_size (typing_spine_weaken_concl sp tyT Hs ist) <= - max (typing_spine_size sp) (typing_size ist.π2). + max (typing_spine_size sp) (typing_size ist.π2.2). Proof. induction sp; simpl; auto. lia. rewrite - !Nat.max_assoc. - specialize (IHsp (PCUICArities.isType_apply i t) Hs). lia. + specialize (IHsp (PCUICArities.isType_apply i t) Hs). + cbn in IHsp. lia. Qed. Ltac sig := unshelve eexists. @@ -1920,7 +1922,7 @@ Lemma typing_spine_app {cf:checker_flags} Σ Γ ty args na A B arg (sp : typing_spine Σ Γ ty args (tProd na A B)) (argd : Σ ;;; Γ |- arg : A) : ∑ sp' : typing_spine Σ Γ ty (args ++ [arg]) (B {0 := arg}), - typing_spine_size sp' <= max (typing_size isty.π2) (max (typing_spine_size sp) (typing_size argd)). + typing_spine_size sp' <= max (typing_size isty.π2.2) (max (typing_spine_size sp) (typing_size argd)). Proof. revert arg argd. depind sp. @@ -1937,7 +1939,7 @@ Proof. specialize (IHsp wf isty _ Harg) as [sp' Hsp']. sig. * econstructor. apply i. auto. auto. exact sp'. - * simpl. lia. + * cbn in Hsp' |- *. lia. Qed. (** Likewise, in Template-Coq, we can append without re-typing the result *) @@ -1988,7 +1990,7 @@ Proof. exists fty, d'. split. lia. destruct hd' as [sp' Hsp]. - destruct (typing_spine_app _ _ _ _ _ _ _ _ wfΣ (s; d1) sp' d3) as [appsp Happ]. + destruct (typing_spine_app _ _ _ _ _ _ _ _ wfΣ (s; (I, d1)) sp' d3) as [appsp Happ]. simpl in Happ. move: appsp Happ. rewrite equ firstn_skipn. intros app happ. exists app. lia. @@ -2009,10 +2011,10 @@ Proof. destruct s0 as [sp Hsp]. unshelve eexists. eapply typing_spine_weaken_concl; eauto. exact (PCUICValidity.validity d'). eapply cumulSpec_cumulAlgo_curry in c; eauto; fvs. - exact (s; d2). cbn. + exact (s; (I, d2)). cbn. epose proof (typing_spine_weaken_concl_size wfΣ sp (validity_term wfΣ d') (cumulSpec_cumulAlgo_curry Σ wfΣ Γ A B (typing_closed_context d') (type_is_open_term d1) (subject_is_open_term d2) c) - (s; d2)). simpl in H0. lia. + (s; (I, d2))). simpl in H0. lia. Qed. (** Finally, for each typing spine built above, assuming we can apply the induction hypothesis @@ -2042,7 +2044,7 @@ Proof. split. eapply cumul_decorate in c; tea. now eapply trans_cumul in c. - specialize (IH _ _ i.π2 H). now exists i.π1. + specialize (IH _ _ i.π2.2 H). now exists i.π1. - simpl; intros. forward IHsp. @@ -2052,8 +2054,8 @@ Proof. forward IHt by lia. eapply cumul_decorate in c; tea. apply trans_cumul in c. - specialize (IH _ _ i.π2) as IHi. - forward IHi by lia. + specialize (IH _ _ i.π2.2) as IHi. + forward IHi by cbn in H |- *; lia. simpl in IHi. destruct IHsp as [T' [Hsp eq]]. destruct eq. subst T'. @@ -2279,7 +2281,7 @@ Proof. rewrite -[List.rev (trans_local _)]map_rev. clear. unfold app_context. change subst_instance_context with SE.subst_instance_context. unfold context. rewrite -map_rev. set (ctx := map _ (List.rev _)). clearbody ctx. - intro HH; pose proof (ctx_inst_impl _ (fun _ _ _ _ => TT.typing _ _ _ _ ) _ _ _ _ HH (fun _ _ H => H.2)); revert X; clear HH. + intro HH; pose proof (ctx_inst_impl _ (fun _ _ _ => TT.typing _ _ _ _ ) _ _ _ HH (fun _ _ H => H.2)); revert X; clear HH. now move: ctx; induction 1; cbn; constructor; auto; rewrite -(List.rev_involutive (map trans_decl Δ)) subst_telescope_subst_context -map_rev -(trans_subst_context [_]) -map_rev -PCUICSpine.subst_telescope_subst_context List.rev_involutive. @@ -2343,7 +2345,7 @@ Proof. now eapply TT.All_local_env_app_inv in X as []. + fold trans. eapply All_map, (All_impl X0). - intros x [s ?]; exists s; intuition auto. + intros x (s & e & Hs & Hts); exists s; split; [apply e|]. intuition auto. + fold trans;subst types. now apply trans_mfix_All2. + now rewrite trans_wf_cofixpoint. diff --git a/pcuic/theories/PCUICTyping.v b/pcuic/theories/PCUICTyping.v index 69328e905..01184e20d 100644 --- a/pcuic/theories/PCUICTyping.v +++ b/pcuic/theories/PCUICTyping.v @@ -154,7 +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) - (ind_inst : ctx_inst typing Σ Γ (p.(pparams) ++ indices) + (ind_inst : ctx_inst (typing Σ) Γ (p.(pparams) ++ indices) (List.rev (subst_instance p.(puinst) (ind_params mdecl ,,, ind_indices idecl)))) (not_cofinite : isCoFinite mdecl.(ind_finite) = false). @@ -185,16 +185,19 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> Σ ;;; Γ |- tSort s : tSort (Universe.super s) | type_Prod : forall na A B s1 s2, + isSortRel s1 na.(binder_relevance) -> Σ ;;; Γ |- A : tSort s1 -> Σ ;;; Γ ,, vass na A |- B : tSort s2 -> Σ ;;; Γ |- tProd na A B : tSort (Universe.sort_of_product s1 s2) | type_Lambda : forall na A t s1 B, + isSortRel s1 na.(binder_relevance) -> Σ ;;; Γ |- A : tSort s1 -> Σ ;;; Γ ,, vass na A |- t : B -> Σ ;;; Γ |- tLambda na A t : tProd na A B | type_LetIn : forall na b B t s1 A, + isSortRel s1 na.(binder_relevance) -> Σ ;;; Γ |- B : tSort s1 -> Σ ;;; Γ |- b : B -> Σ ;;; Γ ,, vdef na b B |- t : A -> @@ -248,8 +251,8 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> wf_local Σ Γ -> fix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> - All (fun d => {s & Σ ;;; Γ |- d.(dtype) : tSort s}) mfix -> - All (fun d => (Σ ;;; Γ ,,, fix_context mfix |- d.(dbody) : lift0 #|fix_context mfix| d.(dtype))) mfix -> + All (on_def_type (lift_typing typing Σ) Γ) mfix -> + All (on_def_body (lift_typing typing Σ) (fix_context mfix) Γ) mfix -> wf_fixpoint Σ mfix -> Σ ;;; Γ |- tFix mfix n : decl.(dtype) @@ -257,8 +260,8 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> wf_local Σ Γ -> cofix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> - All (fun d => {s & Σ ;;; Γ |- d.(dtype) : tSort s}) mfix -> - All (fun d => Σ ;;; Γ ,,, fix_context mfix |- d.(dbody) : lift0 #|fix_context mfix| d.(dtype)) mfix -> + All (on_def_type (lift_typing typing Σ) Γ) mfix -> + All (on_def_body (lift_typing typing Σ) (fix_context mfix) Γ) mfix -> wf_cofixpoint Σ mfix -> Σ ;;; Γ |- tCoFix mfix n : decl.(dtype) @@ -360,7 +363,7 @@ Section CtxInstSize. Context {cf} (typing : global_env_ext -> context -> term -> term -> Type) (typing_size : forall {Σ Γ t T}, typing Σ Γ t T -> size). - Fixpoint ctx_inst_size {Σ Γ args Δ} (c : ctx_inst typing Σ Γ args Δ) : size := + Fixpoint ctx_inst_size {Σ Γ args Δ} (c : ctx_inst (typing Σ) Γ args Δ) : size := match c with | ctx_inst_nil => 0 | ctx_inst_ass na t i inst Δ ty ctxi => (typing_size _ _ _ _ ty) + (ctx_inst_size ctxi) @@ -368,11 +371,10 @@ Section CtxInstSize. end. End CtxInstSize. -Definition typing_size `{checker_flags} {Σ Γ t T} (d : Σ ;;; Γ |- t : T) : size. +Fixpoint typing_size `{cf : checker_flags} {Σ Γ t T} (d : Σ ;;; Γ |- t : T) {struct d} : size. Proof. - revert Σ Γ t T d. - fix typing_size 5. - destruct 1. + specialize (typing_size cf). + destruct d. 10: destruct c0, c1. all: repeat match goal with | H : typing _ _ _ _ |- _ => apply typing_size in H @@ -396,9 +398,9 @@ Proof. (Nat.max (ctx_inst_size _ typing_size ind_inst) (Nat.max d2 (Nat.max d3 (branches_size typing_size brs_ty)))))). - exact (S (Nat.max (Nat.max (All_local_env_size typing_size _ _ a) - (all_size _ (fun x p => (infer_sort_size (typing_sort_size typing_size)) Σ _ _ p) a0)) (all_size _ (fun x p => typing_size Σ _ _ _ p) a1))). + (all_size _ (fun x p => (lift_typing_size typing_size) Σ _ _ _ p) a0)) (all_size _ (fun x p => typing_size Σ _ _ _ p) a1))). - exact (S (Nat.max (Nat.max (All_local_env_size typing_size _ _ a) - (all_size _ (fun x p => (infer_sort_size (typing_sort_size typing_size)) Σ _ _ p) a0)) (all_size _ (fun x p => typing_size Σ _ _ _ p) a1))). + (all_size _ (fun x p => (lift_typing_size typing_size) Σ _ _ _ p) a0)) (all_size _ (fun x p => typing_size Σ _ _ _ p) a1))). Defined. Lemma typing_size_pos `{checker_flags} {Σ Γ t T} (d : Σ ;;; Γ |- t : T) : typing_size d > 0. @@ -515,7 +517,7 @@ Lemma type_Cumul' {cf:checker_flags} {Σ Γ t} T {T'} : Σ ;;; Γ |- T <=s T' -> Σ ;;; Γ |- t : T'. Proof. - intros Ht [s Hs] cum. + intros Ht (s & e & Hs) cum. eapply type_Cumul; eauto. Qed. @@ -533,9 +535,7 @@ Lemma typing_wf_local_size `{checker_flags} {Σ} {Γ t T} (d :Σ ;;; Γ |- t : T) : All_local_env_size (@typing_size _) Σ _ (typing_wf_local d) < typing_size d. Proof. - induction d; try destruct c0, c1; simpl; - change (fun (x : global_env_ext) (x0 : context) (x1 x2 : term) - (x3 : x;;; x0 |- x1 : x2) => typing_size x3) with (@typing_size H); try lia. + induction d; try destruct c0, c1; simpl; lia. Qed. Lemma wf_local_inv `{checker_flags} {Σ Γ'} (w : wf_local Σ Γ') : @@ -563,15 +563,17 @@ Proof. destruct w. - simpl. congruence. - intros [=]. subst d Γ0. - exists w. simpl. destruct l. exists x. exists t0. pose (typing_size_pos t0). + destruct l as (s & e & Hs). + exists w. simpl. exists s. exists Hs. pose (typing_size_pos Hs). cbn. split. + lia. + auto with arith. - intros [=]. subst d Γ0. - exists w. simpl. simpl in l. destruct l as [u h]. + destruct l as (s & e & Hs). + exists w. simpl. simpl in l0. - exists u, l0, h. cbn. - pose (typing_size_pos h). + exists s, l0, Hs. cbn. + pose (typing_size_pos Hs). pose (typing_size_pos l0). intuition eauto. all: try lia. @@ -605,29 +607,32 @@ Lemma typing_ind_env_app_size `{cf : checker_flags} : wf_universe Σ u -> P Σ Γ (tSort u) (tSort (Universe.super u))) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : aname) (t b : term) (s1 s2 : Universe.t), + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (t b : term) (s1 s2 : Universe.t), + isSortRel s1 na.(binder_relevance) -> PΓ Σ Γ -> Σ ;;; Γ |- t : tSort s1 -> P Σ Γ t (tSort s1) -> - Σ ;;; Γ,, vass n t |- b : tSort s2 -> - P Σ (Γ,, vass n t) b (tSort s2) -> P Σ Γ (tProd n t b) (tSort (Universe.sort_of_product s1 s2))) -> + Σ ;;; Γ,, vass na t |- b : tSort s2 -> + P Σ (Γ,, vass na t) b (tSort s2) -> P Σ Γ (tProd na t b) (tSort (Universe.sort_of_product s1 s2))) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : aname) (t b : term) + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (t b : term) (s1 : Universe.t) (bty : term), + isSortRel s1 na.(binder_relevance) -> PΓ Σ Γ -> Σ ;;; Γ |- t : tSort s1 -> P Σ Γ t (tSort s1) -> - Σ ;;; Γ,, vass n t |- b : bty -> P Σ (Γ,, vass n t) b bty -> P Σ Γ (tLambda n t b) (tProd n t bty)) -> + Σ ;;; Γ,, vass na t |- b : bty -> P Σ (Γ,, vass na t) b bty -> P Σ Γ (tLambda na t b) (tProd na t bty)) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : aname) (b b_ty b' : term) + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (b b_ty b' : term) (s1 : Universe.t) (b'_ty : term), + isSortRel s1 na.(binder_relevance) -> PΓ Σ Γ -> Σ ;;; Γ |- b_ty : tSort s1 -> P Σ Γ b_ty (tSort s1) -> Σ ;;; Γ |- b : b_ty -> P Σ Γ b b_ty -> - Σ ;;; Γ,, vdef n b b_ty |- b' : b'_ty -> - P Σ (Γ,, vdef n b b_ty) b' b'_ty -> P Σ Γ (tLetIn n b b_ty b') (tLetIn n b b_ty b'_ty)) -> + Σ ;;; Γ,, vdef na b b_ty |- b' : b'_ty -> + P Σ (Γ,, vdef na b b_ty) b' b'_ty -> P Σ Γ (tLetIn na b b_ty b') (tLetIn na b b_ty b'_ty)) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (t : term) na A B u s, PΓ Σ Γ -> @@ -676,7 +681,7 @@ Lemma typing_ind_env_app_size `{cf : checker_flags} : wf_local Σ (Γ ,,, predctx) -> PΓ Σ (Γ ,,, predctx) -> is_allowed_elimination Σ idecl.(ind_kelim) ps -> - ctx_inst (Prop_conj typing P) Σ Γ (p.(pparams) ++ indices) + ctx_inst (Prop_conj typing P Σ) Γ (p.(pparams) ++ indices) (List.rev (subst_instance p.(puinst) (mdecl.(ind_params) ,,, idecl.(ind_indices)))) -> Σ ;;; Γ |- c : mkApps (tInd ci.(ci_ind) p.(puinst)) (p.(pparams) ++ indices) -> P Σ Γ c (mkApps (tInd ci.(ci_ind) p.(puinst)) (p.(pparams) ++ indices)) -> @@ -812,7 +817,8 @@ Proof. eapply type_local_ctx_impl. eapply ind_sorts. intros ??? HT. apply lift_typing_impl with (1 := HT); intros ? Hs. apply (IH (_; _; _; Hs)). - - apply (onIndices Xg). } + - apply Xg.(ind_relevance_compat). + - apply Xg.(onIndices). } { red in onP |- *. apply All_local_env_impl with (1 := onP); intros ??? HT. apply lift_typing_impl with (1 := HT); intros ? Hs. @@ -823,7 +829,7 @@ Proof. typing_size Hty < typing_size H -> Forall_decls_typing P Σ.1 * P Σ Γ t T). { intros. - specialize (IH (existT _ Σ (existT _ wfΣ (existT _ _ (existT _ _ (existT _ _ Hty)))))). + specialize (IH (Σ; wfΣ; _; _; _; Hty)). simpl in IH. forward IH. constructor 2. simpl. apply H0. @@ -839,7 +845,7 @@ Proof. { intros. eapply (XΓ _ _ _ Hwf); auto. clear -Pdecl wfΣ X14 H0. revert Hwf H0. - induction Hwf; cbn in *; try destruct t2 as [u Hu]; simpl in *; constructor. + induction Hwf; cbn in *; try destruct t2 as (u & e & Hu); simpl in *; constructor. - simpl in *. apply IHHwf. lia. - red. apply (X14 _ _ _ Hu). lia. - simpl in *. apply IHHwf. lia. @@ -892,18 +898,12 @@ Proof. eapply (X8 Σ wfΣ Γ (typing_wf_local H0) ci); eauto. ++ eapply (X14 _ _ _ H); eauto. rewrite /predctx; simpl; lia. ++ eapply (X14 _ _ _ H); eauto. rewrite /predctx; simpl; lia. - ++ eapply (Hwf _ wf_pctx). rewrite /predctx; simpl. - change (fun (x : global_env_ext) (x0 : context) (x1 x2 : term) - (x3 : x;;; x0 |- x1 : x2) => typing_size x3) with (@typing_size cf). - lia. + ++ eapply (Hwf _ wf_pctx). rewrite /predctx; simpl; lia. ++ clear -ind_inst X14. assert (forall (Γ' : context) (t T : term) (Hty : Σ;;; Γ' |- t : T), typing_size Hty <= ctx_inst_size _ (@typing_size _) ind_inst -> P Σ Γ' t T). - { intros. eapply (X14 _ _ _ Hty). simpl. - change (fun (x : global_env_ext) (x0 : context) (x1 x2 : term) - (x3 : x;;; x0 |- x1 : x2) => typing_size x3) with (@typing_size cf). - lia. } + { intros. eapply (X14 _ _ _ Hty). simpl; lia. } clear -X ind_inst. revert ind_inst X. generalize (List.rev (subst_instance (puinst p) (ind_params mdecl ,,, ind_indices idecl))). @@ -921,10 +921,7 @@ Proof. ** destruct r0 as [eq [wfcbc [t t0]]]. constructor. --- split; auto. intros brctxty. repeat split. - +++ eapply (Hwf _ wfcbc); eauto. simpl. - change (fun (x : global_env_ext) (x0 : context) (x1 x2 : term) - (x3 : x;;; x0 |- x1 : x2) => typing_size x3) with (@typing_size cf). - lia. + +++ eapply (Hwf _ wfcbc); eauto. simpl; lia. +++ exact t. +++ unshelve eapply (X14 _ _ _ t _); eauto. simpl. lia. @@ -945,39 +942,30 @@ Proof. -- clear X X0 X1 X2 X3 X4 X5 X6 X7 X8 X9 X11 X12. eapply X10; eauto; clear X10. simpl in *. * assert(forall Γ0 (t T : term) (Hty : Σ;;; Γ0 |- t : T), - typing_size Hty < - S - (all_size _ (fun (x : def term) p => typing_size p) a1) -> - PΓ Σ Γ0). - {intros. eapply (Htywf _ _ _ Hty); eauto. lia. } + typing_size Hty < S (all_size _ (fun x p => typing_size p) a1) -> + PΓ Σ Γ0). + { intros. eapply (Htywf _ _ _ Hty); eauto. lia. } destruct mfix. now rewrite nth_error_nil in e. depelim a1. - eapply (X _ _ _ t). simpl. lia. + eapply (X _ _ _ o). simpl. lia. * assert(forall (t T : term) (Hty : Σ;;; Γ |- t : T), - typing_size Hty < - S (all_size (fun x : def term => - ∑ s : Universe.t, Σ;;; Γ |- dtype x : tSort s) - (fun (x : def term) - (p : ∑ s : Universe.t, Σ;;; Γ |- dtype x : tSort s) => - let 'existT s d := p in typing_size d) a0) -> + typing_size Hty < S (all_size _ (fun x '(s; (e, d)) => typing_size d) a0) -> Forall_decls_typing P Σ.1 * P Σ Γ t T). - intros; eauto. eapply (X14 _ _ _ Hty); eauto. simpl. unfold infer_sort. lia. + intros; eauto. eapply (X14 _ _ _ Hty); eauto. simpl in *. unfold infer_sort. cbn. lia. clear Hwf Htywf X14 a pΓ Hdecls. clear -a0 X. induction a0; constructor. - destruct p as [s Hs]. exists s; split; auto. + destruct p as (s & e & Hs). exists s; repeat split; auto. apply (X (dtype x) (tSort s) Hs). simpl. lia. apply IHa0. intros. eapply (X _ _ Hty); eauto. - simpl. lia. + cbn [all_size]. lia. * simpl in X14. assert(forall Γ0 : context, wf_local Σ Γ0 -> - forall (t T : term) (Hty : Σ;;; Γ0 |- t : T), - typing_size Hty < - S - (all_size _ (fun (x : def term) p => typing_size p) a1) -> - Forall_decls_typing P Σ.1 * P Σ Γ0 t T). - {intros. eapply (X14 _ _ _ Hty); eauto. lia. } + forall (t T : term) (Hty : Σ;;; Γ0 |- t : T), + typing_size Hty < S (all_size _ (fun x p => typing_size p) a1) -> + Forall_decls_typing P Σ.1 * P Σ Γ0 t T). + { intros. eapply (X14 _ _ _ Hty); eauto. lia. } clear X14 Hwf Htywf. clear e decl f a0 Hdecls i. remember (fix_context mfix) as mfixcontext. clear Heqmfixcontext. @@ -991,40 +979,30 @@ Proof. -- clear X X0 X1 X2 X3 X4 X5 X6 X7 X8 X9 X10 X12. eapply X11; eauto; clear X11. simpl in *. * assert(forall Γ0 (t T : term) (Hty : Σ;;; Γ0 |- t : T), - typing_size Hty < - S - (all_size _ (fun (x : def term) p => typing_size p) a1) -> - PΓ Σ Γ0). - {intros. eapply (Htywf _ _ _ Hty); eauto. lia. } + typing_size Hty < S (all_size _ (fun x p => typing_size p) a1) -> + PΓ Σ Γ0). + { intros. eapply (Htywf _ _ _ Hty); eauto. lia. } destruct mfix. now rewrite nth_error_nil in e. depelim a1. - eapply (X _ _ _ t). simpl. lia. + eapply (X _ _ _ o). simpl. lia. * assert(forall (t T : term) (Hty : Σ;;; Γ |- t : T), - typing_size Hty < - S (all_size (fun x : def term => - ∑ s : Universe.t, Σ;;; Γ |- dtype x : tSort s) - (fun (x : def term) - (p : ∑ s : Universe.t, Σ;;; Γ |- dtype x : tSort s) => - let 'existT s d := p in typing_size d) a0) -> - Forall_decls_typing P Σ.1 * P Σ Γ t T). - intros; eauto. eapply (X14 _ _ _ Hty); eauto. simpl. unfold infer_sort. lia. + typing_size Hty < S (all_size _ (fun x '(s; (e, d)) => typing_size d) a0) -> + Forall_decls_typing P Σ.1 * P Σ Γ t T). + intros; eauto. eapply (X14 _ _ _ Hty); eauto. simpl in *. unfold infer_sort. cbn. lia. clear Hwf Htywf X14 a pΓ Hdecls. clear -a0 X. induction a0; constructor. - destruct p as [s Hs]. exists s; split; auto. + destruct p as (s & e & Hs). exists s; repeat split; auto. apply (X (dtype x) (tSort s) Hs). simpl. lia. apply IHa0. intros. eapply (X _ _ Hty); eauto. - simpl. lia. + cbn [all_size]. lia. * simpl in X14. assert(forall Γ0 : context, wf_local Σ Γ0 -> - forall (t T : term) (Hty : Σ;;; Γ0 |- t : T), - typing_size Hty < - S - (all_size (fun x : def term => (Σ;;; Γ ,,, fix_context mfix |- dbody x : lift0 #|fix_context mfix| (dtype x))%type) - (fun (x : def term) p => typing_size p) a1) -> - Forall_decls_typing P Σ.1 * P Σ Γ0 t T). - {intros. eapply (X14 _ _ _ Hty); eauto. lia. } + forall (t T : term) (Hty : Σ;;; Γ0 |- t : T), + typing_size Hty < S (all_size _ (fun x p => typing_size p) a1) -> + Forall_decls_typing P Σ.1 * P Σ Γ0 t T). + { intros. eapply (X14 _ _ _ Hty); eauto. lia. } clear X14 Hwf Htywf. clear e decl c a0 Hdecls i. remember (fix_context mfix) as mfixcontext. clear Heqmfixcontext. @@ -1054,29 +1032,32 @@ Lemma typing_ind_env `{cf : checker_flags} : wf_universe Σ u -> P Σ Γ (tSort u) (tSort (Universe.super u))) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : aname) (t b : term) (s1 s2 : Universe.t), + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (t b : term) (s1 s2 : Universe.t), + isSortRel s1 na.(binder_relevance) -> PΓ Σ Γ -> Σ ;;; Γ |- t : tSort s1 -> P Σ Γ t (tSort s1) -> - Σ ;;; Γ,, vass n t |- b : tSort s2 -> - P Σ (Γ,, vass n t) b (tSort s2) -> P Σ Γ (tProd n t b) (tSort (Universe.sort_of_product s1 s2))) -> + Σ ;;; Γ,, vass na t |- b : tSort s2 -> + P Σ (Γ,, vass na t) b (tSort s2) -> P Σ Γ (tProd na t b) (tSort (Universe.sort_of_product s1 s2))) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : aname) (t b : term) + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (t b : term) (s1 : Universe.t) (bty : term), + isSortRel s1 na.(binder_relevance) -> PΓ Σ Γ -> Σ ;;; Γ |- t : tSort s1 -> P Σ Γ t (tSort s1) -> - Σ ;;; Γ,, vass n t |- b : bty -> P Σ (Γ,, vass n t) b bty -> P Σ Γ (tLambda n t b) (tProd n t bty)) -> + Σ ;;; Γ,, vass na t |- b : bty -> P Σ (Γ,, vass na t) b bty -> P Σ Γ (tLambda na t b) (tProd na t bty)) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : aname) (b b_ty b' : term) + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (b b_ty b' : term) (s1 : Universe.t) (b'_ty : term), + isSortRel s1 na.(binder_relevance) -> PΓ Σ Γ -> Σ ;;; Γ |- b_ty : tSort s1 -> P Σ Γ b_ty (tSort s1) -> Σ ;;; Γ |- b : b_ty -> P Σ Γ b b_ty -> - Σ ;;; Γ,, vdef n b b_ty |- b' : b'_ty -> - P Σ (Γ,, vdef n b b_ty) b' b'_ty -> P Σ Γ (tLetIn n b b_ty b') (tLetIn n b b_ty b'_ty)) -> + Σ ;;; Γ,, vdef na b b_ty |- b' : b'_ty -> + P Σ (Γ,, vdef na b b_ty) b' b'_ty -> P Σ Γ (tLetIn na b b_ty b') (tLetIn na b b_ty b'_ty)) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (t : term) na A B u s, PΓ Σ Γ -> @@ -1121,7 +1102,7 @@ Lemma typing_ind_env `{cf : checker_flags} : P Σ (Γ ,,, predctx) p.(preturn) (tSort ps) -> PΓ Σ (Γ ,,, predctx) -> is_allowed_elimination Σ idecl.(ind_kelim) ps -> - ctx_inst (Prop_conj typing P) Σ Γ (p.(pparams) ++ indices) + ctx_inst (Prop_conj typing P Σ) Γ (p.(pparams) ++ indices) (List.rev (subst_instance p.(puinst) (mdecl.(ind_params) ,,, idecl.(ind_indices)))) -> Σ ;;; Γ |- c : mkApps (tInd ci.(ci_ind) p.(puinst)) (p.(pparams) ++ indices) -> P Σ Γ c (mkApps (tInd ci.(ci_ind) p.(puinst)) (p.(pparams) ++ indices)) -> @@ -1362,8 +1343,8 @@ Section All_local_env. Definition on_local_decl_glob (P : term -> typ_or_sort -> Type) d := match d.(decl_body) with - | Some b => P b (Typ d.(decl_type)) × P d.(decl_type) Sort - | None => P d.(decl_type) Sort + | Some b => P b (Typ d.(decl_type)) × P d.(decl_type) (SortRel d.(decl_name).(binder_relevance)) + | None => P d.(decl_type) (SortRel d.(decl_name).(binder_relevance)) end. Definition lookup_wf_local_decl {Γ P} (wfΓ : All_local_env P Γ) (n : nat) @@ -1387,8 +1368,8 @@ Section All_local_env. (wfΓ : wf_local Σ Γ) {d} (H : on_local_decl_glob (lift_typing typing Σ Γ) d) := match d as d' return (on_local_decl_glob (lift_typing typing Σ Γ) d') -> Type with | {| decl_name := na; decl_body := Some b; decl_type := ty |} => - fun H => (P Σ Γ wfΓ b ty H.1 * P Σ Γ wfΓ _ _ (projT2 (snd H)))%type - | {| decl_name := na; decl_body := None; decl_type := ty |} => fun H => P Σ Γ wfΓ _ _ (projT2 H) + fun H => (P Σ Γ wfΓ b ty H.1 * P Σ Γ wfΓ _ _ H.2.π2.2)%type + | {| decl_name := na; decl_body := None; decl_type := ty |} => fun H => P Σ Γ wfΓ _ _ H.π2.2 end H. Lemma nth_error_All_local_env_over {P Σ Γ n decl} (eq : nth_error Γ n = Some decl) {wfΓ : wf_local Σ Γ} : @@ -1421,20 +1402,20 @@ Section All_local_env. Lemma All_local_env_lift_prod_inv : forall Σ P Q Δ, - All_local_env (lift_typing (fun Σ Γ => Trel_conj (P Σ Γ) (Q Σ Γ)) Σ) Δ -> + All_local_env (lift_typing2 P Q Σ) Δ -> All_local_env (lift_typing P Σ) Δ × All_local_env (lift_typing Q Σ) Δ. Proof using Type. intros Σ P Q Δ h. induction h. - split ; constructor. - - destruct IHh. destruct t0 as [? [? ?]]. - split ; constructor ; auto. - + cbn. eexists. eassumption. - + cbn. eexists. eassumption. - - destruct IHh. destruct t0 as [? [? ?]]. destruct t1. - split ; constructor ; auto. - + cbn. eexists. eassumption. - + cbn. eexists. eassumption. + - destruct IHh as [IHP IHQ]. + split; constructor; tas. + all: eapply lift_typing_impl; tea. + all: now intros T [HP HQ]. + - destruct IHh as [IHP IHQ]. + split; constructor; tas. + all: eapply lift_typing_impl; unfold lift_typing2 in t0, t1; tea. + all: now intros T [HP HQ]. Qed. End All_local_env. diff --git a/pcuic/theories/PCUICValidity.v b/pcuic/theories/PCUICValidity.v index 0729150d3..175db0ef9 100644 --- a/pcuic/theories/PCUICValidity.v +++ b/pcuic/theories/PCUICValidity.v @@ -27,7 +27,7 @@ Section Validity. Lemma isType_weaken_full : weaken_env_prop_full cumulSpec0 (lift_typing typing) (fun Σ Γ t T => isType Σ Γ T). Proof using Type. red. intros. - apply infer_typing_sort_impl with id X2; intros Hs. + apply infer_typing_sort_impl with id X2 => //; intros Hs. unshelve eapply (weaken_env_prop_typing _ _ _ _ _ X1 _ _ (Typ (tSort _))); eauto with pcuic. red. simpl. destruct Σ. eapply Hs. Qed. @@ -52,7 +52,7 @@ Section Validity. isType (Σ, φ) Γ t0 -> isType (Σ', φ) Γ t0. Proof using Type. intros wfΣ wfΣ' ext Γ t Hty. - apply infer_typing_sort_impl with id Hty; intros Hs. + apply infer_typing_sort_impl with id Hty => //; intros Hs. eapply (env_prop_typing weakening_env (Σ, φ)); auto. Qed. @@ -70,8 +70,8 @@ Section Validity. Lemma isType_Sort_inv {Σ : global_env_ext} {Γ s} : wf Σ -> isType Σ Γ (tSort s) -> wf_universe Σ s. Proof using Type. - intros wfΣ [u Hu]. - now eapply inversion_Sort in Hu as [? [? ?]]. + intros wfΣ (u & _ & Hu). + now eapply inversion_Sort in Hu as (? & ? & ?). Qed. Lemma isType_subst_instance_decl {Σ Γ T c decl u} : @@ -82,7 +82,7 @@ Section Validity. isType Σ (subst_instance u Γ) (subst_instance u T). Proof using Type. destruct Σ as [Σ φ]. intros X X0 Hty X1. - eapply infer_typing_sort_impl with _ Hty; intros Hs. + eapply infer_typing_sort_impl with _ Hty; [apply relevance_subst_opt|]; intros Hs. eapply (typing_subst_instance_decl _ _ _ (tSort _)); eauto. Qed. @@ -106,7 +106,7 @@ Section Validity. isType Σ Γ T. Proof using Type. intros wfΣ wfΓ HT. - apply infer_typing_sort_impl with id HT; intros Hs. + apply infer_typing_sort_impl with id HT => //; intros Hs. eapply (weaken_ctx (Γ:=[])); eauto. Qed. @@ -208,17 +208,17 @@ Section Validity. (fun (Γ0 : context) (t : term) (T : typ_or_sort) => match T with | Typ T0 => isType Σ (Γ,,, Γ0) T0 × Σ ;;; (Γ ,,, Γ0) |- t : T0 - | Sort => isType Σ (Γ,,, Γ0) t + | Sort relopt => isTypeRelOpt Σ (Γ,,, Γ0) t relopt end) Δ -> ∑ xs, sorts_local_ctx (lift_typing typing) Σ Γ Δ xs. Proof using Type. induction 1. - exists []; cbn; auto. exact tt. - destruct IHX as [xs Hxs]. - destruct t0 as [s Hs]. + destruct t0 as (s & e & Hs). exists (s :: xs). cbn. split; auto. - - destruct IHX as [xs Hxs]. destruct t0 as [s Hs]. - exists xs; cbn. split; auto. + - destruct IHX as [xs Hxs]. destruct t0 as (s & e & Hs), t1 as [t1a t1b]. + exists xs; cbn. repeat split; [| eexists; split |]; tea. Qed. Import PCUICOnFreeVars. @@ -226,7 +226,7 @@ Section Validity. Theorem validity_env : env_prop (fun Σ Γ t T => isType Σ Γ T) (fun Σ Γ => wf_local Σ Γ × All_local_env - (fun Γ t T => match T with Typ T => (isType Σ Γ T × Σ ;;; Γ |- t : T) | Sort => isType Σ Γ t end) Γ). + (fun Γ t T => match T with Typ T => (isType Σ Γ T × Σ ;;; Γ |- t : T) | Sort relopt => isTypeRelOpt Σ Γ t relopt end) Γ). Proof using Type. apply typing_ind_env; intros; rename_all_hyps. @@ -237,40 +237,42 @@ Section Validity. destruct decl as [na [b|] ty]; cbn -[skipn] in *; destruct hd. + eapply isType_lift; eauto. now apply nth_error_Some_length in heq_nth_error. - + eapply isType_lift; eauto. + + destruct p. eapply isType_lift; eauto. now apply nth_error_Some_length in heq_nth_error. now exists x. - (* Universe *) exists (Universe.super (Universe.super u)). + split; [cbnr|]. constructor; auto. now apply wf_universe_super. - (* Product *) - eexists. + eexists; split; [cbnr|]. eapply isType_Sort_inv in X1; eapply isType_Sort_inv in X3; auto. econstructor; eauto. now apply wf_universe_product. - (* Lambda *) - destruct X3 as [bs tybs]. + destruct X3 as (bs & e & tybs). eapply isType_Sort_inv in X1; auto. exists (Universe.sort_of_product s1 bs). + split; [apply e|]. constructor; auto. - (* Let *) - apply infer_typing_sort_impl with id X5; unfold id in *; intros Hs. + apply infer_typing_sort_impl with id X5 => //; unfold id; intros Hs. eapply type_Cumul. eapply type_LetIn; eauto. econstructor; pcuic. eapply convSpec_cumulSpec, red1_cumulSpec; constructor. - (* Application *) - apply infer_typing_sort_impl with id X3; unfold id in *; intros Hs'. + apply infer_typing_sort_impl with id X3 => //; unfold id in *; intros Hs'. move: (typing_wf_universe wf Hs') => wfs. eapply (substitution0 (n := na) (T := tSort _)); eauto. - apply inversion_Prod in Hs' as [na' [s1 [s2 Hs]]]; tas. intuition. - eapply (weakening_ws_cumul_pb (pb:=Cumul) (Γ' := []) (Γ'' := [vass na A])) in b0; pcuic. - simpl in b0. + apply inversion_Prod in Hs' as (s1 & s2 & e & H1 & H2 & H3); tas. + eapply (weakening_ws_cumul_pb (pb:=Cumul) (Γ' := []) (Γ'' := [vass na A])) in H3; pcuic. + simpl in H3. eapply (type_ws_cumul_pb (pb:=Cumul)); eauto. pcuic. etransitivity; tea. eapply into_ws_cumul_pb => //. @@ -294,13 +296,14 @@ Section Validity. destruct (on_declared_inductive isdecl); pcuic. destruct isdecl. apply onArity in o0. + eapply isType_of_isTypeRel in o0. eapply isType_weakening; eauto. eapply (isType_subst_instance_decl (Γ:=[])); eauto. - (* Constructor type *) destruct (on_declared_constructor isdecl) as [[oni oib] [cs [declc onc]]]. unfold type_of_constructor. - eapply infer_typing_sort_impl with _ (on_ctype onc); intros Hs. + eapply infer_typing_sort_impl with _ (on_ctype onc); [apply relevance_subst_opt|]; intros Hs. eapply instantiate_minductive in Hs; eauto. 2:(destruct isdecl as [[] ?]; eauto). simpl in Hs. @@ -322,15 +325,15 @@ Section Validity. eapply spine_subst_smash in X6; tea. destruct X4. destruct (on_declared_inductive isdecl) as [onmind oib]. - rewrite /ptm. exists ps. + rewrite /ptm. exists ps. split; [cbnr|]. eapply type_mkApps; eauto. eapply type_it_mkLambda_or_LetIn; tea. have typred : isType Σ Γ (it_mkProd_or_LetIn predctx (tSort ps)). { eapply All_local_env_app_inv in a0 as [_ onp]. eapply validity_wf_local in onp as [xs Hs]. - eexists _. + eexists. split; [cbnr|]. eapply type_it_mkProd_or_LetIn_sorts; tea. - exact X3.π2. } + exact X3.π2.2. } have wfps : wf_universe Σ ps. { pcuic. } eapply typing_spine_strengthen; tea. @@ -365,7 +368,7 @@ Section Validity. unshelve eapply isType_mkApps_Ind_inv in X2 as [parsubst [argsubst [sppar sparg lenpars lenargs cu]]]; eauto. 2:eapply isdecl.p1. - eapply infer_typing_sort_impl with _ isdecl'; intros Hs. + eapply infer_typing_sort_impl with _ isdecl'; [apply relevance_subst_opt|]; intros Hs. eapply (typing_subst_instance_decl _ _ _ _ _ _ _ wf isdecl.p1.p1.p1) in Hs; eauto. simpl in Hs. eapply (weaken_ctx Γ) in Hs; eauto. @@ -387,11 +390,11 @@ Section Validity. assumption. - (* Fix *) - eapply nth_error_all in X0 as [s Hs]; eauto. + eapply nth_error_all in X0 as (s & e & Hs); eauto. pcuic. - (* CoFix *) - eapply nth_error_all in X0 as [s Hs]; pcuic. + eapply nth_error_all in X0 as (s & e & Hs); pcuic. - (* Conv *) now exists s. @@ -446,7 +449,7 @@ Lemma type_App' {cf:checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} {Γ t na Σ;;; Γ |- u : A -> Σ;;; Γ |- tApp t u : B {0 := u}. Proof. intros Ht Hu. - have [s Hs] := validity Ht. + destruct (validity Ht) as (s & e & Hs). eapply type_App; eauto. Qed. diff --git a/pcuic/theories/PCUICWellScopedCumulativity.v b/pcuic/theories/PCUICWellScopedCumulativity.v index 656788f21..8fb9b92e8 100644 --- a/pcuic/theories/PCUICWellScopedCumulativity.v +++ b/pcuic/theories/PCUICWellScopedCumulativity.v @@ -192,9 +192,14 @@ Arguments wt_cumul_pb_eq {cf c Σ Γ T U}. Section EqualityLemmas. Context {cf : checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ}. + Lemma isTypeRelOpt_open {Γ T relopt} : isTypeRelOpt Σ Γ T relopt -> on_free_vars (shiftnP #|Γ| xpred0) T. + Proof using wfΣ. + move/isTypeRelOpt_closedPT. now rewrite closedP_shiftnP. + Qed. + Lemma isType_open {Γ T} : isType Σ Γ T -> on_free_vars (shiftnP #|Γ| xpred0) T. Proof using wfΣ. - move/isType_closedPT. now rewrite closedP_shiftnP. + apply isTypeRelOpt_open. Qed. Lemma into_ws_cumul_pb {pb} {Γ : context} {T U} : @@ -438,13 +443,13 @@ 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' -> eq_binder_annot na na' -> + isTypeRel Σ Γ T na.(binder_relevance) -> isTypeRel Σ Γ' T' na'.(binder_relevance) -> + 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'} : eq_binder_annot na na' -> - isType Σ Γ T -> isType Σ Γ' T' -> + isTypeRel Σ Γ T na.(binder_relevance) -> isTypeRel Σ Γ' T' na'.(binder_relevance) -> Σ ;;; Γ |- b : T -> Σ ;;; Γ' |- b' : T' -> Σ ;;; Γ |- b = b' -> conv_cum pb Σ Γ T T' -> @@ -497,8 +502,8 @@ Section WtContextConversion. 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 + | {| decl_name := na; decl_body := None; decl_type := ty |} => isTypeRel Σ Γ ty na.(binder_relevance) + | {| decl_name := na; decl_body := Some b; decl_type := ty |} => isTypeRel Σ Γ ty na.(binder_relevance) × Σ ;;; Γ |- b : ty end. Lemma wf_local_All_fold Γ : @@ -507,7 +512,7 @@ Section WtContextConversion. Proof using Type. split. - induction 1; constructor; auto. - red in t0, t1. cbn. split; auto. + now split. - induction 1; [constructor|]. destruct d as [na [b|] ty]; cbn in p; constructor; intuition auto. Qed. @@ -547,15 +552,15 @@ Section WtContextConversion. intros ???? wt ws eq; pose proof (All2_fold_length wt). destruct eq. - - pose proof (isType_wf_local i). + - pose proof (isType_wf_local (isType_of_isTypeRel i)). eapply wf_local_closed_context in X. - eapply isType_open in i. apply isType_open in i0. + eapply isTypeRelOpt_open in i. apply isTypeRelOpt_open in i0. eapply into_ws_cumul_decls with Δ; eauto with fvs. constructor; auto. rewrite (All2_fold_length ws) //. - - pose proof (isType_wf_local i). + - pose proof (isType_wf_local (isType_of_isTypeRel i)). eapply wf_local_closed_context in X. - eapply isType_open in i. apply isType_open in i0. + eapply isTypeRelOpt_open in i. apply isTypeRelOpt_open in i0. eapply PCUICClosedTyp.subject_closed in t. eapply PCUICClosedTyp.subject_closed in t0. eapply (@closedn_on_free_vars xpred0) in t. diff --git a/pcuic/theories/PCUICWfUniverses.v b/pcuic/theories/PCUICWfUniverses.v index 04b559f69..2f3fb20f8 100644 --- a/pcuic/theories/PCUICWfUniverses.v +++ b/pcuic/theories/PCUICWfUniverses.v @@ -584,7 +584,7 @@ Qed. destruct X as [? [? ?]]. now to_prop. apply IHΔ. apply X. simpl. - destruct X as [? ?]. now to_prop. + destruct X as (? & ? & ?). now to_prop. apply IHΔ. apply X. Qed. @@ -648,7 +648,7 @@ Qed. - intros [h [h' h'']]. destruct n. simpl. move=> [= <-] /=. do 2 eexists; eauto. now simpl; eapply IHΔ. - - destruct s => //. intros [h h']. + - destruct s => //. intros (h & e & h'). destruct n. simpl. move=> [= <-] /=. eexists; eauto. now simpl; eapply IHΔ. Qed. @@ -1029,10 +1029,7 @@ Qed. - split. * induction X; constructor; auto. - destruct tu as [s tu]; exists s; simpl. - now simpl in Hs. - destruct tu as [s tu]; exists s; simpl. - now simpl in Hs. + all: apply infer_typing_sort_impl with id tu => //. * induction X; simpl; auto. rewrite IHX /= /test_decl /=. now move/andP: Hs. rewrite IHX /= /test_decl /=. now move/andP: Hc => [] -> ->. @@ -1044,7 +1041,7 @@ Qed. rewrite heq_nth_error /= in X. red in X. destruct decl as [na [b|] ty]; cbn -[skipn] in *. + now to_prop. - + destruct X as [s Hs]. now to_prop. + + now move: X => [s [] _ /andP[Hc _]]. - apply/andP; split; to_wfu; cbn ; eauto with pcuic. @@ -1064,7 +1061,7 @@ Qed. eapply wf_universes_inst. 2:eauto. all:eauto. simpl in H2. now eapply consistent_instance_ext_wf. - * move: X1 => [s /andP[Hc _]]. + * move: X1 => [s [] _ /andP[Hc _]]. to_prop. eapply wf_universes_inst; eauto. exact (weaken_lookup_on_global_env' Σ.1 _ _ wf H). @@ -1076,7 +1073,7 @@ Qed. eapply consistent_instance_ext_wf; eauto. } pose proof (declared_inductive_inv wf_universes_weaken wf X isdecl). cbn in X1. eapply onArity in X1. cbn in X1. - move: X1 => [s /andP[Hind ?]]. + move: X1 => [s [] _ /andP[Hind ?]]. eapply wf_universes_inst; eauto. exact (weaken_lookup_on_global_env' Σ.1 _ _ wf (proj1 isdecl)). now eapply consistent_instance_ext_wf. @@ -1091,7 +1088,7 @@ Qed. { apply wf_universes_inds. now eapply consistent_instance_ext_wf. } eapply on_ctype in onc. cbn in onc. - move: onc=> [_ /andP[onc _]]. + move: onc=> [_ [] _ /andP[onc _]]. eapply wf_universes_inst; eauto. exact (weaken_lookup_on_global_env' Σ.1 _ _ wf (proj1 (proj1 isdecl))). now eapply consistent_instance_ext_wf. @@ -1103,8 +1100,7 @@ Qed. rewrite wfu /= wfpars wf_universes_mkApps /= forallb_app wfinds /= H /= !andb_true_r. pose proof (declared_inductive_inv wf_universes_weaken wf X isdecl). - destruct X5. destruct onArity as [s Hs]. - move/andP: Hs => [] /= hty hs. + destruct X5. destruct onArity as (s & e & (hty & hs)%andb_and). rewrite ind_arity_eq in hty. rewrite !wf_universes_it_mkProd_or_LetIn in hty. move/and3P: hty => [] wfp wfindis wfisort. @@ -1123,8 +1119,7 @@ Qed. (ind_ctors idecl). { clear -wf ond onConstructors. red in onConstructors. solve_all. destruct X. - do 2 red in on_ctype. destruct on_ctype as [s Hs]. - move/andP: Hs => [] wfty _. + destruct on_ctype as (s & e & (wfty & _)%andb_and). rewrite cstr_eq in wfty. rewrite !wf_universes_it_mkProd_or_LetIn in wfty. move/and3P: wfty => [] _ clargs _. @@ -1179,24 +1174,24 @@ Qed. noconf heq. simpl. rewrite wf_universes_subst. apply wf_extended_subst. - rewrite ind_arity_eq in onArity. destruct onArity as [s' Hs]. + rewrite ind_arity_eq in onArity. destruct onArity as (s' & e & Hs). rewrite wf_universes_it_mkProd_or_LetIn in Hs. now move/andP: Hs => /andP /andP [] /andP []. rewrite wf_universes_lift. eapply wf_sorts_local_ctx_smash in s. eapply wf_sorts_local_ctx_nth_error in s as [? [? H]]; eauto. red in H. destruct x0. now move/andP: H => []. - now destruct H as [s [Hs _]%andb_and]. + now destruct H as (s & e & (Hs & _)%andb_and). - apply/andP; split; auto. solve_all; destruct a0 as (? & _ & ?), b0; rtoProp; tas. - eapply nth_error_all in X0; eauto. - simpl in X0. now move: X0 => [s [Hty /andP[wfty _]]]. + now rewrite wf_universes_lift in H2. + eapply nth_error_all in X0 as (s & e & Hty & (wfty & _)%andb_and); eauto. - apply/andP; split; auto. solve_all; destruct a0 as (? & _ & ?), b0; rtoProp; tas. - eapply nth_error_all in X0; eauto. - simpl in X0. now move: X0 => [s [Hty /andP[wfty _]]]. + now rewrite wf_universes_lift in H2. + eapply nth_error_all in X0 as (s & e & Hty & (wfty & _)%andb_and); eauto. Qed. Lemma typing_wf_universes {Σ : global_env_ext} {Γ t T} : @@ -1218,7 +1213,7 @@ Qed. Lemma isType_wf_universes {Σ Γ T} : wf Σ.1 -> isType Σ Γ T -> wf_universes Σ T. Proof using Type. - intros wfΣ [s Hs]. now eapply typing_wf_universes in Hs as [HT _]%andb_and. + intros wfΣ (s & e & Hs). now eapply typing_wf_universes in Hs as [HT _]%andb_and. Qed. End CheckerFlags. diff --git a/pcuic/theories/Syntax/PCUICDepth.v b/pcuic/theories/Syntax/PCUICDepth.v index f8c5ac716..81d87aa4a 100644 --- a/pcuic/theories/Syntax/PCUICDepth.v +++ b/pcuic/theories/Syntax/PCUICDepth.v @@ -255,7 +255,7 @@ Proof. Defined.*) Definition onctx_rel (P : context -> term -> Type) (Γ Δ : context) := - All_local_env (PCUICInduction.on_local_decl (fun Δ => P (Γ ,,, Δ))) Δ. + All_local_env (lift_wf_term (fun Δ => P (Γ ,,, Δ))) Δ. Definition CasePredProp (P : context -> term -> Type) Γ (p : predicate term) := All (P Γ) p.(pparams) × onctx_rel P Γ (inst_case_context p.(pparams) p.(puinst) p.(pcontext)) × @@ -330,10 +330,10 @@ Lemma term_forall_ctx_list_ind : P Γ (tCase ci p t brs)) -> (forall Γ (s : projection) (t : term), P Γ t -> P Γ (tProj s t)) -> (forall Γ (m : mfixpoint term) (n : nat), - All_local_env (PCUICInduction.on_local_decl (fun Γ' t => P (Γ ,,, Γ') t)) (fix_context m) -> + All_local_env (lift_wf_term (ctx_shifted P Γ)) (fix_context m) -> tFixProp (P Γ) (P (Γ ,,, fix_context m)) m -> P Γ (tFix m n)) -> (forall Γ (m : mfixpoint term) (n : nat), - All_local_env (PCUICInduction.on_local_decl (fun Γ' t => P (Γ ,,, Γ') t)) (fix_context m) -> + All_local_env (lift_wf_term (ctx_shifted P Γ)) (fix_context m) -> tFixProp (P Γ) (P (Γ ,,, fix_context m)) m -> P Γ (tCoFix m n)) -> (* (forall Γ p, P Γ (tPrim p)) -> *) forall Γ (t : term), P Γ t. @@ -368,12 +368,12 @@ Proof. - case: a => [na [b|] ty] /=; rewrite {1}/decl_depth_gen /context_depth_gen /= => Hlt; constructor; auto. + eapply IHΔ => //. unfold context_depth. lia. - + simpl. apply aux => //. red. lia. + + simpl. apply aux => //; red. lia. + simpl. split. - * apply aux => //. red. lia. - * apply aux=> //; red; lia. + * apply aux => //; red. lia. + * apply aux => //; red; lia. + apply IHΔ => //; unfold context_depth; lia. - + apply aux => //. red. lia. } + + simpl. apply aux => //; red. lia. } assert (forall m, list_depth_gen (fun x : def term => depth (dtype x)) m < S (mfixpoint_depth_gen depth m)). { clear. unfold mfixpoint_depth_gen, def_depth_gen. induction m. simpl. auto. simpl. lia. } assert (forall m, list_depth_gen (fun x : def term => depth (dbody x)) m < S (mfixpoint_depth_gen depth m)). diff --git a/pcuic/theories/Syntax/PCUICInduction.v b/pcuic/theories/Syntax/PCUICInduction.v index d21e9d86e..b65fcb5ae 100644 --- a/pcuic/theories/Syntax/PCUICInduction.v +++ b/pcuic/theories/Syntax/PCUICInduction.v @@ -435,13 +435,6 @@ Proof. f_equal; symmetry; apply size_lift. Qed. -Definition on_local_decl (P : context -> term -> Type) - (Γ : context) (t : term) (T : typ_or_sort) := - match T with - | Typ T => (P Γ t * P Γ T)%type - | Sort => P Γ t - end. - (* TODO: remove List.rev *) Lemma list_size_rev {A} size (l : list A) : list_size size (List.rev l) = list_size size l. @@ -451,7 +444,7 @@ Proof. Qed. Definition onctx_rel (P : context -> term -> Type) (Γ Δ : context) := - All_local_env (on_local_decl (fun Δ => P (Γ ,,, Δ))) Δ. + All_local_env (lift_wf_term (ctx_shifted P Γ)) Δ. Definition CasePredProp (P : context -> term -> Type) Γ (p : predicate term) := All (P Γ) p.(pparams) × onctx_rel P Γ (pcontext p) × @@ -482,10 +475,10 @@ Lemma term_forall_ctx_list_ind : P Γ (tCase ci p t brs)) -> (forall Γ (s : projection) (t : term), P Γ t -> P Γ (tProj s t)) -> (forall Γ (m : mfixpoint term) (n : nat), - All_local_env (on_local_decl (fun Γ' t => P (Γ ,,, Γ') t)) (fix_context m) -> + All_local_env (lift_wf_term (fun Γ' t => P (Γ ,,, Γ') t)) (fix_context m) -> tFixProp (P Γ) (P (Γ ,,, fix_context m)) m -> P Γ (tFix m n)) -> (forall Γ (m : mfixpoint term) (n : nat), - All_local_env (on_local_decl (fun Γ' t => P (Γ ,,, Γ') t)) (fix_context m) -> + All_local_env (lift_wf_term (fun Γ' t => P (Γ ,,, Γ') t)) (fix_context m) -> tFixProp (P Γ) (P (Γ ,,, fix_context m)) m -> P Γ (tCoFix m n)) -> (* (forall Γ p, P Γ (tPrim p)) -> *) forall Γ (t : term), P Γ t. @@ -522,9 +515,9 @@ Proof. + simpl. apply aux => //. red. lia. + simpl. split. * apply aux => //. red. lia. - * apply aux=> //; red; lia. + * apply aux => //; red; lia. + apply IHΔ => //; unfold context_size; lia. - + apply aux => //. red. lia. } + + simpl. apply aux => //. red. lia. } assert (forall m, list_size (fun x : def term => size (dtype x)) m < S (mfixpoint_size size m)). { clear. unfold mfixpoint_size, def_size. induction m. simpl. auto. simpl. lia. } assert (forall m, list_size (fun x : def term => size (dbody x)) m < S (mfixpoint_size size m)). diff --git a/pcuic/theories/TemplateToPCUICCorrectness.v b/pcuic/theories/TemplateToPCUICCorrectness.v index 892874926..33db00d38 100644 --- a/pcuic/theories/TemplateToPCUICCorrectness.v +++ b/pcuic/theories/TemplateToPCUICCorrectness.v @@ -326,7 +326,7 @@ Lemma weaken_wf_decl_pred {cf} (Σ Σ' : Ast.Env.global_env) Γ t T : Proof. intros wf ext wf' ong. red in ong |- *. - destruct T; intuition eauto using wf_extends. + destruct T; cbn in ong |- *; intuition eauto using wf_extends. Qed. Lemma trans_lookup {cf} Σ cst : @@ -908,14 +908,14 @@ Section Trans_Global. WfAst.wf Σ (dbody def) ) m ). - { solve_all. } + { solve_all; destruct a0; auto. } assert ( w2 : All (fun def => WfAst.wf Σ (dtype def) × WfAst.wf Σ (dbody def)) mfix' ). - { solve_all. } + { solve_all; destruct b; auto. } pose proof (All2_All_mix_left X X0) as h1. simpl in h1. pose proof (All2_All_mix_left w1 h1) as h2. pose proof (All2_All_mix_right w2 h2) as h3. @@ -931,12 +931,12 @@ Section Trans_Global. w1 : All (fun def => WfAst.wf Σ (dtype def) × WfAst.wf Σ (dbody def)) m ). - { solve_all. } + { solve_all; destruct a0; auto. } assert ( w2 : All (fun def => WfAst.wf Σ (dtype def) * WfAst.wf Σ (dbody def)) mfix' ). - { solve_all. } + { solve_all; destruct b; auto. } pose proof (All2_All_mix_left X X0) as h1. simpl in h1. pose proof (All2_All_mix_left w1 h1) as h2. pose proof (All2_All_mix_right w2 h2) as h3. @@ -1624,28 +1624,32 @@ Section Trans_Global. solve_all. red. rewrite <- !map_dtype. rewrite <- !map_dbody. inversion b0. clear b0. + split. destruct b. intuition eauto. unfold map_def. simpl. congruence. - apply fix_red_body. apply OnOne2_map. repeat toAll. apply (OnOne2_All_mix_left Hwf) in X. solve_all. - red. rewrite <- !map_dtype. rewrite <- !map_dbody. intuition eauto. - + unfold Ast.Env.app_context, trans_local in b. - simpl in a. rewrite -> map_app in b. - unfold app_context. unfold ST.fix_context in b. - rewrite map_rev map_mapi in b. simpl in b. + red. rewrite <- !map_dtype. rewrite <- !map_dbody. + split. destruct b. + intuition eauto. + + unfold Ast.Env.app_context, trans_local in b1. + simpl in a. rewrite -> map_app in b1. + unfold app_context. unfold ST.fix_context in b1. + rewrite map_rev map_mapi in b1. simpl in b1. unfold fix_context. rewrite mapi_map. simpl. - forward b. - { clear b. solve_all. eapply All_app_inv; auto. + forward b1. + { clear b1. solve_all. eapply All_app_inv; auto. apply All_rev. apply All_mapi. clear -Hwf; generalize 0 at 2; induction mfix0; constructor; hnf; simpl; auto. intuition auto. depelim a. simpl. depelim Hwf. simpl in *. intuition auto. + destruct w. now eapply WfAst.wf_lift. depelim a. simpl. depelim Hwf. simpl in *. intuition auto. } - forward b by auto. - eapply (refine_red1_Γ); [|apply b]. + forward b1 by auto. + eapply (refine_red1_Γ); [|apply b1]. f_equal. f_equal. apply mapi_ext; intros [] []. rewrite lift0_p. simpl. rewrite LiftSubst.lift0_p. reflexivity. cbn. rewrite /trans_decl /= /vass. f_equal. @@ -1657,28 +1661,29 @@ Section Trans_Global. solve_all. red. rewrite <- !map_dtype. rewrite <- !map_dbody. inversion b0. clear b0. - intuition eauto. + destruct b; intuition eauto. unfold map_def. simpl. congruence. - apply cofix_red_body. apply OnOne2_map. repeat toAll. apply (OnOne2_All_mix_left Hwf) in X. - solve_all. + solve_all. destruct b. red. rewrite <- !map_dtype. rewrite <- !map_dbody. intuition eauto. - + unfold Ast.Env.app_context, trans_local in b. - simpl in a. rewrite -> map_app in b. - unfold app_context. unfold ST.fix_context in b. - rewrite map_rev map_mapi in b. simpl in b. + + unfold Ast.Env.app_context, trans_local in b1. + simpl in a. rewrite -> map_app in b1. + unfold app_context. unfold ST.fix_context in b1. + rewrite map_rev map_mapi in b1. simpl in b1. unfold fix_context. rewrite mapi_map. simpl. - forward b. - { clear b. solve_all. eapply All_app_inv; auto. + forward b1. + { clear b1. solve_all. eapply All_app_inv; auto. apply All_rev. apply All_mapi. clear -Hwf; generalize 0 at 2; induction mfix0; constructor; hnf; simpl; auto. intuition auto. depelim a. simpl. depelim Hwf. simpl in *. intuition auto. + destruct w. now eapply WfAst.wf_lift. depelim a. simpl. depelim Hwf. simpl in *. intuition auto. } - forward b by auto. - eapply (refine_red1_Γ); [|apply b]. + forward b1 by auto. + eapply (refine_red1_Γ); [|apply b1]. f_equal. f_equal. apply mapi_ext; intros [] []. rewrite lift0_p. simpl. rewrite LiftSubst.lift0_p. reflexivity. cbn. rewrite /trans_decl /vass /=. @@ -1766,8 +1771,8 @@ Proof. - simpl. constructor. - simpl. econstructor. + eapply IHX0. - + simpl. destruct tu. exists x. now eapply Hs. - - simpl. constructor; auto. red. destruct tu. exists x; auto. + + simpl. destruct tu as (s & e & Hty). exists s. split; [apply e|]. now eapply Hs. + - simpl. constructor; auto. red. destruct tu as (s & e & Hty). exists s; auto. simpl. eauto. Qed. @@ -1785,9 +1790,10 @@ Proof. - simpl. constructor. - simpl. econstructor. + eapply IHX. - + simpl. destruct t0. exists x. eapply p. - - simpl. constructor; auto. red. destruct t0. exists x. intuition auto. - red in t1. destruct t1. cbn. eapply p. + + simpl. destruct t0 as (s & e & Hs). exists s. split; [apply e|]. eapply Hs. + - simpl. constructor; auto. red. destruct t0 as (s & e & Hs). exists s. intuition auto. + destruct Hs as [Hs Ht]. + red in t1. destruct t1. cbn. eapply Ht. red; red in t1. destruct t1. eapply t2. Qed. @@ -1908,7 +1914,7 @@ Qed. Lemma map_option_out_check_one_fix {Σ mfix} : let Σ' := trans_global_env Σ in - All (fun def => (WfAst.wf Σ (dtype def) * WfAst.wf Σ (dbody def))) mfix -> + All (WfAst.wf_def Σ) mfix -> forall l, map_option_out (map (fun x => ST.check_one_fix x) mfix) = Some l -> map_option_out (map (fun x => check_one_fix (map_def (trans Σ') (trans Σ') x)) mfix) = Some l. @@ -1937,7 +1943,7 @@ Qed. Lemma map_option_out_check_one_cofix {Σ mfix} : let Σ' := trans_global_env Σ in - All (fun def => (WfAst.wf Σ (dtype def) * WfAst.wf Σ (dbody def))) mfix -> + All (WfAst.wf_def Σ) mfix -> forall l, map_option_out (map (fun x => ST.check_one_cofix x) mfix) = Some l -> map_option_out (map (fun x => check_one_cofix (map_def (trans Σ') (trans Σ') x)) mfix) = Some l. @@ -1957,7 +1963,7 @@ Qed. Lemma trans_wf_fixpoint {cf} {Σ} {wfΣ : Typing.wf Σ} {mfix} : let Σ' := trans_global_env Σ in wf Σ' -> - All (fun def => (WfAst.wf Σ (dtype def) * WfAst.wf Σ (dbody def))) mfix -> + All (WfAst.wf_def Σ) mfix -> ST.wf_fixpoint Σ mfix -> wf_fixpoint (trans_global_env Σ) (map (map_def (trans Σ') (trans Σ')) mfix). Proof. @@ -1978,7 +1984,7 @@ Qed. Lemma trans_wf_cofixpoint {cf} {Σ} {wfΣ : Typing.wf Σ} {mfix} : let Σ' := trans_global_env Σ in wf Σ' -> - All (fun def => (WfAst.wf Σ (dtype def) * WfAst.wf Σ (dbody def))) mfix -> + All (WfAst.wf_def Σ) mfix -> ST.wf_cofixpoint Σ mfix -> wf_cofixpoint (trans_global_env Σ) (map (map_def (trans Σ') (trans Σ')) mfix). Proof. @@ -2048,7 +2054,7 @@ Lemma isType_mkApps_Ind_inv_spine {cf:checker_flags} {Σ : global_env_ext} {Γ i Proof. move=> isdecl isty. pose proof (isType_wf_local isty). - destruct isty as [s Hs]. + destruct isty as (s & e & Hs). eapply invert_type_mkApps_ind in Hs as [sp cu]; tea. erewrite ind_arity_eq in sp. 2: eapply PCUICInductives.oib ; eauto. @@ -2241,12 +2247,12 @@ Proof. eapply refine_type; cbn. * eapply type_App. 2:{ eapply type_Lambda; eauto. eapply type_Rel. econstructor; eauto. - eapply typing_wf_local; eauto. now eexists. reflexivity. } - eapply type_Prod. eauto. + eapply typing_wf_local; eauto. eexists; eauto. reflexivity. } + eapply type_Prod. eauto. eauto. instantiate (1 := s). simpl. eapply (weakening _ _ [_] _ (tSort _)); eauto. constructor; eauto. eapply typing_wf_local; eauto. - now eexists. + eexists; split; [apply H|]. auto. now eapply X2. * unfold subst1. rewrite simpl_subst; auto. now rewrite lift0_p. @@ -2267,12 +2273,12 @@ Proof. * simpl in p. destruct (TypingWf.typing_wf _ wfΣ _ _ _ typrod) as [wfAB _]. econstructor; eauto. - + exists s; eauto. + + exists s; split; cbnr; eauto. + rewrite -/Σ'. change (tProd na (trans Σ' A) (trans _ B)) with (trans Σ' (Ast.tProd na A B)). eapply trans_cumulSpec_typed; tea. eapply TypingWf.typing_all_wf_decl; eauto. - exists s; eauto. + exists s; split; cbnr; eauto. + eapply typing_wf in ty; eauto. destruct ty as [wfhd _]. rewrite trans_subst in IHX1; eauto with wf. eapply IHX1; cycle 1. @@ -2416,7 +2422,8 @@ Proof. rewrite /trans_local map_app in X3. rewrite <- trans_lift. apply X3; auto. -- eapply trans_wf_fixpoint => //. - solve_all; destruct a as [s [Hs IH]], b as [Hs' IH']. + solve_all. destruct a as (s & e & Hs & IH), b as (Hs' & IH'). + split. now eapply TypingWf.typing_wf in Hs. now eapply TypingWf.typing_wf in Hs'. -- destruct decl; reflexivity. @@ -2432,14 +2439,15 @@ Proof. -- now eapply cofix_guard_trans. -- now rewrite nth_error_map H0. -- eapply All_map, (All_impl X0). - intros x [s [Hs Hts]]. now exists s. + intros x (s & Hs & Hts). now exists s. -- apply All_map. eapply All_impl; eauto. intuition eauto 3 with wf. hnf in X3. - rewrite H2. rewrite /trans_local map_length. + rewrite H2. rewrite /on_def_body /trans_local map_length. rewrite trans_local_app in X3. cbn. rewrite <- trans_lift. now apply X3. -- eapply trans_wf_cofixpoint => //. - solve_all; destruct a as [s [Hs IH]], b as [Hs' IH']. + solve_all. destruct a as (s & e & Hs & IH), b as (Hs' & IH'). + split. now eapply TypingWf.typing_wf in Hs. now eapply TypingWf.typing_wf in Hs'. -- destruct decl; reflexivity. @@ -2852,9 +2860,9 @@ Proof. intros IH wfΣ wfΣ'. induction cs; simpl; auto. { now intros; eapply trans_wf_universe. } destruct a as [na [b|] ty] => //; - intros [? ?]; cbn. destruct p. + intros [? ?]; cbn. all: destruct p. intuition auto. - - eapply (IH _ _ _ Sort) in i; auto. + - eapply (IH _ _ _ (SortRel _)) in i; auto. rewrite -trans_local_app //. - eapply (IH _ _ _ (Typ _)) in t0 => //. rewrite -trans_local_app //. @@ -2912,7 +2920,7 @@ Proof. intros. destruct T. * eapply typing_wf; tea. - * destruct X2 as [s Hs]. + * destruct X2 as (s & e & Hs). red. split => //. now eapply typing_wf in Hs; tea. Qed. @@ -2943,7 +2951,8 @@ Proof. cbn. rewrite map_app closedn_ctx_app /=. len. apply/andP; split. rewrite /test_decl /trans_decl /=. - destruct p as [s hs]. now eapply subject_closed in hs. + destruct p as (s & e & hs). + now eapply subject_closed in hs. eapply closedn_ctx_upwards; tea. lia. Qed. @@ -3010,7 +3019,7 @@ Proof. red in o |- *. simpl in *. eapply (X (Σg, cst_universes0) [] t (Typ cst_type0)); auto. red in o0 |- *. simpl in *. - now apply (X (Σg, cst_universes0) [] cst_type0 Sort). + now apply (X (Σg, cst_universes0) [] cst_type0 (Sort None)). * destruct o0 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 *. @@ -3022,14 +3031,16 @@ Proof. { eapply closed_arities_context => //. clear -onu wfΣg X0 IHond X onI. eapply Alli_All; tea; cbv beta. move=> n x /Typing.onArity => o. - apply (X (Σg, Ast.Env.ind_universes m) [] (Ast.Env.ind_type x) Sort X0 o) => //. } + eapply isType_of_isTypeRel. + apply (X (Σg, Ast.Env.ind_universes m) [] (Ast.Env.ind_type x) (SortRel (Ast.Env.ind_relevance x)) 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)). + (ind_type (trans_one_ind_body (trans_global_env Σg) idecl)) + (ind_relevance (trans_one_ind_body (trans_global_env Σg) idecl)). { apply ST.onArity in oni. unfold on_type in *; simpl in *. - now apply (X (Σg, Ast.Env.ind_universes m) [] (Ast.Env.ind_type idecl) Sort). } + now apply (X (Σg, Ast.Env.ind_universes m) [] (Ast.Env.ind_type idecl) (SortRel _) _ oni). } 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. @@ -3047,7 +3058,7 @@ Proof. (cstr_type (trans_constructor_body (trans_global_env Σg) x)). { unfold cstr_type, Ast.Env.cstr_type in on_ctype |- *; simpl in *. red. move: (X (Σg, Ast.Env.ind_universes m) (Ast.Env.arities_context (Ast.Env.ind_bodies m)) - (Ast.Env.cstr_type x) Sort). + (Ast.Env.cstr_type x) (Sort None)). rewrite trans_arities_context. intros H'. apply H' => //. } have ceq : cstr_type (trans_constructor_body (trans_global_env Σg) x) = @@ -3082,9 +3093,9 @@ Proof. (Ast.Env.arities_context (Ast.Env.ind_bodies m)) (Ast.Env.ind_params m)) c)); rewrite /trans_local !map_app in foo. - now eapply (foo ty Sort). + now eapply (foo ty (SortRel (binder_relevance na))). now apply (foo b (Typ ty)). - now apply (foo ty Sort). + now apply (foo ty (SortRel (binder_relevance na))). now apply (foo b (Typ ty)). now apply (foo ty (Typ (Ast.tSort _))). + clear -X0 onu wfΣg IHond X on_cindices. @@ -3134,7 +3145,7 @@ Proof. have wfty : WfAst.wf Σg (Ast.Env.ind_type i). { rewrite nth_error_rev in e. len. rewrite List.rev_involutive in e. eapply nth_error_alli in onI; tea. cbn in onI. - destruct onI as [onI _]. eapply Typing.onArity in onI as [s Hs]. + destruct onI as [onI _]. eapply Typing.onArity in onI as (s & ? & Hs). now eapply typing_wf in Hs. } rewrite /trans_one_ind_body /= /ind_realargs /= /ST.ind_realargs. generalize (trans_destArity Σg [] _ wfty wfΣg). @@ -3153,7 +3164,7 @@ Proof. specialize (on_ctype_variance _ indv). cbn -[Σg]. eapply trans_cstr_respects_variance => //. - { do 2 red in onty. destruct onty as [s hs]. + { do 2 red in onty. destruct onty as (s & e & hs). rewrite ceq in hs. eapply PCUICClosedTyp.subject_closed in hs. len in hs. rewrite -it_mkProd_or_LetIn_app in hs. @@ -3168,7 +3179,7 @@ Proof. now apply closedn_smash_context. } { rewrite ceq in onty. rewrite /cstr_concl in onty. - destruct onty as [s hs]. + destruct onty as (s & e & hs). rewrite -it_mkProd_or_LetIn_app in hs. eapply PCUICSpine.type_it_mkProd_or_LetIn_inv in hs as [Δs [us []]]. eapply typing_expand_lets in t. eapply subject_closed in t; move: t. @@ -3204,12 +3215,13 @@ 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. + --- cbn; apply oni. --- have inds := oni.(ST.onIndices). intros v hv. specialize (inds v hv). eapply trans_ind_respects_variance; tea. move: onarity; cbn -[Σg]. rewrite oni.(ST.ind_arity_eq). rewrite !trans_it_mkProd_or_LetIn. - move=> [] s /subject_closed. rewrite -it_mkProd_or_LetIn_app closedn_it_mkProd_or_LetIn /=. + move=> [] s [] e /subject_closed. rewrite -it_mkProd_or_LetIn_app closedn_it_mkProd_or_LetIn /=. rewrite andb_true_r trans_local_app //. -- red in onP. red. epose proof (Typing.env_prop_wf_local template_to_pcuic (Σg, Ast.Env.ind_universes m) X0 _ onP). @@ -3231,13 +3243,16 @@ Proof. intros Hu. eapply trans_on_global_env; eauto. simpl; intros. epose proof (ST.env_prop_typing template_to_pcuic _ X Γ). - forward X2. - red in X0. destruct T. - now eapply ST.typing_wf_local. - destruct X0 as [s Hs]. eapply ST.typing_wf_local; eauto. - destruct T; simpl in *. + forward X2; + destruct T; [ + red in X0 | destruct X0 as (s & e & Hs) | + red in X0 | destruct X0 as (s & e & Hs) + ]. + 1-2: now eapply ST.typing_wf_local. + all: simpl in *. - eapply X2; eauto. - - destruct X0 as [s Hs]. exists s. + - exists s. + split; [apply e|]. eapply (X2 _ (Ast.tSort s)); eauto. Qed. diff --git a/pcuic/theories/TemplateToPCUICExpanded.v b/pcuic/theories/TemplateToPCUICExpanded.v index a338bbaa0..4f9cb1871 100644 --- a/pcuic/theories/TemplateToPCUICExpanded.v +++ b/pcuic/theories/TemplateToPCUICExpanded.v @@ -227,14 +227,14 @@ Proof with eauto using expanded. + eapply template_to_pcuic_env; eauto. - now (wf_inv wf [[]]; eauto using expanded). - wf_inv wf [[]]. wf_inv w ?. eapply expanded_tFix. - + solve_all. + + solve_all; destruct b1. * rewrite trans_isLambda //. - * revert H2. cbn. now rewrite mapi_cst_map rev_map_spec map_map. + * revert a1. cbn. now rewrite mapi_cst_map rev_map_spec map_map. + solve_all. + destruct args; cbn; congruence. + now rewrite nth_error_map H5. + now simpl_list. - - wf_inv wf ?. econstructor. solve_all. + - wf_inv wf ?. econstructor. solve_all. destruct b0; easy. - wf_inv wf [[[]]]. eapply forall_decls_declared_constructor in H; eauto. 2: now eapply template_to_pcuic_env. eapply expanded_tConstruct_app. eauto. cbn. unfold trans_local. now rewrite map_length context_assumptions_map. solve_all. Qed. @@ -255,7 +255,7 @@ Proof. eapply typing_wf_wf in wf. depelim wf. cbn in o0. depelim o0. cbn. split => //. eapply TypingWf.on_global_decl_impl; tea. cbn. - intros. destruct T => //. red. red in X0. destruct X0. intuition auto. + intros. unfold WfAst.wf_decl_pred. destruct T => //. cbn. red in X0. destruct X0. intuition auto. cbn. split => //. Qed. @@ -312,9 +312,9 @@ Lemma wf_context_sorts {cf} {Σ ctx ctx' cunivs} {wfΣ : Typing.wf_ext Σ} : Proof. induction ctx' in cunivs |- *; cbn; auto. destruct a as [na [b|] ty]. - intros [? []]. constructor; auto. eauto. + intros (Hwfctw & Ht & Hb). constructor; auto. eauto. destruct cunivs => //. - intros [? []]. constructor; eauto. constructor; cbn; eauto. + intros (Hwfctw & Ht & (Hb & Ht')). constructor; eauto. constructor; cbn; eauto. Qed. Lemma expanded_trans_global_env {cf} Σ {wfΣ : Typing.wf_ext Σ} : diff --git a/pcuic/theories/Typing/PCUICClosedTyp.v b/pcuic/theories/Typing/PCUICClosedTyp.v index 04e5f9bdc..3e2abb80d 100644 --- a/pcuic/theories/Typing/PCUICClosedTyp.v +++ b/pcuic/theories/Typing/PCUICClosedTyp.v @@ -195,14 +195,14 @@ Proof. rewrite -> andb_and in IH. intuition auto. eauto using closed_upwards with arith. simpl in *. - repeat red in IH. destruct IH as [s Hs]. + repeat red in IH. destruct IH as (s & e & Hs). rewrite -> andb_and in Hs. intuition auto. eauto using closed_upwards with arith. - rewrite closedn_subst_instance. eapply declared_inductive_inv in X0; eauto. apply onArity in X0. repeat red in X0. - destruct X0 as [s Hs]. rewrite -> andb_and in Hs. + destruct X0 as (s & e & Hs). rewrite -> andb_and in Hs. intuition eauto using closed_upwards with arith. - destruct isdecl as [Hidecl Hcdecl]. @@ -215,7 +215,7 @@ Proof. pose proof X0.(onConstructors) as XX. eapply All2_nth_error_Some in Hcdecl; eauto. destruct Hcdecl as [? [? ?]]. cbn in *. - destruct o as [? ? [s Hs] _]. rewrite -> andb_and in Hs. + destruct o as [? ? (s & e' & Hs) _]. rewrite -> andb_and in Hs. apply proj1 in Hs. rewrite arities_context_length in Hs. eauto using closed_upwards with arith. @@ -277,18 +277,18 @@ Proof. subst types. now rewrite app_context_length fix_context_length in H. eapply nth_error_all in X0; eauto. simpl in X0. intuition auto. rtoProp. - destruct X0 as [s [Hs cl]]. now rewrite andb_true_r in cl. + destruct X0 as (s & e & Hs & cl). now rewrite andb_true_r in cl. - split. solve_all. destruct b. destruct x; simpl in *. unfold test_def. simpl. rtoProp. split. - destruct a as [s [Hs cl]]. + destruct a as (s & e & Hs & cl). now rewrite andb_true_r in cl. rewrite -> app_context_length in *. rewrite -> Nat.add_comm in *. subst types. now rewrite fix_context_length in H3. eapply nth_error_all in X0; eauto. - destruct X0 as [s [Hs cl]]. + destruct X0 as (s & e & Hs & cl). now rewrite andb_true_r in cl. Qed. @@ -380,8 +380,11 @@ Qed. #[global] Hint Extern 10 => progress unfold PCUICTypingDef.typing in * : fvs. +Lemma isTypeRelOpt_closed {cf:checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ.1} {Γ T rel} : isTypeRelOpt Σ Γ T rel -> closedn #|Γ| T. +Proof. intros (s & e & Hs); fvs. Qed. + Lemma isType_closed {cf:checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ.1} {Γ T} : isType Σ Γ T -> closedn #|Γ| T. -Proof. intros [s Hs]; fvs. Qed. +Proof. apply isTypeRelOpt_closed. Qed. #[global] Hint Extern 4 (closedn #|?Γ| ?A = true) => match goal with @@ -465,7 +468,7 @@ Qed. end : fvs. Lemma ctx_inst_closed {cf:checker_flags} (Σ : global_env_ext) Γ i Δ : - wf Σ.1 -> ctx_inst typing Σ Γ i Δ -> All (closedn #|Γ|) i. + wf Σ.1 -> ctx_inst (typing Σ) Γ i Δ -> All (closedn #|Γ|) i. Proof. intros wfΣ; induction 1; auto; constructor; auto; fvs. Qed. @@ -497,8 +500,8 @@ Proof. destruct decl as [ty bo un]. simpl in *. destruct bo as [t|]. - now eapply type_closed in decl'. - - cbn in decl'. destruct decl' as [s h]. - now eapply subject_closed in h. Unshelve. all:tea. + - cbn in decl'. destruct decl' as (s & e & Hs). + now eapply subject_closed in Hs. Unshelve. all:tea. Qed. @@ -535,7 +538,7 @@ Proof. destruct h1 as [Σ' [ext wfΣ' decl']]. red in decl'. destruct decl' as [h ? ? ?]. eapply Alli_nth_error in h. 2: eassumption. - simpl in h. destruct h as [? [? h] ? ? ?]. + simpl in h. destruct h as [? (? & ? & h) ? ? ?]. eapply typecheck_closed in h as [? e]. 2: auto. now move: e => [_ /andP []]. Qed. diff --git a/pcuic/theories/Typing/PCUICContextConversionTyp.v b/pcuic/theories/Typing/PCUICContextConversionTyp.v index 53df27200..44df7bc2c 100644 --- a/pcuic/theories/Typing/PCUICContextConversionTyp.v +++ b/pcuic/theories/Typing/PCUICContextConversionTyp.v @@ -168,15 +168,16 @@ Proof. try solve [econstructor; eauto]. - induction X; constructor; auto. - all: now apply infer_typing_sort_impl with id tu. + all: now apply infer_typing_sort_impl with id tu => //. - pose proof heq_nth_error. eapply (All2_fold_nth_r X0) in H as [d' [Hnth [Hrel Hconv]]]. unshelve eapply nth_error_All_local_env in X; tea. 2:eapply nth_error_Some_length in heq_nth_error; lia. rewrite heq_nth_error /= in X. destruct decl as [na [b|] ty] => /=. - + red in X. cbn in X. destruct X as [Hb Hty]. - destruct Hty as [s Hty]. specialize (Hty _ Hrel). + + red in X. cbn in X. + destruct X as [Hb (s & e & Hty)]. + specialize (Hty _ Hrel). forward Hty by now eapply All_local_env_skipn. eapply type_Cumul with _ s. * econstructor. auto. eauto. @@ -199,7 +200,7 @@ Proof. eapply (weakening_cumulSpec0 (Γ := Δ) (Γ'' := Δ') (M := exist t H) (N := exist ty Hty)); cbn. lia. unshelve eapply (@cumulAlgo_cumulSpec _ _ Cumul). apply into_ws_cumul_pb; eauto. intuition. - + cbn in X. destruct X as [s ondecl]. + + cbn in X. destruct X as (s & e & ondecl). specialize (ondecl _ Hrel). depelim Hconv. forward ondecl by now eapply All_local_env_skipn. @@ -231,7 +232,7 @@ Proof. intuition. - constructor; pcuic. eapply forall_Γ'0. repeat (constructor; pcuic). - constructor; auto. red. eexists; eapply forall_Γ'; auto. + constructor; auto. eexists; split; [|eapply forall_Γ']; auto. - econstructor; pcuic. eapply forall_Γ'0; repeat (constructor; pcuic). - econstructor; pcuic. @@ -262,7 +263,7 @@ Proof. + apply wf_local_closed_context; eauto. * eapply (All_impl X0). intros d Ht. - apply infer_typing_sort_impl with id Ht; now intros [_ IH]. + apply infer_typing_sort_impl with id Ht => //; now intros [_ IH]. * eapply (All_impl X1). intros d [Hs IH]. eapply IH. @@ -270,7 +271,7 @@ Proof. eapply (All_mfix_wf); auto. apply (All_impl X0); simpl. intros d' Ht. - apply infer_typing_sort_impl with id Ht; now intros [_ IH']. + apply infer_typing_sort_impl with id Ht => //; now intros [_ IH']. - econstructor. all:pcuic. * eapply cofix_guard_context_cumulativity; eauto. @@ -279,7 +280,7 @@ Proof. + apply wf_local_closed_context; eauto. * eapply (All_impl X0). intros d Ht. - apply infer_typing_sort_impl with id Ht; now intros [_ IH]. + apply infer_typing_sort_impl with id Ht => //; now intros [_ IH]. * eapply (All_impl X1). intros d [Hs IH]. eapply IH. @@ -287,7 +288,7 @@ Proof. eapply (All_mfix_wf); auto. apply (All_impl X0); simpl. intros d' Ht. - apply infer_typing_sort_impl with id Ht; now intros [_ IH']. + apply infer_typing_sort_impl with id Ht => //; now intros [_ IH']. - econstructor; eauto. pose proof (wf_local_closed_context wfΓ). pose proof (type_closed (forall_Γ' _ X5 X6)). eapply (@closedn_on_free_vars xpred0) in H0. diff --git a/pcuic/theories/Typing/PCUICInstTyp.v b/pcuic/theories/Typing/PCUICInstTyp.v index 6e11638ca..a74da646c 100644 --- a/pcuic/theories/Typing/PCUICInstTyp.v +++ b/pcuic/theories/Typing/PCUICInstTyp.v @@ -370,43 +370,46 @@ Proof. - intros Σ wfΣ Γ wfΓ. auto. induction 1; constructor; tas. - all: apply infer_typing_sort_impl with id tu; auto. + all: apply infer_typing_sort_impl with id tu => //; auto. - intros Σ wfΣ Γ wfΓ n decl e X Δ σ hΔ hσ. simpl. eapply hσ. assumption. - intros Σ wfΣ Γ wfΓ l X H0 Δ σ hΔ hσ. simpl. econstructor. all: assumption. - - intros Σ wfΣ Γ wfΓ na A B s1 s2 X hA ihA hB ihB Δ σ hΔ hσ. + - intros Σ wfΣ Γ wfΓ na A B s1 s2 e X hA ihA hB ihB Δ σ hΔ hσ. autorewrite with sigma. simpl. econstructor. + + apply e. + eapply ihA ; auto. + eapply ihB. * econstructor ; auto. - eexists. eapply ihA ; auto. + eexists. split; [apply e|]. eapply ihA ; auto. * eapply well_subst_Up. 2: assumption. econstructor ; auto. - eexists. eapply ihA. all: auto. - - intros Σ wfΣ Γ wfΓ na A t s1 bty X hA ihA ht iht Δ σ hΔ hσ. + eexists. split; [apply e|]. eapply ihA. all: auto. + - intros Σ wfΣ Γ wfΓ na A t s1 bty e X hA ihA ht iht Δ σ hΔ hσ. autorewrite with sigma. econstructor. + + apply e. + eapply ihA ; auto. + eapply iht. * econstructor ; auto. - eexists. eapply ihA ; auto. + eexists. split; [apply e|]. eapply ihA ; auto. * eapply well_subst_Up. 2: assumption. constructor. 1: assumption. - eexists. eapply ihA. all: auto. - - intros Σ wfΣ Γ wfΓ na b B t s1 A X hB ihB hb ihb ht iht Δ σ hΔ hσ. + eexists. split; [apply e|]. eapply ihA. all: auto. + - intros Σ wfΣ Γ wfΓ na b B t s1 A e X hB ihB hb ihb ht iht Δ σ hΔ hσ. autorewrite with sigma. econstructor. + + apply e. + eapply ihB. all: auto. + eapply ihb. all: auto. + eapply iht. * econstructor. all: auto. - -- eexists. eapply ihB. all: auto. + -- eexists. split; [apply e|]. eapply ihB. all: auto. -- simpl. eapply ihb. all: auto. * eapply well_subst_Up'; try assumption. constructor; auto. - ** exists s1. apply ihB; auto. + ** exists s1. split; [apply e|]. apply ihB; auto. ** apply ihb; auto. - intros Σ wfΣ Γ wfΓ t na A B s u X hty ihty ht iht hu ihu Δ σ hΔ hσ. autorewrite with sigma. @@ -535,9 +538,9 @@ Proof. * now eapply fix_guard_inst. * now rewrite nth_error_map hnth. * solve_all. - apply infer_typing_sort_impl with id a; intros [_ IH]. + apply infer_typing_sort_impl with id a => //; intros [_ IH]. now apply IH. - * solve_all. destruct b as [? b0]. + * solve_all. cbn. destruct b as [? b0]. len. rewrite /types in b0. len in b0. pose proof (inst_fix_context mfix σ). setoid_rewrite <-up_Upn at 1 in H. rewrite H. @@ -556,9 +559,9 @@ Proof. * now eapply cofix_guard_inst. * now rewrite nth_error_map hnth. * solve_all. - apply infer_typing_sort_impl with id a; intros [_ IH]. + apply infer_typing_sort_impl with id a => //; intros [_ IH]. now apply IH. - * solve_all. destruct b as [? b0]. + * solve_all. cbn. destruct b as [? b0]. len. rewrite /types in b0. len in b0. pose proof (inst_fix_context mfix σ). setoid_rewrite <-up_Upn at 1 in H. rewrite H. diff --git a/pcuic/theories/Typing/PCUICNamelessTyp.v b/pcuic/theories/Typing/PCUICNamelessTyp.v index 3d113bf85..3064c9caa 100644 --- a/pcuic/theories/Typing/PCUICNamelessTyp.v +++ b/pcuic/theories/Typing/PCUICNamelessTyp.v @@ -35,11 +35,10 @@ Proof. induction h. - constructor. - simpl. unfold map_decl_anon. cbn. constructor. 1: assumption. - apply infer_typing_sort_impl with id tu; intros Hty. - exact Hs. + apply infer_typing_sort_impl with id tu => //. - simpl. unfold map_decl_anon. cbn. constructor. + assumption. - + apply infer_typing_sort_impl with id tu; intros Hty. exact Hs. + + apply infer_typing_sort_impl with id tu => //. + assumption. Qed. diff --git a/pcuic/theories/Typing/PCUICRenameTyp.v b/pcuic/theories/Typing/PCUICRenameTyp.v index 7fbf364cf..7f3261fd9 100644 --- a/pcuic/theories/Typing/PCUICRenameTyp.v +++ b/pcuic/theories/Typing/PCUICRenameTyp.v @@ -619,14 +619,14 @@ Proof. induction X. - apply a. - rewrite rename_context_snoc /=. constructor; auto. - apply infer_typing_sort_impl with id t0; intros Hs. + apply infer_typing_sort_impl with id t0 => //; intros Hs. eapply (Hs P (Δ' ,,, rename_context f Γ0) (shiftn #|Γ0| f)). split => //. eapply urenaming_ext. { now rewrite app_length -shiftnP_add. } { reflexivity. } now eapply urenaming_context. - rewrite rename_context_snoc /=. constructor; auto. - * apply infer_typing_sort_impl with id t0; intros Hs. + * apply infer_typing_sort_impl with id t0 => //; intros Hs. apply (Hs P (Δ' ,,, rename_context f Γ0) (shiftn #|Γ0| f)). split => //. eapply urenaming_ext. @@ -806,7 +806,7 @@ Proof. - intros Σ wfΣ Γ wfΓ HΓ. split; auto. induction HΓ; constructor; tas. - all: apply infer_typing_sort_impl with id tu; intros Hty. + all: apply infer_typing_sort_impl with id tu => //; intros Hty. all: eauto. - intros Σ wfΣ Γ wfΓ n decl isdecl ihΓ P Δ f hf. @@ -820,8 +820,9 @@ Proof. - intros Σ wfΣ Γ wfΓ l X H0 P Δ f [hΔ hf]. simpl. constructor. all: auto. - - intros Σ wfΣ Γ wfΓ na A B s1 s2 X hA ihA hB ihB P Δ f hf. + - intros Σ wfΣ Γ wfΓ na A B s1 s2 e X hA ihA hB ihB P Δ f hf. rewrite /=. econstructor. + + apply e. + eapply ihA; eauto. + eapply ihB; eauto. simpl. @@ -829,20 +830,22 @@ Proof. eapply renaming_vass. 2: eauto. constructor. * destruct hf as [hΔ hf]. auto. - * simpl. exists s1. eapply ihA; eauto. - - intros Σ wfΣ Γ wfΓ na A t s1 B X hA ihA ht iht P Δ f hf. + * simpl. exists s1. split; [apply e|]; eapply ihA; eauto. + - intros Σ wfΣ Γ wfΓ na A t s1 B e X hA ihA ht iht P Δ f hf. simpl. (* /andP [_ havB]. *) simpl. econstructor. + + apply e. + eapply ihA; eauto. + eapply iht; eauto; simpl. eapply renaming_extP. { now rewrite -(shiftnP_add 1). } eapply renaming_vass. 2: eauto. constructor. * destruct hf as [hΔ hf]. auto. - * simpl. exists s1. eapply ihA; eauto. - - intros Σ wfΣ Γ wfΓ na b B t s1 A X hB ihB hb ihb ht iht P Δ f hf. + * simpl. exists s1. split; [apply e|]; eapply ihA; eauto. + - intros Σ wfΣ Γ wfΓ na b B t s1 A e X hB ihB hb ihb ht iht P Δ f hf. simpl. econstructor. + + apply e. + eapply ihB; tea. + eapply ihb; tea. + eapply iht; tea. @@ -850,7 +853,7 @@ Proof. eapply renaming_vdef. 2: eauto. constructor. * destruct hf. assumption. - * simpl. eexists. eapply ihB; tea. + * simpl. exists s1. split; [apply e|]; eapply ihB; tea. * simpl. eapply ihb; tea. - intros Σ wfΣ Γ wfΓ t na A B s u X hty ihty ht iht hu ihu P Δ f hf. simpl. eapply meta_conv. @@ -1001,7 +1004,7 @@ Proof. * destruct hf; eapply fix_guard_rename; eauto. * rewrite nth_error_map. rewrite hdecl. simpl. reflexivity. * apply All_map, (All_impl ihmfixt). - intros x t. apply infer_typing_sort_impl with id t. + intros x t. apply infer_typing_sort_impl with id t => //. intros [_ IHs]; now eapply IHs. * apply All_map, (All_impl ihmfixb). intros x [Hb IHb]. @@ -1027,7 +1030,7 @@ Proof. * destruct hf; eapply cofix_guard_rename; eauto. * rewrite nth_error_map. rewrite hdecl. simpl. reflexivity. * apply All_map, (All_impl ihmfixt). - intros x t. apply infer_typing_sort_impl with id t. + intros x t. apply infer_typing_sort_impl with id t => //. intros [_ IHs]; now eapply IHs. * apply All_map, (All_impl ihmfixb). intros x [Hb IHb]. diff --git a/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v b/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v index 7464d13cb..464838090 100644 --- a/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v +++ b/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v @@ -222,11 +222,11 @@ Proof using Type. induction 1. + constructor. + simpl. constructor; auto. - eapply infer_typing_sort_impl; tea. - intros Hty. eapply Hs; auto. + eapply infer_typing_sort_impl with _ tu; [apply relevance_subst_opt|]; intros Hty. + eapply Hs; auto. + simpl. constructor; auto. - ++ eapply infer_typing_sort_impl; tea. - intros Hty. eapply Hs; auto. + ++ eapply infer_typing_sort_impl with _ tu; [apply relevance_subst_opt|]; intros Hty. + eapply Hs; auto. ++ apply Hc; auto. - intros n decl eq X u univs wfΣ' H. rewrite subst_instance_lift. @@ -238,16 +238,23 @@ Proof using Type. + econstructor. * aa. * now apply wf_universe_subst_instance. - - intros n t0 b s1 s2 X X0 X1 X2 X3 u univs wfΣ' H. - rewrite product_subst_instance; aa. econstructor. + - intros n t0 b s1 s2 Xe X X0 X1 X2 X3 u univs wfΣ' H. + rewrite product_subst_instance; aa. + econstructor. + + apply relevance_subst; apply Xe. + eapply X1; eauto. + eapply X3; eauto. - - intros n t0 b s1 bty X X0 X1 X2 X3 u univs wfΣ' H. + - intros n t0 b s1 bty Xe X X0 X1 X2 X3 u univs wfΣ' H. econstructor. + + apply relevance_subst; apply Xe. + eapply X1; aa. + eapply X3; aa. - - intros n b b_ty b' s1 b'_ty X X0 X1 X2 X3 X4 X5 u univs wfΣ' H. - econstructor; eauto. eapply X5; aa. + - intros n b b_ty b' s1 b'_ty Xe X X0 X1 X2 X3 X4 X5 u univs wfΣ' H. + econstructor. + + eapply relevance_subst. apply Xe. + + eauto. + + eauto. + + eapply X5; aa. - intros t0 na A B s u X X0 X1 X2 X3 X4 X5 u0 univs wfΣ' H. rewrite subst_instance_subst. cbn. econstructor. + eapply X1; eauto. @@ -322,10 +329,10 @@ Proof using Type. + now eapply fix_guard_subst_instance. + rewrite nth_error_map H0. reflexivity. + apply All_map, (All_impl X); simpl; intuition auto. - eapply infer_typing_sort_impl with (tu := X1). + eapply infer_typing_sort_impl with _ X1; [apply relevance_subst_opt|]. intros [_ Hs]; now apply Hs. + eapply All_map, All_impl; tea. - intros x [X1 X3]. + intros x [X1 X3]. cbn. specialize (X3 u univs wfΣ' H2). rewrite (map_dbody (subst_instance u)) in X3. rewrite subst_instance_lift in X3. @@ -350,10 +357,10 @@ Proof using Type. + now eapply cofix_guard_subst_instance. + rewrite nth_error_map H0. reflexivity. + apply All_map, (All_impl X); simpl; intuition auto. - eapply infer_typing_sort_impl with (tu := X1). + eapply infer_typing_sort_impl with _ X1; [apply relevance_subst_opt|]. intros [_ Hs]; now apply Hs. + eapply All_map, All_impl; tea. - intros x [X1 X3]. + intros x [X1 X3]; cbn. specialize (X3 u univs wfΣ' H2). rewrite (map_dbody (subst_instance u)) in X3. rewrite subst_instance_lift in X3. @@ -473,7 +480,7 @@ Lemma isType_subst_instance_decl Σ Γ T c decl u : isType Σ (subst_instance u Γ) (subst_instance u T). Proof using Type. intros wfΣ look isty cu. - eapply infer_typing_sort_impl with (tu := isty). + eapply infer_typing_sort_impl with _ isty; [apply relevance_subst_opt|]. intros Hs; now eapply (typing_subst_instance_decl _ _ _ (tSort _)). Qed. @@ -492,7 +499,7 @@ Lemma wf_local_subst_instance Σ Γ ext u : Proof using Type. destruct Σ as [Σ φ]. intros X X0 X1. simpl in *. induction X1; cbn; constructor; auto. - 1,2: eapply infer_typing_sort_impl with (tu := t0); intros Hs. + 1,2: eapply infer_typing_sort_impl with _ t0; [apply relevance_subst_opt|]; intros Hs. 3: rename t1 into Hs. all: eapply typing_subst_instance'' in Hs; eauto; apply X. Qed. @@ -506,7 +513,7 @@ Lemma wf_local_subst_instance_decl Σ Γ c decl u : Proof using Type. destruct Σ as [Σ φ]. intros X X0 X1 X2. induction X1; cbn; constructor; auto. - 1,2: eapply infer_typing_sort_impl with (tu := t0); intros Hs. + 1,2: eapply infer_typing_sort_impl with _ t0; [apply relevance_subst_opt|]; intros Hs. 3: rename t1 into Hs. all: eapply typing_subst_instance_decl in Hs; eauto; apply X. Qed. @@ -521,7 +528,7 @@ Qed. pose proof (on_declared_inductive decli) as [onmind oib]. pose proof (onArity oib) as ona. rewrite (oib.(ind_arity_eq)) in ona. - red in ona. destruct ona. + red in ona. destruct ona as (s & e & t). eapply typed_subst_abstract_instance in t. 2:split; simpl; auto. - rewrite !subst_instance_it_mkProd_or_LetIn in t. @@ -543,7 +550,7 @@ Qed. pose proof (on_declared_inductive decli) as [_ oib]. pose proof (onArity oib) as ona. rewrite (oib.(ind_arity_eq)) in ona |- *. - red in ona. destruct ona. + red in ona. destruct ona as (s & e & t). eapply typed_subst_abstract_instance in t; eauto. destruct decli as [declm _]. eapply declared_inductive_wf_global_ext in declm; auto. @@ -555,7 +562,8 @@ Qed. isType Σ Γ T -> subst_instance u T = T. Proof using Type. intros wf_ext u isT. - destruct isT. eapply typed_subst_abstract_instance in t; auto. + destruct isT as (s & e & t). + eapply typed_subst_abstract_instance in t; auto. Qed. End SubstIdentity. diff --git a/pcuic/theories/Typing/PCUICWeakeningEnvTyp.v b/pcuic/theories/Typing/PCUICWeakeningEnvTyp.v index 13f76088f..5879403e4 100644 --- a/pcuic/theories/Typing/PCUICWeakeningEnvTyp.v +++ b/pcuic/theories/Typing/PCUICWeakeningEnvTyp.v @@ -35,8 +35,8 @@ Lemma extends_wf_local `{cf : checker_flags} Σ Γ (wfΓ : wf_local Σ Γ) : Proof. intros X0 Σ' H0. induction X0 in H0 |- *; try econstructor; simpl in *; intuition auto. - - apply infer_typing_sort_impl with id tu. intro; eauto. - - apply infer_typing_sort_impl with id tu. intro; eauto. + - apply infer_typing_sort_impl with id tu => //. intro; eauto. + - apply infer_typing_sort_impl with id tu => //. intro; eauto. Qed. #[global] Hint Resolve extends_wf_local : extends. @@ -85,8 +85,8 @@ Proof. rename_all_hyps; try solve [econstructor; eauto 2 with extends]. - induction X; constructor; eauto 2 with extends. - + apply infer_typing_sort_impl with id tu; intro; auto. - + apply infer_typing_sort_impl with id tu; intro; auto. + + apply infer_typing_sort_impl with id tu => //; intro; auto. + + apply infer_typing_sort_impl with id tu => //; intro; auto. + eapply Hc; eauto. - econstructor; eauto 2 with extends. now apply extends_wf_universe. @@ -135,7 +135,7 @@ Proof. destruct Hdecl as [onI onP onnP]; constructor; eauto. - eapply Alli_impl; eauto. intros. destruct X. unshelve econstructor; eauto. - + unfold on_type in *; intuition eauto. + + unfold on_type_rel in *; intuition eauto. + unfold on_constructors in *. eapply All2_impl; eauto. intros. destruct X as [? ? ? ?]. unshelve econstructor; eauto. @@ -212,7 +212,7 @@ Proof. destruct Hdecl as [onI onP onnP]; constructor; eauto. - eapply Alli_impl; eauto. intros. destruct X. unshelve econstructor; eauto. - + unfold on_type in *; intuition eauto. + + unfold on_type_rel in *; intuition eauto. + unfold on_constructors in *. eapply All2_impl; eauto. intros. destruct X as [? ? ? ?]. unshelve econstructor; eauto. diff --git a/pcuic/theories/Typing/PCUICWeakeningTyp.v b/pcuic/theories/Typing/PCUICWeakeningTyp.v index a3f7dff4f..bd736b87d 100644 --- a/pcuic/theories/Typing/PCUICWeakeningTyp.v +++ b/pcuic/theories/Typing/PCUICWeakeningTyp.v @@ -40,7 +40,8 @@ Proof. + apply wf_local_app; auto. apply (All_local_env_fold (fun Δ => lift_typing typing Σ (Γ ,,, Γ'' ,,, Δ))) in IH. apply IH. + apply weakening_renaming. - - intros Hty. simple apply (infer_typing_sort_impl (P := fun Σ Γ T s => forall P Δ f, renaming _ Σ Δ Γ _ -> Σ;;; Δ |- rename f T : rename f s)) with id Hty; intros Hs. + - intros Hty. + apply (infer_typing_sort_impl (P := fun Σ Γ T s => forall P Δ f, renaming _ Σ Δ Γ _ -> Σ;;; Δ |- rename f T : rename f s)) with id Hty => //; intros Hs. rewrite -/(lift_context #|Γ''| 0 Δ). rewrite Nat.add_0_r !lift_rename. eapply (Hs xpredT). @@ -177,7 +178,7 @@ Qed. Corollary All_mfix_wf {cf:checker_flags} Σ Γ mfix : wf Σ.1 -> wf_local Σ Γ -> - All (fun d : def term => isType Σ Γ (dtype d)) mfix -> + All (on_def_type (lift_typing typing Σ) Γ) mfix -> wf_local Σ (Γ ,,, fix_context mfix). Proof. move=> wfΣ wf a; move: wf. @@ -192,13 +193,13 @@ Proof. + simpl. eapply All_local_env_app. split; auto. * repeat constructor. - apply infer_typing_sort_impl with id p; intros Hs. + apply infer_typing_sort_impl with id p => //; intros Hs. eapply (weakening Σ Γ Δ _ (tSort _)); auto. * specialize (IHa (Δ ,,, [vass (dname x) (lift0 #|Δ| (dtype x))])). rewrite app_length in IHa. simpl in IHa. forward IHa. ** simpl; constructor; auto. - apply infer_typing_sort_impl with id p; intros Hs. + apply infer_typing_sort_impl with id p => //; intros Hs. eapply (weakening Σ Γ Δ _ (tSort _)); auto. ** eapply All_local_env_impl; eauto. simpl; intros. @@ -214,7 +215,7 @@ Proof. intros wfΣ wfΓ wfty. rewrite <- (firstn_skipn n Γ) in wfΓ |- *. assert (n = #|firstn n Γ|). { rewrite firstn_length_le; auto with arith. } - apply infer_typing_sort_impl with id wfty; intros Hs. + apply infer_typing_sort_impl with id wfty => //; intros Hs. rewrite {3}H. eapply (weakening_typing (Γ := skipn n Γ) (Γ' := []) (Γ'' := firstn n Γ) (T := tSort _)); eauto with wf. diff --git a/safechecker/theories/PCUICEqualityDec.v b/safechecker/theories/PCUICEqualityDec.v index 7f48784cc..8a74754d3 100644 --- a/safechecker/theories/PCUICEqualityDec.v +++ b/safechecker/theories/PCUICEqualityDec.v @@ -27,7 +27,7 @@ Proof. Qed. Lemma ctx_inst_on_universes Σ Γ ts Ts : - ctx_inst (fun Σ' _ t' _ => wf_universes Σ' t') Σ Γ ts Ts -> + ctx_inst (fun _ t _ => wf_universes Σ t) Γ ts Ts -> Forall (wf_universes Σ) ts. Proof. induction 1. diff --git a/template-coq/theories/BasicAst.v b/template-coq/theories/BasicAst.v index ce9f97c0c..dc0a91aa9 100644 --- a/template-coq/theories/BasicAst.v +++ b/template-coq/theories/BasicAst.v @@ -211,7 +211,7 @@ Qed. Variant typ_or_sort_ {term} := Typ (T : term) | Sort (relopt : option relevance). Arguments typ_or_sort_ : clear implicits. -Definition SortRel {term} rel : typ_or_sort_ term := Sort (Some rel). +Notation SortRel rel := (Sort (Some rel)). Definition typ_or_sort_map {T T'} (f: T -> T') t := match t with diff --git a/template-coq/theories/EnvironmentTyping.v b/template-coq/theories/EnvironmentTyping.v index ec79e1e88..bd05e444f 100644 --- a/template-coq/theories/EnvironmentTyping.v +++ b/template-coq/theories/EnvironmentTyping.v @@ -411,20 +411,20 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). P (Γ ,,, types) d.(dbody) (Typ (lift0 #|types| d.(dtype))). Definition lift_judgment - (check : global_env_ext -> context -> term -> term -> Type) - (infer_sort : global_env_ext -> context -> term -> option relevance -> Type) : - (global_env_ext -> context -> term -> typ_or_sort -> Type) := - fun Σ Γ t T => + (check : context -> term -> term -> Type) + (infer_sort : context -> term -> option relevance -> Type) : + (context -> term -> typ_or_sort -> Type) := + fun Γ t T => match T with - | Typ T => check Σ Γ t T - | Sort relopt => infer_sort Σ Γ t relopt + | Typ T => check Γ t T + | Sort relopt => infer_sort Γ t relopt end. - Lemma lift_judgment_impl {P Ps Q Qs Σ Σ' Γ Γ' t t' T} : - lift_judgment P Ps Σ Γ t T -> - (forall T, P Σ Γ t T -> Q Σ' Γ' t' T) -> - (forall relopt, Ps Σ Γ t relopt -> Qs Σ' Γ' t' relopt) -> - lift_judgment Q Qs Σ' Γ' t' T. + Lemma lift_judgment_impl {P Ps Q Qs Γ Γ' t t' T} : + lift_judgment P Ps Γ t T -> + (forall T, P Γ t T -> Q Γ' t' T) -> + (forall relopt, Ps Γ t relopt -> Qs Γ' t' relopt) -> + lift_judgment Q Qs Γ' t' T. Proof. intros HT HPQ HPsQs. destruct T; simpl. @@ -434,14 +434,14 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). (* Common uses *) - Definition lift_wf_term wf_term := (lift_judgment (fun Σ Γ t T => wf_term Σ Γ t × wf_term Σ Γ T) (fun Σ Γ t relopt => wf_term Σ Γ t)). + Definition lift_wf_term wf_term := (lift_judgment (fun Γ t T => wf_term Γ t × wf_term Γ T) (fun Γ t relopt => wf_term Γ t)). Definition infer_sort (sorting : global_env_ext -> context -> term -> Universe.t -> Type) := (fun Σ Γ T relopt => { s : Universe.t & isSortRelOpt s relopt × sorting Σ Γ T s }). Notation typing_sort typing := (fun Σ Γ T s => typing Σ Γ T (tSort s)). - Definition lift_typing typing := lift_judgment typing (infer_sort (typing_sort typing)). - Definition lift_sorting checking sorting := lift_judgment checking (infer_sort sorting). + Definition lift_typing typing Σ := lift_judgment (typing Σ) (infer_sort (typing_sort typing) Σ). + Definition lift_sorting checking sorting Σ := lift_judgment (checking Σ) (infer_sort sorting Σ). Notation Prop_conj P Q := (fun Σ Γ t T => P Σ Γ t T × Q Σ Γ t T). @@ -478,93 +478,95 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). Qed. Section TypeLocalOver. - Context (checking : global_env_ext -> context -> term -> term -> Type). - Context (sorting : global_env_ext -> context -> term -> option relevance -> Type). - Context (cproperty : forall (Σ : global_env_ext) (Γ : context), - All_local_env (lift_judgment checking sorting Σ) Γ -> - forall (t T : term), checking Σ Γ t T -> Type). - Context (sproperty : forall (Σ : global_env_ext) (Γ : context), - All_local_env (lift_judgment checking sorting Σ) Γ -> - forall (t : term) (relopt : option relevance), sorting Σ Γ t relopt -> Type). - - Inductive All_local_env_over_gen (Σ : global_env_ext) : - forall (Γ : context), All_local_env (lift_judgment checking sorting Σ) Γ -> Type := + Context (checking : context -> term -> term -> Type). + Context (sorting : context -> term -> option relevance -> Type). + Context (cproperty : forall (Γ : context), + All_local_env (lift_judgment checking sorting) Γ -> + forall (t T : term), checking Γ t T -> Type). + Context (sproperty : forall (Γ : context), + All_local_env (lift_judgment checking sorting) Γ -> + forall (t : term) (relopt : option relevance), sorting Γ t relopt -> Type). + + Inductive All_local_env_over_gen : + forall (Γ : context), All_local_env (lift_judgment checking sorting) Γ -> Type := | localenv_over_nil : - All_local_env_over_gen Σ [] localenv_nil + All_local_env_over_gen [] localenv_nil | localenv_over_cons_abs Γ na t - (all : All_local_env (lift_judgment checking sorting Σ) Γ) : - All_local_env_over_gen Σ Γ all -> - forall (tu : lift_judgment checking sorting Σ Γ t (SortRel na.(binder_relevance))) - (Hs: sproperty Σ Γ all _ _ tu), - All_local_env_over_gen Σ (Γ ,, vass na t) + (all : All_local_env (lift_judgment checking sorting) Γ) : + All_local_env_over_gen Γ all -> + forall (tu : lift_judgment checking sorting Γ t (SortRel na.(binder_relevance))) + (Hs: sproperty Γ all _ _ tu), + All_local_env_over_gen (Γ ,, vass na t) (localenv_cons_abs all tu) | localenv_over_cons_def Γ na b t - (all : All_local_env (lift_judgment checking sorting Σ) Γ) (tb : checking Σ Γ b t) : - All_local_env_over_gen Σ Γ all -> - forall (Hc: cproperty Σ Γ all _ _ tb), - forall (tu : lift_judgment checking sorting Σ Γ t (SortRel na.(binder_relevance))) - (Hs: sproperty Σ Γ all _ _ tu), - All_local_env_over_gen Σ (Γ ,, vdef na b t) + (all : All_local_env (lift_judgment checking sorting) Γ) (tb : checking Γ b t) : + All_local_env_over_gen Γ all -> + forall (Hc: cproperty Γ all _ _ tb), + forall (tu : lift_judgment checking sorting Γ t (SortRel na.(binder_relevance))) + (Hs: sproperty Γ all _ _ tu), + All_local_env_over_gen (Γ ,, vdef na b t) (localenv_cons_def all tu tb). End TypeLocalOver. Derive Signature for All_local_env_over_gen. - Definition All_local_env_over typing property := - (All_local_env_over_gen typing (infer_sort (typing_sort typing)) property (fun Σ Γ H t relopt tu => property _ _ H _ _ tu.π2.2)). + Definition All_local_env_over typing property Σ := + (All_local_env_over_gen (typing Σ) (infer_sort (typing_sort typing) Σ) (property Σ) (fun Γ H t relopt tu => property _ _ H _ _ tu.π2.2)). - Definition All_local_env_over_sorting checking sorting cproperty (sproperty : forall Σ Γ _ t s, sorting Σ Γ t s -> Type) := - (All_local_env_over_gen checking (infer_sort sorting) cproperty (fun Σ Γ H t relopt tu => sproperty Σ Γ H t tu.π1 tu.π2.2)). + Definition All_local_env_over_sorting checking sorting cproperty (sproperty : forall Σ Γ _ t s, sorting Σ Γ t s -> Type) Σ := + (All_local_env_over_gen (checking Σ) (infer_sort sorting Σ) (cproperty Σ) (fun Γ H t relopt tu => sproperty _ Γ H t tu.π1 tu.π2.2)). Section TypeCtxInst. - Context (typing : forall (Σ : global_env_ext) (Γ : context), term -> term -> Type). + Context (typing : forall (Γ : context), term -> term -> Type). (* Γ |- s : Δ, where Δ is a telescope (reverse context) *) - Inductive ctx_inst Σ (Γ : context) : list term -> context -> Type := - | ctx_inst_nil : ctx_inst Σ Γ [] [] + Inductive ctx_inst (Γ : context) : list term -> context -> Type := + | ctx_inst_nil : ctx_inst Γ [] [] | ctx_inst_ass na t i inst Δ : - typing Σ Γ i t -> - ctx_inst Σ Γ inst (subst_telescope [i] 0 Δ) -> - ctx_inst Σ Γ (i :: inst) (vass na t :: Δ) + typing Γ i t -> + ctx_inst Γ inst (subst_telescope [i] 0 Δ) -> + ctx_inst Γ (i :: inst) (vass na t :: Δ) | ctx_inst_def na b t inst Δ : - ctx_inst Σ Γ inst (subst_telescope [b] 0 Δ) -> - ctx_inst Σ Γ inst (vdef na b t :: Δ). + ctx_inst Γ inst (subst_telescope [b] 0 Δ) -> + ctx_inst Γ inst (vdef na b t :: Δ). Derive Signature NoConfusion for ctx_inst. End TypeCtxInst. - Lemma ctx_inst_impl P Q Σ Γ inst Δ : - ctx_inst P Σ Γ inst Δ -> - (forall t T, P Σ Γ t T -> Q Σ Γ t T) -> - ctx_inst Q Σ Γ inst Δ. + Lemma ctx_inst_impl P Q Γ inst Δ : + ctx_inst P Γ inst Δ -> + (forall t T, P Γ t T -> Q Γ t T) -> + ctx_inst Q Γ inst Δ. Proof. intros H HPQ. induction H; econstructor; auto. Qed. Section All_local_env_size. - Context {P : forall (Σ : global_env_ext) (Γ : context), term -> typ_or_sort -> Type}. - Context (Psize : forall (Σ : global_env_ext) Γ t T, P Σ Γ t T -> size). + Context {P : forall (Γ : context), term -> typ_or_sort -> Type}. + Context (Psize : forall Γ t T, P Γ t T -> size). - Fixpoint All_local_env_size_gen base Σ Γ (w : All_local_env (P Σ) Γ) : size := + Fixpoint All_local_env_size_gen base Γ (w : All_local_env P Γ) : size := match w with | localenv_nil => base - | localenv_cons_abs Γ' na t w' p => Psize _ _ _ _ p + All_local_env_size_gen base _ _ w' - | localenv_cons_def Γ' na b t w' pt pb => Psize _ _ _ _ pt + Psize _ _ _ _ pb + All_local_env_size_gen base _ _ w' + | localenv_cons_abs Γ' na t w' p => Psize _ _ _ p + All_local_env_size_gen base _ w' + | localenv_cons_def Γ' na b t w' pt pb => Psize _ _ _ pt + Psize _ _ _ pb + All_local_env_size_gen base _ w' end. - Lemma All_local_env_size_pos base Σ Γ w : base <= All_local_env_size_gen base Σ Γ w. + Lemma All_local_env_size_pos base Γ w : base <= All_local_env_size_gen base Γ w. Proof using Type. induction w. all: simpl ; lia. Qed. End All_local_env_size. + + Notation ctx_shifted P Γ := (fun Δ => P (Γ ,,, Δ)). + Notation All_local_rel_size_gen Psize base := (fun Γ Δ (w : All_local_rel _ Γ Δ) => + All_local_env_size_gen (ctx_shifted Psize Γ) base Δ w). - Notation All_local_rel_size_gen Psize base := (fun Σ Γ Δ (w : All_local_rel _ Γ Δ) => - All_local_env_size_gen (fun Σ Δ => Psize Σ (Γ ,,, Δ)) base Σ Δ w). - - Lemma All_local_env_size_app P Psize base Σ Γ Γ' (wfΓ : All_local_env (P Σ) Γ) (wfΓ' : All_local_rel (P Σ) Γ Γ') : - All_local_env_size_gen Psize base _ _ (All_local_rel_app wfΓ wfΓ') + base = All_local_env_size_gen Psize base _ _ wfΓ + All_local_rel_size_gen Psize base _ _ _ wfΓ'. + Lemma All_local_env_size_app P Psize base Γ Γ' (wfΓ : All_local_env P Γ) (wfΓ' : All_local_rel P Γ Γ') : + All_local_env_size_gen Psize base _ (All_local_rel_app wfΓ wfΓ') + base + = All_local_env_size_gen Psize base _ wfΓ + All_local_rel_size_gen Psize base _ _ wfΓ'. Proof. induction Γ'. - dependent inversion wfΓ'. @@ -582,15 +584,15 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). Qed. Section lift_judgment_size. - Context {checking : global_env_ext -> context -> term -> term -> Type}. - Context {sorting : global_env_ext -> context -> term -> option relevance -> Type}. - Context (csize : forall (Σ : global_env_ext) (Γ : context) (t T : term), checking Σ Γ t T -> size). - Context (ssize : forall (Σ : global_env_ext) (Γ : context) (t : term) (relopt : option relevance), sorting Σ Γ t relopt -> size). - - Definition lift_judgment_size Σ Γ t T (w : lift_judgment checking sorting Σ Γ t T) : size := - match T return lift_judgment checking sorting Σ Γ t T -> size with - | Typ T => csize _ _ _ _ - | Sort relopt => ssize _ _ _ _ + Context {checking : context -> term -> term -> Type}. + Context {sorting : context -> term -> option relevance -> Type}. + Context (csize : forall (Γ : context) (t T : term), checking Γ t T -> size). + Context (ssize : forall (Γ : context) (t : term) (relopt : option relevance), sorting Γ t relopt -> size). + + Definition lift_judgment_size Γ t T (w : lift_judgment checking sorting Γ t T) : size := + match T return lift_judgment checking sorting Γ t T -> size with + | Typ T => csize _ _ _ + | Sort relopt => ssize _ _ _ end w. End lift_judgment_size. @@ -603,9 +605,9 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). Context {typing : global_env_ext -> context -> term -> term -> Type}. Context (typing_size : forall Σ Γ t T, typing Σ Γ t T -> size). - Definition lift_typing_size := lift_judgment_size typing_size (infer_sort_size (typing_sort_size typing_size)). - Definition All_local_env_size := All_local_env_size_gen lift_typing_size 0. - Definition All_local_rel_size := All_local_rel_size_gen lift_typing_size 0. + Definition lift_typing_size Σ := lift_judgment_size (typing_size Σ) (infer_sort_size (typing_sort_size typing_size) Σ). + Definition All_local_env_size Σ := All_local_env_size_gen (lift_typing_size Σ) 0. + Definition All_local_rel_size Σ := All_local_rel_size_gen (lift_typing_size Σ) 0. End Regular. Section Bidirectional. @@ -613,9 +615,9 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). Context (checking_size : forall Σ Γ t T, checking Σ Γ t T -> size). Context (sorting_size : forall Σ Γ t s, sorting Σ Γ t s -> size). - Definition lift_sorting_size := lift_judgment_size checking_size (infer_sort_size sorting_size). - Definition All_local_env_sorting_size := All_local_env_size_gen lift_sorting_size 1. - Definition All_local_rel_sorting_size := All_local_rel_size_gen lift_sorting_size 1. + Definition lift_sorting_size Σ := lift_judgment_size (checking_size Σ) (infer_sort_size sorting_size Σ). + Definition All_local_env_sorting_size Σ := All_local_env_size_gen (lift_sorting_size Σ) 1. + Definition All_local_rel_sorting_size Σ := All_local_rel_size_gen (lift_sorting_size Σ) 1. End Bidirectional. End EnvTyping. @@ -948,7 +950,7 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT sorts_local_ctx Σ (arities_context mdecl.(ind_bodies) ,,, mdecl.(ind_params)) cdecl.(cstr_args) cunivs; on_cindices : - ctx_inst (fun Σ Γ t T => P Σ Γ t (Typ T)) Σ (arities_context mdecl.(ind_bodies) ,,, mdecl.(ind_params) ,,, cdecl.(cstr_args)) + ctx_inst (fun Γ t T => P Σ Γ t (Typ T)) (arities_context mdecl.(ind_bodies) ,,, mdecl.(ind_params) ,,, cdecl.(cstr_args)) cdecl.(cstr_indices) (List.rev (lift_context #|cdecl.(cstr_args)| 0 ind_indices)); diff --git a/template-coq/theories/Typing.v b/template-coq/theories/Typing.v index 2a13837a2..8b9b1850f 100644 --- a/template-coq/theories/Typing.v +++ b/template-coq/theories/Typing.v @@ -805,7 +805,7 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> context_assumptions mdecl.(ind_params) = #|p.(pparams)| -> consistent_instance_ext Σ (ind_universes mdecl) p.(puinst) -> let predctx := case_predicate_context ci.(ci_ind) mdecl idecl p in - ctx_inst typing Σ Γ (p.(pparams) ++ indices) + ctx_inst (typing Σ) Γ (p.(pparams) ++ indices) (List.rev (ind_params mdecl ,,, ind_indices idecl)@[p.(puinst)]) -> Σ ;;; Γ ,,, predctx |- p.(preturn) : tSort ps -> is_allowed_elimination Σ idecl.(ind_kelim) ps -> @@ -915,7 +915,7 @@ Section CtxInstSize. Context (typing : global_env_ext -> context -> term -> term -> Type) (typing_size : forall {Σ Γ t T}, typing Σ Γ t T -> size). - Fixpoint ctx_inst_size {Σ Γ args Δ} (c : ctx_inst typing Σ Γ args Δ) : size := + Fixpoint ctx_inst_size {Σ Γ args Δ} (c : ctx_inst (typing Σ) Γ args Δ) : size := match c with | ctx_inst_nil => 0 | ctx_inst_ass na t i inst Δ ty ctxi => ((typing_size _ _ _ _ ty) + (ctx_inst_size ctxi))%nat @@ -1183,7 +1183,7 @@ Lemma typing_ind_env `{cf : checker_flags} : context_assumptions mdecl.(ind_params) = #|p.(pparams)| -> consistent_instance_ext Σ (ind_universes mdecl) p.(puinst) -> let predctx := case_predicate_context ci.(ci_ind) mdecl idecl p in - ctx_inst (Prop_conj typing P) Σ Γ (p.(pparams) ++ indices) + ctx_inst (Prop_conj typing P Σ) Γ (p.(pparams) ++ indices) (List.rev (ind_params mdecl ,,, ind_indices idecl)@[p.(puinst)]) -> forall pret : Σ ;;; Γ ,,, predctx |- p.(preturn) : tSort ps, P Σ (Γ ,,, predctx) p.(preturn) (tSort ps) -> diff --git a/template-coq/theories/TypingWf.v b/template-coq/theories/TypingWf.v index 5900dc6d8..24ea31e0c 100644 --- a/template-coq/theories/TypingWf.v +++ b/template-coq/theories/TypingWf.v @@ -867,7 +867,7 @@ Section TypingWf. - eapply All_local_env_wf_decls. induction X; constructor; auto; red; - unfold wf_decl_pred, typ_or_sort_default, SortRel; + unfold wf_decl_pred, typ_or_sort_default; intuition auto. - split; wf. apply wf_lift. apply (nth_error_all H X). diff --git a/template-coq/theories/Universes.v b/template-coq/theories/Universes.v index b17f136d8..431e6db08 100644 --- a/template-coq/theories/Universes.v +++ b/template-coq/theories/Universes.v @@ -1163,6 +1163,8 @@ Notation isSortRel s rel := (relevance_of_sort s = rel). Definition isSortRelOpt s relopt := match relopt with None => True | Some rel => isSortRel s rel end. +#[global] Hint Unfold isSortRelOpt : core. + Notation "⟦ u ⟧_ v" := (Universe.to_cuniv v u) (at level 0, format "⟦ u ⟧_ v", v name) : univ_scope. From 09d673aff6d8103b71c1694dd16633f6dc52f33d Mon Sep 17 00:00:00 2001 From: Yann Leray Date: Fri, 6 May 2022 11:21:46 +0200 Subject: [PATCH 03/17] First files of safechecker --- pcuic/theories/Bidirectional/BDToPCUIC.v | 14 +++++++++++--- safechecker/theories/PCUICConsistency.v | 5 +++-- safechecker/theories/PCUICSafeConversion.v | 2 +- safechecker/theories/PCUICSafeReduce.v | 16 ++++++++-------- safechecker/theories/PCUICSafeRetyping.v | 2 +- safechecker/theories/PCUICTypeChecker.v | 6 +++--- 6 files changed, 27 insertions(+), 18 deletions(-) diff --git a/pcuic/theories/Bidirectional/BDToPCUIC.v b/pcuic/theories/Bidirectional/BDToPCUIC.v index e38a82697..baca0c3cf 100644 --- a/pcuic/theories/Bidirectional/BDToPCUIC.v +++ b/pcuic/theories/Bidirectional/BDToPCUIC.v @@ -467,13 +467,21 @@ Proof. assumption. Qed. +Theorem infering_sort_isTypeRelOpt `{checker_flags} (Σ : global_env_ext) Γ t u rel (wfΣ : wf Σ) : + wf_local Σ Γ -> isSortRelOpt u rel -> Σ ;;; Γ |- t ▹□ u -> isTypeRelOpt Σ Γ t rel. +Proof. + intros wfΓ Hrel Ht. + exists u. + split; [assumption|]. + now apply infering_sort_typing. +Qed. + Theorem infering_sort_isType `{checker_flags} (Σ : global_env_ext) Γ t u (wfΣ : wf Σ) : wf_local Σ Γ -> Σ ;;; Γ |- t ▹□ u -> isType Σ Γ t. Proof. intros wfΓ Ht. - exists u. - split; [cbnr|]. - now apply infering_sort_typing. + change (isType _ _ _) with (isTypeRelOpt Σ Γ t None). + now eapply infering_sort_isTypeRelOpt. Qed. Theorem einfering_sort_isType `{checker_flags} (Σ : global_env_ext) Γ t (wfΣ : wf Σ) : diff --git a/safechecker/theories/PCUICConsistency.v b/safechecker/theories/PCUICConsistency.v index 8f3b61e19..8fd80aa20 100644 --- a/safechecker/theories/PCUICConsistency.v +++ b/safechecker/theories/PCUICConsistency.v @@ -148,8 +148,9 @@ Proof. - hnf. constructor. + constructor. - * econstructor; cbn; auto. + * econstructor; cbnr; auto. -- exists (Universe.super Prop_univ). + split; [cbnr|]. constructor; auto. constructor. -- instantiate (1 := []). @@ -171,7 +172,7 @@ Proof. set (False_ty := tInd (mkInd (make_fresh_name Σ) 0) []). assert (typ_false: (Σ', Σ.2);;; [] |- tApp t False_ty : False_ty). { apply validity in cons as typ_prod; auto. - destruct typ_prod. + destruct typ_prod as (s & e & Hs). eapply type_App with (B := tRel 0) (u := False_ty); eauto. eapply type_Ind with (u := []) (mdecl := False_mib) (idecl := False_oib); eauto. - hnf. diff --git a/safechecker/theories/PCUICSafeConversion.v b/safechecker/theories/PCUICSafeConversion.v index b3a733b30..6a9c48d6f 100644 --- a/safechecker/theories/PCUICSafeConversion.v +++ b/safechecker/theories/PCUICSafeConversion.v @@ -2654,7 +2654,7 @@ Qed. intros typ. destruct (hΣ _ wfΣ). apply PCUICValidity.inversion_mkApps in typ as (?&[typ_prod typ_args]). - apply inversion_Prod in typ_prod as (?&?&?&?&?); [|easy]. + apply inversion_Prod in typ_prod as (?&?&?&?&?&?); [|easy]. eapply PCUICSpine.typing_spine_strengthen in typ_args; eauto. 2:{ eapply PCUICArities.isType_Sort. 2:pcuic. eapply wf_universe_product; now eapply typing_wf_universe. } diff --git a/safechecker/theories/PCUICSafeReduce.v b/safechecker/theories/PCUICSafeReduce.v index 22c0a4f1f..8fa6f9d47 100644 --- a/safechecker/theories/PCUICSafeReduce.v +++ b/safechecker/theories/PCUICSafeReduce.v @@ -1378,9 +1378,9 @@ Corollary R_Acc_aux : - easy. - apply inversion_Sort in typ as (?&?&?); auto. exfalso; eapply invert_cumul_sort_ind; eauto. - - apply inversion_Prod in typ as (?&?&?&?&?); auto. + - apply inversion_Prod in typ as (?&?&?&?&?&?); auto. exfalso; eapply invert_cumul_sort_ind; eauto. - - apply inversion_Lambda in typ as (?&?&?&?&?); auto. + - apply inversion_Lambda in typ as (?&?&?&?&?&?); auto. exfalso; eapply invert_cumul_prod_ind; eauto. - unfold isConstruct_app in ctor. now rewrite decompose_app_mkApps in ctor. @@ -1411,9 +1411,9 @@ Corollary R_Acc_aux : - easy. - apply inversion_Sort in typ as (?&?&?); auto. exfalso; eapply invert_cumul_sort_ind; eauto. - - apply inversion_Prod in typ as (?&?&?&?&?); auto. + - apply inversion_Prod in typ as (?&?&?&?&?&?); auto. exfalso; eapply invert_cumul_sort_ind; eauto. - - apply inversion_Lambda in typ as (?&?&?&?&?); auto. + - apply inversion_Lambda in typ as (?&?&?&?&?&?); auto. exfalso; eapply invert_cumul_prod_ind; eauto. - unfold isConstruct_app in ctor. now rewrite decompose_app_mkApps in ctor. @@ -1439,7 +1439,7 @@ Corollary R_Acc_aux : intros wf uf shape wh typ. 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. + eapply All_nth_error, isType_of_isTypeRel in a; eauto. eapply wf_fixpoint_spine in i; eauto. 2: { eapply PCUICSpine.typing_spine_strengthen; eauto. } unfold unfold_fix in uf. @@ -1568,7 +1568,7 @@ Corollary R_Acc_aux : apply welltyped_context in h; auto. simpl in h. rewrite stack_context_appstack in h. destruct h as [T h]. - apply inversion_App in h as (?&?&?&?&?); auto. + apply inversion_App in h as (?&?&?&?&?&?); auto. apply inversion_Sort in t0 as (?&?&?); auto. eapply PCUICConversion.ws_cumul_pb_Sort_Prod_inv; eauto. + unfold zipp. @@ -1582,8 +1582,8 @@ Corollary R_Acc_aux : specialize (h _ wfΣ). apply welltyped_context in h; auto. simpl in h. rewrite stack_context_appstack in h. destruct h as [T h]. - apply inversion_App in h as (?&?&?&?&?); auto. - apply inversion_Prod in t0 as (?&?&?&?&?); auto. + apply inversion_App in h as (?&?&?&?&?&?); auto. + apply inversion_Prod in t0 as (?&?&?&?&?&?); auto. eapply PCUICConversion.ws_cumul_pb_Sort_Prod_inv; eauto. (* + pose proof hΣ. sq. diff --git a/safechecker/theories/PCUICSafeRetyping.v b/safechecker/theories/PCUICSafeRetyping.v index f159b06ea..d26c0c2d7 100644 --- a/safechecker/theories/PCUICSafeRetyping.v +++ b/safechecker/theories/PCUICSafeRetyping.v @@ -394,7 +394,7 @@ Qed. sq. constructor ; tea. inversion X0. - eapply infering_sort_isType; eauto. + eapply infering_sort_isTypeRelOpt; eauto. Defined. Next Obligation. cbn ; intros. destruct s1, s2. diff --git a/safechecker/theories/PCUICTypeChecker.v b/safechecker/theories/PCUICTypeChecker.v index d63905b02..67bee2124 100644 --- a/safechecker/theories/PCUICTypeChecker.v +++ b/safechecker/theories/PCUICTypeChecker.v @@ -195,7 +195,7 @@ Lemma isType_mkApps_Ind_smash_inv {cf:checker_flags} {Σ Γ ind u args} (wfΣ : consistent_instance_ext Σ (ind_universes mdecl) u. Proof. move=> wfΓ isType. - destruct isType as [s Hs]. + destruct isType as (s & e & Hs). eapply invert_type_mkApps_ind in Hs as [tyargs cu]; eauto. set (decli' := on_declared_inductive declm). rename declm into decli. @@ -239,7 +239,7 @@ Lemma substitution_wf_local_rel `{checker_flags} {Σ} {wfΣ : wf Σ} {Γ Γ' s constructor ; cbn. + eapply IHΔ ; tea. + rewrite Nat.add_0_r. - eapply isType_substitution ; tea. + eapply isTypeRel_substitution ; tea. now rewrite -app_context_assoc. + rewrite Nat.add_0_r. eapply substitution ; tea. @@ -250,7 +250,7 @@ Lemma substitution_wf_local_rel `{checker_flags} {Σ} {wfΣ : wf Σ} {Γ Γ' s constructor ; cbn. + eapply IHΔ ; tea. + rewrite Nat.add_0_r. - eapply isType_substitution ; tea. + eapply isTypeRel_substitution ; tea. now rewrite -app_context_assoc. Qed. From e7e08f8a3265e93785f29e06739a0ed927d53579 Mon Sep 17 00:00:00 2001 From: Yann Leray Date: Wed, 4 May 2022 11:42:32 +0200 Subject: [PATCH 04/17] Relevance on terms --- erasure/theories/ESubstitution.v | 2 +- pcuic/_CoqProject.in | 4 +- .../theories/Bidirectional/BDStrengthening.v | 2 +- pcuic/theories/Conversion/PCUICInstConv.v | 154 ++- pcuic/theories/Conversion/PCUICNamelessConv.v | 22 +- .../theories/Conversion/PCUICOnFreeVarsConv.v | 216 ---- pcuic/theories/Conversion/PCUICRenameConv.v | 889 +------------ pcuic/theories/Conversion/PCUICRenameTerm.v | 1141 +++++++++++++++++ .../Conversion/PCUICUnivSubstitutionConv.v | 21 +- .../theories/Conversion/PCUICWeakeningConv.v | 15 +- .../Conversion/PCUICWeakeningEnvConv.v | 4 +- pcuic/theories/PCUICAlpha.v | 38 +- pcuic/theories/PCUICArities.v | 2 +- pcuic/theories/PCUICAst.v | 38 +- pcuic/theories/PCUICConfluence.v | 253 ++-- pcuic/theories/PCUICConvCumInversion.v | 26 +- pcuic/theories/PCUICConversion.v | 17 +- pcuic/theories/PCUICCumulProp.v | 16 +- pcuic/theories/PCUICEquality.v | 226 ++-- pcuic/theories/PCUICExpandLetsCorrectness.v | 15 +- pcuic/theories/PCUICFirstorder.v | 14 +- pcuic/theories/PCUICGlobalEnv.v | 54 +- pcuic/theories/PCUICInductives.v | 2 +- pcuic/theories/PCUICNormal.v | 8 +- pcuic/theories/PCUICParallelReduction.v | 2 +- .../PCUICParallelReductionConfluence.v | 2 +- pcuic/theories/PCUICPrincipality.v | 6 +- pcuic/theories/PCUICProgress.v | 4 +- pcuic/theories/PCUICRelevance.v | 49 + pcuic/theories/PCUICRelevanceTerm.v | 40 + pcuic/theories/PCUICSR.v | 6 +- pcuic/theories/PCUICSafeLemmata.v | 2 +- pcuic/theories/PCUICSigmaCalculus.v | 5 +- pcuic/theories/PCUICSpine.v | 4 +- pcuic/theories/PCUICToTemplateCorrectness.v | 18 +- pcuic/theories/PCUICTyping.v | 71 +- pcuic/theories/PCUICWeakeningEnv.v | 55 + pcuic/theories/PCUICWfUniverses.v | 2 +- pcuic/theories/README.md | 4 +- pcuic/theories/Syntax/PCUICCases.v | 18 + pcuic/theories/Syntax/PCUICClosed.v | 2 +- pcuic/theories/Syntax/PCUICInstDef.v | 7 +- pcuic/theories/Syntax/PCUICPosition.v | 48 +- pcuic/theories/TemplateToPCUICCorrectness.v | 43 +- pcuic/theories/Typing/PCUICClosedTyp.v | 4 +- pcuic/theories/Typing/PCUICInstTyp.v | 6 +- pcuic/theories/Typing/PCUICNamelessTyp.v | 2 +- pcuic/theories/Typing/PCUICRenameTyp.v | 10 +- pcuic/theories/Typing/PCUICWeakeningTyp.v | 2 +- pcuic/theories/utils/PCUICAstUtils.v | 2 +- safechecker/theories/PCUICEqualityDec.v | 12 +- safechecker/theories/PCUICSafeConversion.v | 8 +- safechecker/theories/PCUICWfReduction.v | 4 +- template-coq/theories/BasicAst.v | 5 + template-coq/theories/Environment.v | 23 +- template-coq/theories/common/uGraph.v | 5 - template-coq/theories/utils.v | 4 + 57 files changed, 1999 insertions(+), 1655 deletions(-) delete mode 100644 pcuic/theories/Conversion/PCUICOnFreeVarsConv.v create mode 100644 pcuic/theories/Conversion/PCUICRenameTerm.v create mode 100644 pcuic/theories/PCUICRelevance.v create mode 100644 pcuic/theories/PCUICRelevanceTerm.v diff --git a/erasure/theories/ESubstitution.v b/erasure/theories/ESubstitution.v index 33ff6b5c7..48933e81c 100644 --- a/erasure/theories/ESubstitution.v +++ b/erasure/theories/ESubstitution.v @@ -362,7 +362,7 @@ Proof. rewrite subst_inst_case_context_wf. rewrite test_context_k_closed_on_free_vars_ctx. eapply alpha_eq_on_free_vars. symmetry; eassumption. rewrite (wf_predicate_length_pars wfp). - rewrite (PCUICGlobalEnv.declared_minductive_ind_npars decl). rewrite -closedn_ctx_on_free_vars. + rewrite (PCUICTyping.declared_minductive_ind_npars decl). rewrite -closedn_ctx_on_free_vars. eapply PCUICClosedTyp.closed_cstr_branch_context; tea. epose proof (PCUICCasesContexts.inst_case_branch_context_eq (p := subst_predicate s k p) a). now rewrite H. diff --git a/pcuic/_CoqProject.in b/pcuic/_CoqProject.in index 87ecee475..6849a6cb6 100644 --- a/pcuic/_CoqProject.in +++ b/pcuic/_CoqProject.in @@ -31,7 +31,7 @@ theories/Conversion/PCUICWeakeningEnvConv.v theories/Conversion/PCUICUnivSubstitutionConv.v theories/Conversion/PCUICWeakeningConv.v theories/Conversion/PCUICClosedConv.v -theories/Conversion/PCUICOnFreeVarsConv.v +theories/Conversion/PCUICRenameTerm.v theories/Typing/PCUICNamelessTyp.v theories/Typing/PCUICRenameTyp.v @@ -108,6 +108,8 @@ theories/Bidirectional/BDUnique.v theories/Bidirectional/BDStrengthening.v theories/PCUICWeakeningEnv.v +theories/PCUICRelevance.v +theories/PCUICRelevanceTerm.v # theories/All.v diff --git a/pcuic/theories/Bidirectional/BDStrengthening.v b/pcuic/theories/Bidirectional/BDStrengthening.v index 0aa7f2475..8c21f92ee 100644 --- a/pcuic/theories/Bidirectional/BDStrengthening.v +++ b/pcuic/theories/Bidirectional/BDStrengthening.v @@ -4,7 +4,7 @@ From MetaCoq.Template Require Import config utils monad_utils. From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICGlobalEnv PCUICTactics PCUICInduction PCUICLiftSubst PCUICUnivSubst PCUICEquality PCUICUtils - PCUICPosition PCUICTyping PCUICSigmaCalculus PCUICOnFreeVars PCUICClosed PCUICConfluence PCUICSpine PCUICInductiveInversion PCUICParallelReductionConfluence PCUICWellScopedCumulativity PCUICClosed PCUICRenameDef PCUICInstConv PCUICClosedTyp PCUICWeakeningEnvTyp PCUICRenameTyp PCUICRenameConv PCUICGuardCondition PCUICWeakeningConv. + PCUICPosition PCUICTyping PCUICSigmaCalculus PCUICOnFreeVars PCUICClosed PCUICConfluence PCUICSpine PCUICInductiveInversion PCUICParallelReductionConfluence PCUICWellScopedCumulativity PCUICClosed PCUICRenameDef PCUICInstConv PCUICClosedTyp PCUICWeakeningEnvTyp PCUICRenameTyp PCUICRenameTerm PCUICRenameConv PCUICGuardCondition PCUICWeakeningConv. From MetaCoq.PCUIC Require Import BDTyping BDToPCUIC BDFromPCUIC. diff --git a/pcuic/theories/Conversion/PCUICInstConv.v b/pcuic/theories/Conversion/PCUICInstConv.v index 3562b7037..c69be650d 100644 --- a/pcuic/theories/Conversion/PCUICInstConv.v +++ b/pcuic/theories/Conversion/PCUICInstConv.v @@ -3,10 +3,10 @@ From Coq Require Import Morphisms. From MetaCoq.Template Require Import config utils. From MetaCoq.PCUIC Require Import PCUICAst PCUICOnOne PCUICTactics PCUICAstUtils PCUICCases PCUICInduction PCUICLiftSubst PCUICUnivSubst - PCUICTyping PCUICReduction PCUICCumulativity + PCUICTyping PCUICReduction PCUICRelevance PCUICCumulativity PCUICEquality PCUICGlobalEnv PCUICClosed PCUICClosedConv PCUICClosedTyp PCUICEquality PCUICWeakeningEnvConv PCUICWeakeningEnvTyp PCUICSigmaCalculus PCUICRenameDef PCUICRenameConv PCUICWeakeningConv PCUICInstDef PCUICWeakeningTyp - PCUICGuardCondition PCUICUnivSubstitutionConv PCUICOnFreeVars PCUICOnFreeVarsConv. + PCUICGuardCondition PCUICUnivSubstitutionConv PCUICOnFreeVars PCUICRenameTerm. Require Import ssreflect ssrbool. @@ -1343,6 +1343,15 @@ Lemma closed_subst_ext {Δ σ σ' Γ} : - eapply usubst_ext; eauto. Qed. +Lemma valid_subst_ext {Σ Δ σ σ' Γ} : + valid_subst Σ Γ σ Δ -> + σ =1 σ' -> + valid_subst Σ Γ σ' Δ. + intros [Hcl Hrel] eq; split. + - eapply closed_subst_ext => //. apply Hcl. + - intros; rewrite -eq; apply Hrel => //. +Qed. + Lemma well_subst_ext Σ Δ σ σ' Γ : Σ ;;; Δ ⊢ σ : Γ -> σ =1 σ' -> @@ -1396,6 +1405,24 @@ Proof using Type. - eapply usubst_Up; eauto; intuition. Qed. +Lemma valid_subst_Up {Σ Γ Δ σ na A} : + valid_subst Σ Γ σ Δ -> + is_open_term Δ A.[σ] -> + valid_subst Σ (Γ ,, vass na A) (⇑ σ) (Δ ,, vass na A.[σ]). +Proof. + intros [HΔ h] HAσ; split. + - apply closed_subst_Up; assumption. + - intros [|n] decl e. + * now inversion e. + * cbn -[rshiftk marks_of_context] in *. + rewrite /subst_compose. + specialize (h _ _ e). + replace (σ n).[↑] with (lift0 1 (σ n)) by (sigma => //). + apply weakening_relevance with (Γ := Δ) (Γ' := []) (Γ'' := [vass na A]) => //. + destruct HΔ as (hcl & hop & hu). + eapply hop; tea. +Qed. + Lemma well_subst_Up {Σ : global_env_ext} {wfΣ : wf Σ} {Γ Δ σ na A} : wf_local Σ (Δ ,, vass na A.[σ]) -> Σ ;;; Δ ⊢ σ : Γ -> @@ -1467,6 +1494,25 @@ Proof using Type. - eapply usubst_Up'; eauto; intuition. Qed. +Lemma valid_subst_Up' {Σ Γ Δ σ na t A} : + valid_subst Σ Γ σ Δ -> + is_open_term Δ A.[σ] -> + is_open_term Δ t.[σ] -> + valid_subst Σ (Γ ,, vdef na t A) (⇑ σ) (Δ ,, vdef na t.[σ] A.[σ]). +Proof. + intros [HΔ h] HAσ; split. + - apply closed_subst_Up'; assumption. + - intros [|n] decl e. + * now inversion e. + * cbn -[rshiftk marks_of_context] in *. + rewrite /subst_compose. + specialize (h _ _ e). + replace (σ n).[↑] with (lift0 1 (σ n)) by (sigma => //). + apply weakening_relevance with (Γ := Δ) (Γ' := []) (Γ'' := [vdef na t A]) => //. + destruct HΔ as (hcl & hop & hu). + eapply hop; tea. +Qed. + Lemma well_subst_Up' {Σ : global_env_ext} {wfΣ : wf Σ} {Γ Δ σ na t A} : wf_local Σ (Δ ,, vdef na t.[σ] A.[σ]) -> Σ ;;; Δ ⊢ σ : Γ -> @@ -1544,6 +1590,34 @@ Proof. destruct H0; tea. Defined. +Lemma valid_subst_app {Σ Γ Δ σ Δ'} : + valid_subst Σ Γ σ Δ -> + on_free_vars_ctx (shiftnP #|Δ| xpred0) (inst_context σ Δ') -> + valid_subst Σ (Γ ,,, Δ') (⇑^#|Δ'| σ) (Δ ,,, inst_context σ Δ'). +Proof. + intros [hs hrel] hΔ'; split. + - apply closed_subst_app => //. + - intros. + unfold Upn, subst_consn, subst_compose. + destruct (nth_error Δ' x) eqn:E. + * assert (x < #|Δ'|) by (apply nth_error_Some; intros e; rewrite e in E; discriminate). + rewrite nth_error_idsn_Some => //. + cbn. rewrite nth_error_map nth_error_app1. + { len. } + rewrite -nth_error_map /inst_context -map_map map_decl_name_fold_context_k map_map nth_error_map. + erewrite <- nth_error_app1, H => //. + * assert (x >= #|Δ'|) by apply nth_error_None => //. + rewrite nth_error_idsn_None => //. + len. + rewrite nth_error_app2 in H => //. + replace ((σ _).[↑^#|Δ'|]) with (lift0 #|Δ'| (σ (x - #|Δ'|))) by (sigma => //). + replace #|Δ'| with #|inst_context σ Δ'| in H |- * by len. + apply weakening_relevance with (Γ := Δ) (Γ' := []) (Γ'' := inst_context σ Δ') => //. + + destruct hs as (hcl & hop & hu). + eapply hop; tea. + + apply hrel => //. +Qed. + Lemma well_subst_app {Σ : global_env_ext} {wfΣ : wf Σ} {Γ Δ σ Δ'} : wf_local Σ (Δ ,,, inst_context σ Δ') -> Σ ;;; Δ ⊢ σ : Γ -> @@ -1688,14 +1762,52 @@ Lemma on_free_vars_ctx_inst_case_context_nil on_free_vars_ctx P (inst_case_context pars puinst pctx). Proof. intros. - assert (on_free_vars_ctx P ([] ,,, inst_case_context pars puinst pctx)). - { apply on_free_vars_ctx_inst_case_context. - - cbn. rewrite shiftnP0; eauto. - - eauto. - - eauto. - } - rewrite on_free_vars_ctx_app in H1. solve_all. cbn in *. rewrite shiftnP0 in H2. tea. -Defined. + eapply on_free_vars_ctx_inst_case_context; tea. + - reflexivity. + - now rewrite test_context_k_closed_on_free_vars_ctx. +Qed. + +Lemma relevance_of_term_inst Σ Γ Δ t rel σ : + valid_subst Σ Γ σ Δ -> + is_closed_context Γ -> + is_open_term Γ t -> + isTermRel Σ (marks_of_context Γ) t rel -> + isTermRel Σ (marks_of_context Δ) t.[σ] rel. +Proof. + intros hσ HΓ Ht h. + induction t using term_forall_list_ind in Ht, σ, Γ, Δ, hσ, HΓ, h |- *. + all: cbn in Ht |- *; auto. + - destruct hσ as (hσ & hrel). + rewrite /relevance_of_term /marks_of_context nth_error_map /option_map in h. + destruct (nth_error Γ n) eqn:E. + 2: { unfold shiftnP in *. rewrite orb_false_r in Ht. apply Nat.ltb_lt in Ht. apply nth_error_None in E; lia. } + rewrite -h; now apply hrel. + - assert (is_open_term Γ t1). + { rewrite shiftnP_add in Ht; toProp; apply Ht. } + eapply IHt2 with (Δ := Δ ,, vass n t1.[σ]). + 4: { instantiate (1 := Γ ,, vass n t1). cbn in h; apply h. } + 2: cbn; rewrite alli_app; toProp; tas; cbn; rewrite andb_true_r; len. + 2: rewrite shiftnP_add in Ht; toProp; apply Ht. + eapply valid_subst_Up in hσ; tas. + * eapply valid_subst_ext. 1: apply hσ. now sigma. + * eapply inst_is_open_term; tea. apply hσ. + - rewrite shiftnP_add in Ht; move/andP: Ht => [] H0 /andP[] H1 H2. + eapply IHt3 with (Δ := Δ ,, vdef n t1.[σ] t2.[σ]). + 4: { instantiate (1 := Γ ,, vdef n t1 t2). cbn in h; apply h. } + 2: cbn; rewrite alli_app; toProp; tas; cbn; rewrite andb_true_r; len. + 3: apply H2. + 2: { rewrite /on_free_vars_decl /test_decl /decl_body /decl_type /vdef /option_default. toProp => //. } + eapply valid_subst_Up' in hσ; tas. + * eapply valid_subst_ext. 1: apply hσ. now sigma. + * eapply inst_is_open_term; tea. apply hσ. + * eapply inst_is_open_term; tea. apply hσ. + - toProp; eapply IHt1; tea. apply Ht. + - unfold option_default. rewrite nth_error_map. unfold relevance_of_term in h. + now destruct (nth_error m). + - unfold option_default. rewrite nth_error_map. unfold relevance_of_term in h. + now destruct (nth_error m). +Qed. + Lemma red1_inst {Σ : global_env_ext} {wfΣ : wf Σ} {Γ Δ u v σ} : usubst Γ σ Δ -> @@ -1866,21 +1978,21 @@ Proof using Type. destruct X; destruct e as [? [? [ectx ?]]]. rewrite (All2_fold_length ectx). red. intuition auto; simpl; solve_all. - * induction X0 in a, brs' |- *. - + inversion a. constructor. - + inversion a. subst. simpl. - destruct X1 as [a0 e0], p0. - constructor; eauto. + * induction X0 in e0, brs' |- *. + + inversion e0. constructor. + + inversion e0. subst. simpl. + destruct X1 as [a e1], p0. + constructor; unfold eq_branches, eq_branch in *; eauto. split; eauto. simpl. - rewrite (All2_fold_length a0). - now eapply e1. - - simpl. constructor. - apply All2_length in a as e. rewrite <- e. + rewrite (All2_fold_length a). + now eapply e2. + - simpl. constructor. unfold eq_mfix in *. + apply All2_length in e as eq. rewrite <- eq. generalize #|m|. intro k. eapply All2_map. simpl. solve_all. - - simpl. constructor. - apply All2_length in a as e. rewrite <- e. + - simpl. constructor. unfold eq_mfix in *. + apply All2_length in e as eq. rewrite <- eq. generalize #|m|. intro k. eapply All2_map. simpl. solve_all. Qed. diff --git a/pcuic/theories/Conversion/PCUICNamelessConv.v b/pcuic/theories/Conversion/PCUICNamelessConv.v index dfe0c1648..2bec334d4 100644 --- a/pcuic/theories/Conversion/PCUICNamelessConv.v +++ b/pcuic/theories/Conversion/PCUICNamelessConv.v @@ -123,7 +123,7 @@ Proof. + simpl in *. eapply nameless_eqctx_IH; eauto. + ih. - * revert brs' H3 H0 a. + * revert brs' H3 H0 e1. induction l ; intros brs' h1 h2 h. + destruct brs' ; inversion h. reflexivity. + destruct brs' ; inversion h. subst. @@ -133,21 +133,21 @@ Proof. apply forallb_All in Hac. f_equal. ++ destruct a, b. cbn in *. destruct X1. - depelim h. destruct p0. depelim X0. simpl in *. + depelim h. destruct e1. depelim X0. simpl in *. destruct p0 as []. destruct X4. f_equal; try ih. { eapply nameless_eqctx_IH; eauto; solve_all. } ++ eapply IHl ; tas. now depelim X0. - f_equal ; try solve [ ih ]. - revert mfix' H1 H2 H H0 a. + revert mfix' H1 H2 H H0 e. induction m ; intros m' h1 h2 h3 h4 h. + destruct m' ; inversion h. reflexivity. + destruct m' ; inversion h. subst. inversion X. subst. cbn in h1, h2, h3, h4. destruct_andb. f_equal. - * destruct a, d. cbn in *. destruct X0 as [[[? ?] ?] ?]. + * destruct a, d. cbn in *. destruct X0 as (? & ? & ? & ?). destruct H0 as [Hty Hbod]. unfold test_def in H4, H. cbn in H4, H. destruct_andb. @@ -158,14 +158,14 @@ Proof. -- eassumption. * eapply IHm ; assumption. - f_equal ; try solve [ ih ]. - revert mfix' H1 H2 H H0 a. + revert mfix' H1 H2 H H0 e. induction m ; intros m' h1 h2 h3 h4 h. + destruct m' ; inversion h. reflexivity. + destruct m' ; inversion h. subst. inversion X. subst. cbn in h1, h2, h3, h4. destruct_andb. f_equal. - * destruct a, d. cbn in *. destruct X0 as [[[? ?] ?] ?]. + * destruct a, d. cbn in *. destruct X0 as (? & ? & ? & ?). destruct H0 as [Hty Hbod]. unfold test_def in H4, H. cbn in H4, H. destruct_andb. anonify. @@ -329,20 +329,20 @@ Proof. eapply R_global_instance_nl; eauto. - econstructor; eauto. + red. destruct e; intuition auto; simpl. - * induction a0; constructor; auto. + * induction a; constructor; auto. * eapply eq_context_nl_IH; tea. * apply aux. auto. - + induction a; constructor; auto. - intuition auto; simpl. + + induction e0; constructor; auto. + unfold eq_branch in *; intuition auto; simpl. * eapply eq_context_nl_IH; tea. * destruct x, y; simpl in *. apply aux; auto. - econstructor; eauto. - induction a; constructor; auto. + induction e; constructor; auto. intuition auto. * destruct x, y; simpl in *. apply aux; auto. * destruct x, y; simpl in *. apply aux; auto. - econstructor; eauto. - induction a; constructor; auto. + induction e; constructor; auto. intuition auto. * destruct x, y; simpl in *. apply aux; auto. * destruct x, y; simpl in *. apply aux; auto. diff --git a/pcuic/theories/Conversion/PCUICOnFreeVarsConv.v b/pcuic/theories/Conversion/PCUICOnFreeVarsConv.v deleted file mode 100644 index b7704aca8..000000000 --- a/pcuic/theories/Conversion/PCUICOnFreeVarsConv.v +++ /dev/null @@ -1,216 +0,0 @@ -(* Distributed under the terms of the MIT license. *) -From Coq Require Import Morphisms. -From MetaCoq.Template Require Import config utils. -From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICCases PCUICInduction - PCUICLiftSubst PCUICUnivSubst PCUICSigmaCalculus PCUICClosed - PCUICOnFreeVars PCUICTyping PCUICReduction PCUICGlobalEnv PCUICWeakeningEnvConv PCUICClosedConv - PCUICWeakeningEnvTyp PCUICInstDef PCUICRenameDef PCUICRenameConv. - -Require Import ssreflect ssrbool. -From Equations Require Import Equations. - -Lemma on_free_vars_rename P f t : - on_free_vars P (rename f t) = on_free_vars (P ∘ f) t. -Proof. - induction t in P, f |- * using term_forall_list_ind. - - all: cbn => //. - - - erewrite forallb_map, All_forallb_eq_forallb ; tea. - all: eauto. - - by rewrite IHt1 IHt2 shiftnP_shiftn. - - by rewrite IHt1 IHt2 shiftnP_shiftn. - - by rewrite IHt1 IHt2 IHt3 shiftnP_shiftn. - - by rewrite IHt1 IHt2. - - destruct X as (IHpar&ctx&IHret). - f_equal. - 1: erewrite forallb_map, All_forallb_eq_forallb ; tea ; eauto. - f_equal. - 1: by rewrite IHret shiftnP_shiftn. - f_equal. - 1: by rewrite map_length. - f_equal. - 1: auto. - erewrite forallb_map, All_forallb_eq_forallb ; tea. - 1: reflexivity. - intros b []. - f_equal. - 1: by rewrite map_length. - by rewrite /PCUICSigmaCalculus.rename_branch /= e shiftnP_shiftn. - - erewrite forallb_map, All_forallb_eq_forallb ; tea. - 1: reflexivity. - intros ? [? ebod]. - rewrite /test_def /=. - f_equal. - 1: auto. - by rewrite map_length ebod shiftnP_shiftn. - - erewrite forallb_map, All_forallb_eq_forallb ; tea. - 1: reflexivity. - intros ? [? ebod]. - rewrite /test_def /=. - f_equal. - 1: auto. - by rewrite map_length ebod shiftnP_shiftn. -Qed. - -Lemma shiftn_ext_cond (P : nat -> bool) f f' n : - (forall i, P i -> f i = f' i) -> - forall k, - shiftnP n P k -> - shiftn n f k = shiftn n f' k. -Proof. - intros. - unfold shiftn, shiftnP in H0 |- *. - nat_compare_specs ; cbn in *. - now f_equal. -Qed. - -Lemma rename_ext_cond (P : nat -> bool) f f' t : - ( forall i, P i -> f i = f' i ) -> - on_free_vars P t -> - rename f t = rename f' t. -Proof. - intros H Ht. - revert P f f' H Ht. - elim t using term_forall_list_ind; cbn in |- *; intros; try easy. - - 1-6,8: - try rewrite H; try rewrite H0 ; try rewrite H1 ; try easy; - solve [f_equal; solve_all; eauto using shiftn_ext_cond]. - - - f_equal; solve_all. - * eapply map_predicate_shift_eq_spec; solve_all; eauto using shiftn_ext_cond. - * apply map_branch_shift_eq_spec; solve_all; eauto using shiftn_ext_cond. - - f_equal. - rewrite /test_def in Ht. - solve_all ; eauto using shiftn_ext_cond. - - f_equal. - rewrite /test_def in Ht. - solve_all ; eauto using shiftn_ext_cond. -Qed. - -Lemma rename_on_free_vars n t f : - on_free_vars (shiftnP n xpred0) t -> rename (shiftn n f) t = t. -Proof. - intros. - erewrite rename_ext_cond ; tea. - 1: now apply rename_ren_id. - intros i. - now rewrite /shiftnP /shiftn orb_false_r => -> //=. -Qed. - - -Lemma urename_on_free_vars_shift P Γ Δ f u (Ξ: context) : - let sP := shiftnP #|Γ| P in - urenaming sP Δ Γ f -> - is_closed_context Γ -> - is_closed_context Δ -> - on_free_vars (shiftnP #|Ξ| (shiftnP #|Γ| xpred0)) u -> - on_free_vars (shiftnP #|Ξ| (shiftnP #|Δ| xpred0)) - (rename (shiftn #|Ξ| f) u). -Proof. - intros sP hf HΓ HΔ Hu. rewrite on_free_vars_rename. - eapply on_free_vars_impl. 2: tea. clear Hu. intros n Hn. - apply urenaming_context with (Ξ:=Ξ) in hf. - unfold urenaming in hf. - specialize (hf n). destruct (nth_error (Γ,,, Ξ) n) eqn : Hnth. - - specialize (hf c); cbn in hf. forward hf. - * unfold shiftnP in Hn. unfold sP , shiftnP. - toProp. toProp Hn. destruct Hn. - + intuition. - + right. toProp. toProp H. destruct H; intuition. - * destruct (hf eq_refl) as [decl' [Hfn _]]. - clear hf Hn. unfold sP , shiftnP. rewrite orb_false_r. - assert (shiftn #|Ξ| f n < #|Δ,,, rename_context f Ξ|). - { eapply nth_error_Some'. exists decl'. eauto. } - rewrite app_context_length in H. - rewrite rename_context_length in H. - toProp. clear -H. - repeat rewrite PeanoNat.Nat.ltb_lt. lia. -- rewrite nth_error_None in Hnth. rewrite app_context_length in Hnth. unfold shiftnP in *. toProp Hn. toProp. unfold shiftn. - clear -Hn Hnth. destruct Hn. - * toProp H. intuition. - * toProp H. destruct H; [toProp H |]; intuition. -Defined. - -Lemma urename_is_open_term P Γ Δ f u : let sP := shiftnP #|Γ| P in - urenaming sP Δ Γ f -> is_closed_context Γ -> is_closed_context Δ -> is_open_term Γ u -> is_open_term Δ (rename f u). -Proof. - intros sP hf HΓ HΔ Hu. - unfold is_open_term. - rewrite <- (shiftnP0 (shiftnP #|Δ| xpred0)). - rewrite <- (shiftn0 f). - eapply urename_on_free_vars_shift with (Ξ:=[]); eauto. - rewrite shiftnP0; eauto. -Defined. - - - - -Lemma on_free_vars_ctx_inst_case_context - P (Γ : list context_decl) (pars : list term) - (puinst : Instance.t) (pctx : list context_decl) : - forallb (on_free_vars (shiftnP #|Γ| P)) pars -> - on_free_vars_ctx (closedP #|pars| xpredT) pctx -> - on_free_vars_ctx P Γ -> - on_free_vars_ctx P (Γ,,, inst_case_context pars puinst pctx). -Proof. - intros. rewrite on_free_vars_ctx_app H1 /=. - eapply on_free_vars_ctx_inst_case_context; trea; solve_all. - rewrite test_context_k_closed_on_free_vars_ctx //. -Qed. - - -Lemma rename_context_on_free_vars f n l : -on_free_vars_ctx (closedP n xpredT) l -> -rename_context (shiftn n f) l = l. -Proof. - intro Hclosed. - unfold on_free_vars_ctx in Hclosed. - unfold rename_context, fold_context_k. - induction l; eauto. - cbn in *. rewrite alli_app in Hclosed. toProp Hclosed. - destruct Hclosed as [H Hclosed]. - rewrite mapi_rec_app. rewrite List.distr_rev. - rewrite IHl; eauto. - cbn in *. f_equal. - toProp Hclosed. destruct Hclosed as [Hclosed _]. - destruct a; unfold map_decl; cbn. - unfold on_free_vars_decl in Hclosed. - unfold test_decl in Hclosed. - toProp Hclosed. cbn in Hclosed. - destruct Hclosed as [Hbody Htype]. - f_equal. - - destruct decl_body; eauto; cbn in *. - f_equal. rewrite closedP_shiftnP in Hbody. - rewrite shiftnP_add in Hbody. rewrite shiftn_add. - apply rename_on_free_vars; eauto. - - rewrite closedP_shiftnP in Htype. - rewrite shiftnP_add in Htype. rewrite shiftn_add. - apply rename_on_free_vars; eauto. -Defined. - -Lemma inst_case_predicate_context_rename f p : - on_free_vars_ctx (closedP #|pparams p| xpredT) (pcontext p) -> - inst_case_predicate_context (rename_predicate f p) = rename_context f (inst_case_predicate_context p). -Proof. - intro Hclosed. unfold inst_case_predicate_context. - unfold pparams at 1. cbn. - replace (pcontext p) with - (rename_context (shiftn #|(pparams p)| f) (pcontext p)) at 1. - - rewrite <- rename_inst_case_context. reflexivity. - - apply rename_context_on_free_vars; eauto. -Defined. - -Lemma inst_case_branch_context_rename f p x : -on_free_vars_ctx (closedP #|pparams p| xpredT) (bcontext x) -> -inst_case_branch_context (rename_predicate f p) - (rename_branch f x) = - rename_context f (inst_case_branch_context p x). -Proof. - intro Hclosed. unfold inst_case_branch_context. cbn. - replace (bcontext x) with - (rename_context (shiftn #|(pparams p)| f) (bcontext x)) at 1. - - rewrite <- rename_inst_case_context. reflexivity. - - apply rename_context_on_free_vars; eauto. -Defined. diff --git a/pcuic/theories/Conversion/PCUICRenameConv.v b/pcuic/theories/Conversion/PCUICRenameConv.v index 46956f7ec..f03dace73 100644 --- a/pcuic/theories/Conversion/PCUICRenameConv.v +++ b/pcuic/theories/Conversion/PCUICRenameConv.v @@ -2,11 +2,11 @@ From Coq Require Import Morphisms. From MetaCoq.Template Require Import config utils. From MetaCoq.PCUIC Require Import PCUICAst PCUICOnOne PCUICAstUtils PCUICCases PCUICInduction - PCUICLiftSubst PCUICUnivSubst PCUICCumulativity + PCUICLiftSubst PCUICUnivSubst PCUICRelevance PCUICCumulativity PCUICReduction PCUICGlobalEnv PCUICClosed PCUICEquality PCUICRenameDef PCUICWeakeningEnvConv PCUICSigmaCalculus PCUICClosed PCUICOnFreeVars PCUICGuardCondition PCUICWeakeningEnvTyp PCUICClosedConv PCUICClosedTyp PCUICViews - PCUICTyping. + PCUICTyping PCUICRenameTerm. Require Import ssreflect ssrbool. From Equations Require Import Equations. @@ -20,895 +20,10 @@ Set Keyed Unification. Set Default Goal Selector "!". -Lemma isConstruct_app_rename r t : - isConstruct_app t = isConstruct_app (rename r t). -Proof. - unfold isConstruct_app. unfold decompose_app. generalize (@nil term) at 1. - change (@nil term) with (map (rename r) []). generalize (@nil term). - induction t; simpl; auto. - intros l l0. specialize (IHt1 (t2 :: l) (t2 :: l0)). - now rewrite IHt1. -Qed. - -Lemma is_constructor_rename x l r : is_constructor x l = is_constructor x (map (rename r) l). -Proof. - rewrite /is_constructor nth_error_map. - elim: nth_error_spec => //. - move=> n hnth xl /=. - now apply isConstruct_app_rename. -Qed. - -Lemma isFixLambda_app_rename r t : ~~ isFixLambda_app t -> ~~ isFixLambda_app (rename r t). -Proof. - induction t; simpl; auto. - destruct t1; simpl in *; auto. -Qed. - -Lemma construct_cofix_rename r t : discr_construct_cofix t -> discr_construct_cofix (rename r t). -Proof. - destruct t; simpl; try congruence. -Qed. - -Lemma rename_mkApps : - forall f t l, - rename f (mkApps t l) = mkApps (rename f t) (map (rename f) l). -Proof. - intros f t l. - autorewrite with sigma. f_equal. -Qed. - -Lemma decompose_app_rec_rename r t l : - forall hd args, - decompose_app_rec t l = (hd, args) -> - decompose_app_rec (rename r t) (map (rename r) l) = (rename r hd, map (rename r) args). -Proof. - induction t in l |- *; simpl; try intros hd args [= <- <-]; auto. - intros hd args e. now apply (IHt1 _ _ _ e). -Qed. - -Lemma decompose_app_rename {r t hd args} : - decompose_app t = (hd, args) -> - decompose_app (rename r t) = (rename r hd, map (rename r) args). -Proof. apply decompose_app_rec_rename. Qed. - -Lemma rename_subst_instance : - forall u t f, - rename f (subst_instance u t) = subst_instance u (rename f t). -Proof. - intros u t f. - rewrite /subst_instance /=. - induction t in f |- * using term_forall_list_ind. - all: try solve [ - simpl ; - rewrite ?IHt ?IHt1 ?IHt2 ; - easy - ]. - - simpl. f_equal. induction X. - + reflexivity. - + simpl. f_equal ; easy. - - simpl. rewrite IHt. f_equal. - * unfold map_predicate, rename_predicate; destruct p; simpl; f_equal; solve_all. - * induction X0. - + reflexivity. - + simpl. f_equal. 2: easy. - destruct x. unfold rename_branch, map_branch. simpl in *. - f_equal; solve_all. - - simpl. f_equal. - rewrite map_length. - generalize #|m|. intro k. - induction X. 1: reflexivity. - destruct p, x. unfold map_def in *. - simpl in *. f_equal. all: easy. - - simpl. f_equal. - rewrite map_length. - generalize #|m|. intro k. - induction X. 1: reflexivity. - destruct p, x. unfold map_def in *. - simpl in *. f_equal. all: easy. -Qed. - -Lemma rename_context_snoc0 : - forall f Γ d, - rename_context f (d :: Γ) = - rename_context f Γ ,, rename_decl (shiftn #|Γ| f) d. -Proof. - intros f Γ d. - unfold rename_context. now rewrite fold_context_k_snoc0. -Qed. -#[global] -Hint Rewrite rename_context_snoc0 : sigma. - -Lemma rename_context_snoc r Γ d : rename_context r (Γ ,, d) = rename_context r Γ ,, map_decl (rename (shiftn #|Γ| r)) d. -Proof. - unfold snoc. apply rename_context_snoc0. -Qed. -#[global] -Hint Rewrite rename_context_snoc : sigma. - -Lemma rename_context_alt r Γ : - rename_context r Γ = - mapi (fun k' d => map_decl (rename (shiftn (Nat.pred #|Γ| - k') r)) d) Γ. -Proof. - unfold rename_context. apply: fold_context_k_alt. -Qed. - - -Lemma rename_context_telescope r Γ : List.rev (rename_context r Γ) = rename_telescope r (List.rev Γ). -Proof. - rewrite !rename_context_alt /rename_telescope. - rewrite mapi_rev. - f_equal. apply mapi_ext => k' d. - apply map_decl_ext => t. lia_f_equal. -Qed. - -Lemma rename_subst0 : - forall f l t, - rename f (subst0 l t) = - subst0 (map (rename f) l) (rename (shiftn #|l| f) t). -Proof. - intros f l t. - autorewrite with sigma. - eapply inst_ext. - now rewrite -ren_shiftn up_Upn Upn_comp; len. -Qed. - -Lemma rename_subst10 : - forall f t u, - rename f (t{ 0 := u }) = (rename (shiftn 1 f) t){ 0 := rename f u }. -Proof. - intros f t u. - eapply rename_subst0. -Qed. - -Lemma rename_context_nth_error : - forall f Γ i decl, - nth_error Γ i = Some decl -> - nth_error (rename_context f Γ) i = - Some (rename_decl (shiftn (#|Γ| - S i) f) decl). -Proof. - intros f Γ i decl h. - induction Γ in f, i, decl, h |- *. - - destruct i. all: discriminate. - - destruct i. - + simpl in h. inversion h. subst. clear h. - rewrite rename_context_snoc0. simpl. - f_equal. f_equal. f_equal. lia. - + simpl in h. rewrite rename_context_snoc0. simpl. - eapply IHΓ. eassumption. -Qed. - -Lemma rename_context_decl_body : - forall f Γ i body, - option_map decl_body (nth_error Γ i) = Some (Some body) -> - option_map decl_body (nth_error (rename_context f Γ) i) = - Some (Some (rename (shiftn (#|Γ| - S i) f) body)). -Proof. - intros f Γ i body h. - destruct (nth_error Γ i) eqn: e. 2: discriminate. - simpl in h. - eapply rename_context_nth_error with (f := f) in e. rewrite e. simpl. - destruct c as [na bo ty]. simpl in h. inversion h. subst. - simpl. reflexivity. -Qed. - -Lemma map_vass_map_def g l r : - (mapi (fun i (d : def term) => vass (dname d) (lift0 i (dtype d))) - (map (map_def (rename r) g) l)) = - (mapi (fun i d => map_decl (rename (shiftn i r)) d) - (mapi (fun i (d : def term) => vass (dname d) (lift0 i (dtype d))) l)). -Proof. - rewrite mapi_mapi mapi_map. apply mapi_ext. - intros. unfold map_decl, vass; simpl; f_equal. - sigma. rewrite -Upn_ren. now rewrite shiftn_Upn. -Qed. - -Lemma rename_fix_context r : - forall (mfix : list (def term)), - fix_context (map (map_def (rename r) (rename (shiftn #|mfix| r))) mfix) = - rename_context r (fix_context mfix). -Proof. - intros mfix. unfold fix_context. - rewrite map_vass_map_def rev_mapi. - fold (fix_context mfix). - rewrite (rename_context_alt r (fix_context mfix)). - unfold map_decl. now rewrite mapi_length fix_context_length. -Qed. - -Section Renaming. - -Context `{cf : checker_flags}. - -Lemma eq_term_upto_univ_rename Σ : - forall Re Rle napp u v f, - eq_term_upto_univ_napp Σ Re Rle napp u v -> - eq_term_upto_univ_napp Σ Re Rle napp (rename f u) (rename f v). -Proof using Type. - intros Re Rle napp u v f h. - induction u in v, napp, Rle, f, h |- * using term_forall_list_ind. - all: dependent destruction h. - all: try solve [ - simpl ; constructor ; eauto - ]. - - simpl. constructor. - induction X in a, args' |- *. - + inversion a. constructor. - + inversion a. subst. simpl. constructor. - all: eauto. - - simpl. constructor. all: eauto. - * rewrite /rename_predicate. - destruct X; destruct e as [? [? [ectx ?]]]. - rewrite (All2_fold_length ectx). red. - intuition auto; simpl; solve_all. - * red in X0. solve_all. - rewrite -(All2_fold_length a). - now eapply b0. - - simpl. constructor. - apply All2_length in a as e. rewrite <- e. - generalize #|m|. intro k. - induction X in mfix', a, f, k |- *. - + inversion a. constructor. - + inversion a. subst. - simpl. constructor. - * unfold map_def. intuition eauto. - * eauto. - - simpl. constructor. - apply All2_length in a as e. rewrite <- e. - generalize #|m|. intro k. - induction X in mfix', a, f, k |- *. - + inversion a. constructor. - + inversion a. subst. - simpl. constructor. - * unfold map_def. intuition eauto. - * eauto. -Qed. - -Lemma urenaming_impl : - forall (P P' : nat -> bool) Γ Δ f, - (forall i, P' i -> P i) -> - urenaming P Δ Γ f -> - urenaming P' Δ Γ f. -Proof using Type. - intros P P' Γ Δ f hP h. - intros i decl p e. - specialize (h i decl (hP _ p) e) as [decl' [h1 [h2 h3]]]. - exists decl'. split ; [| split ]; eauto. -Qed. - -Lemma inst_closed σ k t : closedn k t -> t.[⇑^k σ] = t. -Proof using Type. - intros Hs. - induction t in σ, k, Hs |- * using term_forall_list_ind; intros; sigma; - simpl in *; - rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_branch_map_branch, - ?map_length, ?Nat.add_assoc in *; - unfold test_def, map_branch, test_branch_k, test_predicate_k in *; simpl in *; eauto with all; - simpl closed in *; repeat (rtoProp; f_equal; solve_all); try change_Sk. - - - revert Hs. - unfold Upn. - elim (Nat.ltb_spec n k) => //; intros; simpl in *. - now apply subst_idsn_consn_lt. - - specialize (IHt2 σ (S k) H0). rewrite -{2}IHt2. now sigma. - - specialize (IHt2 σ (S k) H0). rewrite -{2}IHt2. now sigma. - - specialize (IHt3 σ (S k) H0). rewrite -{2}IHt3. now sigma. - - eapply map_predicate_shift_id_spec. - * solve_all. - * solve_all. - * setoid_rewrite up_Upn; setoid_rewrite <- Upn_Upn. - fold (shiftf (fun k => inst (⇑^k σ)) k). simpl; solve_all. - - rtoProp. specialize (b0 σ (#|bcontext x| + k) H4). - revert b0. now sigma. - - rtoProp. specialize (b0 σ (#|m| + k) H0). - revert b0. now sigma. - - rewrite -Upn_Upn. eauto. -Qed. - -Lemma rename_closedn : - forall f n t, - closedn n t -> - rename (shiftn n f) t = t. -Proof using Type. - intros f n t e. - autorewrite with sigma. - erewrite <- inst_closed with (σ := ren f) by eassumption. - eapply inst_ext. intro i. - unfold ren, shiftn, Upn, subst_consn, subst_compose, shift, shiftk. - rewrite idsn_length. - destruct (Nat.ltb_spec i n). - - rewrite nth_error_idsn_Some. all: auto. - - rewrite nth_error_idsn_None. 1: lia. - simpl. reflexivity. -Qed. - -Lemma rename_closed : - forall f t, - closed t -> - rename f t = t. -Proof using Type. - intros f t h. - replace (rename f t) with (rename (shiftn 0 f) t). - - apply rename_closedn. assumption. - - now sigma. -Qed. - -Lemma rename_closed_decl k f d : closed_decl k d -> map_decl (rename (shiftn k f)) d = d. -Proof using Type. - rewrite /map_decl. - destruct d as [? [] ?] => /=. - - move/andP=> [] clt clty. - rewrite !rename_closedn //. - - move=> clt. rewrite !rename_closedn //. -Qed. - -Lemma rename_closedn_ctx f n Γ : - closedn_ctx n Γ -> - rename_context (shiftn n f) Γ = Γ. -Proof using Type. - rewrite test_context_k_eq /rename_context. - apply alli_fold_context_k. - intros. rewrite shiftn_add. - intros. apply rename_closed_decl. - now rewrite Nat.add_comm. -Qed. - -Lemma rename_closedn_terms f n ts : - forallb (closedn n) ts -> map (rename (shiftn n f)) ts = ts. -Proof using Type. - solve_all. now apply rename_closedn. -Qed. - -Lemma rename_closed_extended_subst f Γ : - closed_ctx Γ -> - map (rename (shiftn (context_assumptions Γ) f)) (extended_subst Γ 0) = extended_subst Γ 0. -Proof using Type. - intros cl. apply rename_closedn_terms. - now apply (closedn_extended_subst_gen Γ 0 0). -Qed. - -Arguments Nat.sub !_ !_. - -Lemma urenaming_vass : - forall P Γ Δ na A f, - urenaming P Γ Δ f -> - urenaming (shiftnP 1 P) (Γ ,, vass na (rename f A)) (Δ ,, vass na A) (shiftn 1 f). -Proof using Type. - intros P Γ Δ na A f h. unfold urenaming in *. - intros [|i] decl hP e. - - simpl in e. inversion e. subst. clear e. - simpl. eexists. split. 1: reflexivity. - repeat split. - now sigma. - - simpl in e. simpl. - replace (i - 0) with i by lia. - eapply h in e as [decl' [? [? [h1 h2]]]]. - 2:{ unfold shiftnP in hP. simpl in hP. now rewrite Nat.sub_0_r in hP. } - eexists. split. 1: eassumption. - repeat split. - + tas. - + setoid_rewrite shiftn_1_S. - rewrite -rename_compose h1. - now sigma. - + move: h2. - destruct (decl_body decl) => /= //; destruct (decl_body decl') => /= //. - setoid_rewrite shiftn_1_S => [=] h2. - now rewrite -rename_compose h2; sigma. -Qed. - - -Lemma urenaming_vdef : - forall P Γ Δ na b B f, - urenaming P Γ Δ f -> - urenaming (shiftnP 1 P) (Γ ,, vdef na (rename f b) (rename f B)) (Δ ,, vdef na b B) (shiftn 1 f). -Proof using Type. - intros P Γ Δ na b B f h. unfold urenaming in *. - intros [|i] decl hP e. - - simpl in e. inversion e. subst. clear e. - simpl. eexists. split. 1: reflexivity. - repeat split. - + now sigma. - + simpl. now sigma. - - simpl in e. simpl. - replace (i - 0) with i by lia. - eapply h in e as [decl' [? [? [h1 h2]]]]. - 2:{ rewrite /shiftnP /= Nat.sub_0_r // in hP. } - eexists. split. 1: eassumption. - repeat split => //. - + setoid_rewrite shiftn_1_S. - rewrite -rename_compose h1. - now sigma. - + move: h2. - destruct (decl_body decl) => /= //; destruct (decl_body decl') => /= //. - setoid_rewrite shiftn_1_S => [=] h2. - now rewrite -rename_compose h2; sigma. -Qed. - -Lemma urenaming_ext : - forall P P' Γ Δ f g, - P =1 P' -> - f =1 g -> - urenaming P Δ Γ f -> - urenaming P' Δ Γ g. -Proof using Type. - intros P P' Γ Δ f g hP hfg h. - intros i decl p e. - rewrite <-hP in p. - specialize (h i decl p e) as [decl' [? [h1 [h2 h3]]]]. - exists decl'. repeat split => //. - - rewrite <- (hfg i). assumption. - - rewrite <- (hfg i). rewrite <- h2. - eapply rename_ext. intros j. symmetry. apply hfg. - - move: h3. destruct (decl_body decl) => /= //. - rewrite /rshiftk. - destruct (decl_body decl') => /= //. - intros [=]; f_equal. - now setoid_rewrite <- (hfg _). -Qed. - -Lemma renaming_extP P P' Σ Γ Δ f : - P =1 P' -> - renaming P Σ Γ Δ f -> renaming P' Σ Γ Δ f. -Proof using Type. - intros hP; rewrite /renaming. - intros []; split; eauto. - eapply urenaming_ext; eauto. reflexivity. -Qed. - -Lemma urenaming_context : - forall P Γ Δ Ξ f, - urenaming P Δ Γ f -> - urenaming (shiftnP #|Ξ| P) (Δ ,,, rename_context f Ξ) (Γ ,,, Ξ) (shiftn #|Ξ| f). -Proof using Type. - intros P Γ Δ Ξ f h. - induction Ξ as [| [na [bo|] ty] Ξ ih] in Γ, Δ, f, h |- *. - - simpl. eapply urenaming_ext. 3: eassumption. - * now rewrite shiftnP0. - * intros []. all: reflexivity. - - simpl. rewrite rename_context_snoc. - rewrite app_context_cons. simpl. unfold rename_decl. unfold map_decl. simpl. - eapply urenaming_ext. - 3:{ eapply urenaming_vdef; tea. eapply ih; assumption. } - * now rewrite shiftnP_add. - * now rewrite shiftn_add. - - simpl. rewrite rename_context_snoc. - rewrite app_context_cons. simpl. unfold rename_decl. unfold map_decl. simpl. - eapply urenaming_ext. - 3:{eapply urenaming_vass; tea. eapply ih; assumption. } - * now rewrite shiftnP_add. - * now rewrite shiftn_add. -Qed. - -Definition rename_branch := (map_branch_shift rename shiftn). - -(* TODO MOVE *) -Lemma isLambda_rename : - forall t f, - isLambda t -> - isLambda (rename f t). -Proof using Type. - intros t f h. - destruct t. - all: try discriminate. - simpl. reflexivity. -Qed. - -(* TODO MOVE *) -Lemma rename_unfold_fix : - forall mfix idx narg fn f, - unfold_fix mfix idx = Some (narg, fn) -> - unfold_fix (map (map_def (rename f) (rename (shiftn #|mfix| f))) mfix) idx - = Some (narg, rename f fn). -Proof using Type. - intros mfix idx narg fn f h. - unfold unfold_fix in *. rewrite nth_error_map. - case_eq (nth_error mfix idx). - 2: intro neq ; rewrite neq in h ; discriminate. - intros d e. rewrite e in h. - inversion h. clear h. - simpl. - f_equal. f_equal. - rewrite rename_subst0. rewrite fix_subst_length. - f_equal. - unfold fix_subst. rewrite map_length. - generalize #|mfix| at 2 3. intro n. - induction n. - - reflexivity. - - simpl. - f_equal. rewrite IHn. reflexivity. -Qed. - -(* TODO MOVE *) -Lemma rename_unfold_cofix : - forall mfix idx narg fn f, - unfold_cofix mfix idx = Some (narg, fn) -> - unfold_cofix (map (map_def (rename f) (rename (shiftn #|mfix| f))) mfix) idx - = Some (narg, rename f fn). -Proof using Type. - intros mfix idx narg fn f h. - unfold unfold_cofix in *. rewrite nth_error_map. - case_eq (nth_error mfix idx). - 2: intro neq ; rewrite neq in h ; discriminate. - intros d e. rewrite e in h. - inversion h. - simpl. f_equal. f_equal. - rewrite rename_subst0. rewrite cofix_subst_length. - f_equal. - unfold cofix_subst. rewrite map_length. - generalize #|mfix| at 2 3. intro n. - induction n. - - reflexivity. - - simpl. - f_equal. rewrite IHn. reflexivity. -Qed. - -Definition rename_constructor_body mdecl f c := - map_constructor_body #|mdecl.(ind_params)| #|mdecl.(ind_bodies)| - (fun k => rename (shiftn k f)) c. - -(* TODO move *) -Lemma map2_set_binder_name_fold bctx f Γ : - #|bctx| = #|Γ| -> - map2 set_binder_name bctx (fold_context_k f Γ) = - fold_context_k f (map2 set_binder_name bctx Γ). -Proof using Type. - intros hl. - rewrite !fold_context_k_alt mapi_map2 -{1}(map_id bctx). - rewrite -mapi_cst_map map2_mapi. - rewrite map2_length; len => //. - eapply map2i_ext => i x y. - rewrite hl. - destruct y; reflexivity. -Qed. - -Lemma rename_subst : - forall f s n t, - rename (shiftn n f) (subst s n t) = - subst (map (rename f) s) n (rename (shiftn n (shiftn #|s| f)) t). -Proof using Type. - intros f s n t. - autorewrite with sigma. - eapply inst_ext. intro i. unfold Upn. - unfold ren, subst_consn, shiftn, subst_compose. simpl. - rewrite nth_error_map. - destruct (Nat.ltb_spec i n). - - rewrite idsn_lt //. simpl. - destruct (Nat.ltb_spec i n) => //. lia. - - rewrite nth_error_idsn_None //. - destruct (Nat.ltb_spec (i - n) #|s|). - * rewrite nth_error_idsn_None //; try lia. - len. - replace (n + (i - n) - n) with (i - n) by lia. - destruct nth_error eqn:hnth => /=. - ** sigma. apply inst_ext. - intros k. cbn. - elim: (Nat.ltb_spec (n + k) n); try lia. - intros. eapply nth_error_Some_length in hnth. - unfold shiftk. lia_f_equal. - ** eapply nth_error_None in hnth. lia. - * len. - replace (n + (#|s| + f (i - n - #|s|)) - n) with - (#|s| + f (i - n - #|s|)) by lia. - rewrite nth_error_idsn_None; try lia. - destruct nth_error eqn:hnth. - ** eapply nth_error_Some_length in hnth. lia. - ** simpl. - eapply nth_error_None in hnth. - destruct nth_error eqn:hnth'. - + eapply nth_error_Some_length in hnth'. lia. - + simpl. unfold shiftk. - case: Nat.ltb_spec; try lia. - intros. lia_f_equal. - assert (n + (i - n - #|s|) - n = (i - n - #|s|)) as -> by lia. - lia. -Qed. - -Lemma rename_context_subst f s Γ : - rename_context f (subst_context s 0 Γ) = - subst_context (map (rename f) s) 0 (rename_context (shiftn #|s| f) Γ). -Proof using Type. - rewrite !rename_context_alt !subst_context_alt. - rewrite !mapi_mapi. apply mapi_ext => i x. - rewrite /subst_decl !compose_map_decl. - apply map_decl_ext => t. - len. - generalize (Nat.pred #|Γ| - i). - intros. - now rewrite rename_subst. -Qed. - -Lemma rename_shiftnk : - forall f n k t, - rename (shiftn (n + k) f) (lift n k t) = lift n k (rename (shiftn k f) t). -Proof using Type. - intros f n k t. - rewrite !lift_rename. - autorewrite with sigma. - rewrite - !Upn_ren. sigma. - rewrite Upn_compose. - rewrite -Upn_Upn Nat.add_comm Upn_Upn Upn_compose. - now rewrite shiftn_Upn. -Qed. - -Lemma rename_context_lift f n k Γ : - rename_context (shiftn (n + k) f) (lift_context n k Γ) = - lift_context n k (rename_context (shiftn k f) Γ). -Proof using Type. - rewrite !rename_context_alt !lift_context_alt. - rewrite !mapi_mapi. apply mapi_ext => i x. - rewrite /lift_decl !compose_map_decl. - apply map_decl_ext => t; len. - generalize (Nat.pred #|Γ| - i). - intros. - rewrite shiftn_add. - rewrite (Nat.add_comm n k) Nat.add_assoc Nat.add_comm. - now rewrite rename_shiftnk shiftn_add. -Qed. - -Lemma rename_inds f ind pu bodies : map (rename f) (inds ind pu bodies) = inds ind pu bodies. -Proof using Type. - unfold inds. - induction #|bodies|; simpl; auto. f_equal. - apply IHn. -Qed. - - -End Renaming. - -#[global] Instance rename_context_ext : Proper (`=1` ==> Logic.eq ==> Logic.eq) rename_context. -Proof. - intros f g Hfg x y ->. - apply fold_context_k_ext => i t. - now rewrite Hfg. -Qed. - - Section Renaming2. Context `{cf : checker_flags}. - -Lemma rename_context_subst_instance f u Γ : - rename_context f (subst_instance u Γ) = - subst_instance u (rename_context f Γ). -Proof using Type. unfold rename_context. - rewrite fold_context_k_map // [subst_instance _ _]map_fold_context_k. - now setoid_rewrite rename_subst_instance. -Qed. - -Lemma rename_context_subst_k f s k Γ : - rename_context (shiftn k f) (subst_context s k Γ) = - subst_context (map (rename f) s) k (rename_context (shiftn (k + #|s|) f) Γ). -Proof using Type. - rewrite /rename_context /subst_context. - rewrite !fold_context_k_compose. - apply fold_context_k_ext => i t. - rewrite shiftn_add. - now rewrite rename_subst !shiftn_add Nat.add_assoc. -Qed. - -Lemma rename_cstr_args mdecl f cdecl : - cstr_args (rename_constructor_body mdecl f cdecl) = - rename_context (shiftn (#|mdecl.(ind_params)| + #|ind_bodies mdecl|) f) (cstr_args cdecl). -Proof using Type. - simpl. unfold rename_context. - apply fold_context_k_ext => i t. - now rewrite shiftn_add Nat.add_assoc. -Qed. - -Lemma rename_reln f ctx n acc : - forallb (closedn (n + #|ctx|)) acc -> - map (rename (shiftn (n + #|ctx|) f)) (reln acc n ctx) = - reln acc n ctx. -Proof using Type. - induction ctx in n, acc |- *; simpl; auto. - - intros clacc. solve_all. now rewrite rename_closedn. - - intros clacc. - destruct a as [? [] ?]. - * rewrite Nat.add_succ_r. - change (S (n + #|ctx|)) with (S n + #|ctx|). - rewrite Nat.add_1_r IHctx // /= -Nat.add_succ_r //. - * rewrite Nat.add_succ_r Nat.add_1_r. rewrite (IHctx (S n)) /= // -Nat.add_succ_r //. - simpl. rewrite clacc andb_true_r. - eapply Nat.ltb_lt. lia. -Qed. - -Lemma rename_to_extended_list f ctx : - map (rename (shiftn #|ctx| f)) (to_extended_list ctx) = to_extended_list ctx. -Proof using Type. - unfold to_extended_list, to_extended_list_k. - now apply (rename_reln _ _ 0). -Qed. - -Lemma to_extended_list_rename f ctx : - to_extended_list (rename_context f ctx) = to_extended_list ctx. -Proof using Type. - unfold to_extended_list, to_extended_list_k. - now rewrite (reln_fold _ _ 0). -Qed. - -Lemma rename_context_set_binder_name f ctx ctx' : - #|ctx| = #|ctx'| -> - rename_context f (map2 set_binder_name ctx ctx') = - map2 set_binder_name ctx (rename_context f ctx'). -Proof using Type. - induction ctx in ctx' |- *; destruct ctx'; simpl; auto. - rewrite !rename_context_snoc /= /snoc. - intros [=]. f_equal; auto. - rewrite /set_binder_name /map_decl /=; f_equal. - - rewrite map2_length // H0 //. - - rewrite map2_length // H0 //. -Qed. - -Lemma rename_inst_case_context f pars puinst ctx : - rename_context f (inst_case_context pars puinst ctx) = - inst_case_context (map (rename f) pars) puinst - (rename_context (shiftn #|pars| f) ctx). -Proof using Type. - rewrite /inst_case_context. - now rewrite rename_context_subst map_rev rename_context_subst_instance; len. -Qed. - -Lemma rename_inst_case_context_wf f pars puinst ctx : - test_context_k (fun k : nat => on_free_vars (closedP k xpredT)) #|pars| ctx -> - rename_context f (inst_case_context pars puinst ctx) = - inst_case_context (map (rename f) pars) puinst ctx. -Proof using Type. - rewrite test_context_k_closed_on_free_vars_ctx => H. - rewrite rename_inst_case_context. f_equal. - now rewrite rename_closedn_ctx // closedn_ctx_on_free_vars. -Qed. - -Lemma rename_cstr_branch_context f ind mdecl cdecl : - closed_ctx (ind_params mdecl) -> - rename_context (shiftn (context_assumptions (ind_params mdecl)) f) (cstr_branch_context ind mdecl cdecl) = - cstr_branch_context ind mdecl (rename_constructor_body mdecl f cdecl). -Proof using Type. - intros clctx. - rewrite /cstr_branch_context. - rewrite /expand_lets_ctx /expand_lets_k_ctx. - rewrite rename_context_subst. - rewrite rename_closed_extended_subst //. f_equal. - rewrite shiftn_add Nat.add_comm. len. - rewrite rename_context_lift. f_equal. - rewrite rename_context_subst_k rename_inds /=; len. - setoid_rewrite <-Nat.add_assoc. - now setoid_rewrite <-shiftn_add at 2. -Qed. - -Lemma rename_case_branch_context_gen ind mdecl f p bctx cdecl : - closed_ctx (ind_params mdecl) -> - #|bctx| = #|cstr_args cdecl| -> - #|pparams p| = context_assumptions (ind_params mdecl) -> - rename_context f (case_branch_context ind mdecl p bctx cdecl) = - case_branch_context ind mdecl (rename_predicate f p) bctx - (rename_constructor_body mdecl f cdecl). -Proof using Type. - intros clpars. unfold case_branch_context, case_branch_context_gen, pre_case_branch_context_gen. - cbn -[fold_context_k]. - intros hlen hlen'. - rewrite rename_context_set_binder_name; len => //. f_equal. - rewrite rename_inst_case_context; f_equal. - rewrite hlen' rename_cstr_branch_context //. -Qed. - -Lemma forget_types_mapi_context (f : nat -> term -> term) (ctx : context) : - forget_types (mapi_context f ctx) = forget_types ctx. -Proof using Type. - now rewrite /forget_types map_mapi_context /= mapi_cst_map. -Qed. - -Lemma forget_types_map_context (f : term -> term) (ctx : context) : - forget_types (map_context f ctx) = forget_types ctx. -Proof using Type. - now rewrite /forget_types map_map_compose. -Qed. - -Lemma rename_closed_constructor_body mdecl cdecl f : - closed_constructor_body mdecl cdecl -> - rename_constructor_body mdecl f cdecl = cdecl. -Proof using Type. - rewrite /closed_constructor_body /rename_constructor_body /map_constructor_body. - move/andP=> [] /andP [] clctx clind clty. - destruct cdecl; cbn -[fold_context_k] in *; f_equal. - + move: clctx. rewrite test_context_k_eq. - apply alli_fold_context_k => i d cldecl. - rewrite rename_closed_decl //. - red; rewrite -cldecl; lia_f_equal. - + solve_all. rewrite rename_closedn //. - red; rewrite -H. lia_f_equal. - + now rewrite rename_closedn. -Qed. - -Lemma rename_mkLambda_or_LetIn f d t : - rename f (mkLambda_or_LetIn d t) = - mkLambda_or_LetIn (rename_decl f d) (rename (shiftn 1 f) t). -Proof using Type. - destruct d as [na [] ty]; rewrite /= /mkLambda_or_LetIn /=; f_equal. -Qed. - -Lemma rename_it_mkLambda_or_LetIn f ctx t : - rename f (it_mkLambda_or_LetIn ctx t) = - it_mkLambda_or_LetIn (rename_context f ctx) (rename (shiftn #|ctx| f) t). -Proof using Type. - move: t. - induction ctx; simpl => t. - - now rewrite shiftn0. - - rewrite /= IHctx rename_context_snoc /snoc /=. f_equal. - now rewrite rename_mkLambda_or_LetIn /= shiftn_add; len. -Qed. - -Lemma rename_mkProd_or_LetIn f d t : - rename f (mkProd_or_LetIn d t) = - mkProd_or_LetIn (rename_decl f d) (rename (shiftn 1 f) t). -Proof using Type. - destruct d as [na [] ty]; rewrite /= /mkProd_or_LetIn /=; f_equal. -Qed. - -Lemma rename_it_mkProd_or_LetIn f ctx t : - rename f (it_mkProd_or_LetIn ctx t) = - it_mkProd_or_LetIn (rename_context f ctx) (rename (shiftn #|ctx| f) t). -Proof using Type. - move: t. - induction ctx; simpl => t. - - now rewrite shiftn0. - - rewrite /= IHctx rename_context_snoc /snoc /=. f_equal. - now rewrite rename_mkProd_or_LetIn /= shiftn_add; len. -Qed. - -Lemma rename_wf_predicate mdecl idecl f p : - wf_predicate mdecl idecl p -> - wf_predicate mdecl idecl (rename_predicate f p). -Proof using Type. - intros []. split => //. now len. -Qed. - -Lemma rename_wf_branch cdecl f br : - wf_branch cdecl br -> - wf_branch cdecl (rename_branch f br). -Proof using Type. intros x; exact x. Qed. - -Lemma rename_wf_branches cdecl f brs : - wf_branches cdecl brs -> - wf_branches cdecl (map (rename_branch f) brs). -Proof using Type. - unfold wf_branches, wf_branches_gen. - intros h. solve_all. -Qed. - -Lemma rename_compose f f' x : rename f (rename f' x) = rename (f ∘ f') x. -Proof using Type. now rewrite (rename_compose _ _ _). Qed. - -Lemma rename_predicate_set_pparams f p params : - rename_predicate f (set_pparams p params) = - set_pparams (rename_predicate f p) (map (rename f) params). -Proof using Type. reflexivity. Qed. - -(* Lemma rename_predicate_set_pcontext f p pcontext' : - #|pcontext'| = #|p.(pcontext)| -> - rename_predicate f (set_pcontext p pcontext') = - set_pcontext (rename_predicate f p) - (mapi_context (fun k => rename (shiftn k f)) pcontext'). -Proof. rewrite /rename_predicate /= /set_pcontext. simpl. intros ->. reflexivity. Qed. *) - -Lemma rename_predicate_set_preturn f p pret : - rename_predicate f (set_preturn p pret) = - set_preturn (rename_predicate f p) (rename (shiftn #|pcontext p| f) pret). -Proof using Type. reflexivity. Qed. - -Lemma rename_extended_subst f Γ : - map (rename (shiftn (context_assumptions Γ) f)) (extended_subst Γ 0) = extended_subst (rename_context f Γ) 0. -Proof using Type. - induction Γ as [|[na [b|] ty] Γ]; auto; rewrite rename_context_snoc /=; len. - - rewrite rename_subst0. - rewrite IHΓ. len. f_equal. f_equal. - now rewrite shiftn_add Nat.add_comm rename_shiftnk. - - f_equal; auto. - rewrite !(lift_extended_subst _ 1). - rewrite map_map_compose. - setoid_rewrite <- (shiftn_add 1 (context_assumptions Γ)). - setoid_rewrite rename_shiftn. - rewrite -map_map_compose. now f_equal. -Qed. - Lemma rename_iota_red : forall f p pars args br, #|skipn pars args| = context_assumptions br.(bcontext) -> diff --git a/pcuic/theories/Conversion/PCUICRenameTerm.v b/pcuic/theories/Conversion/PCUICRenameTerm.v new file mode 100644 index 000000000..7ac7bef32 --- /dev/null +++ b/pcuic/theories/Conversion/PCUICRenameTerm.v @@ -0,0 +1,1141 @@ +(* Distributed under the terms of the MIT license. *) +From Coq Require Import Morphisms. +From MetaCoq.Template Require Import config utils. +From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICCases PCUICInduction + PCUICViews PCUICLiftSubst PCUICUnivSubst PCUICSigmaCalculus PCUICClosed + PCUICOnFreeVars PCUICRenameDef + PCUICEquality PCUICOnOne PCUICRelevance. + (* PCUICTyping PCUICReduction PCUICGlobalEnv PCUICWeakeningEnvConv PCUICClosedConv *) + (* PCUICWeakeningEnvTyp PCUICInstDef PCUICRenameConv. *) + +Require Import ssreflect ssrbool. +From Equations Require Import Equations. +Require Import Equations.Prop.DepElim. +Set Equations With UIP. + +(** * Type preservation for σ-calculus operations *) + +Open Scope sigma_scope. +Set Keyed Unification. + +Set Default Goal Selector "!". + +Lemma isConstruct_app_rename r t : + isConstruct_app t = isConstruct_app (rename r t). +Proof. + unfold isConstruct_app. unfold decompose_app. generalize (@nil term) at 1. + change (@nil term) with (map (rename r) []). generalize (@nil term). + induction t; simpl; auto. + intros l l0. specialize (IHt1 (t2 :: l) (t2 :: l0)). + now rewrite IHt1. +Qed. + +Lemma is_constructor_rename x l r : is_constructor x l = is_constructor x (map (rename r) l). +Proof. + rewrite /is_constructor nth_error_map. + elim: nth_error_spec => //. + move=> n hnth xl /=. + now apply isConstruct_app_rename. +Qed. + +Lemma isFixLambda_app_rename r t : ~~ isFixLambda_app t -> ~~ isFixLambda_app (rename r t). +Proof. + induction t; simpl; auto. + destruct t1; simpl in *; auto. +Qed. + +Lemma construct_cofix_rename r t : discr_construct_cofix t -> discr_construct_cofix (rename r t). +Proof. + destruct t; simpl; try congruence. +Qed. + +Lemma rename_mkApps : + forall f t l, + rename f (mkApps t l) = mkApps (rename f t) (map (rename f) l). +Proof. + intros f t l. + autorewrite with sigma. f_equal. +Qed. + +Lemma decompose_app_rec_rename r t l : + forall hd args, + decompose_app_rec t l = (hd, args) -> + decompose_app_rec (rename r t) (map (rename r) l) = (rename r hd, map (rename r) args). +Proof. + induction t in l |- *; simpl; try intros hd args [= <- <-]; auto. + intros hd args e. now apply (IHt1 _ _ _ e). +Qed. + +Lemma decompose_app_rename {r t hd args} : + decompose_app t = (hd, args) -> + decompose_app (rename r t) = (rename r hd, map (rename r) args). +Proof. apply decompose_app_rec_rename. Qed. + +Lemma rename_subst_instance : + forall u t f, + rename f (subst_instance u t) = subst_instance u (rename f t). +Proof. + intros u t f. + rewrite /subst_instance /=. + induction t in f |- * using term_forall_list_ind. + all: try solve [ + simpl ; + rewrite ?IHt ?IHt1 ?IHt2 ; + easy + ]. + - simpl. f_equal. induction X. + + reflexivity. + + simpl. f_equal ; easy. + - simpl. rewrite IHt. f_equal. + * unfold map_predicate, rename_predicate; destruct p; simpl; f_equal; solve_all. + * induction X0. + + reflexivity. + + simpl. f_equal. 2: easy. + destruct x. unfold rename_branch, map_branch. simpl in *. + f_equal; solve_all. + - simpl. f_equal. + rewrite map_length. + generalize #|m|. intro k. + induction X. 1: reflexivity. + destruct p, x. unfold map_def in *. + simpl in *. f_equal. all: easy. + - simpl. f_equal. + rewrite map_length. + generalize #|m|. intro k. + induction X. 1: reflexivity. + destruct p, x. unfold map_def in *. + simpl in *. f_equal. all: easy. +Qed. + +Lemma rename_context_snoc0 : + forall f Γ d, + rename_context f (d :: Γ) = + rename_context f Γ ,, rename_decl (shiftn #|Γ| f) d. +Proof. + intros f Γ d. + unfold rename_context. now rewrite fold_context_k_snoc0. +Qed. +#[global] +Hint Rewrite rename_context_snoc0 : sigma. + +Lemma rename_context_snoc r Γ d : rename_context r (Γ ,, d) = rename_context r Γ ,, map_decl (rename (shiftn #|Γ| r)) d. +Proof. + unfold snoc. apply rename_context_snoc0. +Qed. +#[global] +Hint Rewrite rename_context_snoc : sigma. + +Lemma rename_context_alt r Γ : + rename_context r Γ = + mapi (fun k' d => map_decl (rename (shiftn (Nat.pred #|Γ| - k') r)) d) Γ. +Proof. + unfold rename_context. apply: fold_context_k_alt. +Qed. + + +Lemma rename_context_telescope r Γ : List.rev (rename_context r Γ) = rename_telescope r (List.rev Γ). +Proof. + rewrite !rename_context_alt /rename_telescope. + rewrite mapi_rev. + f_equal. apply mapi_ext => k' d. + apply map_decl_ext => t. lia_f_equal. +Qed. + +Lemma rename_subst0 : + forall f l t, + rename f (subst0 l t) = + subst0 (map (rename f) l) (rename (shiftn #|l| f) t). +Proof. + intros f l t. + autorewrite with sigma. + eapply inst_ext. + now rewrite -ren_shiftn up_Upn Upn_comp; len. +Qed. + +Lemma rename_subst10 : + forall f t u, + rename f (t{ 0 := u }) = (rename (shiftn 1 f) t){ 0 := rename f u }. +Proof. + intros f t u. + eapply rename_subst0. +Qed. + +Lemma rename_context_nth_error : + forall f Γ i decl, + nth_error Γ i = Some decl -> + nth_error (rename_context f Γ) i = + Some (rename_decl (shiftn (#|Γ| - S i) f) decl). +Proof. + intros f Γ i decl h. + induction Γ in f, i, decl, h |- *. + - destruct i. all: discriminate. + - destruct i. + + simpl in h. inversion h. subst. clear h. + rewrite rename_context_snoc0. simpl. + f_equal. f_equal. f_equal. lia. + + simpl in h. rewrite rename_context_snoc0. simpl. + eapply IHΓ. eassumption. +Qed. + +Lemma rename_context_decl_body : + forall f Γ i body, + option_map decl_body (nth_error Γ i) = Some (Some body) -> + option_map decl_body (nth_error (rename_context f Γ) i) = + Some (Some (rename (shiftn (#|Γ| - S i) f) body)). +Proof. + intros f Γ i body h. + destruct (nth_error Γ i) eqn: e. 2: discriminate. + simpl in h. + eapply rename_context_nth_error with (f := f) in e. rewrite e. simpl. + destruct c as [na bo ty]. simpl in h. inversion h. subst. + simpl. reflexivity. +Qed. + +Lemma map_vass_map_def g l r : + (mapi (fun i (d : def term) => vass (dname d) (lift0 i (dtype d))) + (map (map_def (rename r) g) l)) = + (mapi (fun i d => map_decl (rename (shiftn i r)) d) + (mapi (fun i (d : def term) => vass (dname d) (lift0 i (dtype d))) l)). +Proof. + rewrite mapi_mapi mapi_map. apply mapi_ext. + intros. unfold map_decl, vass; simpl; f_equal. + sigma. rewrite -Upn_ren. now rewrite shiftn_Upn. +Qed. + +Lemma rename_fix_context r : + forall (mfix : list (def term)), + fix_context (map (map_def (rename r) (rename (shiftn #|mfix| r))) mfix) = + rename_context r (fix_context mfix). +Proof. + intros mfix. unfold fix_context. + rewrite map_vass_map_def rev_mapi. + fold (fix_context mfix). + rewrite (rename_context_alt r (fix_context mfix)). + unfold map_decl. now rewrite mapi_length fix_context_length. +Qed. + + + +Section Renaming. + +Context `{cf : checker_flags}. + +Lemma eq_term_upto_univ_rename Σ : + forall Re Rle napp u v f, + eq_term_upto_univ_napp Σ Re Rle napp u v -> + eq_term_upto_univ_napp Σ Re Rle napp (rename f u) (rename f v). +Proof using Type. + intros Re Rle napp u v f h. + induction u in v, napp, Rle, f, h |- * using term_forall_list_ind. + all: dependent destruction h. + all: try solve [ + simpl ; constructor ; eauto + ]. + - simpl. constructor. + induction X in a, args' |- *. + + inversion a. constructor. + + inversion a. subst. simpl. constructor. + all: eauto. + - simpl. constructor. all: eauto. + * rewrite /rename_predicate. + destruct X; destruct e as [? [? [ectx ?]]]. + rewrite (All2_fold_length ectx). red. + intuition auto; simpl; solve_all. + * red in X0. unfold eq_branches, eq_branch in *. solve_all. + rewrite -(All2_fold_length a1). + now eapply b0. + - simpl. constructor. unfold eq_mfix in *. + apply All2_length in e as eq. rewrite <- eq. + generalize #|m|. intro k. + induction X in mfix', e, f, k |- *. + + inversion e. constructor. + + inversion e. subst. + simpl. constructor. + * unfold map_def. intuition eauto. + * eauto. + - simpl. constructor. unfold eq_mfix in *. + apply All2_length in e as eq. rewrite <- eq. + generalize #|m|. intro k. + induction X in mfix', e, f, k |- *. + + inversion e. constructor. + + inversion e. subst. + simpl. constructor. + * unfold map_def. intuition eauto. + * eauto. +Qed. + +Lemma urenaming_impl : + forall (P P' : nat -> bool) Γ Δ f, + (forall i, P' i -> P i) -> + urenaming P Δ Γ f -> + urenaming P' Δ Γ f. +Proof using Type. + intros P P' Γ Δ f hP h. + intros i decl p e. + specialize (h i decl (hP _ p) e) as [decl' [h1 [h2 h3]]]. + exists decl'. split ; [| split ]; eauto. +Qed. + +Lemma inst_closed σ k t : closedn k t -> t.[⇑^k σ] = t. +Proof using Type. + intros Hs. + induction t in σ, k, Hs |- * using term_forall_list_ind; intros; sigma; + simpl in *; + rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_branch_map_branch, + ?map_length, ?Nat.add_assoc in *; + unfold test_def, map_branch, test_branch_k, test_predicate_k in *; simpl in *; eauto with all; + simpl closed in *; repeat (rtoProp; f_equal; solve_all); try change_Sk. + + - revert Hs. + unfold Upn. + elim (Nat.ltb_spec n k) => //; intros; simpl in *. + now apply subst_idsn_consn_lt. + - specialize (IHt2 σ (S k) H0). rewrite -{2}IHt2. now sigma. + - specialize (IHt2 σ (S k) H0). rewrite -{2}IHt2. now sigma. + - specialize (IHt3 σ (S k) H0). rewrite -{2}IHt3. now sigma. + - eapply map_predicate_shift_id_spec. + * solve_all. + * solve_all. + * setoid_rewrite up_Upn; setoid_rewrite <- Upn_Upn. + fold (shiftf (fun k => inst (⇑^k σ)) k). simpl; solve_all. + - rtoProp. specialize (b0 σ (#|bcontext x| + k) H4). + revert b0. now sigma. + - rtoProp. specialize (b0 σ (#|m| + k) H0). + revert b0. now sigma. + - rewrite -Upn_Upn. eauto. +Qed. + +Lemma rename_closedn : + forall f n t, + closedn n t -> + rename (shiftn n f) t = t. +Proof using Type. + intros f n t e. + autorewrite with sigma. + erewrite <- inst_closed with (σ := ren f) by eassumption. + eapply inst_ext. intro i. + unfold ren, shiftn, Upn, subst_consn, subst_compose, shift, shiftk. + rewrite idsn_length. + destruct (Nat.ltb_spec i n). + - rewrite nth_error_idsn_Some. all: auto. + - rewrite nth_error_idsn_None. 1: lia. + simpl. reflexivity. +Qed. + +Lemma rename_closed : + forall f t, + closed t -> + rename f t = t. +Proof using Type. + intros f t h. + replace (rename f t) with (rename (shiftn 0 f) t). + - apply rename_closedn. assumption. + - now sigma. +Qed. + +Lemma rename_closed_decl k f d : closed_decl k d -> map_decl (rename (shiftn k f)) d = d. +Proof using Type. + rewrite /map_decl. + destruct d as [? [] ?] => /=. + - move/andP=> [] clt clty. + rewrite !rename_closedn //. + - move=> clt. rewrite !rename_closedn //. +Qed. + +Lemma rename_closedn_ctx f n Γ : + closedn_ctx n Γ -> + rename_context (shiftn n f) Γ = Γ. +Proof using Type. + rewrite test_context_k_eq /rename_context. + apply alli_fold_context_k. + intros. rewrite shiftn_add. + intros. apply rename_closed_decl. + now rewrite Nat.add_comm. +Qed. + +Lemma rename_closedn_terms f n ts : + forallb (closedn n) ts -> map (rename (shiftn n f)) ts = ts. +Proof using Type. + solve_all. now apply rename_closedn. +Qed. + +Lemma rename_closed_extended_subst f Γ : + closed_ctx Γ -> + map (rename (shiftn (context_assumptions Γ) f)) (extended_subst Γ 0) = extended_subst Γ 0. +Proof using Type. + intros cl. apply rename_closedn_terms. + now apply (closedn_extended_subst_gen Γ 0 0). +Qed. + +Arguments Nat.sub !_ !_. + +Lemma urenaming_vass : + forall P Γ Δ na A f, + urenaming P Γ Δ f -> + urenaming (shiftnP 1 P) (Γ ,, vass na (rename f A)) (Δ ,, vass na A) (shiftn 1 f). +Proof using Type. + intros P Γ Δ na A f h. unfold urenaming in *. + intros [|i] decl hP e. + - simpl in e. inversion e. subst. clear e. + simpl. eexists. split. 1: reflexivity. + repeat split. + now sigma. + - simpl in e. simpl. + replace (i - 0) with i by lia. + eapply h in e as [decl' [? [? [h1 h2]]]]. + 2:{ unfold shiftnP in hP. simpl in hP. now rewrite Nat.sub_0_r in hP. } + eexists. split. 1: eassumption. + repeat split. + + tas. + + setoid_rewrite shiftn_1_S. + rewrite -rename_compose h1. + now sigma. + + move: h2. + destruct (decl_body decl) => /= //; destruct (decl_body decl') => /= //. + setoid_rewrite shiftn_1_S => [=] h2. + now rewrite -rename_compose h2; sigma. +Qed. + + +Lemma urenaming_vdef : + forall P Γ Δ na b B f, + urenaming P Γ Δ f -> + urenaming (shiftnP 1 P) (Γ ,, vdef na (rename f b) (rename f B)) (Δ ,, vdef na b B) (shiftn 1 f). +Proof using Type. + intros P Γ Δ na b B f h. unfold urenaming in *. + intros [|i] decl hP e. + - simpl in e. inversion e. subst. clear e. + simpl. eexists. split. 1: reflexivity. + repeat split. + + now sigma. + + simpl. now sigma. + - simpl in e. simpl. + replace (i - 0) with i by lia. + eapply h in e as [decl' [? [? [h1 h2]]]]. + 2:{ rewrite /shiftnP /= Nat.sub_0_r // in hP. } + eexists. split. 1: eassumption. + repeat split => //. + + setoid_rewrite shiftn_1_S. + rewrite -rename_compose h1. + now sigma. + + move: h2. + destruct (decl_body decl) => /= //; destruct (decl_body decl') => /= //. + setoid_rewrite shiftn_1_S => [=] h2. + now rewrite -rename_compose h2; sigma. +Qed. + +Lemma urenaming_ext : + forall P P' Γ Δ f g, + P =1 P' -> + f =1 g -> + urenaming P Δ Γ f -> + urenaming P' Δ Γ g. +Proof using Type. + intros P P' Γ Δ f g hP hfg h. + intros i decl p e. + rewrite <-hP in p. + specialize (h i decl p e) as [decl' [? [h1 [h2 h3]]]]. + exists decl'. repeat split => //. + - rewrite <- (hfg i). assumption. + - rewrite <- (hfg i). rewrite <- h2. + eapply rename_ext. intros j. symmetry. apply hfg. + - move: h3. destruct (decl_body decl) => /= //. + rewrite /rshiftk. + destruct (decl_body decl') => /= //. + intros [=]; f_equal. + now setoid_rewrite <- (hfg _). +Qed. + +Lemma renaming_extP P P' Σ Γ Δ f : + P =1 P' -> + renaming P Σ Γ Δ f -> renaming P' Σ Γ Δ f. +Proof using Type. + intros hP; rewrite /renaming. + intros []; split; eauto. + eapply urenaming_ext; eauto. reflexivity. +Qed. + +Lemma urenaming_context : + forall P Γ Δ Ξ f, + urenaming P Δ Γ f -> + urenaming (shiftnP #|Ξ| P) (Δ ,,, rename_context f Ξ) (Γ ,,, Ξ) (shiftn #|Ξ| f). +Proof using Type. + intros P Γ Δ Ξ f h. + induction Ξ as [| [na [bo|] ty] Ξ ih] in Γ, Δ, f, h |- *. + - simpl. eapply urenaming_ext. 3: eassumption. + * now rewrite shiftnP0. + * intros []. all: reflexivity. + - simpl. rewrite rename_context_snoc. + rewrite app_context_cons. simpl. unfold rename_decl. unfold map_decl. simpl. + eapply urenaming_ext. + 3:{ eapply urenaming_vdef; tea. eapply ih; assumption. } + * now rewrite shiftnP_add. + * now rewrite shiftn_add. + - simpl. rewrite rename_context_snoc. + rewrite app_context_cons. simpl. unfold rename_decl. unfold map_decl. simpl. + eapply urenaming_ext. + 3:{eapply urenaming_vass; tea. eapply ih; assumption. } + * now rewrite shiftnP_add. + * now rewrite shiftn_add. +Qed. + +Definition rename_branch := (map_branch_shift rename shiftn). + +(* TODO MOVE *) +Lemma isLambda_rename : + forall t f, + isLambda t -> + isLambda (rename f t). +Proof using Type. + intros t f h. + destruct t. + all: try discriminate. + simpl. reflexivity. +Qed. + +(* TODO MOVE *) +Lemma rename_unfold_fix : + forall mfix idx narg fn f, + unfold_fix mfix idx = Some (narg, fn) -> + unfold_fix (map (map_def (rename f) (rename (shiftn #|mfix| f))) mfix) idx + = Some (narg, rename f fn). +Proof using Type. + intros mfix idx narg fn f h. + unfold unfold_fix in *. rewrite nth_error_map. + case_eq (nth_error mfix idx). + 2: intro neq ; rewrite neq in h ; discriminate. + intros d e. rewrite e in h. + inversion h. clear h. + simpl. + f_equal. f_equal. + rewrite rename_subst0. rewrite fix_subst_length. + f_equal. + unfold fix_subst. rewrite map_length. + generalize #|mfix| at 2 3. intro n. + induction n. + - reflexivity. + - simpl. + f_equal. rewrite IHn. reflexivity. +Qed. + +(* TODO MOVE *) +Lemma rename_unfold_cofix : + forall mfix idx narg fn f, + unfold_cofix mfix idx = Some (narg, fn) -> + unfold_cofix (map (map_def (rename f) (rename (shiftn #|mfix| f))) mfix) idx + = Some (narg, rename f fn). +Proof using Type. + intros mfix idx narg fn f h. + unfold unfold_cofix in *. rewrite nth_error_map. + case_eq (nth_error mfix idx). + 2: intro neq ; rewrite neq in h ; discriminate. + intros d e. rewrite e in h. + inversion h. + simpl. f_equal. f_equal. + rewrite rename_subst0. rewrite cofix_subst_length. + f_equal. + unfold cofix_subst. rewrite map_length. + generalize #|mfix| at 2 3. intro n. + induction n. + - reflexivity. + - simpl. + f_equal. rewrite IHn. reflexivity. +Qed. + +Definition rename_constructor_body mdecl f c := + map_constructor_body #|mdecl.(ind_params)| #|mdecl.(ind_bodies)| + (fun k => rename (shiftn k f)) c. + +(* TODO move *) +Lemma map2_set_binder_name_fold bctx f Γ : + #|bctx| = #|Γ| -> + map2 set_binder_name bctx (fold_context_k f Γ) = + fold_context_k f (map2 set_binder_name bctx Γ). +Proof using Type. + intros hl. + rewrite !fold_context_k_alt mapi_map2 -{1}(map_id bctx). + rewrite -mapi_cst_map map2_mapi. + rewrite map2_length; len => //. + eapply map2i_ext => i x y. + rewrite hl. + destruct y; reflexivity. +Qed. + +Lemma rename_subst : + forall f s n t, + rename (shiftn n f) (subst s n t) = + subst (map (rename f) s) n (rename (shiftn n (shiftn #|s| f)) t). +Proof using Type. + intros f s n t. + autorewrite with sigma. + eapply inst_ext. intro i. unfold Upn. + unfold ren, subst_consn, shiftn, subst_compose. simpl. + rewrite nth_error_map. + destruct (Nat.ltb_spec i n). + - rewrite idsn_lt //. simpl. + destruct (Nat.ltb_spec i n) => //. lia. + - rewrite nth_error_idsn_None //. + destruct (Nat.ltb_spec (i - n) #|s|). + * rewrite nth_error_idsn_None //; try lia. + len. + replace (n + (i - n) - n) with (i - n) by lia. + destruct nth_error eqn:hnth => /=. + ** sigma. apply inst_ext. + intros k. cbn. + elim: (Nat.ltb_spec (n + k) n); try lia. + intros. eapply nth_error_Some_length in hnth. + unfold shiftk. lia_f_equal. + ** eapply nth_error_None in hnth. lia. + * len. + replace (n + (#|s| + f (i - n - #|s|)) - n) with + (#|s| + f (i - n - #|s|)) by lia. + rewrite nth_error_idsn_None; try lia. + destruct nth_error eqn:hnth. + ** eapply nth_error_Some_length in hnth. lia. + ** simpl. + eapply nth_error_None in hnth. + destruct nth_error eqn:hnth'. + + eapply nth_error_Some_length in hnth'. lia. + + simpl. unfold shiftk. + case: Nat.ltb_spec; try lia. + intros. lia_f_equal. + assert (n + (i - n - #|s|) - n = (i - n - #|s|)) as -> by lia. + lia. +Qed. + +Lemma rename_context_subst f s Γ : + rename_context f (subst_context s 0 Γ) = + subst_context (map (rename f) s) 0 (rename_context (shiftn #|s| f) Γ). +Proof using Type. + rewrite !rename_context_alt !subst_context_alt. + rewrite !mapi_mapi. apply mapi_ext => i x. + rewrite /subst_decl !compose_map_decl. + apply map_decl_ext => t. + len. + generalize (Nat.pred #|Γ| - i). + intros. + now rewrite rename_subst. +Qed. + +Lemma rename_shiftnk : + forall f n k t, + rename (shiftn (n + k) f) (lift n k t) = lift n k (rename (shiftn k f) t). +Proof using Type. + intros f n k t. + rewrite !lift_rename. + autorewrite with sigma. + rewrite - !Upn_ren. sigma. + rewrite Upn_compose. + rewrite -Upn_Upn Nat.add_comm Upn_Upn Upn_compose. + now rewrite shiftn_Upn. +Qed. + +Lemma rename_context_lift f n k Γ : + rename_context (shiftn (n + k) f) (lift_context n k Γ) = + lift_context n k (rename_context (shiftn k f) Γ). +Proof using Type. + rewrite !rename_context_alt !lift_context_alt. + rewrite !mapi_mapi. apply mapi_ext => i x. + rewrite /lift_decl !compose_map_decl. + apply map_decl_ext => t; len. + generalize (Nat.pred #|Γ| - i). + intros. + rewrite shiftn_add. + rewrite (Nat.add_comm n k) Nat.add_assoc Nat.add_comm. + now rewrite rename_shiftnk shiftn_add. +Qed. + +Lemma rename_inds f ind pu bodies : map (rename f) (inds ind pu bodies) = inds ind pu bodies. +Proof using Type. + unfold inds. + induction #|bodies|; simpl; auto. f_equal. + apply IHn. +Qed. + +End Renaming. + +#[global] Instance rename_context_ext : Proper (`=1` ==> Logic.eq ==> Logic.eq) rename_context. +Proof. + intros f g Hfg x y ->. + apply fold_context_k_ext => i t. + now rewrite Hfg. +Qed. + + +Lemma rename_context_subst_instance f u Γ : + rename_context f (subst_instance u Γ) = + subst_instance u (rename_context f Γ). +Proof using Type. unfold rename_context. + rewrite fold_context_k_map // [subst_instance _ _]map_fold_context_k. + now setoid_rewrite rename_subst_instance. +Qed. + +Lemma rename_context_subst_k f s k Γ : + rename_context (shiftn k f) (subst_context s k Γ) = + subst_context (map (rename f) s) k (rename_context (shiftn (k + #|s|) f) Γ). +Proof using Type. + rewrite /rename_context /subst_context. + rewrite !fold_context_k_compose. + apply fold_context_k_ext => i t. + rewrite shiftn_add. + now rewrite rename_subst !shiftn_add Nat.add_assoc. +Qed. + +Lemma rename_cstr_args mdecl f cdecl : + cstr_args (rename_constructor_body mdecl f cdecl) = + rename_context (shiftn (#|mdecl.(ind_params)| + #|ind_bodies mdecl|) f) (cstr_args cdecl). +Proof using Type. + simpl. unfold rename_context. + apply fold_context_k_ext => i t. + now rewrite shiftn_add Nat.add_assoc. +Qed. + +Lemma rename_reln f ctx n acc : + forallb (closedn (n + #|ctx|)) acc -> + map (rename (shiftn (n + #|ctx|) f)) (reln acc n ctx) = + reln acc n ctx. +Proof using Type. + induction ctx in n, acc |- *; simpl; auto. + - intros clacc. solve_all. now rewrite rename_closedn. + - intros clacc. + destruct a as [? [] ?]. + * rewrite Nat.add_succ_r. + change (S (n + #|ctx|)) with (S n + #|ctx|). + rewrite Nat.add_1_r IHctx // /= -Nat.add_succ_r //. + * rewrite Nat.add_succ_r Nat.add_1_r. rewrite (IHctx (S n)) /= // -Nat.add_succ_r //. + simpl. rewrite clacc andb_true_r. + eapply Nat.ltb_lt. lia. +Qed. + +Lemma rename_to_extended_list f ctx : + map (rename (shiftn #|ctx| f)) (to_extended_list ctx) = to_extended_list ctx. +Proof using Type. + unfold to_extended_list, to_extended_list_k. + now apply (rename_reln _ _ 0). +Qed. + +Lemma to_extended_list_rename f ctx : + to_extended_list (rename_context f ctx) = to_extended_list ctx. +Proof using Type. + unfold to_extended_list, to_extended_list_k. + now rewrite (reln_fold _ _ 0). +Qed. + +Lemma rename_context_set_binder_name f ctx ctx' : + #|ctx| = #|ctx'| -> + rename_context f (map2 set_binder_name ctx ctx') = + map2 set_binder_name ctx (rename_context f ctx'). +Proof using Type. + induction ctx in ctx' |- *; destruct ctx'; simpl; auto. + rewrite !rename_context_snoc /= /snoc. + intros [=]. f_equal; auto. + rewrite /set_binder_name /map_decl /=; f_equal. + - rewrite map2_length // H0 //. + - rewrite map2_length // H0 //. +Qed. + +Lemma rename_inst_case_context f pars puinst ctx : + rename_context f (inst_case_context pars puinst ctx) = + inst_case_context (map (rename f) pars) puinst + (rename_context (shiftn #|pars| f) ctx). +Proof using Type. + rewrite /inst_case_context. + now rewrite rename_context_subst map_rev rename_context_subst_instance; len. +Qed. + +Lemma rename_inst_case_context_wf f pars puinst ctx : + test_context_k (fun k : nat => on_free_vars (closedP k xpredT)) #|pars| ctx -> + rename_context f (inst_case_context pars puinst ctx) = + inst_case_context (map (rename f) pars) puinst ctx. +Proof using Type. + rewrite test_context_k_closed_on_free_vars_ctx => H. + rewrite rename_inst_case_context. f_equal. + now rewrite rename_closedn_ctx // closedn_ctx_on_free_vars. +Qed. + +Lemma rename_cstr_branch_context f ind mdecl cdecl : + closed_ctx (ind_params mdecl) -> + rename_context (shiftn (context_assumptions (ind_params mdecl)) f) (cstr_branch_context ind mdecl cdecl) = + cstr_branch_context ind mdecl (rename_constructor_body mdecl f cdecl). +Proof using Type. + intros clctx. + rewrite /cstr_branch_context. + rewrite /expand_lets_ctx /expand_lets_k_ctx. + rewrite rename_context_subst. + rewrite rename_closed_extended_subst //. f_equal. + rewrite shiftn_add Nat.add_comm. len. + rewrite rename_context_lift. f_equal. + rewrite rename_context_subst_k rename_inds /=; len. + setoid_rewrite <-Nat.add_assoc. + now setoid_rewrite <-shiftn_add at 2. +Qed. + +Lemma rename_case_branch_context_gen ind mdecl f p bctx cdecl : + closed_ctx (ind_params mdecl) -> + #|bctx| = #|cstr_args cdecl| -> + #|pparams p| = context_assumptions (ind_params mdecl) -> + rename_context f (case_branch_context ind mdecl p bctx cdecl) = + case_branch_context ind mdecl (rename_predicate f p) bctx + (rename_constructor_body mdecl f cdecl). +Proof using Type. + intros clpars. unfold case_branch_context, case_branch_context_gen, pre_case_branch_context_gen. + cbn -[fold_context_k]. + intros hlen hlen'. + rewrite rename_context_set_binder_name; len => //. f_equal. + rewrite rename_inst_case_context; f_equal. + rewrite hlen' rename_cstr_branch_context //. +Qed. + +Lemma forget_types_mapi_context (f : nat -> term -> term) (ctx : context) : + forget_types (mapi_context f ctx) = forget_types ctx. +Proof using Type. + now rewrite /forget_types map_mapi_context /= mapi_cst_map. +Qed. + +Lemma forget_types_map_context (f : term -> term) (ctx : context) : + forget_types (map_context f ctx) = forget_types ctx. +Proof using Type. + now rewrite /forget_types map_map_compose. +Qed. + +Lemma rename_closed_constructor_body mdecl cdecl f : + closed_constructor_body mdecl cdecl -> + rename_constructor_body mdecl f cdecl = cdecl. +Proof using Type. + rewrite /closed_constructor_body /rename_constructor_body /map_constructor_body. + move/andP=> [] /andP [] clctx clind clty. + destruct cdecl; cbn -[fold_context_k] in *; f_equal. + + move: clctx. rewrite test_context_k_eq. + apply alli_fold_context_k => i d cldecl. + rewrite rename_closed_decl //. + red; rewrite -cldecl; lia_f_equal. + + solve_all. rewrite rename_closedn //. + red; rewrite -H. lia_f_equal. + + now rewrite rename_closedn. +Qed. + +Lemma rename_mkLambda_or_LetIn f d t : + rename f (mkLambda_or_LetIn d t) = + mkLambda_or_LetIn (rename_decl f d) (rename (shiftn 1 f) t). +Proof using Type. + destruct d as [na [] ty]; rewrite /= /mkLambda_or_LetIn /=; f_equal. +Qed. + +Lemma rename_it_mkLambda_or_LetIn f ctx t : + rename f (it_mkLambda_or_LetIn ctx t) = + it_mkLambda_or_LetIn (rename_context f ctx) (rename (shiftn #|ctx| f) t). +Proof using Type. + move: t. + induction ctx; simpl => t. + - now rewrite shiftn0. + - rewrite /= IHctx rename_context_snoc /snoc /=. f_equal. + now rewrite rename_mkLambda_or_LetIn /= shiftn_add; len. +Qed. + +Lemma rename_mkProd_or_LetIn f d t : + rename f (mkProd_or_LetIn d t) = + mkProd_or_LetIn (rename_decl f d) (rename (shiftn 1 f) t). +Proof using Type. + destruct d as [na [] ty]; rewrite /= /mkProd_or_LetIn /=; f_equal. +Qed. + +Lemma rename_it_mkProd_or_LetIn f ctx t : + rename f (it_mkProd_or_LetIn ctx t) = + it_mkProd_or_LetIn (rename_context f ctx) (rename (shiftn #|ctx| f) t). +Proof using Type. + move: t. + induction ctx; simpl => t. + - now rewrite shiftn0. + - rewrite /= IHctx rename_context_snoc /snoc /=. f_equal. + now rewrite rename_mkProd_or_LetIn /= shiftn_add; len. +Qed. + +Lemma rename_wf_predicate mdecl idecl f p : + wf_predicate mdecl idecl p -> + wf_predicate mdecl idecl (rename_predicate f p). +Proof using Type. + intros []. split => //. now len. +Qed. + +Lemma rename_wf_branch cdecl f br : + wf_branch cdecl br -> + wf_branch cdecl (rename_branch f br). +Proof using Type. intros x; exact x. Qed. + +Lemma rename_wf_branches cdecl f brs : + wf_branches cdecl brs -> + wf_branches cdecl (map (rename_branch f) brs). +Proof using Type. + unfold wf_branches, wf_branches_gen. + intros h. solve_all. +Qed. + +Lemma rename_predicate_set_pparams f p params : + rename_predicate f (set_pparams p params) = + set_pparams (rename_predicate f p) (map (rename f) params). +Proof using Type. reflexivity. Qed. + +(* Lemma rename_predicate_set_pcontext f p pcontext' : + #|pcontext'| = #|p.(pcontext)| -> + rename_predicate f (set_pcontext p pcontext') = + set_pcontext (rename_predicate f p) + (mapi_context (fun k => rename (shiftn k f)) pcontext'). +Proof. rewrite /rename_predicate /= /set_pcontext. simpl. intros ->. reflexivity. Qed. *) + +Lemma rename_predicate_set_preturn f p pret : + rename_predicate f (set_preturn p pret) = + set_preturn (rename_predicate f p) (rename (shiftn #|pcontext p| f) pret). +Proof using Type. reflexivity. Qed. + +Lemma rename_extended_subst f Γ : + map (rename (shiftn (context_assumptions Γ) f)) (extended_subst Γ 0) = extended_subst (rename_context f Γ) 0. +Proof using Type. + induction Γ as [|[na [b|] ty] Γ]; auto; rewrite rename_context_snoc /=; len. + - rewrite rename_subst0. + rewrite IHΓ. len. f_equal. f_equal. + now rewrite shiftn_add Nat.add_comm rename_shiftnk. + - f_equal; auto. + rewrite !(lift_extended_subst _ 1). + rewrite map_map_compose. + setoid_rewrite <- (shiftn_add 1 (context_assumptions Γ)). + setoid_rewrite rename_shiftn. + rewrite -map_map_compose. now f_equal. +Qed. + +Lemma on_free_vars_rename P f t : + on_free_vars P (rename f t) = on_free_vars (P ∘ f) t. +Proof. + induction t in P, f |- * using term_forall_list_ind. + + all: cbn => //. + + - erewrite forallb_map, All_forallb_eq_forallb ; tea. + all: eauto. + - by rewrite IHt1 IHt2 shiftnP_shiftn. + - by rewrite IHt1 IHt2 shiftnP_shiftn. + - by rewrite IHt1 IHt2 IHt3 shiftnP_shiftn. + - by rewrite IHt1 IHt2. + - destruct X as (IHpar&ctx&IHret). + f_equal. + 1: erewrite forallb_map, All_forallb_eq_forallb ; tea ; eauto. + f_equal. + 1: by rewrite IHret shiftnP_shiftn. + f_equal. + 1: by rewrite map_length. + f_equal. + 1: auto. + erewrite forallb_map, All_forallb_eq_forallb ; tea. + 1: reflexivity. + intros b []. + f_equal. + 1: by rewrite map_length. + by rewrite /PCUICSigmaCalculus.rename_branch /= e shiftnP_shiftn. + - erewrite forallb_map, All_forallb_eq_forallb ; tea. + 1: reflexivity. + intros ? [? ebod]. + rewrite /test_def /=. + f_equal. + 1: auto. + by rewrite map_length ebod shiftnP_shiftn. + - erewrite forallb_map, All_forallb_eq_forallb ; tea. + 1: reflexivity. + intros ? [? ebod]. + rewrite /test_def /=. + f_equal. + 1: auto. + by rewrite map_length ebod shiftnP_shiftn. +Qed. + +Lemma shiftn_ext_cond (P : nat -> bool) f f' n : + (forall i, P i -> f i = f' i) -> + forall k, + shiftnP n P k -> + shiftn n f k = shiftn n f' k. +Proof. + intros. + unfold shiftn, shiftnP in H0 |- *. + nat_compare_specs ; cbn in *. + now f_equal. +Qed. + +Lemma rename_ext_cond (P : nat -> bool) f f' t : + ( forall i, P i -> f i = f' i ) -> + on_free_vars P t -> + rename f t = rename f' t. +Proof. + intros H Ht. + revert P f f' H Ht. + elim t using term_forall_list_ind; cbn in |- *; intros; try easy. + + 1-6,8: + try rewrite H; try rewrite H0 ; try rewrite H1 ; try easy; + solve [f_equal; solve_all; eauto using shiftn_ext_cond]. + + - f_equal; solve_all. + * eapply map_predicate_shift_eq_spec; solve_all; eauto using shiftn_ext_cond. + * apply map_branch_shift_eq_spec; solve_all; eauto using shiftn_ext_cond. + - f_equal. + rewrite /test_def in Ht. + solve_all ; eauto using shiftn_ext_cond. + - f_equal. + rewrite /test_def in Ht. + solve_all ; eauto using shiftn_ext_cond. +Qed. + +Lemma rename_on_free_vars n t f : + on_free_vars (shiftnP n xpred0) t -> rename (shiftn n f) t = t. +Proof. + intros. + erewrite rename_ext_cond ; tea. + 1: now apply rename_ren_id. + intros i. + now rewrite /shiftnP /shiftn orb_false_r => -> //=. +Qed. + + +Lemma urename_on_free_vars_shift P Γ Δ f u (Ξ: context) : + let sP := shiftnP #|Γ| P in + urenaming sP Δ Γ f -> + is_closed_context Γ -> + is_closed_context Δ -> + on_free_vars (shiftnP #|Ξ| (shiftnP #|Γ| xpred0)) u -> + on_free_vars (shiftnP #|Ξ| (shiftnP #|Δ| xpred0)) + (rename (shiftn #|Ξ| f) u). +Proof. + intros sP hf HΓ HΔ Hu. rewrite on_free_vars_rename. + eapply on_free_vars_impl. 2: tea. clear Hu. intros n Hn. + apply urenaming_context with (Ξ:=Ξ) in hf. + unfold urenaming in hf. + specialize (hf n). destruct (nth_error (Γ,,, Ξ) n) eqn : Hnth. + - specialize (hf c); cbn in hf. forward hf. + * unfold shiftnP in Hn. unfold sP , shiftnP. + toProp. toProp Hn. destruct Hn. + + intuition. + + right. toProp. toProp H. destruct H; intuition. + * destruct (hf eq_refl) as [decl' [Hfn _]]. + clear hf Hn. unfold sP , shiftnP. rewrite orb_false_r. + assert (shiftn #|Ξ| f n < #|Δ,,, rename_context f Ξ|). + { eapply nth_error_Some'. exists decl'. eauto. } + rewrite app_context_length in H. + rewrite rename_context_length in H. + toProp. clear -H. + repeat rewrite PeanoNat.Nat.ltb_lt. lia. +- rewrite nth_error_None in Hnth. rewrite app_context_length in Hnth. unfold shiftnP in *. toProp Hn. toProp. unfold shiftn. + clear -Hn Hnth. destruct Hn. + * toProp H. intuition. + * toProp H. destruct H; [toProp H |]; intuition. +Defined. + +Lemma urename_is_open_term P Γ Δ f u : let sP := shiftnP #|Γ| P in + urenaming sP Δ Γ f -> is_closed_context Γ -> is_closed_context Δ -> is_open_term Γ u -> is_open_term Δ (rename f u). +Proof. + intros sP hf HΓ HΔ Hu. + unfold is_open_term. + rewrite <- (shiftnP0 (shiftnP #|Δ| xpred0)). + rewrite <- (shiftn0 f). + eapply urename_on_free_vars_shift with (Ξ:=[]); eauto. + rewrite shiftnP0; eauto. +Defined. + + +Lemma on_free_vars_ctx_inst_case_context_app + P (Γ : list context_decl) (pars : list term) + (puinst : Instance.t) (pctx : list context_decl) : + forallb (on_free_vars (shiftnP #|Γ| P)) pars -> + on_free_vars_ctx (closedP #|pars| xpredT) pctx -> + on_free_vars_ctx P Γ -> + on_free_vars_ctx P (Γ,,, inst_case_context pars puinst pctx). +Proof. + intros. rewrite on_free_vars_ctx_app H1 /=. + eapply on_free_vars_ctx_inst_case_context; trea; solve_all. + rewrite test_context_k_closed_on_free_vars_ctx //. +Qed. + + +Lemma rename_context_on_free_vars f n l : +on_free_vars_ctx (closedP n xpredT) l -> +rename_context (shiftn n f) l = l. +Proof. + intro Hclosed. + unfold on_free_vars_ctx in Hclosed. + unfold rename_context, fold_context_k. + induction l; eauto. + cbn in *. rewrite alli_app in Hclosed. toProp Hclosed. + destruct Hclosed as [H Hclosed]. + rewrite mapi_rec_app. rewrite List.distr_rev. + rewrite IHl; eauto. + cbn in *. f_equal. + toProp Hclosed. destruct Hclosed as [Hclosed _]. + destruct a; unfold map_decl; cbn. + unfold on_free_vars_decl in Hclosed. + unfold test_decl in Hclosed. + toProp Hclosed. cbn in Hclosed. + destruct Hclosed as [Hbody Htype]. + f_equal. + - destruct decl_body; eauto; cbn in *. + f_equal. rewrite closedP_shiftnP in Hbody. + rewrite shiftnP_add in Hbody. rewrite shiftn_add. + apply rename_on_free_vars; eauto. + - rewrite closedP_shiftnP in Htype. + rewrite shiftnP_add in Htype. rewrite shiftn_add. + apply rename_on_free_vars; eauto. +Defined. + +Lemma inst_case_predicate_context_rename f p : + on_free_vars_ctx (closedP #|pparams p| xpredT) (pcontext p) -> + inst_case_predicate_context (rename_predicate f p) = rename_context f (inst_case_predicate_context p). +Proof. + intro Hclosed. unfold inst_case_predicate_context. + unfold pparams at 1. cbn. + replace (pcontext p) with + (rename_context (shiftn #|(pparams p)| f) (pcontext p)) at 1. + - rewrite <- rename_inst_case_context. reflexivity. + - apply rename_context_on_free_vars; eauto. +Defined. + +Lemma inst_case_branch_context_rename f p x : +on_free_vars_ctx (closedP #|pparams p| xpredT) (bcontext x) -> +inst_case_branch_context (rename_predicate f p) + (rename_branch f x) = + rename_context f (inst_case_branch_context p x). +Proof. + intro Hclosed. unfold inst_case_branch_context. cbn. + replace (bcontext x) with + (rename_context (shiftn #|(pparams p)| f) (bcontext x)) at 1. + - rewrite <- rename_inst_case_context. reflexivity. + - apply rename_context_on_free_vars; eauto. +Defined. + + +Lemma rename_relevance_of_term P f Σ Δ Γ t rel : + let sP := shiftnP #|Γ| P in + urenaming sP Δ Γ f -> + is_open_term Γ t -> + isTermRel Σ (marks_of_context Γ) t rel -> + isTermRel Σ (marks_of_context Δ) (rename f t) rel. +Proof. + revert f Γ Δ. + induction t using term_forall_list_ind; cbnr; intros f Γ Δ Hur (* wfΓ wfΔ *) wft H; auto. + - unfold relevance_of_term, option_default in *. + rewrite nth_error_map /option_map in H. + rewrite nth_error_map /option_map. + destruct (nth_error Γ n) eqn:E. + 2: { unfold shiftnP in *. rewrite orb_false_r in wft. apply Nat.ltb_lt in wft. apply nth_error_None in E; lia. } + eapply Hur in E as (decl & -> & <- & _) => //. + clear -wft. unfold shiftnP in *. toProp. rewrite orb_false_r in wft. now left. + - cbn in wft; rtoProp. + eapply IHt2 with (Δ := Δ,, vass n (rename _ t1)). + 3: instantiate (1 := Γ,, vass n t1); apply H. + * eapply urenaming_impl. 1: intro; rewrite shiftnP_S; eauto. apply urenaming_vass; eauto. + * eapply on_free_vars_impl; tea. 1: intro; rewrite -shiftnP_S; eauto. + - cbn in wft; rtoProp. + eapply IHt3 with (Δ := Δ,, vdef n (rename _ t1) (rename _ t2)). + 3: instantiate (1 := Γ,, vdef n t1 t2); apply H. + * eapply urenaming_impl. 1: intro; rewrite shiftnP_S; eauto. apply urenaming_vdef; eauto. + * eapply on_free_vars_impl; tea. 1: intro; rewrite -shiftnP_S; eauto. + - eapply IHt1; tea. now toProp wft. + - unfold relevance_of_term, option_default in *. + rewrite nth_error_map. unfold option_map; cbn. destruct (nth_error m) eqn:E; tas. + - unfold relevance_of_term, option_default in *. + rewrite nth_error_map. unfold option_map; cbn. destruct (nth_error m) eqn:E; tas. +Qed. diff --git a/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v b/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v index fa3bff8c1..800525136 100644 --- a/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v +++ b/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v @@ -3,7 +3,7 @@ From Coq Require Import ssreflect CRelationClasses. From MetaCoq.Template Require Import utils config Universes uGraph. From MetaCoq.PCUIC Require Import PCUICAst PCUICOnOne PCUICAstUtils PCUICInduction PCUICLiftSubst PCUICEquality PCUICUnivSubst - PCUICCases PCUICCumulativity PCUICTyping + PCUICCases PCUICRelevance PCUICCumulativity PCUICTyping PCUICReduction PCUICWeakeningEnv PCUICClosed PCUICPosition PCUICGuardCondition. @@ -188,7 +188,9 @@ Proof. * eapply subst_equal_inst_inst => //. * solve_all. reflexivity. * eapply X => //. - - solve_all. reflexivity. + - rewrite /eq_branch /id /=. split. + * reflexivity. + * apply e => //. Qed. #[global] @@ -879,6 +881,7 @@ Proof. repeat split; simpl; eauto; solve_all. * eapply precompose_subst_instance. eapply R_universe_instance_impl; eauto. + - destruct X6; split => //. apply e => //. Qed. Lemma leq_term_subst_instance {cf : checker_flags} Σ : SubstUnivPreserved (leq_term Σ). @@ -1687,6 +1690,20 @@ Qed. Definition wf_global_ext {cf : checker_flags} Σ ext := wf_ext_wk (Σ, ext). + +Lemma relevance_of_term_subst_instance Σ Γ u t rel : + isTermRel Σ Γ t rel -> isTermRel Σ Γ t@[u] rel. +Proof. + revert Γ. + induction t using term_forall_list_ind; intros Γ Hirrel; cbnr; auto. + - unfold subst_instance, relevance_of_term, option_default in *. + destruct (nth_error m) eqn:E; + rewrite nth_error_map E; now cbn. + - unfold subst_instance, relevance_of_term, option_default in *. + destruct (nth_error m) eqn:E; + rewrite nth_error_map E; now cbn. +Qed. + Require Import Morphisms. Require Import ssreflect. Set SimplIsCbn. diff --git a/pcuic/theories/Conversion/PCUICWeakeningConv.v b/pcuic/theories/Conversion/PCUICWeakeningConv.v index 57889ec69..d6793f50f 100644 --- a/pcuic/theories/Conversion/PCUICWeakeningConv.v +++ b/pcuic/theories/Conversion/PCUICWeakeningConv.v @@ -2,9 +2,9 @@ From Coq Require Import Morphisms. From MetaCoq.Template Require Import config utils. From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICCases PCUICInduction - PCUICLiftSubst PCUICTyping PCUICCumulativity + PCUICLiftSubst PCUICTyping PCUICCumulativity PCUICRelevance PCUICClosed PCUICReduction - PCUICSigmaCalculus PCUICRenameDef PCUICRenameConv PCUICOnFreeVars + PCUICSigmaCalculus PCUICRenameDef PCUICRenameTerm PCUICRenameConv PCUICOnFreeVars PCUICClosedConv PCUICClosedTyp. Require Import ssreflect ssrbool. @@ -93,6 +93,17 @@ Proof. apply rename_ext => k. simpl. now repeat nat_compare_specs. Qed. +Lemma weakening_relevance Σ Γ Γ' Γ'' : + forall t r, is_open_term (Γ ,,, Γ') t -> + isTermRel Σ (marks_of_context (Γ ,,, Γ')) t r -> + isTermRel Σ (marks_of_context (Γ ,,, Γ'' ,,, lift_context #|Γ''| 0 Γ')) (lift #|Γ''| #|Γ'| t) r. +Proof. + intros. + rewrite lift_rename. + eapply rename_relevance_of_term with (P := xpred0); tea. + apply weakening_renaming. +Qed. + (* Variant lookup_decl_spec Γ Δ i : option context_decl -> Type := | lookup_head d : i < #|Δ| -> nth_error Δ i = Some d -> lookup_decl_spec Γ Δ i (Some d) diff --git a/pcuic/theories/Conversion/PCUICWeakeningEnvConv.v b/pcuic/theories/Conversion/PCUICWeakeningEnvConv.v index 6c95a7348..1cd60da16 100644 --- a/pcuic/theories/Conversion/PCUICWeakeningEnvConv.v +++ b/pcuic/theories/Conversion/PCUICWeakeningEnvConv.v @@ -118,11 +118,11 @@ Proof using P Pcmp cf. - inversion 1; subst; constructor. eapply All2_impl'; tea. eapply All_impl; eauto. - cbn. intros x [? ?] y [[[? ?] ?] ?]. repeat split; eauto. + cbn. intros x [? ?] y (? & ? & ? & ?). repeat split; eauto. - inversion 1; subst; constructor. eapply All2_impl'; tea. eapply All_impl; eauto. - cbn. intros x [? ?] y [[[? ?] ?] ?]. repeat split; eauto. + cbn. intros x [? ?] y (? & ? & ? & ?). repeat split; eauto. Qed. Lemma weakening_env_red1 Σ Σ' Γ M N : diff --git a/pcuic/theories/PCUICAlpha.v b/pcuic/theories/PCUICAlpha.v index b4739323d..deb0de897 100644 --- a/pcuic/theories/PCUICAlpha.v +++ b/pcuic/theories/PCUICAlpha.v @@ -83,11 +83,11 @@ Section Alpha. (* TODO MOVE *) Lemma nth_error_weak_context : - forall Γ Δ i d, + forall {A} Γ Δ i (d : A), nth_error Δ i = Some d -> nth_error (Γ ,,, Δ) i = Some d. Proof using Type. - intros Γ Δ i d h. + intros ? Γ Δ i d h. rewrite -> nth_error_app_context_lt. - assumption. - apply nth_error_Some_length in h. assumption. @@ -844,13 +844,13 @@ Section Alpha. - intros mfix n decl types hguard hnth hwf ihmfix ihmfixb wffix Δ v e e'; invs e. eapply All2_nth_error_Some in hnth as hnth' ; eauto. destruct hnth' as [decl' [hnth' hh]]. - destruct hh as [[[ety ebo] era] eqann]. + destruct hh as (eqann & era & ety & ebo). assert (hwf' : wf_local Σ (Γ ,,, fix_context mfix')). { apply PCUICWeakeningTyp.All_mfix_wf; auto. eapply (All2_All_mix_left ihmfix) in X. clear -X. induction X; constructor; simpl; auto. - destruct r as [Hty [[[eqty eqbod] eqrarg] eqann]]. + destruct r as [Hty (eqann & eqrarg & eqty & eqbod)]. rewrite /on_def_type -eqann. apply infer_typing_sort_impl with id Hty => //; intros [Hs IH]. apply IH; eauto. reflexivity. } @@ -867,7 +867,7 @@ Section Alpha. eapply (All2_All_mix_left ihmfix) in X. clear -X. induction X; try constructor; simpl; intros n; auto. - destruct r as [Hty [[[eqty eqbod] eqrarg] eqann]]. + destruct r as [Hty (eqann & eqrarg & eqty & eqbod)]. eapply eq_context_upto_cat. + constructor; constructor; auto. eapply eq_term_upto_univ_empty_impl; eauto. @@ -891,19 +891,19 @@ Section Alpha. * eapply (All2_All_mix_left ihmfix) in X. clear -X. induction X; constructor; simpl; auto. - destruct r as [Hty [[[eqty eqbod] eqrarg] eqann]]. + destruct r as [Hty (eqann & eqrarg & eqty & eqbod)]. rewrite /on_def_type -eqann. apply infer_typing_sort_impl with id Hty => //; intros [Hs IH]. apply IH; eauto. reflexivity. - * solve_all. + * unfold eq_mfix in X; solve_all. destruct b0 as [Hb IHb]. - specialize (IHb (Γ ,,, types) _ b2). + specialize (IHb (Γ ,,, types) _ b1). forward IHb. { apply eq_context_upto_cat; reflexivity. } eapply context_conversion; eauto. eapply (type_Cumul' (lift0 #|fix_context mfix| (dtype x))); auto. eapply isType_of_isTypeRel. apply infer_typing_sort_impl with id a1 => //; intros [Hs IH]. - specialize (IH _ _ a0 e'). + specialize (IH _ _ a3 e'). rewrite <-H. eapply (weakening _ _ _ _ (tSort _)); eauto. eapply eq_context_conversion; tea. now symmetry. @@ -917,12 +917,12 @@ Section Alpha. unfold wf_fixpoint, wf_fixpoint_gen. move/andP => [] hm ho. apply/andP; split. - { solve_all. move: b b2. + { unfold eq_mfix in X; solve_all. move: b b1. generalize (dbody x) (dbody y). clear=> t t' isL eq. destruct t => //. now depelim eq. } move: ho; enough (map check_one_fix mfix = map check_one_fix mfix') as ->; auto. - apply upto_names_check_fix. solve_all. + apply upto_names_check_fix. unfold eq_mfix in X. solve_all. + eapply All_nth_error in ihmfix; tea. eapply isType_of_isTypeRel. now apply infer_typing_sort_impl with id ihmfix => //; intros []. @@ -932,13 +932,13 @@ Section Alpha. - intros mfix n decl types hguard hnth hwf ihmfix ihmfixb wffix Δ v e e'; invs e. eapply All2_nth_error_Some in hnth as hnth' ; eauto. destruct hnth' as [decl' [hnth' hh]]. - destruct hh as [[[ety ebo] era] eqann]. + destruct hh as (eqann & era & ety & ebo). assert (hwf' : wf_local Σ (Γ ,,, fix_context mfix')). { apply PCUICWeakeningTyp.All_mfix_wf; auto. eapply (All2_All_mix_left ihmfix) in X. clear -X. induction X; constructor; simpl; auto. - destruct r as [Hty [[[eqty eqbod] eqrarg] eqann]]. + destruct r as [Hty (eqann & eqrarg & eqty & eqbod)]. rewrite /on_def_type -eqann. apply infer_typing_sort_impl with id Hty => //; intros [Hs IH]. apply IH; eauto. reflexivity. } @@ -955,7 +955,7 @@ Section Alpha. eapply (All2_All_mix_left ihmfix) in X. clear -X. induction X; try constructor; simpl; intros n; auto. - destruct r as [Hty [[[eqty eqbod] eqrarg] eqann]]. + destruct r as [Hty (eqann & eqrarg & eqty & eqbod)]. eapply eq_context_upto_cat. + constructor; constructor; auto. eapply eq_term_upto_univ_empty_impl; eauto. @@ -979,19 +979,19 @@ Section Alpha. * eapply (All2_All_mix_left ihmfix) in X. clear -X. induction X; constructor; simpl; auto. - destruct r as [Hty [[[eqty eqbod] eqrarg] eqann]]. + destruct r as [Hty (eqann & eqrarg & eqty & eqbod)]. rewrite /on_def_type -eqann. apply infer_typing_sort_impl with id Hty => //; intros [Hs IH]. apply IH; eauto. reflexivity. - * solve_all. + * unfold eq_mfix in X; solve_all. destruct b0 as [Hb IHb]. - specialize (IHb (Γ ,,, types) _ b2). + specialize (IHb (Γ ,,, types) _ b1). forward IHb. { apply eq_context_upto_cat; reflexivity. } eapply context_conversion; eauto. eapply (type_Cumul' (lift0 #|fix_context mfix| (dtype x))); auto. eapply isType_of_isTypeRel. apply infer_typing_sort_impl with id a1 => //; intros [Hs IH]. - specialize (IH _ _ a0 e'). + specialize (IH _ _ a3 e'). rewrite <-H. eapply (weakening _ _ _ _ (tSort _)); eauto. eapply eq_context_conversion; tea. now symmetry. @@ -1003,7 +1003,7 @@ Section Alpha. * revert wffix. unfold wf_cofixpoint, wf_cofixpoint_gen. enough (map check_one_cofix mfix = map check_one_cofix mfix') as ->; auto. - apply upto_names_check_cofix. solve_all. + apply upto_names_check_cofix. unfold eq_mfix in X; solve_all. + eapply All_nth_error in ihmfix; tea. eapply isType_of_isTypeRel. now apply infer_typing_sort_impl with id ihmfix => //; intros []. diff --git a/pcuic/theories/PCUICArities.v b/pcuic/theories/PCUICArities.v index 00aa92213..d5e2340f5 100644 --- a/pcuic/theories/PCUICArities.v +++ b/pcuic/theories/PCUICArities.v @@ -477,7 +477,7 @@ Section WfEnv. apply (type_mkProd_or_LetIn {| decl_body := None |}) => /=; eauto. Qed. - Lemma app_context_push Γ Δ Δ' d : (Γ ,,, Δ ,,, Δ') ,, d = (Γ ,,, Δ ,,, (Δ' ,, d)). + Lemma app_context_push {A} Γ Δ Δ' (d: A) : (Γ ,,, Δ ,,, Δ') ,, d = (Γ ,,, Δ ,,, (Δ' ,, d)). Proof using Type. reflexivity. Qed. diff --git a/pcuic/theories/PCUICAst.v b/pcuic/theories/PCUICAst.v index 61996e02b..7e999a7c5 100644 --- a/pcuic/theories/PCUICAst.v +++ b/pcuic/theories/PCUICAst.v @@ -42,6 +42,27 @@ Derive NoConfusion for predicate. Arguments predicate : clear implicits. Arguments mk_predicate {_}. +Section test_predicate. + Context {term: Type}. + + Definition test_predicate (instp : Instance.t -> bool) (p : term -> bool) + (pred : predicate term) := + instp pred.(puinst) && forallb p pred.(pparams) && + test_context p pred.(pcontext) && p pred.(preturn). + + Definition test_predicate_k (instp : Instance.t -> bool) + (p : nat -> term -> bool) k (pred : predicate term) := + instp pred.(puinst) && forallb (p k) pred.(pparams) && + test_context_k p #|pred.(pparams)| pred.(pcontext) && + p (#|pred.(pcontext)| + k) pred.(preturn). + + Definition test_predicate_ku (instp : nat -> Instance.t -> bool) + (p : nat -> term -> bool) k (pred : predicate term) := + instp k pred.(puinst) && forallb (p k) pred.(pparams) && + test_context (p #|pred.(puinst)|) pred.(pcontext) && + p k pred.(preturn). +End test_predicate. + Section map_predicate. Context {term term' : Type}. Context (uf : Instance.t -> Instance.t). @@ -100,23 +121,6 @@ Section map_predicate_k. Lemma map_k_puinst k (p : predicate term) : uf (puinst p) = puinst (map_predicate_k k p). Proof using Type. reflexivity. Qed. - - Definition test_predicate (instp : Instance.t -> bool) (p : term -> bool) - (pred : predicate term) := - instp pred.(puinst) && forallb p pred.(pparams) && - test_context p pred.(pcontext) && p pred.(preturn). - - Definition test_predicate_k (instp : Instance.t -> bool) - (p : nat -> term -> bool) k (pred : predicate term) := - instp pred.(puinst) && forallb (p k) pred.(pparams) && - test_context_k p #|pred.(pparams)| pred.(pcontext) && - p (#|pred.(pcontext)| + k) pred.(preturn). - - Definition test_predicate_ku (instp : nat -> Instance.t -> bool) - (p : nat -> term -> bool) k (pred : predicate term) := - instp k pred.(puinst) && forallb (p k) pred.(pparams) && - test_context (p #|pred.(puinst)|) pred.(pcontext) && - p k pred.(preturn). End map_predicate_k. diff --git a/pcuic/theories/PCUICConfluence.v b/pcuic/theories/PCUICConfluence.v index d8796b8d7..9446647bc 100644 --- a/pcuic/theories/PCUICConfluence.v +++ b/pcuic/theories/PCUICConfluence.v @@ -348,7 +348,7 @@ Proof. eexists; split. eapply case_red_brs; tea. econstructor; eauto; try reflexivity. - eapply OnOne2_All2; tea => /=; intuition eauto; try reflexivity. + eapply OnOne2_All2; tea => /=; unfold eq_branch in *; intuition eauto; try reflexivity. now rewrite b. - destruct (IHh _ e) as [x [redl redr]]. @@ -388,11 +388,7 @@ Proof. (d'.(dname), d'.(dbody), d'.(rarg)) ) mfix0 mfix' * - All2 (fun x y => - eq_term_upto_univ Σ' Re Re (dtype x) (dtype y) * - eq_term_upto_univ Σ' Re Re (dbody x) (dbody y) * - (rarg x = rarg y) * - (eq_binder_annot (dname x) (dname y)))%type mfix1 mfix'). + eq_mfix (eq_term_upto_univ Σ' Re Re) mfix1 mfix'). { induction X. - destruct p as [[p1 p2] p3]. eapply p2 in e as hh. destruct hh as [? [? ?]]. @@ -423,11 +419,7 @@ Proof. (d.(dname), d.(dtype), d.(rarg)) = (d'.(dname), d'.(dtype), d'.(rarg)) ) mfix0 mfix' * - All2 (fun x y => - eq_term_upto_univ Σ' Re Re (dtype x) (dtype y) * - eq_term_upto_univ Σ' Re Re (dbody x) (dbody y) * - (rarg x = rarg y) * - eq_binder_annot (dname x) (dname y))%type mfix1 mfix'). + eq_mfix (eq_term_upto_univ Σ' Re Re) mfix1 mfix'). { (* Maybe we should use a lemma using firstn or skipn to keep fix_context intact. Anything general? *) @@ -459,12 +451,7 @@ Proof. (fun d d' : def term => red1 Σ (Δ ,,, fix_context L) (dbody d) (dbody d') × (dname d, dtype d, rarg d) = (dname d', dtype d', rarg d')) mfix0 mfix' - × All2 - (fun x y : def term => - ((eq_term_upto_univ Σ' Re Re (dtype x) (dtype y) - × eq_term_upto_univ Σ' Re Re (dbody x) (dbody y)) × - (rarg x = rarg y)) * - eq_binder_annot (dname x) (dname y)) mfix1 mfix') _ _). + × eq_mfix (eq_term_upto_univ Σ' Re Re) mfix1 mfix') _ _). - intros L x y l [[p1 p2] p3]. assert ( e' : eq_context_upto Σ' Re Rle (Γ ,,, fix_context L) (Δ ,,, fix_context L) @@ -498,11 +485,7 @@ Proof. (d.(dname), d.(dbody), d.(rarg)) = (d'.(dname), d'.(dbody), d'.(rarg)) ) mfix0 mfix' * - All2 (fun x y => - eq_term_upto_univ Σ' Re Re (dtype x) (dtype y) * - eq_term_upto_univ Σ' Re Re (dbody x) (dbody y) * - (rarg x = rarg y) * - eq_binder_annot (dname x) (dname y))%type mfix1 mfix' + eq_mfix (eq_term_upto_univ Σ' Re Re) mfix1 mfix' ). { induction X. - destruct p as [[p1 p2] p3]. @@ -534,11 +517,7 @@ Proof. (d.(dname), d.(dtype), d.(rarg)) = (d'.(dname), d'.(dtype), d'.(rarg)) ) mfix0 mfix' * - All2 (fun x y => - eq_term_upto_univ Σ' Re Re (dtype x) (dtype y) * - eq_term_upto_univ Σ' Re Re (dbody x) (dbody y) * - (rarg x = rarg y) * - eq_binder_annot (dname x) (dname y))%type mfix1 mfix'). + eq_mfix (eq_term_upto_univ Σ' Re Re) mfix1 mfix'). { (* Maybe we should use a lemma using firstn or skipn to keep fix_context intact. Anything general? *) @@ -569,12 +548,7 @@ Proof. (fun d d' : def term => red1 Σ (Δ ,,, fix_context L) (dbody d) (dbody d') × (dname d, dtype d, rarg d) = (dname d', dtype d', rarg d')) mfix0 mfix' - × All2 - (fun x y : def term => - ((eq_term_upto_univ Σ' Re Re (dtype x) (dtype y) - × eq_term_upto_univ Σ' Re Re (dbody x) (dbody y)) × - rarg x = rarg y) * - eq_binder_annot (dname x) (dname y)) mfix1 mfix')) _ _). + × eq_mfix (eq_term_upto_univ Σ' Re Re) mfix1 mfix')) _ _). - intros L x y l [[p1 p2] p3]. assert ( e' : eq_context_upto Σ' Re Rle (Γ ,,, fix_context L) (Δ ,,, fix_context L) @@ -907,7 +881,7 @@ Proof. - dependent destruction e. apply eq_term_upto_univ_mkApps_l_inv in e0 as [? [? [[h1 h2] h3]]]. subst. dependent destruction h1. - eapply All2_nth_error_Some in a as [t' [hnth [eqctx eqbod]]]; tea. + eapply All2_nth_error_Some in e1 as [t' [hnth [eqctx eqbod]]]; tea. have lenctxass := eq_context_gen_context_assumptions eqctx. have lenctx := All2_fold_length eqctx. eexists. split. @@ -934,10 +908,10 @@ Proof. dependent destruction h1. unfold unfold_fix in H. case_eq (nth_error mfix idx) ; - try (intros e ; rewrite e in H ; discriminate H). - intros d e. rewrite e in H. inversion H. subst. clear H. - eapply All2_nth_error_Some in a as hh ; try eassumption. - destruct hh as [d' [e' [[[? ?] erarg] eann]]]. + try (intros eq ; rewrite eq in H ; discriminate H). + intros d eq. rewrite eq in H. inversion H. subst. clear H. + eapply All2_nth_error_Some in e as hh ; try eassumption. + destruct hh as (d' & e' & eann & erarg & ? & ?). unfold is_constructor in H0. case_eq (nth_error args (rarg d)) ; try (intros bot ; rewrite bot in H0 ; discriminate H0). @@ -954,7 +928,7 @@ Proof. * eapply eq_term_upto_univ_substs ; eauto. -- eapply (eq_term_upto_univ_leq _ _ _ 0) ; eauto with arith. -- unfold fix_subst. - apply All2_length in a as el. rewrite <- el. + apply All2_length in e as el. rewrite <- el. generalize #|mfix|. intro n. induction n. ++ constructor. @@ -966,8 +940,8 @@ Proof. dependent destruction h1. unfold unfold_cofix in H. destruct (nth_error mfix idx) eqn:hnth; noconf H. - eapply All2_nth_error_Some in a0 as hh ; tea. - destruct hh as [d' [e' [[[? ?] erarg] eann]]]. + eapply All2_nth_error_Some in e0 as hh ; tea. + destruct hh as (d' & e' & eann & erarg & ? & ?). eexists. split. + eapply red_cofix_case. unfold unfold_cofix. rewrite e'. reflexivity. @@ -977,7 +951,7 @@ Proof. eapply (eq_term_upto_univ_leq _ _ _ 0); auto with arith. typeclasses eauto. unfold cofix_subst. - apply All2_length in a0 as el. rewrite <- el. + apply All2_length in e0 as el. rewrite <- el. generalize #|mfix|. intro n. induction n. * constructor. @@ -988,10 +962,10 @@ Proof. dependent destruction h1. unfold unfold_cofix in H. case_eq (nth_error mfix idx) ; - try (intros e ; rewrite e in H ; discriminate H). - intros d e. rewrite e in H. inversion H. subst. clear H. + try (intros eq ; rewrite eq in H ; discriminate H). + intros d eq. rewrite eq in H. inversion H. subst. clear H. eapply All2_nth_error_Some in e as hh ; try eassumption. - destruct hh as [d' [e' [[[? ?] erarg] eann]]]. + destruct hh as (d' & e' & eann & erarg & ? & ?). eexists. split. + eapply red_cofix_proj. unfold unfold_cofix. rewrite e'. reflexivity. @@ -1001,7 +975,7 @@ Proof. eapply (eq_term_upto_univ_leq _ _ _ 0); auto with arith. typeclasses eauto. unfold cofix_subst. - apply All2_length in a as el. rewrite <- el. + apply All2_length in e as el. rewrite <- el. generalize #|mfix|. intro n. induction n. * constructor. @@ -1056,16 +1030,16 @@ Proof. All2 (eq_term_upto_univ Σ' Re Re) params' args ). { destruct p, p' as []; cbn in *. - induction X in a0, pparams, pparams0, X |- *. - - depelim a0. + induction X in a, pparams, pparams0, X |- *. + - depelim a. eapply p in e as hh ; eauto. destruct hh as [? [? ?]]. eexists. split. + constructor; tea. + constructor. all: eauto. + tc. - - depelim a0. - destruct (IHX _ a0) as [? [? ?]]. + - depelim a. + destruct (IHX _ a) as [? [? ?]]. eexists. split. + eapply OnOne2_tl. eassumption. + constructor. all: eauto. @@ -1101,14 +1075,11 @@ Proof. assert (h : ∑ brs0, OnOne2 (fun br br' => on_Trel_eq (red1 Σ (Γ ,,, inst_case_branch_context p' br)) bbody bcontext br br') brs' brs0 * - All2 (fun x y => - eq_context_gen eq eq (bcontext x) (bcontext y) * - (eq_term_upto_univ Σ' Re Re (bbody x) (bbody y)) - )%type brs'0 brs0 + eq_branches (eq_term_upto_univ Σ' Re Re) brs'0 brs0 ). - { induction X in a, brs' |- *. + { induction X in e1, brs' |- *. - destruct p0 as [p2 p3]. - dependent destruction a. destruct p0 as [h1 h2]. + dependent destruction e1. destruct e1 as [h1 h2]. eapply p2 in h2 as hh ; eauto. destruct hh as [? [? ?]]. eapply (red1_eq_context_upto_l (Re:=Re) (Rle:=Rle) (Δ := Γ ,,, inst_case_branch_context p' y)) in r; cycle -1. @@ -1125,8 +1096,8 @@ Proof. + constructor. all: eauto. split ; eauto. cbn. transitivity (bcontext hd) ; eauto. now rewrite p3. simpl. now transitivity x. - - dependent destruction a. - destruct (IHX _ a) as [? [? ?]]. + - dependent destruction e1. + destruct (IHX _ e2) as [? [? ?]]. eexists. split. + eapply OnOne2_tl. eassumption. + constructor. all: eauto. @@ -1167,17 +1138,13 @@ Proof. (d0.(dname), d0.(dbody), d0.(rarg)) = (d1.(dname), d1.(dbody), d1.(rarg)) ) mfix' mfix * - All2 (fun x y => - eq_term_upto_univ Σ' Re Re x.(dtype) y.(dtype) * - eq_term_upto_univ Σ' Re Re x.(dbody) y.(dbody) * - (x.(rarg) = y.(rarg)) * - eq_binder_annot (dname x) (dname y))%type mfix1 mfix + eq_mfix (eq_term_upto_univ Σ' Re Re) mfix1 mfix ). - { induction X in a, mfix' |- *. + { induction X in e, mfix' |- *. - destruct p as [[p1 p2] p3]. - dependent destruction a. - destruct p as [[[h1 h2] h3] h4]. - eapply p2 in h1 as hh ; eauto. + dependent destruction e. + destruct p as (hann & hrarg & hty & hbod). + eapply p2 in hty as hh ; eauto. destruct hh as [? [? ?]]. eexists. split. + constructor. @@ -1187,8 +1154,8 @@ Proof. simpl. inversion p3. repeat split ; eauto. + tc. - - dependent destruction a. destruct p as [[[h1 h2] h3] h4]. - destruct (IHX _ a) as [? [? ?]]. + - dependent destruction e. destruct p as (hann & hrarg & hty & hbod). + destruct (IHX _ e) as [? [? ?]]. eexists. split. + eapply OnOne2_tl. eassumption. + constructor. all: eauto. @@ -1203,13 +1170,9 @@ Proof. red1 Σ (Γ ,,, fix_context mfix0) x.(dbody) y.(dbody) × (dname x, dtype x, rarg x) = (dname y, dtype y, rarg y) ) mfix' mfix * - All2 (fun x y => - eq_term_upto_univ Σ' Re Re x.(dtype) y.(dtype) * - eq_term_upto_univ Σ' Re Re x.(dbody) y.(dbody) * - (x.(rarg) = y.(rarg)) * - eq_binder_annot (dname x) (dname y)) mfix1 mfix + eq_mfix (eq_term_upto_univ Σ' Re Re) mfix1 mfix ). - { revert mfix' a. + { revert mfix' e. refine (OnOne2_ind_l _ (fun L x y => (red1 Σ (Γ ,,, fix_context L) (dbody x) (dbody y) × (forall Rle napp (u' : term), RelationClasses.Reflexive Re -> @@ -1223,34 +1186,26 @@ Proof. ∑ v' : term, red1 Σ (Γ ,,, fix_context L) u' v' × eq_term_upto_univ_napp Σ' Re Rle napp (dbody y) v' )) - × (dname x, dtype x, rarg x) = (dname y, dtype y, rarg y)) (fun L mfix0 mfix1 o => forall mfix', All2 - (fun x y : def term => - ((eq_term_upto_univ Σ' Re Re (dtype x) (dtype y) - × eq_term_upto_univ Σ' Re Re (dbody x) (dbody y)) × - rarg x = rarg y) * eq_binder_annot (dname x) (dname y)) mfix0 mfix' -> ∑ mfix : list (def term), + × (dname x, dtype x, rarg x) = (dname y, dtype y, rarg y)) (fun L mfix0 mfix1 o => forall mfix', + eq_mfix (eq_term_upto_univ Σ' Re Re) mfix0 mfix' -> ∑ mfix : list (def term), ( OnOne2 (fun x y : def term => red1 Σ (Γ ,,, fix_context L) (dbody x) (dbody y) × (dname x, dtype x, rarg x) = (dname y, dtype y, rarg y)) mfix' mfix ) * - ( All2 - (fun x y : def term => - ((eq_term_upto_univ Σ' Re Re (dtype x) (dtype y) - × eq_term_upto_univ Σ' Re Re (dbody x) (dbody y)) × - (rarg x = rarg y)) * eq_binder_annot (dname x) (dname y)) mfix1 mfix )) _ _ _ _ X). + ( eq_mfix (eq_term_upto_univ Σ' Re Re) mfix1 mfix )) _ _ _ _ X). - clear X. intros L x y l [[p1 p2] p3] mfix' h. - dependent destruction h. destruct p as [[[h1 h2] h3] h4]. - eapply p2 in h2 as hh ; eauto. + dependent destruction h. destruct p as (hann & hrarg & hty & hbod). + eapply p2 in hbod as hh ; eauto. destruct hh as [? [? ?]]. eexists. split. + constructor. constructor. instantiate (1 := mkdef _ _ _ _ _). simpl. eauto. reflexivity. - + constructor. constructor; simpl. all: eauto. + + constructor; auto. inversion p3. - simpl. repeat split ; eauto. destruct y0. simpl in *. - congruence. + simpl. repeat split ; eauto. - clear X. intros L x l l' h ih mfix' ha. - dependent destruction ha. destruct p as [[h1 h2] h3]. + dependent destruction ha. destruct p as (hann & hrarg & hty & hbod). destruct (ih _ ha) as [? [? ?]]. eexists. split. + eapply OnOne2_tl. eauto. @@ -1262,12 +1217,7 @@ Proof. red1 Σ (Γ ,,, fix_context mfix') x.(dbody) y.(dbody) × (dname x, dtype x, rarg x) = (dname y, dtype y, rarg y) ) mfix' mfix × - All2 (fun x y => - eq_term_upto_univ Σ' Re Re x.(dtype) y.(dtype) * - eq_term_upto_univ Σ' Re Re x.(dbody) y.(dbody) * - (x.(rarg) = y.(rarg)) * - eq_binder_annot (dname x) (dname y) - ) mfix1 mfix %type + eq_mfix (eq_term_upto_univ Σ' Re Re) mfix1 mfix ). { clear X. assert (hc : eq_context_upto Σ' @@ -1277,9 +1227,9 @@ Proof. ). { eapply eq_context_upto_cat. - eapply eq_context_upto_refl; assumption. - - clear -he hle tRe tRle hR a. induction a. + - clear -he hle tRe tRle hR e. induction e. + constructor. - + destruct r as [[[? ?] ?] ?]. + + destruct r as (? & ? & ? & ?). eapply All2_eq_context_upto. eapply All2_rev. eapply All2_mapi. @@ -1291,17 +1241,17 @@ Proof. eapply eq_term_upto_univ_impl; eauto. all:typeclasses eauto. * eapply All2_impl ; eauto. - intros ? ? [[[? ?] ?] ?] i. split; [split|]. + intros ? ? (? & ? & ? & ?) i. split; [split|]. -- assumption. -- cbn. constructor. -- cbn. eapply eq_term_upto_univ_lift. eapply eq_term_upto_univ_impl; eauto. typeclasses eauto. } - clear a. + clear e. eapply OnOne2_impl_exist_and_All ; try eassumption. - clear o a0. - intros x x' y [r e] [[[? ?] ?] ?]. + clear o e0. + intros x x' y [r e] (? & ? & ? & ?). inversion e. clear e. eapply red1_eq_context_upto_l in r as [? [? ?]]. 7: eassumption. all: tea. @@ -1309,10 +1259,10 @@ Proof. instantiate (1 := mkdef _ _ _ _ _). simpl. intuition eauto. intuition eauto. + - now simpl. + - etransitivity ; eauto. - rewrite H1. eauto. - eapply eq_term_upto_univ_trans; tea. - - etransitivity ; eauto. - - now simpl. } destruct h as [? [? ?]]. eexists. split. @@ -1325,17 +1275,13 @@ Proof. (d0.(dname), d0.(dbody), d0.(rarg)) = (d1.(dname), d1.(dbody), d1.(rarg)) ) mfix' mfix * - All2 (fun x y => - eq_term_upto_univ Σ' Re Re x.(dtype) y.(dtype) * - eq_term_upto_univ Σ' Re Re x.(dbody) y.(dbody) * - (x.(rarg) = y.(rarg)) * - eq_binder_annot (dname x) (dname y))%type mfix1 mfix + eq_mfix (eq_term_upto_univ Σ' Re Re) mfix1 mfix ). - { induction X in a, mfix' |- *. + { induction X in e, mfix' |- *. - destruct p as [[p1 p2] p3]. - dependent destruction a. - destruct p as [[[h1 h2] h3] h4]. - eapply p2 in h1 as hh ; eauto. + dependent destruction e. + destruct p as (hann & hrarg & hty & hbod). + eapply p2 in hty as hh ; eauto. destruct hh as [? [? ?]]. eexists. split. + constructor. @@ -1345,8 +1291,8 @@ Proof. simpl. inversion p3. repeat split ; eauto. + tc. - - dependent destruction a. destruct p as [[h1 h2] h3]. - destruct (IHX _ a) as [? [? ?]]. + - dependent destruction e. destruct p as (hann & hrarg & hty & hbod). + destruct (IHX _ e) as [? [? ?]]. eexists. split. + eapply OnOne2_tl. eassumption. + constructor. all: eauto. @@ -1361,14 +1307,9 @@ Proof. red1 Σ (Γ ,,, fix_context mfix0) x.(dbody) y.(dbody) × (dname x, dtype x, rarg x) = (dname y, dtype y, rarg y) ) mfix' mfix * - All2 (fun x y => - eq_term_upto_univ Σ' Re Re x.(dtype) y.(dtype) * - eq_term_upto_univ Σ' Re Re x.(dbody) y.(dbody) * - (x.(rarg) = y.(rarg)) * - eq_binder_annot (dname x) (dname y) - ) mfix1 mfix + eq_mfix (eq_term_upto_univ Σ' Re Re) mfix1 mfix ). - { revert mfix' a. + { revert mfix' e. refine (OnOne2_ind_l _ (fun L x y => (red1 Σ (Γ ,,, fix_context L) (dbody x) (dbody y) × (forall Rle napp (u' : term), RelationClasses.Reflexive Re -> @@ -1382,23 +1323,16 @@ Proof. ∑ v' : term, red1 Σ (Γ ,,, fix_context L) u' v' × eq_term_upto_univ_napp Σ' Re Rle napp (dbody y) v')) - × (dname x, dtype x, rarg x) = (dname y, dtype y, rarg y)) (fun L mfix0 mfix1 o => forall mfix', All2 - (fun x y : def term => - ((eq_term_upto_univ Σ' Re Re (dtype x) (dtype y) - × eq_term_upto_univ Σ' Re Re (dbody x) (dbody y)) × - rarg x = rarg y) * eq_binder_annot (dname x) (dname y)) mfix0 mfix' -> ∑ mfix : list (def term), + × (dname x, dtype x, rarg x) = (dname y, dtype y, rarg y)) (fun L mfix0 mfix1 o => forall mfix', + eq_mfix (eq_term_upto_univ Σ' Re Re) mfix0 mfix' -> ∑ mfix : list (def term), ( OnOne2 (fun x y : def term => red1 Σ (Γ ,,, fix_context L) (dbody x) (dbody y) × (dname x, dtype x, rarg x) = (dname y, dtype y, rarg y)) mfix' mfix ) * - ( All2 - (fun x y : def term => - ((eq_term_upto_univ Σ' Re Re (dtype x) (dtype y) - × eq_term_upto_univ Σ' Re Re (dbody x) (dbody y)) × - rarg x = rarg y) * eq_binder_annot (dname x) (dname y)) mfix1 mfix )) _ _ _ _ X). + ( eq_mfix (eq_term_upto_univ Σ' Re Re) mfix1 mfix )) _ _ _ _ X). - clear X. intros L x y l [[p1 p2] p3] mfix' h. - dependent destruction h. destruct p as [[[h1 h2] h3] h4]. - eapply p2 in h2 as hh ; eauto. + dependent destruction h. destruct p as (hann & hrarg & hty & hbod). + eapply p2 in hbod as hh ; eauto. destruct hh as [? [? ?]]. noconf p3. hnf in H. noconf H. eexists. split; simpl. @@ -1408,7 +1342,7 @@ Proof. + constructor. all: eauto. simpl. repeat split ; eauto; congruence. - clear X. intros L x l l' h ih mfix' ha. - dependent destruction ha. destruct p as [[h1 h2] h3]. + dependent destruction ha. destruct p as (hann & hrarg & hty & hbod). destruct (ih _ ha) as [? [? ?]]. eexists. split. + eapply OnOne2_tl. eauto. @@ -1420,12 +1354,7 @@ Proof. red1 Σ (Γ ,,, fix_context mfix') x.(dbody) y.(dbody) × (dname x, dtype x, rarg x) = (dname y, dtype y, rarg y) ) mfix' mfix × - All2 (fun x y => - eq_term_upto_univ Σ' Re Re x.(dtype) y.(dtype) * - eq_term_upto_univ Σ' Re Re x.(dbody) y.(dbody) * - (x.(rarg) = y.(rarg)) * - eq_binder_annot (dname x) (dname y) - ) mfix1 mfix + eq_mfix (eq_term_upto_univ Σ' Re Re) mfix1 mfix ). { clear X. assert (hc : eq_context_upto Σ' @@ -1435,9 +1364,9 @@ Proof. ). { eapply eq_context_upto_cat. - eapply eq_context_upto_refl; assumption. - - clear -he he' hle hle' hR a. induction a. + - clear -he he' hle hle' hR e. induction e. + constructor. - + destruct r as [[[? ?] ?] ?]. + + destruct r as (? & ? & ? & ?). eapply All2_eq_context_upto. eapply All2_rev. eapply All2_mapi. @@ -1449,27 +1378,27 @@ Proof. eapply eq_term_upto_univ_impl; eauto. all:typeclasses eauto. * eapply All2_impl ; eauto. - intros ? ? [[[? ?] ?] ?] i. split; [split|]. + intros ? ? (? & ? & ? & ?) i. split; [split|]. -- assumption. -- cbn. constructor. -- cbn. eapply eq_term_upto_univ_lift. eapply eq_term_upto_univ_impl; eauto. all:typeclasses eauto. } - clear a. + clear e. eapply OnOne2_impl_exist_and_All ; try eassumption. - clear o a0. - intros x x' y [r e] [[? ?] ?]. + clear o e0. + intros x x' y [r e] (? & ? & ? & ?). inversion e. clear e. eapply red1_eq_context_upto_l in r as [? [? ?]]. 7: eassumption. all: tea. eexists. instantiate (1 := mkdef _ _ _ _ _). simpl. intuition eauto. + - congruence. + - etransitivity ; eauto. - rewrite H1. eauto. - eapply eq_term_upto_univ_trans ; tea. - - etransitivity ; eauto. - - congruence. } destruct h as [? [? ?]]. eexists. split. @@ -2999,7 +2928,7 @@ Section RedConfluence. Proof using Type. move/on_free_vars_ctx_All_fold => a eqctx. apply on_free_vars_ctx_All_fold. - eapply eq_context_upto_names_gen in eqctx. + eapply eq_context_upto_names_upto_names_gen in eqctx. eapply All2_fold_All_fold_mix_left in eqctx; tea. cbn in eqctx. induction eqctx. - constructor; auto. @@ -3209,7 +3138,7 @@ Section RedConfluence. exists body'; split => //. rewrite -lift0_inst. econstructor; eauto. destruct (nth_error Δ x) eqn:hnth' => //. - eapply eq_context_upto_names_gen in eqctx'. + eapply eq_context_upto_names_upto_names_gen in eqctx'. eapply All2_fold_nth in eqctx' as [d' [hnth'' [eqctx'' eqd]]]; tea. depelim eqd. subst. noconf eq. subst. noconf eq. rewrite hnth'' //. @@ -3550,7 +3479,7 @@ Section RedConfluence. Proof using Type. move=> Hctx. eapply context_pres_let_bodies_red1. - eapply eq_context_upto_names_gen in Hctx. + eapply eq_context_upto_names_upto_names_gen in Hctx. eapply All2_fold_impl; tea => /= _ _ ? ? [] /=; rewrite /pres_let_bodies /= //; intros; congruence. Qed. @@ -3779,23 +3708,23 @@ Proof. apply: term_on_free_vars_ind; intros; depelim eqt. all:simpl; auto. all:try solve [solve_all]. - - destruct e as [? [? [? ?]]]. - rewrite -(All2_fold_length a1). - rewrite -(All2_length a0). + - destruct e as [? [? [? ?]]]. unfold eq_branches, eq_branch in e0. + rewrite -(All2_fold_length a0). + rewrite -(All2_length a). solve_all. rewrite test_context_k_closed_on_free_vars_ctx. eapply eq_context_upto_names_on_free_vars; tea. - now eapply eq_context_upto_names_gen in a1. + now eapply eq_context_upto_names_upto_names_gen in a0. rewrite test_context_k_closed_on_free_vars_ctx. - destruct a. + destruct a1. eapply eq_context_upto_names_on_free_vars; tea. - now eapply eq_context_upto_names_gen in a2. - destruct a as [hctx ihctx hb ihb]. + now eapply eq_context_upto_names_upto_names_gen in a2. + destruct a1 as [hctx ihctx hb ihb]. rewrite -(All2_fold_length a2). now eapply ihb. - - rewrite -(All2_length a). solve_all. + - rewrite -(All2_length e). unfold eq_mfix in e. solve_all. apply/andP; split; eauto. len in b2. eapply b2. eauto. - - rewrite -(All2_length a). solve_all. + - rewrite -(All2_length e). unfold eq_mfix in e. solve_all. apply/andP; split; eauto. len in b2. eapply b2. eauto. Qed. diff --git a/pcuic/theories/PCUICConvCumInversion.v b/pcuic/theories/PCUICConvCumInversion.v index eb371d645..345979774 100644 --- a/pcuic/theories/PCUICConvCumInversion.v +++ b/pcuic/theories/PCUICConvCumInversion.v @@ -372,9 +372,9 @@ Section fixed. eapply on_free_vars_ctx_inst_case_context; tea => //. now rewrite test_context_k_closed_on_free_vars_ctx. + eapply red_on_free_vars in r1; tea. - { len. rewrite (All2_fold_length a5). + { len. rewrite (All2_fold_length a4). now setoid_rewrite shiftnP_add in p1. } - len. rewrite -shiftnP_add (All2_fold_length a5). + len. rewrite -shiftnP_add (All2_fold_length a4). eapply on_ctx_free_vars_inst_case_context; auto. 1:now rewrite test_context_k_closed_on_free_vars_ctx. now erewrite -> on_free_vars_ctx_on_ctx_free_vars. } @@ -387,7 +387,7 @@ Section fixed. all:eapply into_closed_red; eauto. - rename a0 into brsa1. rename a2 into brsa2. - rename a3 into brseq. + rename e0 into brseq. clear -wfΣ decli brsa1 brsa2 brseq clΓ wfp wfp' a a1 p0 p5 p4 p9 r3 eqpars. induction brseq in brs, brs', brsa1, brsa2, p4, p9 |- *; depelim brsa1; depelim brsa2; [constructor|]. @@ -455,7 +455,7 @@ Section fixed. eapply whnf_red_inv in r2; eauto. depelim r1. depelim r2. - depelim eq. + depelim eq. unfold eq_mfix in e. split; [easy|]. cbn in cll, clr, clx, cly. have clmfixctx : is_closed_context (Γ ,,, fix_context mfix). @@ -464,10 +464,10 @@ Section fixed. { rewrite on_free_vars_ctx_app clΓ /=. apply on_free_vars_fix_context; solve_all. } solve_all. move: clmfixctx clmfixctx'. - clear -wfΣ a a0 a1 clΓ. + clear -wfΣ a a0 e clΓ. cut (#|mfix| = #|mfix'|); - [|now apply All2_length in a; apply All2_length in a0; apply All2_length in a1]. - revert a a0 a1. + [|now apply All2_length in a; apply All2_length in a0; apply All2_length in e]. + revert a a0 e. generalize mfix at 1 2 4 5 6. generalize mfix' at 1 2 4 5. intros ctx_fix ctx_fix'. @@ -475,7 +475,7 @@ Section fixed. induction all in mfix, mfix', all1, all2, all |- *; depelim all1; depelim all2; subst; [constructor|]. constructor; [|now auto]. - destruct r as ((?&(((? & ?) & ?)&?))&?), p as (?&?&?&?&?), p0 as (?&?&?&?&?). + destruct r as ((? & (? & ? & ? & ?))&?), p as (?&?&?&?&?), p0 as (?&?&?&?&?). split; auto; try congruence. - eapply ws_cumul_pb_alt_closed; exists (dtype x), (dtype y). split; eauto. all:eapply into_closed_red; eauto. @@ -515,7 +515,7 @@ Section fixed. eapply whnf_red_inv in r2; eauto. depelim r1. depelim r2. - depelim eq. + depelim eq. unfold eq_mfix in e. split; [easy|]. cbn in cll, clr, clx, cly. have clmfixctx : is_closed_context (Γ ,,, fix_context mfix). @@ -524,10 +524,10 @@ Section fixed. { rewrite on_free_vars_ctx_app clΓ /=. apply on_free_vars_fix_context; solve_all. } solve_all. move: clmfixctx clmfixctx'. - clear -wfΣ a a0 a1 clΓ. + clear -wfΣ a a0 e clΓ. cut (#|mfix| = #|mfix'|); - [|now apply All2_length in a; apply All2_length in a0; apply All2_length in a1]. - revert a a0 a1. + [|now apply All2_length in a; apply All2_length in a0; apply All2_length in e]. + revert a a0 e. generalize mfix at 1 2 4 5 6. generalize mfix' at 1 2 4 5. intros ctx_fix ctx_fix'. @@ -535,7 +535,7 @@ Section fixed. induction all in mfix, mfix', all1, all2, all |- *; depelim all1; depelim all2; subst; [constructor|]. constructor; [|now auto]. - destruct r as ((?&(((? & ?) & ?)&?))&?), p as (?&?&?&?&?), p0 as (?&?&?&?&?). + destruct r as ((? & (? & ? & ? & ?))&?), p as (?&?&?&?&?), p0 as (?&?&?&?&?). split; auto; try congruence. - eapply ws_cumul_pb_alt_closed; exists (dtype x), (dtype y). split; eauto. all:eapply into_closed_red; eauto. diff --git a/pcuic/theories/PCUICConversion.v b/pcuic/theories/PCUICConversion.v index e286bf69e..a3422eedb 100644 --- a/pcuic/theories/PCUICConversion.v +++ b/pcuic/theories/PCUICConversion.v @@ -172,7 +172,7 @@ Section CumulSpecIsCumulAlgo. + intuition. + try_with_nil. intuition. * try_with_nil. - * unfold tCaseBrsProp in X0. eapply All2_All_mix_left in X0. 2: tea. + * unfold eq_branches, eq_branch in e0. unfold tCaseBrsProp in X0. eapply All2_All_mix_left in X0. 2: tea. eapply All2_impl. 1: eassumption. cbn. intuition. try_with_nil. - apply cumul_mkApps; eauto. eapply cumul_Fix. unfold tFixProp in *. @@ -1990,7 +1990,7 @@ Section ConvRedConv. eapply OnOne2_app. constructor; auto. cbn. split; auto. eapply red1_eq_context_upto_names; tea. rewrite /inst_case_branch_context /=. - now eapply eq_context_upto_names_gen, eq_context_gen_inst_case_context. } + now eapply eq_context_upto_names_upto_names_gen, eq_context_gen_inst_case_context. } rewrite [is_open_term _ _]is_open_case_split onp onc /= //. Qed. @@ -2012,7 +2012,7 @@ Section ConvRedConv. now rewrite shiftnP_add. } rewrite !test_context_k_closed_on_free_vars_ctx in cl *. eapply eq_context_upto_names_on_free_vars; tea. - eapply eq_context_upto_names_gen. + eapply eq_context_upto_names_upto_names_gen. now symmetry. - now move: op => /= /andP[] => ->. Qed. @@ -2037,7 +2037,7 @@ Section ConvRedConv. now eapply eq_context_gen_inst_case_context. } eapply ws_cumul_pb_is_closed_context in cv. eapply eq_context_upto_names_on_free_vars; tea. - eapply eq_context_upto_names_gen. + eapply eq_context_upto_names_upto_names_gen. now eapply eq_context_gen_inst_case_context. } induction h. - apply ws_cumul_pb_compare; tas. @@ -2081,10 +2081,7 @@ Section ConvRedConv. (if b then tFix else tCoFix) mfix idx. Lemma eq_term_fix_or_cofix b mfix idx mfix' : - All2 (fun x y : def term => - ((eq_term_upto_univ Σ.1 (eq_universe Σ) (eq_universe Σ) (dtype x) (dtype y) - × eq_term_upto_univ Σ.1 (eq_universe Σ) (eq_universe Σ) (dbody x) (dbody y)) × rarg x = rarg y) * - eq_binder_annot (dname x) (dname y)) mfix mfix' -> + eq_mfix (eq_term_upto_univ Σ.1 (eq_universe Σ) (eq_universe Σ)) mfix mfix' -> eq_term Σ Σ (fix_or_cofix b mfix idx) (fix_or_cofix b mfix' idx). Proof using Type. destruct b; constructor; auto. @@ -3736,7 +3733,7 @@ Proof. eapply Heqxy.2; eauto. + eapply Hpret; eauto. ++ rewrite test_context_k_closed_on_free_vars_ctx in Hcontext. - unfold inst_case_predicate_context. apply PCUICOnFreeVarsConv.on_free_vars_ctx_inst_case_context ; eauto. + unfold inst_case_predicate_context. apply PCUICRenameTerm.on_free_vars_ctx_inst_case_context_app ; eauto. ++ rewrite shiftnP_add in Hreturn. rewrite <- inst_case_predicate_context_length in Hreturn. rewrite <- app_length in Hreturn. eassumption. ++ rewrite shiftnP_add in Hreturn'. rewrite <- (All2_fold_length Hpcon) in Hreturn'. @@ -3747,7 +3744,7 @@ Proof. apply (All2_All_mix_right Hbrs') in Hbrsbrs'. clear Hbrs'. eapply All2_impl. 1: tea. cbn; intros x y [[Hx Heqxy ] Hy]. split; try apply Heqxy.1. clear Hbrsbrs'. rewrite test_context_k_closed_on_free_vars_ctx in Hx. toProp Hx. rewrite test_context_k_closed_on_free_vars_ctx in Hy. toProp Hy. eapply Heqxy.2; eauto. - + apply PCUICOnFreeVarsConv.on_free_vars_ctx_inst_case_context ; eauto; intuition. + + apply PCUICRenameTerm.on_free_vars_ctx_inst_case_context_app ; eauto; intuition. + rewrite shiftnP_add in Hx. erewrite <- inst_case_branch_context_length in Hx. rewrite <- app_length in Hx. intuition. + rewrite shiftnP_add in Hy. rewrite <- (All2_fold_length Heqxy.1.1) in Hy. erewrite <- inst_case_branch_context_length in Hy. diff --git a/pcuic/theories/PCUICCumulProp.v b/pcuic/theories/PCUICCumulProp.v index b02adeff5..86c939378 100644 --- a/pcuic/theories/PCUICCumulProp.v +++ b/pcuic/theories/PCUICCumulProp.v @@ -848,7 +848,7 @@ Proof using Type. { destruct X as [nf [nf' [r r' e]]]. exists nf, nf'. split; try constructor; auto; fvs. } eexists _, _; split; intuition auto. clear clΓ clT. induction T using PCUICInduction.term_forall_list_ind; cbn; intros; - try solve [constructor; eauto; solve_all]. + try solve [constructor; try unfold eq_mfix; eauto; solve_all]. - cbn. constructor. destruct s; split; reflexivity. - constructor. eapply PCUICEquality.eq_term_upto_univ_impl in IHT1; eauto. @@ -864,7 +864,7 @@ Proof using Type. apply IHT. eapply All2_map. eapply All_All2; tea. cbn. - intuition auto. rewrite /id. reflexivity. + intuition auto. rewrite /eq_branch /id. split => //. reflexivity. Qed. Lemma R_eq_univ_prop_consistent_instances Σ univs u u' : @@ -1251,7 +1251,7 @@ Proof using Hcf Hcf'. eapply All2_symP => //. typeclasses eauto. eapply app_inj in e as [eql ->] => //. move: (All2_length eqpars). - move: (All2_length a0). lia. fvs. now eapply subject_is_open_term in scrut_ty. + move: (All2_length a). lia. fvs. now eapply subject_is_open_term in scrut_ty. now apply subject_is_open_term in X6. - eapply inversion_Proj in X3 as (u' & mdecl' & idecl' & cdecl' & pdecl' & args' & inv); auto. @@ -1295,26 +1295,24 @@ Proof using Hcf Hcf'. - eapply inversion_Fix in X2 as (decl' & fixguard' & Hnth & types' & bodies & wffix & cum); auto. eapply cumul_cumul_prop in cum; eauto. eapply cumul_prop_trans; eauto. - eapply All2_nth_error in a; eauto. - destruct a as [[[a _] _] _]. + eapply All2_nth_error in e as (eqann & eqrarg & eqty & eqbod); eauto. constructor; [fvs|..]. { eapply nth_error_all in X0 as (? & e & dty & ?); tea. now apply subject_is_open_term in dty. } { now eapply cumul_prop_is_open in cum as []. } eapply eq_term_eq_term_prop_impl; eauto. - now symmetry in a. + now symmetry in eqty. - eapply inversion_CoFix in X2 as (decl' & fixguard' & Hnth & types' & bodies & wfcofix & cum); auto. eapply cumul_cumul_prop in cum; eauto. eapply cumul_prop_trans; eauto. - eapply All2_nth_error in a; eauto. - destruct a as [[[a _] _] _]. + eapply All2_nth_error in e as (eqann & eqrarg & eqty & eqbod); eauto. constructor; [fvs|..]. { eapply nth_error_all in X0 as (? & ? & dty & ?); tea. now apply subject_is_open_term in dty. } { now eapply cumul_prop_is_open in cum as []. } eapply eq_term_eq_term_prop_impl; eauto. - now symmetry in a. + now symmetry in eqty. Qed. End no_prop_leq_type. diff --git a/pcuic/theories/PCUICEquality.v b/pcuic/theories/PCUICEquality.v index ff3b500b6..5e6a614f4 100644 --- a/pcuic/theories/PCUICEquality.v +++ b/pcuic/theories/PCUICEquality.v @@ -1,7 +1,7 @@ (* Distributed under the terms of the MIT license. *) From Coq Require Import CMorphisms. From MetaCoq.Template Require Import LibHypsNaming config utils Reflect. -From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICInduction +From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICInduction PCUICCases PCUICLiftSubst PCUICReflect. Require Import ssreflect ssrbool. @@ -128,7 +128,9 @@ Notation eq_context_upto_names := (All2 (compare_decls eq eq)). Notation eq_context_gen eq_term leq_term := (All2_fold (fun _ _ => compare_decls eq_term leq_term)). -Lemma eq_context_upto_names_gen Γ Γ' : eq_context_upto_names Γ Γ' <~> eq_context_gen eq eq Γ Γ'. +Notation eq_context_upto_names_gen := (eq_context_gen eq eq). + +Lemma eq_context_upto_names_upto_names_gen Γ Γ' : eq_context_upto_names Γ Γ' <~> eq_context_upto_names_gen Γ Γ'. Proof. split; intros e; depind e; constructor; auto. Qed. @@ -262,10 +264,25 @@ Proof. Qed. Definition eq_predicate (eq_term : term -> term -> Type) Re p p' := - All2 eq_term p.(pparams) p'.(pparams) * - (R_universe_instance Re p.(puinst) p'.(puinst) * - ((eq_context_gen eq eq p.(pcontext) p'.(pcontext)) * - eq_term p.(preturn) p'.(preturn))). + All2 eq_term p.(pparams) p'.(pparams) × + R_universe_instance Re p.(puinst) p'.(puinst) × + eq_context_upto_names_gen p.(pcontext) p'.(pcontext) × + eq_term p.(preturn) p'.(preturn). + +Definition eq_branch (eq_term : term -> term -> Type) br br' := + eq_context_upto_names_gen br.(bcontext) br'.(bcontext) × + eq_term br.(bbody) br'.(bbody). + +Definition eq_branches (eq_term : term -> term -> Type) brs brs' := + All2 (eq_branch eq_term) brs brs'. + +Definition eq_mfix (eq_term : term -> term -> Type) mfix mfix' := + All2 (fun d d' => + eq_binder_annot d.(dname) d'.(dname) × + d.(rarg) = d'.(rarg) × + eq_term d.(dtype) d'.(dtype) × + eq_term d.(dbody) d'.(dbody) + ) mfix mfix'. (** ** Syntactic ws_cumul_pb up-to universes We don't look at printing annotations *) @@ -332,10 +349,7 @@ Inductive eq_term_upto_univ_napp Σ (Re Rle : Universe.t -> Universe.t -> Prop) | eq_Case : forall indn p p' c c' brs brs', eq_predicate (eq_term_upto_univ_napp Σ Re Re 0) Re p p' -> Σ ⊢ c <==[ Re , 0 ] c' -> - All2 (fun x y => - eq_context_gen eq eq (bcontext x) (bcontext y) * - (Σ ⊢ x.(bbody) <==[ Re , 0 ] y.(bbody)) - ) brs brs' -> + eq_branches (eq_term_upto_univ_napp Σ Re Re 0) brs brs' -> Σ ⊢ tCase indn p c brs <==[ Rle , napp ] tCase indn p' c' brs' | eq_Proj : forall p c c', @@ -343,21 +357,11 @@ Inductive eq_term_upto_univ_napp Σ (Re Rle : Universe.t -> Universe.t -> Prop) Σ ⊢ tProj p c <==[ Rle , napp ] tProj p c' | eq_Fix : forall mfix mfix' idx, - All2 (fun x y => - (Σ ⊢ x.(dtype) <==[ Re , 0 ] y.(dtype)) * - (Σ ⊢ x.(dbody) <==[ Re , 0 ] y.(dbody)) * - (x.(rarg) = y.(rarg)) * - eq_binder_annot x.(dname) y.(dname) - )%type mfix mfix' -> + eq_mfix (eq_term_upto_univ_napp Σ Re Re 0) mfix mfix' -> Σ ⊢ tFix mfix idx <==[ Rle , napp ] tFix mfix' idx | eq_CoFix : forall mfix mfix' idx, - All2 (fun x y => - (Σ ⊢ x.(dtype) <==[ Re , 0 ] y.(dtype)) * - (Σ ⊢ x.(dbody) <==[ Re , 0 ] y.(dbody)) * - (x.(rarg) = y.(rarg)) * - eq_binder_annot x.(dname) y.(dname) - ) mfix mfix' -> + eq_mfix (eq_term_upto_univ_napp Σ Re Re 0) mfix mfix' -> Σ ⊢ tCoFix mfix idx <==[ Rle , napp ] tCoFix mfix' idx (* | eq_Prim i : eq_term_upto_univ_napp Σ Re Rle napp (tPrim i) (tPrim i) *) @@ -499,6 +503,34 @@ Proof. eapply All2_same; reflexivity. Qed. +#[global] +Polymorphic Instance eq_branch_refl Re : + CRelationClasses.Reflexive Re -> + CRelationClasses.Reflexive (eq_branch Re). +Proof. + intros hre. + intros br. unfold eq_branch; intuition auto; reflexivity. +Qed. + +#[global] +Polymorphic Instance eq_branches_refl Re : + CRelationClasses.Reflexive Re -> + CRelationClasses.Reflexive (eq_branches Re). +Proof. + intros hre. + intros brs. apply All2_same; reflexivity. +Qed. + +#[global] +Polymorphic Instance eq_mfix_refl Re : + CRelationClasses.Reflexive Re -> + CRelationClasses.Reflexive (eq_mfix Re). +Proof. + intros hre. + intros mfix. unfold eq_mfix; eapply All2_same. + intuition auto; try reflexivity. +Qed. + #[global] Polymorphic Instance eq_term_upto_univ_refl Σ Re Rle napp : RelationClasses.Reflexive Re -> @@ -511,14 +543,13 @@ Proof. all: try constructor. all: eauto. all: try solve [eapply All_All2 ; eauto]. all: try solve [eapply Forall2_same ; eauto]. - all: try solve [unfold eq_predicate; solve_all; eapply All_All2; eauto]. - apply R_global_instance_refl; auto. - apply R_global_instance_refl; auto. - destruct X as [? [? ?]]. unfold eq_predicate; solve_all. eapply All_All2; eauto. reflexivity. eapply onctx_eq_ctx in a0; eauto. - - eapply All_All2; eauto; simpl; intuition eauto. + - eapply All_All2; unfold eq_branch; eauto; simpl; intuition eauto. eapply onctx_eq_ctx in a; eauto. - eapply All_All2; eauto; simpl; intuition eauto. - eapply All_All2; eauto; simpl; intuition eauto. @@ -579,27 +610,27 @@ Proof. + constructor. + destruct r as [h1 h2]. eapply h1 in h2 ; auto. - solve_all. destruct e as (r & ? & eq & ?). - econstructor; rewrite ?e; unfold eq_predicate in *; solve_all; eauto. + econstructor; rewrite ?e; unfold eq_predicate, eq_branches, eq_branch in *; solve_all; eauto. eapply All2_sym; solve_all. unfold R_universe_instance in r |- *. eapply Forall2_symP; eauto. - eapply onctx_eq_ctx_sym in a1; eauto. + eapply onctx_eq_ctx_sym in eq; eauto. eapply All2_sym; solve_all. - eapply onctx_eq_ctx_sym in a0; eauto. + eapply onctx_eq_ctx_sym in a; eauto. - econstructor. eapply All2_All_mix_left in X as h; eauto. - clear a X. + clear e X. induction h. + constructor. - + destruct r as [[h1 h2] [[[h3 h4] h5] h6]]. - eapply h1 in h3; auto. constructor; auto. + + destruct r as ((h1 & h2) & (han & hrarg & hty & hbod)). + eapply h1 in hty; auto. constructor; auto. - econstructor. eapply All2_All_mix_left in X as h; eauto. - clear a X. + clear e X. induction h. + constructor. - + destruct r as [[h1 h2] [[[h3 h4] h5] h6]]. eapply h1 in h3 ; auto. - constructor; auto. + + destruct r as ((h1 & h2) & (han & hrarg & hty & hbod)). + eapply h1 in hty; auto. constructor; auto. Qed. #[global] @@ -612,6 +643,35 @@ Proof. intros p. unfold eq_predicate; intuition auto; try now symmetry. Qed. +#[global] +Polymorphic Instance eq_branch_sym Re : + CRelationClasses.Symmetric Re -> + CRelationClasses.Symmetric (eq_branch Re). +Proof. + intros hre. + intros br. unfold eq_branch; intuition auto; try now symmetry. +Qed. + +#[global] +Polymorphic Instance eq_branches_sym Re : + CRelationClasses.Symmetric Re -> + CRelationClasses.Symmetric (eq_branches Re). +Proof. + intros hre. + intros brs. unfold eq_branches; intuition auto; try now symmetry. +Qed. + +#[global] +Polymorphic Instance eq_mfix_sym Re : + CRelationClasses.Symmetric Re -> + CRelationClasses.Symmetric (eq_mfix Re). +Proof. + intros hre. + intros mfix. unfold eq_mfix; + intros; apply All2_sym, All2_impl with (1 := X). + intuition auto; try now symmetry. +Qed. + #[global] Instance compare_term_sym {cf} Σ φ : Symmetric (compare_term Conv Σ φ). Proof. @@ -663,6 +723,37 @@ Proof. etransitivity; tea. Qed. +#[global] +Polymorphic Instance eq_branch_trans Re : + CRelationClasses.Transitive Re -> + CRelationClasses.Transitive (eq_branch Re). +Proof. + intros hre. + intros br. unfold eq_branch; intuition auto; try now etransitivity. + etransitivity; tea. +Qed. + +#[global] +Polymorphic Instance eq_branches_trans Re : + CRelationClasses.Transitive Re -> + CRelationClasses.Transitive (eq_branches Re). +Proof. + intros hre. + intros brs. unfold eq_branches; intros. + eapply All2_trans; tea. tc. +Qed. + +#[global] +Polymorphic Instance eq_mfix_trans Re : + CRelationClasses.Transitive Re -> + CRelationClasses.Transitive (eq_mfix Re). +Proof. + intros hre. + intros mfix. unfold eq_mfix; intros. + eapply All2_trans; tea. + intros d d' d'' [] []; repeat split; now etransitivity. +Qed. + #[global] Polymorphic Instance eq_term_upto_univ_trans Σ Re Rle napp : RelationClasses.Transitive Re -> @@ -695,7 +786,7 @@ Proof. - dependent destruction e2. unfold eq_predicate in *. !!solve_all. - econstructor; unfold eq_predicate in *; solve_all; eauto. + econstructor; unfold eq_predicate, eq_branches, eq_branch in *; solve_all; eauto. * clear -he hh1 hh2. revert hh1 hh2. generalize (pparams p), p'.(pparams), p'0.(pparams). intros l l' l''. @@ -704,10 +795,10 @@ Proof. * etransitivity; eauto. * eapply onctx_eq_ctx_trans in hh; eauto. intros ???? -> ->; eauto. - * clear -H he a a0. - induction a in a0, brs'0 |- *. + * clear -H he e0 e4. + induction e0 in e4, brs'0 |- *. + assumption. - + depelim a0. destruct p. constructor; auto. + + depelim e4. destruct p. constructor; auto. destruct r as [[h1 h1'] [h2 h3]]. split. eapply onctx_eq_ctx_trans in h1; eauto. @@ -716,19 +807,19 @@ Proof. - dependent destruction e2. econstructor. eapply All2_All_mix_left in X as h; eauto. - clear a X. - induction h in a0, mfix'0 |- *. + clear e X. unfold eq_mfix in *. + induction h in e0, mfix'0 |- *. + assumption. - + dependent destruction a0. constructor ; eauto. + + dependent destruction e0. constructor ; eauto. intuition eauto. transitivity (rarg y); auto. - dependent destruction e2. econstructor. eapply All2_All_mix_left in X as h; eauto. - clear a X. - induction h in a0, mfix'0 |- *. + clear e X. unfold eq_mfix in *. + induction h in e0, mfix'0 |- *. + assumption. - + dependent destruction a0. constructor ; eauto. + + dependent destruction e0. constructor ; eauto. intuition eauto. transitivity (rarg y); auto. Qed. @@ -935,16 +1026,16 @@ Proof. eapply R_global_instance_impl. 5:eauto. all:auto. - inversion 1; subst; constructor. eapply R_global_instance_impl. 5:eauto. all:eauto. - - inversion 1; subst; constructor; unfold eq_predicate in *; eauto; solve_all. + - inversion 1; subst; constructor; unfold eq_predicate, eq_branches, eq_branch in *; eauto; solve_all. * eapply R_universe_instance_impl'; eauto. - inversion 1; subst; constructor. eapply All2_impl'; tea. eapply All_impl; eauto. - cbn. intros x [? ?] y [[[? ?] ?] ?]. repeat split; eauto. + cbn. intros x [? ?] y (? & ? & ? & ?). repeat split; eauto. - inversion 1; subst; constructor. eapply All2_impl'; tea. eapply All_impl; eauto. - cbn. intros x [? ?] y [[[? ?] ?] ?]. repeat split; eauto. + cbn. intros x [? ?] y (? & ? & ? & ?). repeat split; eauto. Qed. #[global] @@ -966,16 +1057,16 @@ Proof. eapply R_global_instance_empty_impl. 4:eauto. all:eauto. - inversion 1; subst; constructor. eapply R_global_instance_empty_impl. 4:eauto. all:eauto. - - inversion 1; subst; constructor; unfold eq_predicate in *; solve_all. + - inversion 1; subst; constructor; unfold eq_predicate, eq_branches, eq_branch in *; solve_all. * eapply R_universe_instance_impl'; eauto. - inversion 1; subst; constructor. eapply All2_impl'; tea. eapply All_impl; eauto. - cbn. intros x [? ?] y [[[? ?] ?] ?]. repeat split; eauto. + cbn. intros x [? ?] y (? & ? & ? & ?). repeat split; eauto. - inversion 1; subst; constructor. eapply All2_impl'; tea. eapply All_impl; eauto. - cbn. intros x [? ?] y [[[? ?] ?] ?]. repeat split; eauto. + cbn. intros x [? ?] y (? & ? & ? & ?). repeat split; eauto. Qed. #[global] @@ -1021,15 +1112,15 @@ Proof. all: dependent destruction e. all: try solve [cbn ; constructor ; try lih ; try assumption; solve_all]. - cbn. destruct e as (? & ? & e & ?). - constructor; unfold eq_predicate in *; simpl; !!solve_all. + constructor; unfold eq_predicate, eq_branches, eq_branch in *; simpl; !!solve_all. * rewrite -?(All2_fold_length e). eapply hh0; eauto. - * rewrite (All2_fold_length a). now eapply hh4. - - cbn. constructor. - pose proof (All2_length a). + * rewrite (All2_fold_length hh5). now eapply hh4. + - cbn. constructor. unfold eq_mfix in *. + pose proof (All2_length e). solve_all. rewrite H. eauto. - - cbn. constructor. - pose proof (All2_length a). + - cbn. constructor. unfold eq_mfix in *. + pose proof (All2_length e). solve_all. rewrite H. eauto. Qed. @@ -1100,14 +1191,14 @@ Proof. - cbn. constructor. solve_all. - cbn. destruct e as (? & ? & e & ?). - constructor ; unfold eq_predicate; simpl; try sih ; solve_all. - * rewrite -(All2_fold_length e). eapply e1; eauto. - * rewrite (All2_fold_length a). now eapply b0. - - cbn. constructor ; try sih ; eauto. - pose proof (All2_length a). + constructor ; unfold eq_predicate, eq_branches, eq_branch in *; simpl; try sih ; solve_all. + * rewrite -(All2_fold_length e). eapply e2; eauto. + * rewrite (All2_fold_length a0). now eapply b0. + - cbn. constructor ; unfold eq_mfix in *; try sih ; eauto. + pose proof (All2_length e). solve_all. now rewrite H. - - cbn. constructor ; try sih ; eauto. - pose proof (All2_length a). + - cbn. constructor ; unfold eq_mfix in *; try sih ; eauto. + pose proof (All2_length e). solve_all. now rewrite H. Qed. @@ -1651,17 +1742,6 @@ Proof. induction 1; intros; constructor; intuition auto. all:try solve [now symmetry]. all:eauto using R_global_instance_flip. - - eapply All2_sym. solve_all. - * eapply eq_context_sym; try tc. tas. - * now eapply eq_term_upto_univ_sym. - - eapply All2_sym. solve_all. - now eapply eq_term_upto_univ_sym. - now eapply eq_term_upto_univ_sym. - now symmetry. - - eapply All2_sym. solve_all. - now eapply eq_term_upto_univ_sym. - now eapply eq_term_upto_univ_sym. - now symmetry. Qed. Lemma eq_univ_make : diff --git a/pcuic/theories/PCUICExpandLetsCorrectness.v b/pcuic/theories/PCUICExpandLetsCorrectness.v index 849f5a1b7..c7de784cc 100644 --- a/pcuic/theories/PCUICExpandLetsCorrectness.v +++ b/pcuic/theories/PCUICExpandLetsCorrectness.v @@ -57,7 +57,7 @@ Proof. reflexivity. Qed. -Lemma trans_local_app Γ Δ : trans_local (SE.app_context Γ Δ) = trans_local Γ ,,, trans_local Δ. +Lemma trans_local_app Γ Δ : trans_local (Γ ,,, Δ) = trans_local Γ ,,, trans_local Δ. Proof. now rewrite /trans_local map_app. Qed. @@ -2312,6 +2312,7 @@ Proof. red in X0. do 2 apply All2_map. eapply All2_All_mix_left in X3; tea. eapply All2_impl; tea; cbv beta. + unfold eq_branch in X3 |- *. intuition auto. rewrite !trans_bcontext. eapply eq_context_gen_smash_context. @@ -2325,8 +2326,8 @@ Proof. now eapply trans_eq_context_gen. cbn. eapply All2_rev. solve_all. eauto using subrelation_refl. cbn. eauto using subrelation_refl. - - constructor. solve_all; eauto using subrelation_refl. - - constructor; solve_all; eauto using subrelation_refl. + - constructor; unfold eq_mfix in X0 |- *; solve_all; eauto using subrelation_refl. + - constructor; unfold eq_mfix in X0 |- *; solve_all; eauto using subrelation_refl. Qed. Lemma trans_compare_term {cf} {Σ : global_env} {pb ϕ T U} : @@ -2524,7 +2525,7 @@ Lemma trans_mfix_All {cf} Σ Γ mfix idx : (Γ : SE.context) (b ty : PCUICAst.term) => ST.typing Σ Γ b ty × TT.typing (H := cf' cf) (trans_global Σ) (trans_local Γ) (trans b) (trans ty)) Σ) - (SE.app_context Γ (SE.fix_context mfix)) -> + (Γ ,,, (SE.fix_context mfix)) -> TTwf_local (trans_global Σ) (trans_local Γ ,,, fix_context (map (map_def trans trans) mfix)). Proof. @@ -2532,10 +2533,10 @@ Proof. rewrite -(trans_fix_context (shiftnP #|Γ| xpred0) _ idx) //. match goal with |- TTwf_local _ ?A => - replace A with (trans_local (SE.app_context Γ (SE.fix_context mfix))) + replace A with (trans_local (Γ ,,, (SE.fix_context mfix))) end. 2: { - unfold trans_local, SE.app_context. + unfold trans_local, app_context. now rewrite map_app. } @@ -3634,7 +3635,7 @@ Proof. generalize (pparams p ++ indices). change T.PCUICEnvironment.subst_instance_context with subst_instance_context. rewrite -/context. - generalize (ind_params mdecl ,,, ind_indices idecl)@[puinst p] as Δ. + generalize (ind_params mdecl ,,, ind_indices idecl : context)@[puinst p] as Δ. intros c; revert Γ. induction c using ctx_length_rev_ind. * intros Γ l wf. intros c; depelim c. constructor. diff --git a/pcuic/theories/PCUICFirstorder.v b/pcuic/theories/PCUICFirstorder.v index 3c466a053..ceac959b3 100644 --- a/pcuic/theories/PCUICFirstorder.v +++ b/pcuic/theories/PCUICFirstorder.v @@ -191,7 +191,9 @@ Lemma isType_context_conversion {Σ : global_env_ext} {wfΣ : wf Σ} {Γ Δ} {T} wf_local Σ Δ -> isType Σ Δ T. Proof using Type. - intros [s Hs]. exists s. eapply context_conversion; tea. now eapply ws_cumul_ctx_pb_forget. + intros HT wf eq. + apply infer_sort_impl with id HT => //; intros Hs. + eapply closed_context_conversion; tea. Qed. Lemma typing_spine_arity_spine {Σ : global_env_ext} {wfΣ : wf Σ} Γ Δ args T' i u pars : @@ -217,7 +219,7 @@ Proof using Type. now eapply isType_open. -- 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 isType_of_isTypeRel, i0. ++ rewrite /subst1 PCUICLiftSubst.subst_it_mkProd_or_LetIn. autorewrite with subst. cbn. eapply X. len. lia. eapply typing_spine_strengthen. eauto. @@ -232,7 +234,7 @@ Proof using Type. eapply 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. eapply isType_tProd in i1 as [tdom _] => //. Qed. Lemma leb_spect : forall x y : nat, BoolSpecSet (x <= y) (y < x) (x <=? y). @@ -323,7 +325,7 @@ Proof using Type. now eapply isType_tLetIn_red in isty; pcuic. - depelim hi. solve_discr. specialize (i1 hd). specialize (IHhsp i1). - destruct (validity t) as [s Hs]. eapply inversion_mkApps in Hs as [? [hi _]]. + destruct (validity t) as (s & e & Hs). eapply inversion_mkApps in Hs as [? [hi _]]. eapply inversion_Ind in hi as [mdecl [idecl [decli [? ?]]]]. econstructor; tea. 2:{ eapply IHhsp. eapply isType_apply in isty; tea. } now eapply isType_ws_cumul_pb_refl. eauto. @@ -531,9 +533,9 @@ Proof using Type. - destruct t; inversion_clear Hvalue. + exfalso. eapply inversion_Sort in Hty as (? & ? & Hcumul); eauto. now eapply invert_cumul_sort_ind in Hcumul. - + exfalso. eapply inversion_Prod in Hty as (? & ? & ? & ? & Hcumul); eauto. + + exfalso. eapply inversion_Prod in Hty as (? & ? & ? & ? & ? & Hcumul); eauto. now eapply invert_cumul_sort_ind in Hcumul. - + exfalso. eapply inversion_Lambda in Hty as (? & ? & ? & ? & Hcumul); eauto. + + exfalso. eapply inversion_Lambda in Hty as (? & ? & ? & ? & ? & Hcumul); eauto. now eapply invert_cumul_prod_ind in Hcumul. + exfalso. eapply inversion_Ind in Hty as (? & ? & ? & ? & ? & ?); eauto. eapply PCUICInductives.declared_inductive_type in d. diff --git a/pcuic/theories/PCUICGlobalEnv.v b/pcuic/theories/PCUICGlobalEnv.v index 99f935d00..b3accc1db 100644 --- a/pcuic/theories/PCUICGlobalEnv.v +++ b/pcuic/theories/PCUICGlobalEnv.v @@ -1,8 +1,36 @@ (* Distributed under the terms of the MIT license. *) From Coq Require Import ProofIrrelevance. From MetaCoq.Template Require Import config utils uGraph. -From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils - PCUICReflect PCUICTyping. +From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils. + +Require Import ssreflect. + + +Lemma lookup_on_global_env {cf : checker_flags} Pcmp P (Σ : global_env) c decl : + on_global_env Pcmp P Σ -> + lookup_env Σ c = Some decl -> + { Σ' : global_env & [× extends Σ' Σ, on_global_env Pcmp P Σ' & + on_global_decl Pcmp P (Σ', universes_decl_of_decl decl) c decl] }. +Proof using Type. + destruct Σ as [univs Σ]; unfold on_global_env, lookup_env; cbn. + intros [cu Σp]. + induction Σp; simpl. congruence. + destruct (eqb_specT c kn); subst. + - intros [= ->]. + exists ({| universes := univs; declarations := Σ |}). + split. + * red; cbn. split; [split;[lsets|csets]|]. + exists [(kn, decl)] => //. + * split => //. + * apply o0. + - intros hl. destruct (IHΣp hl) as [Σ' []]. + exists Σ'. + split=> //. + destruct e as [eu ed]. red; cbn in *. + split; [auto|]. + destruct ed as [Σ'' ->]. + exists (Σ'' ,, (kn, d)) => //. +Qed. (** Injectivity of declared_*, inversion lemmas on declared global references and universe consistency of the global environment. @@ -66,9 +94,9 @@ Proof. now intros []. Qed. Coercion declared_projection_constructor : declared_projection >-> declared_constructor. Section DeclaredInv. - Context {cf:checker_flags} {Σ} {wfΣ : wf Σ}. + Context {cf:checker_flags} {Σ} {Pcmp P} {wfΣ : on_global_env Pcmp P Σ}. - Lemma declared_minductive_ind_npars {mdecl ind} : + Lemma declared_minductive_ind_npars_gen {mdecl ind} : declared_minductive Σ ind mdecl -> ind_npars mdecl = context_assumptions mdecl.(ind_params). Proof using wfΣ. @@ -82,8 +110,8 @@ Section DeclaredInv. End DeclaredInv. -Definition wf_global_uctx_invariants {cf:checker_flags} {P} Σ : - on_global_env cumulSpec0 P Σ -> +Definition wf_global_uctx_invariants {cf:checker_flags} {Pcmp P} Σ : + on_global_env Pcmp P Σ -> global_uctx_invariants (global_uctx Σ). Proof. intros HΣ. split. @@ -105,8 +133,8 @@ Proof. right. now apply LevelSet.union_spec. Qed. -Definition wf_ext_global_uctx_invariants {cf:checker_flags} {P} Σ : - on_global_env_ext cumulSpec0 P Σ -> +Definition wf_ext_global_uctx_invariants {cf:checker_flags} {Pcmp P} Σ : + on_global_env_ext Pcmp P Σ -> global_uctx_invariants (global_ext_uctx Σ). Proof. intros HΣ. split. @@ -122,18 +150,18 @@ Proof. split; apply LevelSet.union_spec; right; apply XX. Qed. -Lemma wf_consistent {cf:checker_flags} Σ {P} : - on_global_env cumulSpec0 P Σ -> consistent (global_constraints Σ). +Lemma wf_consistent {cf:checker_flags} Σ {Pcmp P} : + on_global_env Pcmp P Σ -> consistent (global_constraints Σ). Proof. destruct Σ. intros [cu ong]. apply cu. Qed. -Definition global_ext_uctx_consistent {cf:checker_flags} {P} Σ - : on_global_env_ext cumulSpec0 P Σ -> consistent (global_ext_uctx Σ).2. +Definition global_ext_uctx_consistent {cf:checker_flags} {Pcmp P} Σ + : on_global_env_ext Pcmp P Σ -> consistent (global_ext_uctx Σ).2. Proof. intros HΣ. cbn. unfold global_ext_constraints. - unfold wf_ext, on_global_env_ext in HΣ. + unfold on_global_env_ext in HΣ. destruct HΣ as (_ & _ & _ & HH & _). apply HH. Qed. diff --git a/pcuic/theories/PCUICInductives.v b/pcuic/theories/PCUICInductives.v index 3803b814a..8011fa280 100644 --- a/pcuic/theories/PCUICInductives.v +++ b/pcuic/theories/PCUICInductives.v @@ -4,7 +4,7 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICTactics PCUICInduc PCUICLiftSubst PCUICUnivSubst PCUICTyping PCUICWeakeningEnvConv PCUICWeakeningEnvTyp PCUICWeakeningConv PCUICWeakeningTyp PCUICSigmaCalculus PCUICInstDef PCUICInstConv PCUICContextSubst - PCUICRenameDef PCUICRenameConv PCUICRenameTyp + PCUICRenameDef PCUICRenameTerm PCUICRenameConv PCUICRenameTyp PCUICSubstitution PCUICOnFreeVars PCUICClosed PCUICClosedConv PCUICClosedTyp PCUICCumulativity PCUICGeneration PCUICReduction PCUICWellScopedCumulativity PCUICEquality PCUICConfluence PCUICParallelReductionConfluence diff --git a/pcuic/theories/PCUICNormal.v b/pcuic/theories/PCUICNormal.v index 990faf564..1deb138cf 100644 --- a/pcuic/theories/PCUICNormal.v +++ b/pcuic/theories/PCUICNormal.v @@ -1565,10 +1565,10 @@ Proof. destruct e0 as (?&?&?). eapply whne_fixapp. + unfold unfold_fix. - rewrite e1. + rewrite e2. reflexivity. + rewrite <- e. - destruct p. rewrite e3. reflexivity. + destruct p as (? & ? & ?). rewrite e4. reflexivity. + eapply IHwh; eauto. - destruct args using MCList.rev_ind; [|rewrite mkApps_app in x; discriminate x]. now rewrite nth_error_nil in e0. @@ -1601,8 +1601,8 @@ Proof. unfold unfold_fix in *. destruct (nth_error mfix' idx) eqn:nth; [|easy]. eapply All2_nth_error_Some_r in nth; eauto. - destruct nth as (?&?&((? & ?)&?)&?). - rewrite e e2 in y. + destruct nth as (? & ? & ? & ? & ? & ?). + rewrite e0 e2 in y. eapply All2_nth_error_None; eauto. - apply eq_term_upto_univ_napp_mkApps_l_inv in eq as (?&?&(?&?)&->). depelim e. diff --git a/pcuic/theories/PCUICParallelReduction.v b/pcuic/theories/PCUICParallelReduction.v index 37a92f5cf..e8899607e 100644 --- a/pcuic/theories/PCUICParallelReduction.v +++ b/pcuic/theories/PCUICParallelReduction.v @@ -4,7 +4,7 @@ From MetaCoq.Template Require Import config utils. From MetaCoq.PCUIC Require Import PCUICUtils PCUICOnOne PCUICAst PCUICAstUtils PCUICTactics PCUICDepth PCUICCases PCUICLiftSubst PCUICUnivSubst PCUICReduction PCUICTyping PCUICSigmaCalculus PCUICWeakeningEnvConv PCUICInduction - PCUICRenameDef PCUICRenameConv PCUICInstDef PCUICInstConv PCUICOnFreeVars + PCUICRenameDef PCUICRenameTerm PCUICRenameConv PCUICInstDef PCUICInstConv PCUICOnFreeVars PCUICWeakeningConv PCUICWeakeningTyp PCUICSubstitution. Require Import ssreflect ssrbool. diff --git a/pcuic/theories/PCUICParallelReductionConfluence.v b/pcuic/theories/PCUICParallelReductionConfluence.v index 2240ffa25..36c827a0f 100644 --- a/pcuic/theories/PCUICParallelReductionConfluence.v +++ b/pcuic/theories/PCUICParallelReductionConfluence.v @@ -4,7 +4,7 @@ From MetaCoq.Template Require Import config utils. From MetaCoq.PCUIC Require Import PCUICAst PCUICOnOne PCUICAstUtils PCUICTactics PCUICSize PCUICLiftSubst PCUICSigmaCalculus PCUICUnivSubst PCUICTyping PCUICReduction PCUICReflect PCUICInduction PCUICClosed PCUICClosedConv PCUICClosedTyp PCUICDepth PCUICOnFreeVars - PCUICRenameDef PCUICRenameConv PCUICInstDef PCUICInstConv PCUICWeakeningConv PCUICWeakeningTyp + PCUICRenameDef PCUICRenameTerm PCUICRenameConv PCUICInstDef PCUICInstConv PCUICWeakeningConv PCUICWeakeningTyp PCUICViews PCUICParallelReduction. diff --git a/pcuic/theories/PCUICPrincipality.v b/pcuic/theories/PCUICPrincipality.v index 02ad31605..7a380e2c3 100644 --- a/pcuic/theories/PCUICPrincipality.v +++ b/pcuic/theories/PCUICPrincipality.v @@ -703,8 +703,7 @@ Proof. eapply (All_impl X0); pcuicfo. apply infer_typing_sort_impl with id X2 => //; now intros []. eapply (All_impl X1); pcuicfo; now destruct X2. - eapply All2_nth_error in a; eauto. - destruct a as [[[eqty _] _] _]. + eapply All2_nth_error in e as (_ & _ & eqty & _); eauto. constructor. eapply eq_term_empty_leq_term in eqty. now eapply leq_term_empty_leq_term. @@ -716,8 +715,7 @@ Proof. eapply (All_impl X0); pcuicfo. apply infer_typing_sort_impl with id X2 => //; now intros []. eapply (All_impl X1); pcuicfo; now destruct X2. - eapply All2_nth_error in a; eauto. - destruct a as [[[eqty _] _] _]. + eapply All2_nth_error in e as (_ & _ & eqty & _); eauto. constructor. apply eq_term_empty_leq_term in eqty. now eapply leq_term_empty_leq_term. Qed. diff --git a/pcuic/theories/PCUICProgress.v b/pcuic/theories/PCUICProgress.v index c31e32f6f..12f3e4ecb 100644 --- a/pcuic/theories/PCUICProgress.v +++ b/pcuic/theories/PCUICProgress.v @@ -511,7 +511,7 @@ Proof. unfold cstr_concl in hsp. cbn in hsp. len in hsp. rewrite H in hsp. clear H. eapply (declared_constructor_ind_decl declc). clear H. eapply typing_spine_length in hsp. len in hsp. unfold cstr_arity. - now rewrite (PCUICGlobalEnv.declared_minductive_ind_npars declc). + now rewrite (PCUICTyping.declared_minductive_ind_npars declc). Qed. Lemma value_mkApps_inv' Σ f args : @@ -769,7 +769,7 @@ Proof. - inversion Heval; subst; clear Heval. all:cbn in Hty; solve_all. all: now econstructor; eauto with wcbv. - inversion Heval; subst; clear Heval. all:cbn in Hty; solve_all. all: try now econstructor; eauto with wcbv. - eapply eval_iota. eapply eval_mkApps_Construct; tea. now econstructor. unfold cstr_arity. rewrite e0. - rewrite (PCUICGlobalEnv.declared_minductive_ind_npars d). + rewrite (PCUICTyping.declared_minductive_ind_npars d). now rewrite -(declared_minductive_ind_npars d) /cstr_arity. all:tea. eapply All_All2_refl. solve_all. now eapply value_final. - inversion Heval; subst; clear Heval. all:cbn in Hty; solve_all. all: now econstructor; eauto with wcbv. diff --git a/pcuic/theories/PCUICRelevance.v b/pcuic/theories/PCUICRelevance.v new file mode 100644 index 000000000..8c9da0cf5 --- /dev/null +++ b/pcuic/theories/PCUICRelevance.v @@ -0,0 +1,49 @@ +(* Distributed under the terms of the MIT license. *) +From MetaCoq.Template Require Import config utils. +From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils. + +Require Import ssreflect. + +Definition relevance_of_constant Σ kn := + match lookup_constant Σ kn with + | Some decl => decl.(cst_relevance) + | None => Relevant + end. + +Definition relevance_of_constructor Σ ind (k : nat) := + match lookup_inductive Σ ind with + | Some (_, idecl) => idecl.(ind_relevance) + | None => Relevant + end. + +Definition relevance_of_projection Σ (p: projection) := + match lookup_projection Σ p with + | Some (_, _, _, pdecl) => pdecl.(proj_relevance) + | None => Relevant + end. + +Definition mark_context := list relevance. +Definition marks_of_context (Γ: context) := List.map (binder_relevance ∘ decl_name) Γ. + +Fixpoint relevance_of_term (Σ : global_env) (Γ : mark_context) (t: term) : relevance := + match t with + | tRel n => + option_default id (nth_error Γ n) Relevant + | tLambda na A t => relevance_of_term Σ (Γ ,, na.(binder_relevance)) t + | tLetIn na b B t => relevance_of_term Σ (Γ ,, na.(binder_relevance)) t + | tApp u _ => relevance_of_term Σ Γ u + | tConst k _ => relevance_of_constant Σ k + | tConstruct ind i _ => relevance_of_constructor Σ ind i + | tCase ci _ _ _ => ci.(ci_relevance) + | tProj p _ => relevance_of_projection Σ p + | tFix mfix n | tCoFix mfix n => option_default (binder_relevance ∘ dname) (nth_error mfix n) Relevant + | tVar _ | tEvar _ _ | tSort _ | tProd _ _ _ | tInd _ _ => Relevant + end. + +Notation isTermRel Σ Γ t rel := (relevance_of_term Σ Γ t = rel). + +Lemma mark_of_context_app Γ Γ' Δ Δ' : + marks_of_context Γ = marks_of_context Γ' -> marks_of_context Δ = marks_of_context Δ' -> marks_of_context (Γ,,, Δ) = marks_of_context (Γ',,, Δ'). +Proof. + rewrite /marks_of_context !map_app; congruence. +Qed. diff --git a/pcuic/theories/PCUICRelevanceTerm.v b/pcuic/theories/PCUICRelevanceTerm.v new file mode 100644 index 000000000..7e89970de --- /dev/null +++ b/pcuic/theories/PCUICRelevanceTerm.v @@ -0,0 +1,40 @@ +From MetaCoq.Template Require Import config utils. +From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICRelevance PCUICInduction PCUICCases PCUICSigmaCalculus PCUICLiftSubst PCUICWeakeningEnv. + +Require Import ssreflect. + + +Lemma mark_inst_case_context params puinst (pctx : context) : + marks_of_context (inst_case_context params puinst pctx) = marks_of_context pctx. +Proof. + unfold marks_of_context, inst_case_context, subst_context, subst_instance, subst_instance_context, map_context. + rewrite -map_map map_decl_name_fold_context_k !map_map. now cbn. +Qed. + +Lemma mark_inst_case_predicate_context p : + marks_of_context (inst_case_predicate_context p) = marks_of_context p.(pcontext). +Proof. apply mark_inst_case_context. Qed. + +Lemma mark_inst_case_branch_context p br : + marks_of_context (inst_case_branch_context p br) = marks_of_context br.(bcontext). +Proof. apply mark_inst_case_context. Qed. + +Lemma extends_irrelevant {cf : checker_flags} {Pcmp P} Σ Σ' Γ t : + on_global_env Pcmp P Σ' -> + extends Σ Σ' -> + isTermRel Σ Γ t Irrelevant -> + isTermRel Σ' Γ t Irrelevant. +Proof. + induction t in Γ |- * using term_forall_list_ind; + cbn; intros ext Hirr; auto. + - unfold relevance_of_constant. + destruct lookup_constant eqn:H => //. + erewrite extends_lookup_constant; eauto. + - unfold relevance_of_constructor. + destruct lookup_inductive eqn:H => //. + erewrite extends_lookup_inductive; eauto. + - unfold relevance_of_projection. + destruct lookup_projection eqn:H => //. + erewrite extends_lookup_projection; eauto. +Qed. + diff --git a/pcuic/theories/PCUICSR.v b/pcuic/theories/PCUICSR.v index b01ff0533..fc46b5f18 100644 --- a/pcuic/theories/PCUICSR.v +++ b/pcuic/theories/PCUICSR.v @@ -1274,8 +1274,8 @@ Lemma map2_set_binder_name_expand_lets nas Γ Δ : Proof. move=> hlen. rewrite /expand_lets_ctx /expand_lets_k_ctx. - rewrite PCUICRenameConv.map2_set_binder_name_fold ?lengths //. - rewrite PCUICRenameConv.map2_set_binder_name_fold ?lengths //. + rewrite PCUICRenameTerm.map2_set_binder_name_fold ?lengths //. + rewrite PCUICRenameTerm.map2_set_binder_name_fold ?lengths //. Qed. Lemma closed_red1_eq_context_upto_names {Σ Γ Γ'} {t u} : @@ -2264,7 +2264,7 @@ Proof. eapply All2_app. - rewrite /case_branch_context /case_branch_context_gen /pre_case_branch_context /pre_case_branch_context_gen /inst_case_context. cbn. - symmetry. rewrite /subst_context PCUICRenameConv.map2_set_binder_name_fold ?lengths. + symmetry. rewrite /subst_context PCUICRenameTerm.map2_set_binder_name_fold ?lengths. now apply wf_branch_length. rewrite -/(subst_context _ _ _). relativize #|cstr_args cdecl|. diff --git a/pcuic/theories/PCUICSafeLemmata.v b/pcuic/theories/PCUICSafeLemmata.v index a83c24ebc..b0932f2c2 100644 --- a/pcuic/theories/PCUICSafeLemmata.v +++ b/pcuic/theories/PCUICSafeLemmata.v @@ -181,7 +181,7 @@ Section Lemmata. eapply ws_cumul_pb_Lambda_r. assumption. Qed. - Lemma snoc_app_context {Γ Δ d} : (Γ ,,, (d :: Δ)) = (Γ ,,, Δ) ,,, [d]. + Lemma snoc_app_context {A Γ Δ} {d: A} : (Γ ,,, (d :: Δ)) = (Γ ,,, Δ) ,,, [d]. Proof using Type. reflexivity. Qed. diff --git a/pcuic/theories/PCUICSigmaCalculus.v b/pcuic/theories/PCUICSigmaCalculus.v index b32443626..19e149f11 100644 --- a/pcuic/theories/PCUICSigmaCalculus.v +++ b/pcuic/theories/PCUICSigmaCalculus.v @@ -380,7 +380,7 @@ Qed. #[global] Hint Rewrite mapi_context_compose : map. -Lemma rename_compose f f' : rename f ∘ rename f' =1 rename (f ∘ f'). +Lemma rename_compose_eq1 f f' : rename f ∘ rename f' =1 rename (f ∘ f'). Proof. intros x. induction x in f, f' |- * using term_forall_list_ind; simpl; @@ -398,6 +398,9 @@ Proof. * len. rewrite b. apply rename_ext, shiftn_compose. Qed. +Lemma rename_compose f f' x : rename f (rename f' x) = rename (f ∘ f') x. +Proof using Type. apply (rename_compose_eq1 f f' x). Qed. + Lemma map_predicate_shift_map_predicate_shift {T} {fn : (nat -> T) -> term -> term} {shift : nat -> (nat -> T) -> nat -> T} diff --git a/pcuic/theories/PCUICSpine.v b/pcuic/theories/PCUICSpine.v index 450d3c12d..f5676bcef 100644 --- a/pcuic/theories/PCUICSpine.v +++ b/pcuic/theories/PCUICSpine.v @@ -227,7 +227,7 @@ Proof. apply IHctxi. Qed. -Record spine_subst {cf:checker_flags} Σ Γ inst s Δ := mkSpineSubst { +Record spine_subst {cf:checker_flags} Σ Γ inst s (Δ : context) := mkSpineSubst { spine_dom_wf : wf_local Σ Γ; spine_codom_wf : wf_local Σ (Γ ,,, Δ); inst_ctx_subst :> context_subst Δ inst s; @@ -1620,7 +1620,7 @@ Section WfEnv. now rewrite lift_context_app Nat.add_0_r. Qed. - Lemma spine_subst_to_extended_list_k {Δ Γ} : + Lemma spine_subst_to_extended_list_k {Δ Γ : context} : wf_local Σ (Γ ,,, Δ) -> spine_subst Σ (Γ ,,, Δ) (reln [] 0 Δ) (all_rels Δ 0 #|Δ|) (lift_context #|Δ| 0 Δ). diff --git a/pcuic/theories/PCUICToTemplateCorrectness.v b/pcuic/theories/PCUICToTemplateCorrectness.v index 1cf4342ff..b917afc94 100644 --- a/pcuic/theories/PCUICToTemplateCorrectness.v +++ b/pcuic/theories/PCUICToTemplateCorrectness.v @@ -70,7 +70,7 @@ Proof. reflexivity. Qed. -Lemma trans_local_app Γ Δ : trans_local (SE.app_context Γ Δ) = trans_local Γ ,,, trans_local Δ. +Lemma trans_local_app Γ Δ : trans_local (Γ ,,, Δ) = trans_local Γ ,,, trans_local Δ. Proof. now rewrite /trans_local map_app. Qed. @@ -1421,10 +1421,10 @@ Proof. eapply ih ; auto end ]. - 1,6,7: try solve [ constructor; solve_all ]. + 1,6,7: try solve [ constructor; try unfold eq_mfix in X0; solve_all ]. all: try solve [ constructor; now eapply trans_R_global_instance ]. - eapply (TermEquality.eq_term_upto_univ_mkApps _ _ _ _ _ [_] _ [_]); simpl; eauto. - - destruct X1 as [Hpars [Huinst [Hctx Hret]]]. + - destruct X1 as [Hpars [Huinst [Hctx Hret]]]. unfold eq_branches, eq_branch in X3. destruct X as [IHpars [IHctx IHret]]. constructor; cbn; auto. solve_all. red in X0. @@ -1563,7 +1563,7 @@ Lemma trans_mfix_All {cf} Σ Γ mfix: (Γ : SE.context) (b ty : PCUICAst.term) => ST.typing Σ Γ b ty × TT.typing (trans_global Σ) (trans_local Γ) (trans b) (trans ty)) Σ) - (SE.app_context Γ (SE.fix_context mfix)) -> + (Γ ,,, (SE.fix_context mfix)) -> TTwf_local (trans_global Σ) (trans_local Γ ,,, TT.fix_context (map (map_def trans trans) mfix)). Proof. @@ -1571,10 +1571,10 @@ Proof. rewrite <- trans_fix_context. match goal with |- TTwf_local _ ?A => - replace A with (trans_local (SE.app_context Γ (SE.fix_context mfix))) + replace A with (trans_local (Γ ,,, (SE.fix_context mfix))) end. 2: { - unfold trans_local, SE.app_context. + unfold trans_local, app_context. now rewrite map_app. } @@ -1592,11 +1592,11 @@ Qed. Lemma trans_mfix_All2 {cf} Σ Γ mfix xfix: All (fun d : def PCUICAst.term => - (ST.typing Σ (SE.app_context Γ (SE.fix_context xfix)) + (ST.typing Σ (Γ ,,, (SE.fix_context xfix)) (dbody d) (S.lift0 #|SE.fix_context xfix| (dtype d))) × TT.typing (trans_global Σ) - (trans_local (SE.app_context Γ (SE.fix_context xfix))) + (trans_local (Γ ,,, (SE.fix_context xfix))) (trans (dbody d)) (trans (S.lift0 #|SE.fix_context xfix| @@ -1612,7 +1612,7 @@ Proof. - constructor. - simpl; constructor. + destruct p as []. - unfold app_context, SE.app_context in *. + unfold app_context in *. unfold trans_local in t0. rewrite map_app trans_fix_context in t0. rewrite trans_dbody trans_lift trans_dtype in t0. diff --git a/pcuic/theories/PCUICTyping.v b/pcuic/theories/PCUICTyping.v index 01184e20d..ba3ee229f 100644 --- a/pcuic/theories/PCUICTyping.v +++ b/pcuic/theories/PCUICTyping.v @@ -1,7 +1,7 @@ (* Distributed under the terms of the MIT license. *) From MetaCoq.Template Require Import config utils. From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils - PCUICLiftSubst PCUICUnivSubst PCUICEquality PCUICUtils PCUICPosition. + PCUICLiftSubst PCUICUnivSubst PCUICEquality PCUICUtils PCUICPosition PCUICGlobalEnv. From MetaCoq.PCUIC Require Export PCUICCumulativitySpec. From MetaCoq.PCUIC Require Export PCUICCases. @@ -156,7 +156,7 @@ Variant case_side_conditions `{checker_flags} wf_local_fun typing Σ Γ ci p ps (allowed_elim : is_allowed_elimination Σ idecl.(ind_kelim) ps) (ind_inst : ctx_inst (typing Σ) Γ (p.(pparams) ++ indices) (List.rev (subst_instance p.(puinst) - (ind_params mdecl ,,, ind_indices idecl)))) + (ind_params mdecl ,,, ind_indices idecl : context)))) (not_cofinite : isCoFinite mdecl.(ind_finite) = false). Variant case_branch_typing `{checker_flags} wf_local_fun typing Σ Γ (ci:case_info) p ps mdecl idecl ptm brs := @@ -682,7 +682,7 @@ Lemma typing_ind_env_app_size `{cf : checker_flags} : PΓ Σ (Γ ,,, predctx) -> is_allowed_elimination Σ idecl.(ind_kelim) ps -> ctx_inst (Prop_conj typing P Σ) Γ (p.(pparams) ++ indices) - (List.rev (subst_instance p.(puinst) (mdecl.(ind_params) ,,, idecl.(ind_indices)))) -> + (List.rev (subst_instance p.(puinst) (mdecl.(ind_params) ,,, idecl.(ind_indices) : context))) -> Σ ;;; Γ |- c : mkApps (tInd ci.(ci_ind) p.(puinst)) (p.(pparams) ++ indices) -> P Σ Γ c (mkApps (tInd ci.(ci_ind) p.(puinst)) (p.(pparams) ++ indices)) -> isCoFinite mdecl.(ind_finite) = false -> @@ -906,7 +906,7 @@ Proof. { intros. eapply (X14 _ _ _ Hty). simpl; lia. } clear -X ind_inst. revert ind_inst X. - generalize (List.rev (subst_instance (puinst p) (ind_params mdecl ,,, ind_indices idecl))). + generalize (List.rev (subst_instance (puinst p) (ind_params mdecl ,,, ind_indices idecl : context))). generalize (pparams p ++ indices). intros l c ctxi IH; induction ctxi; constructor; eauto. * split; tas. eapply (IH _ _ _ t0); simpl; auto. lia. @@ -1103,7 +1103,7 @@ Lemma typing_ind_env `{cf : checker_flags} : PΓ Σ (Γ ,,, predctx) -> is_allowed_elimination Σ idecl.(ind_kelim) ps -> ctx_inst (Prop_conj typing P Σ) Γ (p.(pparams) ++ indices) - (List.rev (subst_instance p.(puinst) (mdecl.(ind_params) ,,, idecl.(ind_indices)))) -> + (List.rev (subst_instance p.(puinst) (mdecl.(ind_params) ,,, idecl.(ind_indices) : context))) -> Σ ;;; Γ |- c : mkApps (tInd ci.(ci_ind) p.(puinst)) (p.(pparams) ++ indices) -> P Σ Γ c (mkApps (tInd ci.(ci_ind) p.(puinst)) (p.(pparams) ++ indices)) -> isCoFinite mdecl.(ind_finite) = false -> @@ -1195,61 +1195,25 @@ Section All_local_env. apply IHX. simpl in *. lia. Qed. - Lemma lookup_on_global_env P (Σ : global_env) c decl : - on_global_env cumulSpec0 P Σ -> - lookup_env Σ c = Some decl -> - { Σ' : global_env & [× extends Σ' Σ, on_global_env cumulSpec0 P Σ' & - on_global_decl cumulSpec0 P (Σ', universes_decl_of_decl decl) c decl] }. - Proof using Type. - destruct Σ as [univs Σ]; rewrite /on_global_env /lookup_env; cbn. - intros [cu Σp]. - induction Σp; simpl. congruence. - destruct (eqb_specT c kn); subst. - - intros [= ->]. - exists ({| universes := univs; declarations := Σ |}). - split. - * red; cbn. split; [split;[lsets|csets]|]. - exists [(kn, decl)] => //. - * split => //. - * apply o0. - - intros hl. destruct (IHΣp hl) as [Σ' []]. - exists Σ'. - split=> //. - destruct e as [eu ed]. red; cbn in *. - split; [auto|]. - destruct ed as [Σ'' ->]. - exists (Σ'' ,, (kn, d)) => //. - Qed. - Lemma All_local_env_app (P : context -> term -> typ_or_sort -> Type) l l' : All_local_env P l * All_local_env (fun Γ t T => P (l ,,, Γ) t T) l' -> All_local_env P (l ,,, l'). - Proof using Type. - induction l'; simpl; auto. intuition. - intuition. destruct a. destruct decl_body. - inv b. econstructor; eauto. inv b; econstructor; eauto. - Qed. + Proof using Type. intros [H H']. apply All_local_rel_app => //. Qed. Lemma All_local_env_app_inv (P : context -> term -> typ_or_sort -> Type) l l' : All_local_env P (l ,,, l') -> All_local_env P l * All_local_env (fun Γ t T => P (l ,,, Γ) t T) l'. - Proof using Type. - apply All_local_app_rel. - Qed. + Proof using Type. apply All_local_app_rel. Qed. Definition wf_local_rel_app_inv {Σ Γ1 Γ2 Γ3} : wf_local_rel Σ Γ1 (Γ2 ,,, Γ3) -> wf_local_rel Σ Γ1 Γ2 * wf_local_rel Σ (Γ1 ,,, Γ2) Γ3. Proof. - intros h. apply All_local_env_app_inv in h as [h1 h2]. + intros h. apply All_local_app_rel in h as [h1 h2]. split. - exact h1. - - eapply All_local_env_impl. 1: exact h2. - intros Γ t [T|] h. - all: cbn in *. - all: change PCUICEnvironment.app_context with app_context in *. - all: rewrite <- app_context_assoc. - all: auto. + - eapply All_local_env_impl with (1 := h2). + intros; now rewrite -app_context_assoc. Defined. Lemma All_local_env_lookup {P Γ n} {decl} : @@ -1268,14 +1232,10 @@ Section All_local_env. wf_local_rel Σ Γ1 Γ2 -> wf_local_rel Σ (Γ1 ,,, Γ2) Γ3 -> wf_local_rel Σ Γ1 (Γ2 ,,, Γ3). Proof. - intros h1 h2. eapply All_local_env_app. - split. + intros h1 h2. eapply All_local_rel_app. - assumption. - - eapply All_local_env_impl. - + eassumption. - + change PCUICEnvironment.app_context with app_context. - intros Γ t []; cbn; - now rewrite app_context_assoc. + - apply All_local_env_impl with (1 := h2). + intros; now rewrite app_context_assoc. Defined. Definition wf_local_app {Σ Γ1 Γ2} : @@ -1419,3 +1379,8 @@ Section All_local_env. Qed. End All_local_env. + +Lemma declared_minductive_ind_npars {cf:checker_flags} {Σ} {wfΣ : wf Σ} {mdecl ind} : + declared_minductive Σ ind mdecl -> + ind_npars mdecl = context_assumptions mdecl.(ind_params). +Proof. apply (declared_minductive_ind_npars_gen (wfΣ := wfΣ)). Qed. diff --git a/pcuic/theories/PCUICWeakeningEnv.v b/pcuic/theories/PCUICWeakeningEnv.v index 954c66c65..4d85bb618 100644 --- a/pcuic/theories/PCUICWeakeningEnv.v +++ b/pcuic/theories/PCUICWeakeningEnv.v @@ -282,6 +282,61 @@ Proof using P Pcmp cf. Qed. Hint Resolve extends_lookup : extends. +Lemma extends_lookup_constant Σ Σ' kn decl : + wf Σ' -> extends Σ Σ' -> + lookup_constant Σ kn = Some decl -> + lookup_constant Σ' kn = Some decl. +Proof. + intros wfΣ' ext look. unfold lookup_constant in *. + destruct lookup_env eqn:H => //. + erewrite extends_lookup; tea. +Qed. +Hint Resolve extends_lookup_constant : extends. + +Lemma extends_lookup_minductive Σ Σ' mind mdecl : + wf Σ' -> extends Σ Σ' -> + lookup_minductive Σ mind = Some mdecl -> + lookup_minductive Σ' mind = Some mdecl. +Proof. + intros wfΣ' ext look. unfold lookup_minductive in *. + destruct lookup_env eqn:H => //. + erewrite extends_lookup; tea. +Qed. +Hint Resolve extends_lookup_minductive : extends. + +Lemma extends_lookup_inductive Σ Σ' ind midecl : + wf Σ' -> extends Σ Σ' -> + lookup_inductive Σ ind = Some midecl -> + lookup_inductive Σ' ind = Some midecl. +Proof. + intros wfΣ' ext look. unfold lookup_inductive in *. + destruct lookup_minductive eqn:H => //. + erewrite extends_lookup_minductive; tea. +Qed. +Hint Resolve extends_lookup_inductive : extends. + +Lemma extends_lookup_constructor Σ Σ' ind k micdecl : + wf Σ' -> extends Σ Σ' -> + lookup_constructor Σ ind k = Some micdecl -> + lookup_constructor Σ' ind k = Some micdecl. +Proof. + intros wfΣ' ext look. unfold lookup_constructor in *. + destruct lookup_inductive eqn:H => //. + erewrite extends_lookup_inductive; tea. +Qed. +Hint Resolve extends_lookup_constructor : extends. + +Lemma extends_lookup_projection Σ Σ' p micpdecl : + wf Σ' -> extends Σ Σ' -> + lookup_projection Σ p = Some micpdecl -> + lookup_projection Σ' p = Some micpdecl. +Proof. + intros wfΣ' ext look. unfold lookup_projection in *. + destruct lookup_constructor eqn:H => //. + erewrite extends_lookup_constructor; tea. +Qed. +Hint Resolve extends_lookup_projection : extends. + Lemma weakening_env_declared_constant : forall (Σ : global_env) cst (decl : constant_body), declared_constant Σ cst decl -> diff --git a/pcuic/theories/PCUICWfUniverses.v b/pcuic/theories/PCUICWfUniverses.v index 2f3fb20f8..454ae5428 100644 --- a/pcuic/theories/PCUICWfUniverses.v +++ b/pcuic/theories/PCUICWfUniverses.v @@ -709,7 +709,7 @@ Qed. Qed. - Lemma test_context_app p Γ Δ : + Lemma test_context_app p (Γ Δ : context) : test_context p (Γ ,,, Δ) = test_context p Γ && test_context p Δ. Proof using Type. induction Δ; simpl; auto. diff --git a/pcuic/theories/README.md b/pcuic/theories/README.md index a1be9d7c2..6e11765eb 100644 --- a/pcuic/theories/README.md +++ b/pcuic/theories/README.md @@ -167,7 +167,7 @@ | [PCUICWeakeningEnvConv] | Stability of conversion/cumulativity by global environment extension | [PCUICUnivSubstitutionConv] | Stability of conversion/cumulativity by substitution of universe variables | [PCUICClosedConv] | Helper lemmas on the closedness predicate -| [PCUICOnFreeVarsConv] | Helper lemmas for renamings and free variables +| [PCUICRenameTerm] | Helper lemmas for renamings and free variables [PCUICRenameConv]: ./Conversion/PCUICRenameConv.v [PCUICWeakeningConv]: ./Conversion/PCUICWeakeningConv.v @@ -175,7 +175,7 @@ [PCUICWeakeningEnvConv]: ./Conversion/PCUICWeakeningEnvConv.v [PCUICUnivSubstitutionConv]: ./Conversion/PCUICUnivSubstitutionConv.v [PCUICClosedConv]: ./Conversion/PCUICClosedConv.v -[PCUICOnFreeVarsConv]: ./Conversion/PCUICOnFreeVarsConv.v +[PCUICRenameTerm]: ./Conversion/PCUICRenameTerm.v ## Stability of Typing diff --git a/pcuic/theories/Syntax/PCUICCases.v b/pcuic/theories/Syntax/PCUICCases.v index dbc1ef550..9800bfc02 100644 --- a/pcuic/theories/Syntax/PCUICCases.v +++ b/pcuic/theories/Syntax/PCUICCases.v @@ -166,6 +166,24 @@ Lemma case_branch_type_fst ci mdecl idecl p br ptm c cdecl : (case_branch_context ci mdecl p (forget_types br.(bcontext)) cdecl). Proof. reflexivity. Qed. +Definition on_predicate P Pu Pctx Γ p := + [× All (P Γ) p.(pparams), Pu p.(puinst), Pctx p.(pcontext) & P (Γ ,,, inst_case_predicate_context p) p.(preturn)]. + +Definition on_predicate2 P Pu Pctx Γ (p p': predicate term) := + [× All2 (P Γ) p.(pparams) p'.(pparams), Pu p.(puinst) p'.(puinst), Pctx p.(pcontext) p'.(pcontext) & P (Γ ,,, inst_case_predicate_context p) p.(preturn) p'.(preturn)]. + +Definition on_branch P Pctx Γ p (br: branch term) := + Pctx br.(bcontext) × P (Γ ,,, inst_case_branch_context p br) br.(bbody). + +Definition on_branch2 P Pctx Γ p (br br': branch term) := + Pctx br.(bcontext) br'.(bcontext) × P (Γ ,,, inst_case_branch_context p br) br.(bbody) br'.(bbody). + +Definition on_branches P Pctx Γ p brs := + All (on_branch P Pctx Γ p) brs. + +Definition on_branches2 P Pctx Γ p brs brs' := + All2 (on_branch2 P Pctx Γ p) brs brs'. + (* Definition case_branches_types_gen ind mdecl idecl params puinst ptm : list (context * term) := mapi (case_branch_type_gen ind mdecl idecl params puinst ptm) idecl.(ind_ctors). diff --git a/pcuic/theories/Syntax/PCUICClosed.v b/pcuic/theories/Syntax/PCUICClosed.v index 693bcf2d7..7f970a044 100644 --- a/pcuic/theories/Syntax/PCUICClosed.v +++ b/pcuic/theories/Syntax/PCUICClosed.v @@ -42,7 +42,7 @@ Proof. rewrite /=. rewrite Nat.add_comm. bool_congr. Qed. -Lemma test_context_k_app p n Γ Γ' : +Lemma test_context_k_app p n (Γ Γ' : context) : test_context_k p n (Γ ,,, Γ') = test_context_k p n Γ && test_context_k p (n + #|Γ|) Γ'. Proof. diff --git a/pcuic/theories/Syntax/PCUICInstDef.v b/pcuic/theories/Syntax/PCUICInstDef.v index 44b579724..9950c858c 100644 --- a/pcuic/theories/Syntax/PCUICInstDef.v +++ b/pcuic/theories/Syntax/PCUICInstDef.v @@ -2,7 +2,7 @@ From Coq Require Import Morphisms. From MetaCoq.Template Require Import config utils. From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICCases PCUICInduction - PCUICLiftSubst PCUICUnivSubst + PCUICLiftSubst PCUICUnivSubst PCUICRelevance PCUICTyping PCUICEquality PCUICOnFreeVars PCUICSigmaCalculus PCUICRenameDef. @@ -60,6 +60,11 @@ Definition closed_subst (Γ : context) σ (Δ : context) := (forall x decl, nth_error Γ x = Some decl -> is_open_term Δ (σ x)) × usubst Γ σ Δ. +(* Substitution accounting relevance marks for reduction / cumulativity *) +Definition valid_subst Σ Γ σ Δ := + closed_subst Γ σ Δ × + (forall x decl, nth_error Γ x = Some decl -> isTermRel Σ (marks_of_context Δ) (σ x) decl.(decl_name).(binder_relevance)). + (* Well-typedness of a substitution *) Definition well_subst {cf} Σ (Γ : context) σ (Δ : context) := diff --git a/pcuic/theories/Syntax/PCUICPosition.v b/pcuic/theories/Syntax/PCUICPosition.v index 4e62c7a09..27d191dc8 100644 --- a/pcuic/theories/Syntax/PCUICPosition.v +++ b/pcuic/theories/Syntax/PCUICPosition.v @@ -208,50 +208,50 @@ Proof. + dependent destruction e. simpl in *. destruct (nth_error (pparams p0) n) as [par|] eqn:enth. 2: discriminate. destruct e. - induction a0 in n, par, enth, ih, vp |- *. 1: rewrite enth. 1: assumption. + induction a in n, par, enth, ih, vp |- *. 1: rewrite enth. 1: assumption. destruct n. * simpl in *. apply some_inj in enth. subst. intuition eauto. - * simpl in *. eapply IHa0. all: eauto. + * simpl in *. eapply IHa. all: eauto. + dependent destruction e. simpl in *. eapply ih; eauto. apply e. + dependent destruction e. simpl in *. destruct nth_error eqn:nth; [|congruence]. - eapply All2_nth_error_Some in a; eauto. - destruct a as (?&->&_&eq). + eapply All2_nth_error_Some in e1; eauto. + destruct e1 as (?&->&_&eq). eauto. + dependent destruction e. simpl in *. - destruct (nth_error mfix n) as [[na ty bo ra]|] eqn:e. 2: discriminate. - induction a in n, na, ty, bo, ra, e, ih, vp |- *. - 1:{ rewrite e. assumption. } + destruct (nth_error mfix n) as [[na ty bo ra]|] eqn:E. 2: discriminate. + induction e in n, na, ty, bo, ra, E, ih, vp |- *. + 1:{ rewrite E. assumption. } destruct n. - * simpl in *. apply some_inj in e. subst. + * simpl in *. apply some_inj in E. subst. destruct y as [na' ty' bo' ra']. simpl in *. intuition eauto. - * simpl in *. eapply IHa. all: eauto. + * simpl in *. eapply IHe. all: eauto. + dependent destruction e. simpl in *. - destruct (nth_error mfix n) as [[na ty bo ra]|] eqn:e. 2: discriminate. - induction a in n, na, ty, bo, ra, e, ih, vp |- *. - 1:{ rewrite e. assumption. } + destruct (nth_error mfix n) as [[na ty bo ra]|] eqn:E. 2: discriminate. + induction e in n, na, ty, bo, ra, E, ih, vp |- *. + 1:{ rewrite E. assumption. } destruct n. - * simpl in *. apply some_inj in e. subst. + * simpl in *. apply some_inj in E. subst. destruct y as [na' ty' bo' ra']. simpl in *. intuition eauto. - * simpl in *. eapply IHa. all: eauto. + * simpl in *. eapply IHe. all: eauto. + dependent destruction e. simpl in *. - destruct (nth_error mfix n) as [[na ty bo ra]|] eqn:e. 2: discriminate. - induction a in n, na, ty, bo, ra, e, ih, vp |- *. - 1:{ rewrite e. assumption. } + destruct (nth_error mfix n) as [[na ty bo ra]|] eqn:E. 2: discriminate. + induction e in n, na, ty, bo, ra, E, ih, vp |- *. + 1:{ rewrite E. assumption. } destruct n. - * simpl in *. apply some_inj in e. subst. + * simpl in *. apply some_inj in E. subst. destruct y as [na' ty' bo' ra']. simpl in *. intuition eauto. - * simpl in *. eapply IHa. all: eauto. + * simpl in *. eapply IHe. all: eauto. + dependent destruction e. simpl in *. - destruct (nth_error mfix n) as [[na ty bo ra]|] eqn:e. 2: discriminate. - induction a in n, na, ty, bo, ra, e, ih, vp |- *. - 1:{ rewrite e. assumption. } + destruct (nth_error mfix n) as [[na ty bo ra]|] eqn:E. 2: discriminate. + induction e in n, na, ty, bo, ra, E, ih, vp |- *. + 1:{ rewrite E. assumption. } destruct n. - * simpl in *. apply some_inj in e. subst. + * simpl in *. apply some_inj in E. subst. destruct y as [na' ty' bo' ra']. simpl in *. intuition eauto. - * simpl in *. eapply IHa. all: eauto. + * simpl in *. eapply IHe. all: eauto. Qed. Lemma eq_term_valid_pos : diff --git a/pcuic/theories/TemplateToPCUICCorrectness.v b/pcuic/theories/TemplateToPCUICCorrectness.v index 33db00d38..5c63d5ad4 100644 --- a/pcuic/theories/TemplateToPCUICCorrectness.v +++ b/pcuic/theories/TemplateToPCUICCorrectness.v @@ -879,8 +879,8 @@ Section Trans_Global. | ih : forall Rle u, _ |- _ => now eapply ih end - ]. - unfold trans_predicate, eq_predicate; cbn. + ]; + unfold trans_predicate, eq_predicate, trans_branch, eq_branches, eq_branch; cbn. repeat split; solve_all; try typeclasses eauto with typeclass_instances core. * rewrite map2_map2_bias_left; len. eapply All2_length in hpctx. len in hpctx. @@ -1137,7 +1137,7 @@ Section Trans_Global. (fun (x : aname) (y : context_decl) => set_binder_name x (map_decl (subst_instance (Ast.puinst p)) y))). rewrite -PCUICUnivSubstitutionConv.map2_map_r. - rewrite -[fold_context_k _ _]PCUICRenameConv.map2_set_binder_name_fold; len. + rewrite -[fold_context_k _ _]PCUICRenameTerm.map2_set_binder_name_fold; len. rewrite /trans_local map_map2 map2_trans. rewrite -PCUICUnivSubstitutionConv.map2_map_r. f_equal. rewrite [map _ _]trans_subst_context map_rev. @@ -1310,7 +1310,7 @@ Section Trans_Global. now rewrite trans_subst_context map_rev trans_local_subst_instance. Qed. - Lemma trans_local_app Γ Δ : trans_local Σ' (Ast.Env.app_context Γ Δ) = trans_local Σ' Γ ,,, trans_local Σ' Δ. + Lemma trans_local_app Γ Δ : trans_local Σ' (Γ ,,, Δ) = trans_local Σ' Γ ,,, trans_local Σ' Δ. Proof using Σ. now rewrite /trans_local map_app. Qed. @@ -1634,7 +1634,7 @@ Section Trans_Global. red. rewrite <- !map_dtype. rewrite <- !map_dbody. split. destruct b. intuition eauto. - + unfold Ast.Env.app_context, trans_local in b1. + + unfold app_context, trans_local in b1. simpl in a. rewrite -> map_app in b1. unfold app_context. unfold ST.fix_context in b1. rewrite map_rev map_mapi in b1. simpl in b1. @@ -1668,7 +1668,7 @@ Section Trans_Global. apply (OnOne2_All_mix_left Hwf) in X. solve_all. destruct b. red. rewrite <- !map_dtype. rewrite <- !map_dbody. intuition eauto. - + unfold Ast.Env.app_context, trans_local in b1. + + unfold app_context, trans_local in b1. simpl in a. rewrite -> map_app in b1. unfold app_context. unfold ST.fix_context in b1. rewrite map_rev map_mapi in b1. simpl in b1. @@ -2591,10 +2591,10 @@ Lemma trans_cumul_ctx_rel {cf} {Σ : Ast.Env.global_env_ext} Γ Δ Δ' : let Σ' := trans_global Σ in Typing.wf Σ -> wf Σ' -> ST.TemplateConversion.cumul_ctx_rel Typing.cumul_gen Σ Γ Δ Δ' -> - closed_ctx (trans_local Σ' (Ast.Env.app_context Γ Δ)) -> - closed_ctx (trans_local Σ' (Ast.Env.app_context Γ Δ')) -> - All (WfAst.wf_decl Σ) (Ast.Env.app_context Γ Δ) -> - All (WfAst.wf_decl Σ) (Ast.Env.app_context Γ Δ') -> + closed_ctx (trans_local Σ' (Γ ,,, Δ)) -> + closed_ctx (trans_local Σ' (Γ ,,, Δ')) -> + All (WfAst.wf_decl Σ) (Γ ,,, Δ) -> + All (WfAst.wf_decl Σ) (Γ ,,, Δ') -> cumul_ctx_rel cumulSpec0 Σ' (trans_local Σ' Γ) (trans_local Σ' Δ) (trans_local Σ' Δ'). Proof. intros Σ' wfΣ wfΣ'. @@ -2702,17 +2702,13 @@ Lemma trans_cstr_respects_variance {cf} Σ mdecl v cdecl : All (WfAst.wf Σ) (Ast.Env.cstr_indices cdecl) -> (closed_ctx (trans_local (trans_global_env Σ) - (Ast.Env.app_context - (Ast.Env.app_context (ST.ind_arities mdecl) - (Ast.Env.smash_context [] (Ast.Env.ind_params mdecl))) + (ST.ind_arities mdecl ,,, Ast.Env.smash_context [] (Ast.Env.ind_params mdecl) ,,, (Ast.Env.expand_lets_ctx (Ast.Env.ind_params mdecl) (Ast.Env.smash_context [] (Ast.Env.cstr_args cdecl)))))) -> All (fun x : Ast.term => closedn (Ast.Env.context_assumptions (Ast.Env.cstr_args cdecl) + Ast.Env.context_assumptions (Ast.Env.ind_params mdecl) + #|ST.ind_arities mdecl|) - (trans Σ' (Ast.Env.expand_lets - (Ast.Env.app_context (Ast.Env.ind_params mdecl) - (Ast.Env.cstr_args cdecl)) x))) + (trans Σ' (Ast.Env.expand_lets (Ast.Env.ind_params mdecl ,,, Ast.Env.cstr_args cdecl) x))) (Ast.Env.cstr_indices cdecl) -> ST.cstr_respects_variance Typing.cumul_gen Σ mdecl v cdecl -> cstr_respects_variance cumulSpec0 Σ' mdecl' v cdecl'. @@ -2801,9 +2797,7 @@ Lemma trans_ind_respects_variance {cf} Σ mdecl v idecl : wf Σ' -> wf_inductive_body Σ idecl -> All (WfAst.wf_decl Σ) (Ast.Env.ind_params mdecl) -> - closed_ctx (trans_local (trans_global_env Σ) - (Ast.Env.app_context (Ast.Env.ind_params mdecl) - (Ast.Env.ind_indices idecl))) -> + closed_ctx (trans_local (trans_global_env Σ) (Ast.Env.ind_params mdecl ,,, Ast.Env.ind_indices idecl)) -> ST.ind_respects_variance Typing.cumul_gen Σ mdecl v (Ast.Env.ind_indices idecl) -> ind_respects_variance cumulSpec0 Σ' mdecl' v (ind_indices idecl'). Proof. @@ -3088,10 +3082,7 @@ Proof. induction (Ast.Env.cstr_args x) in cs |- *; destruct cs; simpl; auto; destruct a as [na [b|] ty]; simpl in *; auto; split; intuition eauto; - specialize (foo - (Ast.Env.app_context (Ast.Env.app_context - (Ast.Env.arities_context (Ast.Env.ind_bodies m)) (Ast.Env.ind_params m)) - c)); + 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 eapply (foo ty (SortRel (binder_relevance na))). now apply (foo b (Typ ty)). @@ -3106,9 +3097,7 @@ Proof. rewrite -trans_arities_context. induction 1; simpl; constructor; auto; have foo := (X (Σg, Ast.Env.ind_universes m) _ _ _ X0); - specialize (foo (Ast.Env.app_context (Ast.Env.app_context - (Ast.Env.arities_context (Ast.Env.ind_bodies m)) - (Ast.Env.ind_params m)) (Ast.Env.cstr_args x))); + specialize (foo (Ast.Env.arities_context (Ast.Env.ind_bodies m) ,,, Ast.Env.ind_params m ,,, Ast.Env.cstr_args x)); rewrite /trans_local !map_app in foo. now apply (foo i (Typ t)). now rewrite (trans_subst_telescope _ [i] 0) in IHon_cindices. @@ -3208,7 +3197,7 @@ Proof. unfold ST.on_projection, on_projection. cbn -[inds Σg]. rewrite context_assumptions_map. rewrite -[trans_local _ _ ++ _]trans_local_app -(trans_smash_context _ []) nth_error_map. - rewrite /Ast.Env.app_context. destruct nth_error => // /=. + rewrite /app_context. destruct nth_error => // /=. rewrite /trans_projection_body /=. move=> [] /= hb -> ->. cbn -[Σg]. split => //. rewrite trans_subst trans_inds. cbn -[Σg]. f_equal. diff --git a/pcuic/theories/Typing/PCUICClosedTyp.v b/pcuic/theories/Typing/PCUICClosedTyp.v index 3e2abb80d..05c889ece 100644 --- a/pcuic/theories/Typing/PCUICClosedTyp.v +++ b/pcuic/theories/Typing/PCUICClosedTyp.v @@ -272,7 +272,7 @@ Proof. destruct x; simpl in *. unfold test_def. simpl. rtoProp. split. - rewrite -> app_context_length in *. rewrite -> Nat.add_comm in *. + rewrite -> app_context_length in H, H3. rewrite -> Nat.add_comm in *. eapply closedn_lift_inv in H3; eauto. lia. subst types. now rewrite app_context_length fix_context_length in H. @@ -285,7 +285,7 @@ Proof. split. destruct a as (s & e & Hs & cl). now rewrite andb_true_r in cl. - rewrite -> app_context_length in *. rewrite -> Nat.add_comm in *. + rewrite -> app_context_length in H3, H4. rewrite -> Nat.add_comm in *. subst types. now rewrite fix_context_length in H3. eapply nth_error_all in X0; eauto. destruct X0 as (s & e & Hs & cl). diff --git a/pcuic/theories/Typing/PCUICInstTyp.v b/pcuic/theories/Typing/PCUICInstTyp.v index a74da646c..b21a7717c 100644 --- a/pcuic/theories/Typing/PCUICInstTyp.v +++ b/pcuic/theories/Typing/PCUICInstTyp.v @@ -6,7 +6,7 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICTactics PCUICCases PCUICTyping PCUICReduction PCUICCumulativity PCUICEquality PCUICGlobalEnv PCUICClosed PCUICClosedConv PCUICClosedTyp PCUICEquality PCUICWeakeningEnvConv PCUICWeakeningEnvTyp PCUICSigmaCalculus PCUICRenameDef PCUICRenameConv PCUICWeakeningConv PCUICWeakeningTyp PCUICInstDef PCUICInstConv - PCUICGuardCondition PCUICUnivSubstitutionConv PCUICOnFreeVars PCUICOnFreeVarsConv PCUICClosedTyp PCUICClosedTyp. + PCUICGuardCondition PCUICUnivSubstitutionConv PCUICOnFreeVars PCUICRenameTerm PCUICClosedTyp PCUICClosedTyp. Require Import ssreflect ssrbool. From Equations Require Import Equations. @@ -189,7 +189,7 @@ Proof. eapply inst_is_open_term; eauto. +++ rewrite map_length. rewrite inst_context_on_free_vars ; eauto. ++ unfold PCUICCases.inst_case_predicate_context. - apply on_free_vars_ctx_inst_case_context; eauto. + apply on_free_vars_ctx_inst_case_context_app; eauto. ++ unfold PCUICCases.inst_case_predicate_context. unfold is_open_term. rewrite app_length. rewrite <- shiftnP_add. @@ -223,7 +223,7 @@ Proof. eapply inst_is_open_term; eauto. +++ rewrite map_length. tea. + unfold PCUICCases.inst_case_predicate_context. - apply on_free_vars_ctx_inst_case_context; eauto. + apply on_free_vars_ctx_inst_case_context_app; eauto. + unfold PCUICCases.inst_case_predicate_context. unfold is_open_term. rewrite app_length. rewrite <- shiftnP_add. diff --git a/pcuic/theories/Typing/PCUICNamelessTyp.v b/pcuic/theories/Typing/PCUICNamelessTyp.v index 3064c9caa..e277737d9 100644 --- a/pcuic/theories/Typing/PCUICNamelessTyp.v +++ b/pcuic/theories/Typing/PCUICNamelessTyp.v @@ -4,7 +4,7 @@ From MetaCoq.Template Require Import config utils. From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICInduction PCUICLiftSubst PCUICEquality PCUICReduction PCUICTyping PCUICPosition PCUICUnivSubst PCUICNamelessDef PCUICGuardCondition PCUICNamelessConv PCUICConversion - PCUICWellScopedCumulativity PCUICOnFreeVars PCUICOnFreeVarsConv PCUICConfluence PCUICClosedTyp PCUICClosed + PCUICWellScopedCumulativity PCUICOnFreeVars PCUICRenameTerm PCUICConfluence PCUICClosedTyp PCUICClosed PCUICSigmaCalculus (* for context manipulations *). Require Import Equations.Prop.DepElim. Require Import ssreflect ssrbool. diff --git a/pcuic/theories/Typing/PCUICRenameTyp.v b/pcuic/theories/Typing/PCUICRenameTyp.v index 7f3261fd9..57818c4f8 100644 --- a/pcuic/theories/Typing/PCUICRenameTyp.v +++ b/pcuic/theories/Typing/PCUICRenameTyp.v @@ -4,7 +4,7 @@ From MetaCoq.Template Require Import config utils. From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICTactics PCUICCases PCUICInduction PCUICLiftSubst PCUICUnivSubst PCUICCumulativity PCUICTyping PCUICReduction PCUICGlobalEnv PCUICClosed PCUICEquality PCUICRenameDef - PCUICSigmaCalculus PCUICClosed PCUICOnFreeVars PCUICOnFreeVarsConv PCUICGuardCondition PCUICTyping + PCUICSigmaCalculus PCUICClosed PCUICOnFreeVars PCUICRenameTerm PCUICGuardCondition PCUICTyping PCUICWeakeningEnvConv PCUICWeakeningEnvTyp PCUICClosedConv PCUICClosedTyp PCUICRenameConv. Require Import ssreflect ssrbool. @@ -264,7 +264,7 @@ Proof. apply urenaming_context; eauto. ++ rewrite test_context_k_closed_on_free_vars_ctx in Hcontext. unfold inst_case_predicate_context. - apply on_free_vars_ctx_inst_case_context; eauto. + apply on_free_vars_ctx_inst_case_context_app; eauto. ++ unfold inst_case_predicate_context. unfold is_open_term. rewrite app_length. rewrite <- shiftnP_add. @@ -276,7 +276,7 @@ Proof. rewrite inst_case_predicate_context_length. rewrite (All2_fold_length Xcontext). eassumption. ++ rewrite test_context_k_closed_on_free_vars_ctx in Hcontext. - unfold inst_case_predicate_context. apply on_free_vars_ctx_inst_case_context; eauto. + unfold inst_case_predicate_context. apply on_free_vars_ctx_inst_case_context_app; eauto. +++ eapply All_forallb. apply All_map. apply forallb_All in Hp; eapply All_impl. 1: tea. cbn; intros. eapply urename_is_open_term; eauto. +++ unfold pparams. cbn. rewrite map_length. exact Hcontext. @@ -299,7 +299,7 @@ Proof. apply urenaming_context; eauto. + rewrite test_context_k_closed_on_free_vars_ctx in Hx. unfold inst_case_predicate_context. - apply on_free_vars_ctx_inst_case_context; eauto. + apply on_free_vars_ctx_inst_case_context_app; eauto. + unfold inst_case_predicate_context. unfold is_open_term. rewrite app_length. rewrite <- shiftnP_add. @@ -311,7 +311,7 @@ Proof. rewrite inst_case_branch_context_length. rewrite (All2_fold_length Hbcontext). eassumption. + rewrite test_context_k_closed_on_free_vars_ctx in Hcontext. - unfold inst_case_predicate_context. apply on_free_vars_ctx_inst_case_context; eauto. + unfold inst_case_predicate_context. apply on_free_vars_ctx_inst_case_context_app; eauto. ++ eapply All_forallb. apply All_map. apply forallb_All in Hp; eapply All_impl. 1: tea. cbn; intros. eapply urename_is_open_term; eauto. ++ unfold pparams. rewrite test_context_k_closed_on_free_vars_ctx in Hx. diff --git a/pcuic/theories/Typing/PCUICWeakeningTyp.v b/pcuic/theories/Typing/PCUICWeakeningTyp.v index bd736b87d..634ca702d 100644 --- a/pcuic/theories/Typing/PCUICWeakeningTyp.v +++ b/pcuic/theories/Typing/PCUICWeakeningTyp.v @@ -4,7 +4,7 @@ From MetaCoq.Template Require Import config utils. From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICCases PCUICInduction PCUICLiftSubst PCUICUnivSubst PCUICTyping PCUICCumulativity PCUICClosed - PCUICSigmaCalculus PCUICRenameDef PCUICRenameConv PCUICRenameTyp PCUICOnFreeVars + PCUICSigmaCalculus PCUICRenameDef PCUICRenameTerm PCUICRenameTyp PCUICOnFreeVars PCUICClosedConv PCUICClosedTyp PCUICWeakeningConv. Require Import ssreflect ssrbool. diff --git a/pcuic/theories/utils/PCUICAstUtils.v b/pcuic/theories/utils/PCUICAstUtils.v index b91b6b3b7..33101fc82 100644 --- a/pcuic/theories/utils/PCUICAstUtils.v +++ b/pcuic/theories/utils/PCUICAstUtils.v @@ -1,5 +1,5 @@ (* Distributed under the terms of the MIT license. *) -From MetaCoq.Template Require Import utils uGraph Reflect. +From MetaCoq.Template Require Import utils Reflect. From MetaCoq.PCUIC Require Import PCUICAst PCUICSize. Require Import ssreflect. diff --git a/safechecker/theories/PCUICEqualityDec.v b/safechecker/theories/PCUICEqualityDec.v index 8a74754d3..adbdde3c4 100644 --- a/safechecker/theories/PCUICEqualityDec.v +++ b/safechecker/theories/PCUICEqualityDec.v @@ -657,7 +657,7 @@ Proof. 2:{ constructor. intro bot. apply f. inversion bot. subst. inversion X3. subst. constructor; auto. } constructor. inversion e2. subst. - constructor ; auto. + constructor; unfold eq_branches, eq_branch in X3 |- *; auto. - cbn - [eqb]. eqspecs. equspec equ he. equspec lequ hle. ih => //. constructor. constructor ; assumption. @@ -686,12 +686,12 @@ Proof. } cbn - [eqb]. eqspecs. 2:{ constructor. intro bot. inversion bot. subst. - apply n. inversion X0. subst. destruct X2 as [[? ?] ?]. + apply n. inversion X0. subst. destruct X2 as (? & ? & ? & ?). assumption. } destruct (eqb_annot_reflect (dname a) (dname d)). 2:{ constructor. intro bot; inversion bot. subst. - apply n. inversion X0. subst. destruct X2 as [[? ?] ?]. + apply n. inversion X0. subst. destruct X2 as (? & ? & ? & ?). assumption. } cbn - [eqb]. destruct (IHm X1 mfix) => //. @@ -725,12 +725,12 @@ Proof. } cbn - [eqb]. eqspecs. 2:{ constructor. intro bot. inversion bot. subst. - apply n. inversion X0. subst. destruct X2 as [[? ?] ?]. + apply n. inversion X0. subst. destruct X2 as (? & ? & ? & ?). assumption. } destruct (eqb_annot_reflect (dname a) (dname d)). 2:{ constructor. intro bot; inversion bot. subst. - apply n. inversion X0. subst. destruct X2 as [[? ?] ?]. + apply n. inversion X0. subst. destruct X2 as (? & ? & ? & ?). assumption. } cbn - [eqb]. destruct (IHm X1 mfix) => //. @@ -920,7 +920,7 @@ Proof. apply forallb_Forall in H. eapply Forall_impl; eauto. move => ? /wf_universe_reflect ?; eauto. + eapply onctx_eq_ctx in a0; eauto. - - eapply forallb_All in wt; eapply All_mix in X0; try apply wt. + - unfold eq_branches, eq_branch. eapply forallb_All in wt; eapply All_mix in X0; try apply wt. clear wt. eapply All_All2; eauto; simpl; intuition eauto. + eapply onctx_eq_ctx in a0; eauto. + eapply b0; eauto. apply andb_and in a as [? a]. exact a. diff --git a/safechecker/theories/PCUICSafeConversion.v b/safechecker/theories/PCUICSafeConversion.v index 6a9c48d6f..a51b821b4 100644 --- a/safechecker/theories/PCUICSafeConversion.v +++ b/safechecker/theories/PCUICSafeConversion.v @@ -1942,7 +1942,7 @@ Qed. split; auto. rewrite e; cbn. rewrite nth_error_app_ge; auto. now rewrite Nat.sub_diag; cbn. } rewrite (wf_predicate_length_pars wf_pred). - now rewrite (PCUICGlobalEnv.declared_minductive_ind_npars decli). + now rewrite (PCUICTyping.declared_minductive_ind_npars decli). } rewrite test_context_k_closed_on_free_vars_ctx. destruct hcase. @@ -1972,7 +1972,7 @@ Qed. split; auto. rewrite e; cbn. rewrite nth_error_app_ge; auto. now rewrite Nat.sub_diag; cbn. } rewrite (wf_predicate_length_pars wf_pred). - now rewrite (PCUICGlobalEnv.declared_minductive_ind_npars decli). + now rewrite (PCUICTyping.declared_minductive_ind_npars decli). Unshelve. all: eauto. Qed. @@ -3066,7 +3066,7 @@ Qed. { eapply (closed_ind_predicate_context); tea. eapply declared_minductive_closed; tea. exact decli'. } rewrite (wf_predicate_length_pars wf_pred). - now rewrite (PCUICGlobalEnv.declared_minductive_ind_npars decli). + now rewrite (PCUICTyping.declared_minductive_ind_npars decli). } rewrite test_context_k_closed_on_free_vars_ctx. exists mdecl, idecl. @@ -3089,7 +3089,7 @@ Qed. { eapply (closed_ind_predicate_context); tea. eapply declared_minductive_closed; tea. exact decli'. } rewrite (wf_predicate_length_pars wf_pred0). - now rewrite (PCUICGlobalEnv.declared_minductive_ind_npars decli). + now rewrite (PCUICTyping.declared_minductive_ind_npars decli). Qed. Definition forallb2_proper A B (R R' : A -> B -> bool) l l': diff --git a/safechecker/theories/PCUICWfReduction.v b/safechecker/theories/PCUICWfReduction.v index 6e37b0b2d..9584bab0e 100644 --- a/safechecker/theories/PCUICWfReduction.v +++ b/safechecker/theories/PCUICWfReduction.v @@ -149,10 +149,10 @@ Section fix_sigma. intros r. generalize_eqs r. intros ->. revert t ts. induction r. - - intros t ts ->. + - intros t ts e. rewrite e in r. destruct (term_subterm_red1 r) as [t' [[red1 [ts' Hts']]]]. exists t'; split; auto. split; auto. now constructor. exists ts'; auto. - - intros t ts ->. specialize (IHr1 t ts eq_refl) as [t' [[yt' [ts' Hts]]]]. + - intros t ts e. rewrite e in r1, r2. specialize (IHr1 t ts e) as [t' [[yt' [ts' Hts]]]]. specialize (IHr2 t' ts'). forward IHr2. now rewrite Hts. destruct IHr2 as [t'' [[zt' [ts'' Hts'']]]]. diff --git a/template-coq/theories/BasicAst.v b/template-coq/theories/BasicAst.v index dc0a91aa9..9c8717cfe 100644 --- a/template-coq/theories/BasicAst.v +++ b/template-coq/theories/BasicAst.v @@ -305,6 +305,11 @@ Definition snoc {A} (Γ : list A) (d : A) := d :: Γ. Notation " Γ ,, d " := (snoc Γ d) (at level 20, d at next level). +Definition app_context {A} (Γ Γ': list A) := Γ' ++ Γ. + +Notation "Γ ,,, Γ'" := (app_context Γ Γ') (at level 25, Γ' at next level, left associativity). + + Definition ondecl {A} (P : A -> Type) (d : context_decl A) := P d.(decl_type) × option_default P d.(decl_body) unit. diff --git a/template-coq/theories/Environment.v b/template-coq/theories/Environment.v index ae0879d50..eef08fe8f 100644 --- a/template-coq/theories/Environment.v +++ b/template-coq/theories/Environment.v @@ -373,12 +373,6 @@ Module Environment (T : Term). Definition program : Type := global_env * term. - (* TODO MOVE AstUtils factorisation *) - - Definition app_context (Γ Γ' : context) : context := Γ' ++ Γ. - Notation "Γ ,,, Γ'" := - (app_context Γ Γ') (at level 25, Γ' at next level, left associativity). - (** Make a lambda/let-in string of abstractions from a context [Γ], ending with term [t]. *) Definition mkLambda_or_LetIn d t := @@ -499,26 +493,29 @@ Module Environment (T : Term). Proof. unfold arities_context. now rewrite rev_map_length. Qed. #[global] Hint Rewrite arities_context_length : len. - Lemma app_context_nil_l Γ : [] ,,, Γ = Γ. + Lemma app_context_nil_l {T} (Γ: list T) : [] ,,, Γ = Γ. Proof. unfold app_context. rewrite app_nil_r. reflexivity. Qed. - Lemma app_context_assoc Γ Γ' Γ'' : Γ ,,, (Γ' ,,, Γ'') = Γ ,,, Γ' ,,, Γ''. + Lemma app_context_assoc {T} (Γ Γ' Γ'': list T) : Γ ,,, (Γ' ,,, Γ'') = Γ ,,, Γ' ,,, Γ''. Proof. unfold app_context; now rewrite app_assoc. Qed. - Lemma app_context_cons Γ Γ' A : Γ ,,, (Γ' ,, A) = (Γ ,,, Γ') ,, A. + Lemma app_context_cons (Γ Γ': context) A : Γ ,,, (Γ' ,, A) = (Γ ,,, Γ') ,, A. Proof. exact (app_context_assoc _ _ [A]). Qed. - Lemma app_context_length Γ Γ' : #|Γ ,,, Γ'| = #|Γ'| + #|Γ|. + Lemma app_context_length {T} (Γ Γ': list T) : #|Γ ,,, Γ'| = #|Γ'| + #|Γ|. Proof. unfold app_context. now rewrite app_length. Qed. - #[global] Hint Rewrite app_context_length : len. + + Lemma app_context_length' (Γ Γ': context) : #|Γ ,,, Γ'| = #|Γ'| + #|Γ|. + Proof. apply app_context_length. Qed. + #[global] Hint Rewrite app_context_length' : len. - Lemma nth_error_app_context_ge v Γ Γ' : + Lemma nth_error_app_context_ge {T} v (Γ Γ': list T) : #|Γ'| <= v -> nth_error (Γ ,,, Γ') v = nth_error Γ (v - #|Γ'|). Proof. apply nth_error_app_ge. Qed. - Lemma nth_error_app_context_lt v Γ Γ' : + Lemma nth_error_app_context_lt {T} v (Γ Γ': list T) : v < #|Γ'| -> nth_error (Γ ,,, Γ') v = nth_error Γ' v. Proof. apply nth_error_app_lt. Qed. diff --git a/template-coq/theories/common/uGraph.v b/template-coq/theories/common/uGraph.v index f915c8dc5..53e566145 100644 --- a/template-coq/theories/common/uGraph.v +++ b/template-coq/theories/common/uGraph.v @@ -7,11 +7,6 @@ Import ConstraintType. Import MCMonadNotation. - -Arguments Z.add : simpl nomatch. -Arguments Nat.leb : simpl nomatch. -Arguments Nat.eqb : simpl nomatch. - Definition Z_of_bool (b : bool) : Z := match b with | true => 1 diff --git a/template-coq/theories/utils.v b/template-coq/theories/utils.v index 16e27e7f3..c1f72bdf1 100644 --- a/template-coq/theories/utils.v +++ b/template-coq/theories/utils.v @@ -30,3 +30,7 @@ Notation "A * B" := (prod A B) : type_scope2. Global Open Scope type_scope2. Global Open Scope metacoq_scope. + +Arguments Z.add : simpl nomatch. +Arguments Nat.leb : simpl nomatch. +Arguments Nat.eqb : simpl nomatch. From 6bacff97363d05387a44fc5df3055cb74ff55e86 Mon Sep 17 00:00:00 2001 From: Yann Leray Date: Tue, 17 May 2022 16:31:32 +0200 Subject: [PATCH 05/17] Compatibility between relevance from syntax and relevance from type --- pcuic/_CoqProject.in | 1 + pcuic/theories/Bidirectional/BDFromPCUIC.v | 3 +- .../theories/Bidirectional/BDStrengthening.v | 2 +- pcuic/theories/Bidirectional/BDTyping.v | 2 + pcuic/theories/Bidirectional/BDUnique.v | 2 +- pcuic/theories/PCUICAlpha.v | 2 +- pcuic/theories/PCUICCumulProp.v | 258 +++++-------- pcuic/theories/PCUICElimination.v | 146 +++----- pcuic/theories/PCUICExpandLetsCorrectness.v | 17 +- pcuic/theories/PCUICInversion.v | 1 + .../PCUICParallelReductionConfluence.v | 6 +- pcuic/theories/PCUICPrincipality.v | 55 ++- pcuic/theories/PCUICProgress.v | 4 +- pcuic/theories/PCUICRelevance.v | 6 +- pcuic/theories/PCUICRelevanceTerm.v | 4 +- pcuic/theories/PCUICRelevanceTyp.v | 344 ++++++++++++++++++ pcuic/theories/PCUICSR.v | 20 +- pcuic/theories/PCUICSpine.v | 4 +- pcuic/theories/PCUICSubstitution.v | 20 +- pcuic/theories/PCUICToTemplateCorrectness.v | 7 +- pcuic/theories/PCUICTyping.v | 20 +- pcuic/theories/PCUICValidity.v | 21 +- pcuic/theories/PCUICWcbvEval.v | 4 +- pcuic/theories/PCUICWfUniverses.v | 21 +- pcuic/theories/TemplateToPCUICCorrectness.v | 30 +- pcuic/theories/TemplateToPCUICExpanded.v | 5 +- pcuic/theories/Typing/PCUICClosedTyp.v | 54 +-- pcuic/theories/Typing/PCUICInstTyp.v | 2 +- pcuic/theories/Typing/PCUICRenameTyp.v | 2 +- .../Typing/PCUICUnivSubstitutionTyp.v | 3 +- pcuic/theories/Typing/PCUICWeakeningEnvTyp.v | 18 +- pcuic/theories/Typing/PCUICWeakeningTyp.v | 15 +- template-coq/theories/Environment.v | 2 +- template-coq/theories/EnvironmentTyping.v | 82 ++++- template-coq/theories/EtaExpand.v | 10 +- template-coq/theories/Typing.v | 19 +- template-coq/theories/TypingWf.v | 17 +- template-coq/theories/Universes.v | 55 +-- template-coq/theories/common/uGraph.v | 2 +- 39 files changed, 788 insertions(+), 498 deletions(-) create mode 100644 pcuic/theories/PCUICRelevanceTyp.v diff --git a/pcuic/_CoqProject.in b/pcuic/_CoqProject.in index 6849a6cb6..a96b60531 100644 --- a/pcuic/_CoqProject.in +++ b/pcuic/_CoqProject.in @@ -110,6 +110,7 @@ theories/Bidirectional/BDStrengthening.v theories/PCUICWeakeningEnv.v theories/PCUICRelevance.v theories/PCUICRelevanceTerm.v +theories/PCUICRelevanceTyp.v # theories/All.v diff --git a/pcuic/theories/Bidirectional/BDFromPCUIC.v b/pcuic/theories/Bidirectional/BDFromPCUIC.v index 51e865279..d69536e61 100644 --- a/pcuic/theories/Bidirectional/BDFromPCUIC.v +++ b/pcuic/theories/Bidirectional/BDFromPCUIC.v @@ -217,7 +217,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 ? ? sortrel Hinst ty_c Cumc ? ? ? ty_br. apply conv_infer_sort in Cump as (?&?&?) ; auto. apply conv_infer_ind in Cumc as (?&?&[]) ; auto. @@ -228,6 +228,7 @@ Proof. * by eapply All_local_app_rel. * eapply is_allowed_elimination_monotone. all: eassumption. + * eapply geq_relevance; tea. * 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. diff --git a/pcuic/theories/Bidirectional/BDStrengthening.v b/pcuic/theories/Bidirectional/BDStrengthening.v index 8c21f92ee..488516a3c 100644 --- a/pcuic/theories/Bidirectional/BDStrengthening.v +++ b/pcuic/theories/Bidirectional/BDStrengthening.v @@ -752,7 +752,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/BDTyping.v b/pcuic/theories/Bidirectional/BDTyping.v index f99a9d3c9..e873e5b1b 100644 --- a/pcuic/theories/Bidirectional/BDTyping.v +++ b/pcuic/theories/Bidirectional/BDTyping.v @@ -85,6 +85,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 -> @@ -400,6 +401,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 857760357..eb10e01a3 100644 --- a/pcuic/theories/Bidirectional/BDUnique.v +++ b/pcuic/theories/Bidirectional/BDUnique.v @@ -193,7 +193,7 @@ Proof using wfΣ. - intros ? T' ty_T'. inversion ty_T' ; subst. - 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. diff --git a/pcuic/theories/PCUICAlpha.v b/pcuic/theories/PCUICAlpha.v index deb0de897..94df50c65 100644 --- a/pcuic/theories/PCUICAlpha.v +++ b/pcuic/theories/PCUICAlpha.v @@ -690,7 +690,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 sortrel HIHctxi Hc IHc iscof ptm wfbrs Hbrs Δ v e e'; invs e. have eqp := X1. eassert (ctx_inst _ _ _ _) as Hctxi by now eapply ctx_inst_impl with (2 := HIHctxi). eassert (PCUICEnvTyping.ctx_inst _ _ _ _) as IHctxi. diff --git a/pcuic/theories/PCUICCumulProp.v b/pcuic/theories/PCUICCumulProp.v index 86c939378..460b3ff50 100644 --- a/pcuic/theories/PCUICCumulProp.v +++ b/pcuic/theories/PCUICCumulProp.v @@ -22,8 +22,6 @@ Implicit Types (Σ : global_env_ext). Section no_prop_leq_type. Context `{cf : checker_flags}. -Variable Hcf : prop_sub_type = false. -Variable Hcf' : check_univs. Lemma cumul_sort_confluence {Σ} {wfΣ : wf Σ} {Γ A u v} : Σ ;;; Γ ⊢ A ≤ tSort u -> @@ -83,30 +81,30 @@ Proof using Type. now eapply wf_local_closed_context. Qed. -Lemma is_prop_bottom {Σ Γ T s s'} : - wf_ext Σ -> +Lemma is_prop_bottom {Σ Γ T s s'} (Hcf : prop_sub_type = false) : + wf Σ -> Σ ;;; Γ ⊢ T ≤ tSort s -> Σ ;;; Γ ⊢ T ≤ tSort s' -> Universe.is_prop s -> Universe.is_prop s'. -Proof using Hcf Hcf'. +Proof using Type. intros wfΣ hs hs'. destruct (cumul_sort_confluence hs hs') as [x' [conv [leq leq']]]. intros isp. eapply leq_universe_prop_r in leq; eauto. - unshelve eapply (leq_universe_prop_no_prop_sub_type _ _ _ _ _ _ leq'); eauto. + unshelve eapply (leq_universe_prop_no_prop_sub_type _ _ _ _ leq'); eauto. Qed. Lemma is_sprop_bottom {Σ Γ T s s'} : - wf_ext Σ -> + wf Σ -> Σ ;;; Γ ⊢ T ≤ tSort s -> Σ ;;; Γ ⊢ T ≤ tSort s' -> Universe.is_sprop s -> Universe.is_sprop s'. -Proof using Hcf'. +Proof using Type. intros wfΣ hs hs'. destruct (cumul_sort_confluence hs hs') as [x' [conv [leq leq']]]. intros isp. eapply leq_universe_sprop_r in leq; eauto. - unshelve eapply (leq_universe_sprop_l _ _ _ _ _ leq'); eauto. + unshelve eapply (leq_universe_sprop_l _ _ _ leq'); eauto. Qed. Lemma prop_sort_eq {Σ Γ u u'} : Universe.is_prop u -> Universe.is_prop u' -> @@ -140,33 +138,7 @@ Proof using Type. now depelim eqvv'. Qed. -Lemma is_prop_superE {Σ l} : wf_ext Σ -> Universe.is_prop (Universe.super l) -> False. -Proof using Hcf'. - intros wfΣ. - eapply is_prop_gt; eauto. - eapply leq_universe_refl. -Qed. - -Lemma is_sprop_superE {Σ l} : wf_ext Σ -> Universe.is_sprop (Universe.super l) -> False. -Proof using Type. - intros wfΣ. destruct l => //. -Qed. - -Lemma is_prop_prod {s s'} : Universe.is_prop s' -> Universe.is_prop (Universe.sort_of_product s s'). -Proof using Type. - intros isp. - unfold Universe.sort_of_product. rewrite isp. auto. -Qed. - -Lemma is_sprop_prod {s s'} : Universe.is_sprop s' -> Universe.is_sprop (Universe.sort_of_product s s'). -Proof using Type. - intros isp. - unfold Universe.sort_of_product. rewrite isp orb_true_r. auto. -Qed. - -Definition eq_univ_prop (u v : Universe.t) := - (Universe.is_prop u <-> Universe.is_prop v) /\ - (Universe.is_sprop u <-> Universe.is_sprop v). +Definition eq_univ_prop := eq_universe_ prop_sub_type (fun (_ : CS.t) _ _ => True) CS.empty. Definition eq_term_prop (Σ : global_env) napp := PCUICEquality.eq_term_upto_univ_napp Σ eq_univ_prop eq_univ_prop napp. @@ -195,102 +167,67 @@ Inductive cumul_prop `{checker_flags} (Σ : global_env_ext) (Γ : context) : ter where " Σ ;;; Γ |- t ~~ u " := (cumul_prop Σ Γ t u) : type_scope. Lemma eq_term_prop_impl Σ Re Rle t u : - wf_ext Σ -> forall n, PCUICEquality.eq_term_upto_univ_napp Σ.1 Re Rle n t u -> subrelation Re eq_univ_prop -> subrelation Rle eq_univ_prop -> eq_term_prop Σ n t u. Proof using Type. - intros wfΣ n eq. + intros n eq. intros. eapply PCUICEquality.eq_term_upto_univ_impl in eq. eauto. all:auto. Qed. -Lemma leq_universe_prop_spec Σ u1 u2 : - check_univs -> - wf_ext Σ -> - leq_universe Σ u1 u2 -> - match u1, u2 with - | Universe.lProp, Universe.lProp => True - | Universe.lSProp, Universe.lSProp => True - | Universe.lProp, Universe.lSProp => False - | Universe.lSProp, Universe.lProp => False - | Universe.lProp, Universe.lType _ => prop_sub_type - | Universe.lSProp, Universe.lType _ => False - | Universe.lType l, Universe.lType l' => True - | Universe.lType _, _ => False - end. -Proof using Type. - intros cu wf leq. - apply wf_ext_consistent in wf. - apply (leq_universe_props _ _ _ cu wf leq). -Qed. - Lemma subrelation_eq_universe_eq_prop Σ : - wf_ext Σ -> subrelation (eq_universe Σ) eq_univ_prop. -Proof using Hcf Hcf'. - intros wfΣ x y eq'. red. - split; intros. - eapply eq_universe_leq_universe in eq'. - eapply leq_universe_prop_spec in eq'; auto. - destruct x, y; simpl in *; auto; cong. - eapply eq_universe_leq_universe in eq'. - eapply leq_universe_prop_spec in eq'; auto. - destruct x, y; simpl in *; auto; cong. +Proof using Type. + intros s s' eq'. red. + destruct s, s' => //. Qed. Lemma subrelation_leq_universe_eq_prop Σ : - wf_ext Σ -> subrelation (leq_universe Σ) eq_univ_prop. -Proof using Hcf Hcf'. - intros wfΣ x y eq'. red. - split; intros; - eapply leq_universe_prop_spec in eq'; auto; - destruct x, y; simpl in *; auto; cong. +Proof using Type. + intros s s' eq'. red. + destruct s, s' => //. Qed. Lemma eq_term_eq_term_prop_impl Σ t u : - wf_ext Σ -> forall n, PCUICEquality.eq_term_upto_univ_napp Σ.1 (eq_universe Σ) (eq_universe Σ) n t u -> eq_term_prop Σ n t u. -Proof using Hcf Hcf'. - intros wfΣ n eq. eapply eq_term_prop_impl; eauto. +Proof using Type. + intros n eq. eapply eq_term_prop_impl; eauto. now apply subrelation_eq_universe_eq_prop. now apply subrelation_eq_universe_eq_prop. Qed. Lemma leq_term_eq_term_prop_impl Σ t u : - wf_ext Σ -> forall n, PCUICEquality.eq_term_upto_univ_napp Σ.1 (eq_universe Σ) (leq_universe Σ) n t u -> eq_term_prop Σ n t u. -Proof using Hcf Hcf'. - intros wfΣ n eq. eapply eq_term_prop_impl; eauto. +Proof using Type. + intros n eq. eapply eq_term_prop_impl; eauto. now apply subrelation_eq_universe_eq_prop. now apply subrelation_leq_universe_eq_prop. Qed. Lemma cumul_cumul_prop Σ Γ A B : - wf_ext Σ -> Σ ;;; Γ ⊢ A ≤ B -> Σ ;;; Γ |- A ~~ B. -Proof using Hcf Hcf'. - intros wfΣ. induction 1. +Proof using Type. + induction 1. - constructor => //. now apply leq_term_eq_term_prop_impl in c. - econstructor 2; eauto. - econstructor 3; eauto. Qed. Lemma conv_cumul_prop Σ Γ A B : - wf_ext Σ -> Σ ;;; Γ ⊢ A = B -> Σ ;;; Γ |- A ~~ B. -Proof using Hcf Hcf'. - intros wfΣ. induction 1. +Proof using Type. + induction 1. - constructor => //. now apply eq_term_eq_term_prop_impl in c. - econstructor 2; eauto. - econstructor 3; eauto. @@ -326,16 +263,26 @@ Proof using Type. econstructor 2; eauto. Qed. -Lemma cumul_prop_props {Σ Γ u u'} {wfΣ : wf Σ}: - Universe.is_prop u -> +Lemma cumul_sort_inv {Σ Γ u u'} {wfΣ : wf Σ} : Σ ;;; Γ |- tSort u ~~ tSort u' -> - Universe.is_prop u'. + eq_univ_prop u u'. Proof using Type. - intros isp equiv. + intros equiv. eapply cumul_prop_alt in equiv as [nf [nf' [redl redr eq]]]. eapply invert_red_sort in redl. apply invert_red_sort in redr. subst. - depelim eq. red in e. intuition auto. + depelim eq. assumption. +Qed. + +Lemma cumul_prop_props {Σ Γ u u'} {wfΣ : wf Σ} (Hcf : prop_sub_type = false): + Universe.is_prop u -> + Σ ;;; Γ |- tSort u ~~ tSort u' -> + Universe.is_prop u'. +Proof using Type. + intros isp e. + apply cumul_sort_inv in e. + red in e. rewrite Hcf in e. + destruct u, u' => //. Qed. Lemma cumul_sprop_props {Σ Γ u u'} {wfΣ : wf Σ} : @@ -343,26 +290,24 @@ Lemma cumul_sprop_props {Σ Γ u u'} {wfΣ : wf Σ} : Σ ;;; Γ |- tSort u ~~ tSort u' -> Universe.is_sprop u'. Proof using Type. - intros isp equiv. - eapply cumul_prop_alt in equiv as [nf [nf' [redl redr eq]]]. - eapply invert_red_sort in redl. apply invert_red_sort in redr. - subst. - depelim eq. red in e. intuition auto. + intros isp e. + apply cumul_sort_inv in e. + destruct u, u' => //. Qed. Instance refl_eq_univ_prop : RelationClasses.Reflexive eq_univ_prop. Proof using Type. - intros x. red. intuition. + now intros []. Qed. Instance sym_eq_univ_prop : RelationClasses.Symmetric eq_univ_prop. Proof using Type. - intros x y; unfold eq_univ_prop; intuition. + now intros [] []. Qed. Instance trans_eq_univ_prop : RelationClasses.Transitive eq_univ_prop. Proof using Type. - intros x y; unfold eq_univ_prop; intuition. + now intros [] [] []. Qed. Lemma LevelExprSet_For_all (P : LevelExpr.t -> Prop) (u : LevelAlgExpr.t) : @@ -383,7 +328,7 @@ Proof using Type. now eapply InA_In_eq. Qed. -Lemma univ_epxrs_elements_map g s : +Lemma univ_exprs_elements_map g s : forall e, In e (LevelExprSet.elements (NonEmptySetFacts.map g s)) <-> In e (map g (LevelExprSet.elements s)). Proof using Type. @@ -430,7 +375,7 @@ Qed. Lemma univ_is_prop_make x : Universe.is_prop (Universe.make x) = false. Proof using Type. - destruct x; simpl; auto. + reflexivity. Qed. (* Lemma is_prop_subst_level_expr u1 u2 s : @@ -451,11 +396,7 @@ Qed. *) Instance substuniv_eq_univ_prop : SubstUnivPreserving eq_univ_prop. Proof using Type. - intros s u1 u2 hu. - red in hu. - eapply Forall2_map_inv in hu. - rewrite /subst_instance_univ. - destruct s; red; simpl; auto; try intuition reflexivity. + intros [] u1 u2 hu; cbnr. Qed. Lemma cumul_prop_sym Σ Γ T U : @@ -494,43 +435,43 @@ Qed. Global Instance cumul_prop_transitive Σ Γ : wf Σ -> CRelationClasses.Transitive (cumul_prop Σ Γ). Proof using Type. intros. red. intros. now eapply cumul_prop_trans. Qed. -Lemma cumul_prop_cum_l {Σ Γ A T B} {wfΣ : wf_ext Σ} : +Lemma cumul_prop_cum_l {Σ Γ A T B} {wfΣ : wf Σ} : Σ ;;; Γ |- A ~~ T -> Σ ;;; Γ ⊢ A ≤ B -> Σ ;;; Γ |- B ~~ T. -Proof using Hcf Hcf'. +Proof using Type. intros HT cum. eapply cumul_cumul_prop in cum; auto. eapply CRelationClasses.transitivity ; eauto. eapply cumul_prop_sym; eauto. Qed. -Lemma cumul_prop_cum_r {Σ Γ A T B} {wfΣ : wf_ext Σ} : +Lemma cumul_prop_cum_r {Σ Γ A T B} {wfΣ : wf Σ} : Σ ;;; Γ |- A ~~ T -> Σ ;;; Γ ⊢ B ≤ A -> Σ ;;; Γ |- B ~~ T. -Proof using Hcf Hcf'. +Proof using Type. intros HT cum. eapply cumul_cumul_prop in cum; auto. eapply CRelationClasses.transitivity ; eauto. Qed. -Lemma cumul_prop_conv_l {Σ Γ A T B} {wfΣ : wf_ext Σ} : +Lemma cumul_prop_conv_l {Σ Γ A T B} {wfΣ : wf Σ} : Σ ;;; Γ |- A ~~ T -> Σ ;;; Γ ⊢ A = B -> Σ ;;; Γ |- B ~~ T. -Proof using Hcf Hcf'. +Proof using Type. intros HT cum. eapply conv_cumul_prop in cum; auto. eapply CRelationClasses.transitivity ; eauto. eapply cumul_prop_sym; eauto. Qed. -Lemma cumul_prop_conv_r {Σ Γ A T B} {wfΣ : wf_ext Σ} : +Lemma cumul_prop_conv_r {Σ Γ A T B} {wfΣ : wf Σ} : Σ ;;; Γ |- A ~~ T -> Σ ;;; Γ ⊢ B = A -> Σ ;;; Γ |- B ~~ T. -Proof using Hcf Hcf'. +Proof using Type. intros HT cum. eapply conv_cumul_prop in cum; auto. eapply CRelationClasses.transitivity ; eauto. @@ -685,7 +626,7 @@ Proof using Type. reflexivity. Qed. -Lemma cumul_prop_args {Σ Γ args args'} {wfΣ : wf_ext Σ} : +Lemma cumul_prop_args {Σ Γ args args'} {wfΣ : wf Σ} : All2 (cumul_prop Σ Γ) args args' -> ∑ nf nf', [× All2 (closed_red Σ Γ) args nf, All2 (closed_red Σ Γ) args' nf' & All2 (eq_term_prop Σ 0) nf nf']. @@ -709,31 +650,31 @@ Proof using Type. now rewrite closedP_on_free_vars. Qed. -Lemma red_conv_prop {Σ Γ T U} {wfΣ : wf_ext Σ} : +Lemma red_conv_prop {Σ Γ T U} {wfΣ : wf Σ} : Σ ;;; Γ ⊢ T ⇝ U -> Σ ;;; Γ |- T ~~ U. -Proof using Hcf Hcf'. +Proof using Type. move/(red_ws_cumul_pb (pb:=Conv)). now apply conv_cumul_prop. Qed. -Lemma substitution_red_terms_conv_prop {Σ Γ Δ Γ' s s' M} {wfΣ : wf_ext Σ} : +Lemma substitution_red_terms_conv_prop {Σ Γ Δ Γ' s s' M} {wfΣ : wf Σ} : is_closed_context (Γ ,,, Δ ,,, Γ') -> is_open_term (Γ ,,, Δ ,,, Γ') M -> untyped_subslet Γ s Δ -> red_terms Σ Γ s s' -> Σ ;;; (Γ ,,, subst_context s 0 Γ') |- (subst s #|Γ'| M) ~~ (subst s' #|Γ'| M). -Proof using Hcf Hcf'. +Proof using Type. intros. apply red_conv_prop. eapply closed_red_red_subst; tea. Qed. -Lemma context_conversion_cumul_prop {Σ Γ Δ M N} {wfΣ : wf_ext Σ} : +Lemma context_conversion_cumul_prop {Σ Γ Δ M N} {wfΣ : wf Σ} : Σ ;;; Γ |- M ~~ N -> Σ ⊢ Γ = Δ -> Σ ;;; Δ |- M ~~ N. -Proof using Hcf Hcf'. +Proof using Type. induction 1; intros. - constructor => //. eauto with fvs. now rewrite -(All2_fold_length X). now rewrite -(All2_fold_length X). @@ -759,7 +700,7 @@ Qed. harder as it requires a more involved proof about reduction being "preserved" when converting contexts using cumul_prop rather than standard conversion. *) -Lemma substitution_untyped_cumul_prop_cumul {Σ Γ Δ Δ' s s' M} {wfΣ : wf_ext Σ} : +Lemma substitution_untyped_cumul_prop_cumul {Σ Γ Δ Δ' s s' M} {wfΣ : wf Σ} : is_closed_context (Γ ,,, Δ) -> is_closed_context (Γ ,,, Δ') -> is_open_term (Γ ,,, Δ) M -> @@ -767,7 +708,7 @@ Lemma substitution_untyped_cumul_prop_cumul {Σ Γ Δ Δ' s s' M} {wfΣ : wf_ext untyped_subslet Γ s' Δ' -> All2 (cumul_prop Σ Γ) s s' -> Σ ;;; Γ |- subst0 s M ~~ subst0 s' M. -Proof using Hcf Hcf'. +Proof using Type. intros clctx clctx' clM subs subs' Heq. assert (lens' := All2_length Heq). destruct (cumul_prop_args Heq) as (nf & nf' & [redl redr eq]) => //. @@ -830,8 +771,7 @@ Proof using Type. eapply All2_Forall2, All2_map. unfold subst_instance. eapply All2_map. eapply All2_refl. - intros x. red. - rewrite !is_prop_subst_instance_level /=. split; reflexivity. + intros x. reflexivity. Qed. Lemma cumul_prop_subst_instance {Σ Γ univs u u' T} {wfΣ : wf Σ} : @@ -906,12 +846,12 @@ Qed. Hint Resolve conv_ctx_prop_refl : core. -Lemma cumul_prop_tProd {Σ : global_env_ext} {Γ na t ty na' t' ty'} {wfΣ : wf_ext Σ} : +Lemma cumul_prop_tProd {Σ : global_env_ext} {Γ na t ty na' t' ty'} {wfΣ : wf Σ} : eq_binder_annot na na' -> eq_term Σ.1 Σ t t' -> Σ ;;; Γ ,, vass na t |- ty ~~ ty' -> Σ ;;; Γ |- tProd na t ty ~~ tProd na' t' ty'. -Proof using Hcf Hcf'. +Proof using Type. intros eqann eq cum. eapply cumul_prop_alt in cum as (nf & nf' & [redl redr eq']). eapply cumul_prop_alt. eexists (tProd na t nf), (tProd na' t' nf'); split; eauto. @@ -929,13 +869,13 @@ Proof using Hcf Hcf'. eapply eq_term_eq_term_prop_impl; auto. Qed. -Lemma cumul_prop_tLetIn (Σ : global_env_ext) {Γ na t d ty na' t' d' ty'} {wfΣ : wf_ext Σ} : +Lemma cumul_prop_tLetIn (Σ : global_env_ext) {Γ na t d ty na' t' d' ty'} {wfΣ : wf Σ} : eq_binder_annot na na' -> eq_term Σ.1 Σ t t' -> eq_term Σ.1 Σ d d' -> Σ ;;; Γ ,, vdef na d t |- ty ~~ ty' -> Σ ;;; Γ |- tLetIn na d t ty ~~ tLetIn na' d' t' ty'. -Proof using Hcf Hcf'. +Proof using Type. intros eqann eq eq' cum. eapply cumul_prop_alt in cum as (nf & nf' & [redl redr eq'']). eapply cumul_prop_alt. @@ -958,14 +898,14 @@ Proof using Hcf Hcf'. transitivity nf'. auto. now eapply eq_term_eq_term_prop_impl. Qed. -Lemma cumul_prop_mkApps {Σ Γ f args f' args'} {wfΣ : wf_ext Σ} : +Lemma cumul_prop_mkApps {Σ Γ f args f' args'} {wfΣ : wf Σ} : is_closed_context Γ -> is_open_term Γ f -> is_open_term Γ f' -> eq_term Σ.1 Σ f f' -> All2 (cumul_prop Σ Γ) args args' -> Σ ;;; Γ |- mkApps f args ~~ mkApps f' args'. -Proof using Hcf Hcf'. +Proof using Type. intros clΓ clf clf' eq eq'. eapply cumul_prop_alt. eapply cumul_prop_args in eq' as (nf & nf' & [redl redr eq']). @@ -987,7 +927,7 @@ Proof using Type. split; fvs. eapply closed_red_refl; fvs. apply eq_term_upto_univ_refl; typeclasses eauto. Qed. -Lemma eq_term_prop_mkApps_inv {Σ ind u args ind' u' args'} {wfΣ : wf_ext Σ} : +Lemma eq_term_prop_mkApps_inv {Σ ind u args ind' u' args'} {wfΣ : wf Σ} : forall n, eq_term_prop Σ n (mkApps (tInd ind u) args) (mkApps (tInd ind' u') args') -> All2 (eq_term_prop Σ 0) args args'. Proof using Type. @@ -1005,7 +945,7 @@ Proof using Type. red. apply H0. Qed. -Lemma cumul_prop_mkApps_Ind_inv {Σ Γ ind u args ind' u' args'} {wfΣ : wf_ext Σ} : +Lemma cumul_prop_mkApps_Ind_inv {Σ Γ ind u args ind' u' args'} {wfΣ : wf Σ} : Σ ;;; Γ |- mkApps (tInd ind u) args ~~ mkApps (tInd ind' u') args' -> All2 (cumul_prop Σ Γ) args args'. Proof using Type. @@ -1077,17 +1017,15 @@ Qed. Lemma typing_leq_term_prop (Σ : global_env_ext) Γ t t' T T' : wf Σ.1 -> Σ ;;; Γ |- t : T -> - on_udecl Σ.1 Σ.2 -> Σ ;;; Γ |- t' : T' -> forall n, leq_term_napp Σ n t' t -> Σ ;;; Γ |- T ~~ T'. -Proof using Hcf Hcf'. +Proof using Type. intros wfΣ Ht. revert Σ wfΣ Γ t T Ht t' T'. eapply (typing_ind_env (fun Σ Γ t T => forall t' T' : term, - on_udecl Σ.1 Σ.2 -> Σ;;; Γ |- t' : T' -> forall n, leq_term_napp Σ n t' t -> Σ ;;; Γ |- T ~~ T')%type @@ -1095,15 +1033,15 @@ Proof using Hcf Hcf'. 1-13:match goal with [ H : leq_term_napp _ _ _ _ |- _ ] => depelim H - end; assert (wf_ext Σ) by (split; assumption). + end. - 14:{ assert (wf_ext Σ) by (split; assumption). specialize (X1 _ _ H X5 _ X6). + 14:{ specialize (X1 _ _ X5 _ X6). eapply cumul_prop_cum_l; tea. eapply cumulSpec_cumulAlgo_curry in X4; tea; fvs. } 6:{ eapply inversion_App in X6 as (na' & A' & B' & hf & ha & cum); auto. - specialize (X3 _ _ H hf _ X7_1). - specialize (X5 _ _ H ha _ (eq_term_upto_univ_napp_leq X7_2)). + specialize (X3 _ _ hf _ X7_1). + specialize (X5 _ _ ha _ (eq_term_upto_univ_napp_leq X7_2)). eapply cumul_cumul_prop in cum; auto. transitivity (B' {0 := u0}) => //. eapply cumul_prop_prod_inv in X3 => //. @@ -1128,30 +1066,30 @@ Proof using Hcf Hcf'. apply cumul_cumul_prop in Hs => //. eapply cumul_prop_trans; eauto. destruct (cumul_prop_is_open Hs) as []. - constructor => //. constructor. symmetry. - split; split; intros H'. 1,2:now eapply is_prop_superE in H'. - 1,2:now eapply is_sprop_superE in H'. + constructor => //. constructor. + now destruct u, s. - eapply inversion_Prod in X4 as (s1' & s2' & e' & Ha & Hb & Hs); auto. - specialize (X1 _ _ H0 Ha). + specialize (X1 _ _ Ha). specialize (X1 _ (eq_term_upto_univ_napp_leq X5_1)). eapply context_conversion in Hb. 3:{ constructor. apply conv_ctx_refl. constructor. eassumption. constructor. eauto. } 2:{ constructor; eauto. now exists s1. } - specialize (X3 _ _ H0 Hb _ X5_2). + specialize (X3 _ _ Hb _ X5_2). eapply cumul_cumul_prop in Hs => //. eapply cumul_prop_trans; eauto. constructor; fvs. constructor. - split. - * split; intros Hs'; apply is_prop_sort_prod in Hs'; eapply is_prop_prod; eapply cumul_prop_props; eauto. - now eapply cumul_prop_sym; eauto. - * split; intros Hs'; apply is_sprop_sort_prod in Hs'; eapply is_sprop_prod; eapply cumul_sprop_props; eauto. - now eapply cumul_prop_sym; eauto. + apply cumul_sort_inv in X1, X3. + clear -X1 X3. + all: destruct s2; cbn. + all: destruct s2'; trivial. + all: destruct s1; trivial. + all: destruct s1'; trivial. - eapply inversion_Lambda in X4 as (s & B & e' & dom & bod & cum). - specialize (X1 _ _ H0 dom _ (eq_term_upto_univ_napp_leq X5_1)). - specialize (X3 t0 B H0). + specialize (X1 _ _ dom _ (eq_term_upto_univ_napp_leq X5_1)). + specialize (X3 t0 B). assert(conv_context cumulAlgo_gen Σ (Γ ,, vass na0 ty) (Γ ,, vass na t)). { repeat constructor; pcuic. } forward X3 by eapply context_conversion; eauto; pcuic. @@ -1160,11 +1098,11 @@ Proof using Hcf Hcf'. eapply cumul_prop_tProd; eauto. now symmetry. now symmetry. auto. - eapply inversion_LetIn in X6 as (s1' & A & e' & dom & bod & codom & cum); auto. - specialize (X1 _ _ H0 dom _ (eq_term_upto_univ_napp_leq X7_2)). - specialize (X3 _ _ H0 bod _ (eq_term_upto_univ_napp_leq X7_1)). + specialize (X1 _ _ dom _ (eq_term_upto_univ_napp_leq X7_2)). + specialize (X3 _ _ bod _ (eq_term_upto_univ_napp_leq X7_1)). assert(conv_context cumulAlgo_gen Σ (Γ ,, vdef na0 t ty) (Γ ,, vdef na b b_ty)). { repeat constructor; pcuic. } - specialize (X5 u A H0). + specialize (X5 u A). forward X5 by eapply context_conversion; eauto; pcuic. specialize (X5 _ X7_3). eapply cumul_cumul_prop in cum; eauto. @@ -1228,7 +1166,7 @@ Proof using Hcf Hcf'. clear X8. destruct (declared_inductive_inj isdecl isdecl'). subst. destruct data. - specialize (X7 _ _ H5 scrut_ty _ (eq_term_upto_univ_napp_leq X10)). + specialize (X7 _ _ scrut_ty _ (eq_term_upto_univ_napp_leq X10)). eapply cumul_prop_sym => //. destruct e as [eqpars [eqinst [eqpctx eqpret]]]. rewrite /ptm. @@ -1256,7 +1194,7 @@ Proof using Hcf Hcf'. - eapply inversion_Proj in X3 as (u' & mdecl' & idecl' & cdecl' & pdecl' & args' & inv); auto. intuition auto. - specialize (X2 _ _ H0 a0 _ (eq_term_upto_univ_napp_leq X4)). + specialize (X2 _ _ a0 _ (eq_term_upto_univ_napp_leq X4)). eapply eq_term_upto_univ_napp_leq in X4. eapply cumul_cumul_prop in b; eauto. eapply cumul_prop_trans; eauto. @@ -1271,10 +1209,10 @@ Proof using Hcf Hcf'. * cbn -[projection_context on_free_vars_ctx]. eapply is_closed_context_weaken; tas. fvs. now eapply wf_local_closed_context in X3. * cbn -[projection_context on_free_vars_ctx]. - eapply is_closed_context_weaken; tas. fvs. now eapply wf_local_closed_context in X6. + eapply is_closed_context_weaken; tas. fvs. now eapply wf_local_closed_context in X5. * epose proof (declared_projection_closed a). len. rewrite on_free_vars_subst_instance. simpl; len. - rewrite (declared_minductive_ind_npars a) in H1. + rewrite (declared_minductive_ind_npars a) in H0. rewrite closedn_on_free_vars //. eapply closed_upwards; tea. lia. * epose proof (projection_subslet Σ _ _ _ _ _ _ _ _ _ isdecl wfΣ X1 (validity X1)). now eapply subslet_untyped_subslet. @@ -1289,7 +1227,7 @@ Proof using Hcf Hcf'. cbn -[projection_context on_free_vars_ctx]; eapply is_closed_context_weaken => //; fvs. epose proof (declared_projection_closed a). simpl; len. - rewrite (declared_minductive_ind_npars a) in H1. + rewrite (declared_minductive_ind_npars a) in H0. rewrite closedn_on_free_vars //. eapply closed_upwards; tea. lia. - eapply inversion_Fix in X2 as (decl' & fixguard' & Hnth & types' & bodies & wffix & cum); auto. diff --git a/pcuic/theories/PCUICElimination.v b/pcuic/theories/PCUICElimination.v index 598f1b593..229ec6622 100644 --- a/pcuic/theories/PCUICElimination.v +++ b/pcuic/theories/PCUICElimination.v @@ -45,7 +45,7 @@ Definition Informative `{cf : checker_flags} (Σ : global_env_ext) (ind : induct forall mdecl idecl, declared_inductive (fst Σ) ind mdecl idecl -> forall Γ args u n (Σ' : global_env_ext), - wf_ext Σ' -> + wf Σ' -> extends_decls Σ Σ' -> Is_proof Σ' Γ (mkApps (tConstruct ind n u) args) -> #|ind_ctors idecl| <= 1 /\ @@ -119,23 +119,20 @@ Qed. Lemma elim_restriction_works_kelim1 {cf : checker_flags} {Σ : global_env_ext} {Γ T ci p c brs mdecl idecl} : - check_univs -> - wf_ext Σ -> + wf Σ.1 -> declared_inductive Σ ci.(ci_ind) mdecl idecl -> Σ ;;; Γ |- tCase ci p c brs : T -> (Is_proof Σ Γ (tCase ci p c brs) -> False) -> ind_kelim idecl = IntoAny \/ ind_kelim idecl = IntoSetPropSProp. Proof. - intros cu wfΣ. intros. + intros wfΣ **. assert (HT := X). eapply inversion_Case in X as [mdecl' [idecl' [isdecl' [indices [data cum]]]]]; eauto. destruct data. eapply declared_inductive_inj in isdecl' as []. 2:exact H. subst. enough (~ (Universe.is_prop ps \/ Universe.is_sprop ps)). - { clear -cu wfΣ allowed_elim H1. - apply wf_ext_consistent in wfΣ as (val&sat). + { clear -allowed_elim H1. unfold is_allowed_elimination, is_lSet, eq_universe, eq_levelalg in *. - rewrite cu in allowed_elim. destruct (ind_kelim idecl); auto; destruct ps; cbn in *; try discriminate; intuition congruence. } intros Huf. apply H0. @@ -146,10 +143,9 @@ Proof. repeat split; auto. split; auto. split; auto. clear brs_ty. eapply type_mkApps. rewrite /ptm. eapply type_it_mkLambda_or_LetIn; tea. - assert (wf Σ) by apply wfΣ. pose proof (PCUICInductiveInversion.isType_mkApps_Ind_smash H (validity scrut_ty)). - forward X1. apply (wf_predicate_length_pars wf_pred). - simpl in X1. destruct X1 as [sppars [spargs cu']]. + forward X0. apply (wf_predicate_length_pars wf_pred). + simpl in X0. destruct X0 as [sppars [spargs cu']]. assert (eqctx' : All2 (PCUICEquality.compare_decls eq eq) (Γ,,, case_predicate_context' ci mdecl idecl p) (Γ,,, predctx)). @@ -165,8 +161,8 @@ Proof. unshelve epose proof (typing_spine_case_predicate (ps:=ps) _ H cons _ sppars). 1-2:shelve. * pcuic. * now eapply PCUICWfUniverses.typing_wf_universe in pret_ty. - * rewrite -smash_context_subst_context_let_expand in X2. - specialize (X2 spargs scrut_ty). + * rewrite -smash_context_subst_context_let_expand in X1. + specialize (X1 spargs scrut_ty). eapply typing_spine_strengthen; tea; revgoals. + eapply ws_cumul_pb_it_mkProd_or_LetIn. eapply ws_cumul_ctx_pb_rel_app. @@ -271,29 +267,16 @@ Proof. simpl; unfold snoc; rewrite subst_context_snoc; econstructor; eauto. Qed. -Lemma wf_ext_wf {cf:checker_flags} Σ : wf_ext Σ -> wf Σ. -Proof. move=> wfe; apply wfe. Qed. -#[global] -Hint Resolve wf_ext_wf : core. - Lemma is_propositional_subst_instance u s : is_propositional (subst_instance_univ u s) = is_propositional s. Proof. destruct s => //. Qed. -Lemma leq_universe_propositional_l {cf:checker_flags} ϕ u1 u2 : - check_univs -> - prop_sub_type = false -> - consistent ϕ -> +(* TODO: Move to Universes.v *) +Lemma leq_universe_propositional_l {cf:checker_flags} (Hcf : prop_sub_type = false) ϕ u1 u2 : leq_universe ϕ u1 u2 -> is_propositional u1 -> u1 = u2. Proof. - intros Hcf ps cu le. - unfold is_propositional. - destruct (Universe.is_prop u1) eqn:eq. - apply leq_universe_prop_no_prop_sub_type in le; auto. - simpl. now destruct u1, u2. - simpl. intros sp. - apply leq_universe_sprop_l in le; auto. - now destruct u1, u2. + destruct u2 => //; destruct u1 => //. + intros le; hnf in le; congruence. Qed. Lemma isType_ws_cumul_ctx_pb {cf Σ Γ Δ T} {wfΣ : wf Σ}: @@ -308,8 +291,7 @@ Proof. Qed. Lemma typing_spine_proofs {cf:checker_flags} Σ Γ Δ ind u args' args T' s : - check_univs -> - wf_ext Σ -> + wf Σ.1 -> Σ ;;; Γ |- T' : tSort s -> typing_spine Σ Γ (it_mkProd_or_LetIn Δ (mkApps (tInd ind u) args')) args T' -> ((All_local_assum (fun Γ' t => @@ -323,7 +305,7 @@ Lemma typing_spine_proofs {cf:checker_flags} Σ Γ Δ ind u args' args T' s : is_propositional (subst_instance_univ u idecl.(ind_sort)) -> s = subst_instance_univ u idecl.(ind_sort)))))%type. Proof. - intros checku wfΣ Ht. + intros wfΣ Ht. induction Δ using PCUICInduction.ctx_length_rev_ind in Γ, args', args, T', Ht |- *; simpl; intros sp. - dependent elimination sp as [spnil _ _ e|spcons isty isty' e _ sp]. split; [repeat constructor|]. @@ -346,11 +328,9 @@ Proof. eapply leq_universe_prop_r in sp; auto. rewrite (is_prop_subst_instance_univ ui') in sp => //. now destruct (ind_sort idecl). - apply wfΣ. eapply leq_universe_sprop_r in sp; auto. rewrite (is_sprop_subst_instance_univ ui') in sp => //. now destruct (ind_sort idecl). - apply wfΣ. intros propsub props. rewrite is_propositional_subst_instance in props. apply leq_universe_propositional_l in sp; eauto. subst s. @@ -505,15 +485,13 @@ Qed. that elimination to any type is allowed. *) Lemma Is_proof_mkApps_tConstruct `{cf : checker_flags} (Σ : global_env_ext) Γ ind n u mdecl idecl args : - check_univs -> - wf_ext Σ -> + wf Σ.1 -> declared_inductive (fst Σ) ind mdecl idecl -> (ind_kelim idecl <> IntoPropSProp /\ ind_kelim idecl <> IntoSProp) -> Is_proof Σ Γ (mkApps (tConstruct ind n u) args) -> #|ind_ctors idecl| <= 1 /\ ∥ All (Is_proof Σ Γ) (skipn (ind_npars mdecl) args) ∥. Proof. - intros checkunivs HΣ decli kelim [tyc [tycs [hc [hty hp]]]]. - assert (wfΣ : wf Σ) by apply HΣ. + intros wfΣ decli kelim [tyc [tycs [hc [hty hp]]]]. eapply inversion_mkApps in hc as [? [hc hsp]]; auto. eapply inversion_Construct in hc as [mdecl' [idecl' [cdecl' [wfΓ [declc [cu cum']]]]]]; auto. destruct (on_declared_constructor declc) as [[oi oib] [cs [Hnth onc]]]. @@ -605,13 +583,11 @@ Proof. Qed. Lemma elim_restriction_works_kelim `{cf : checker_flags} (Σ : global_env_ext) ind mind idecl : - check_univs -> - wf_ext Σ -> + wf Σ -> declared_inductive (fst Σ) ind mind idecl -> (ind_kelim idecl <> IntoPropSProp /\ ind_kelim idecl <> IntoSProp) -> Informative Σ ind. Proof. - intros cu HΣ H indk. - assert (wfΣ : wf Σ) by apply HΣ. + intros wfΣ H indk. destruct (on_declared_inductive H) as [[]]; eauto. intros ?. intros. eapply declared_inductive_inj in H as []; eauto; subst idecl0 mind. @@ -621,8 +597,7 @@ Proof. Qed. Lemma elim_restriction_works `{cf : checker_flags} (Σ : global_env_ext) Γ T (ci : case_info) p c brs mind idecl : - check_univs -> - wf_ext Σ -> + wf Σ.1 -> declared_inductive (fst Σ) ci mind idecl -> Σ ;;; Γ |- tCase ci p c brs : T -> (Is_proof Σ Γ (tCase ci p c brs) -> False) -> Informative Σ ci.(ci_ind). @@ -662,7 +637,7 @@ Proof. Qed. Lemma elim_restriction_works_proj `{cf : checker_flags} (Σ : global_env_ext) Γ p c mind idecl T : - check_univs -> wf_ext Σ -> + wf Σ.1 -> declared_inductive (fst Σ) p.(proj_ind) mind idecl -> Σ ;;; Γ |- tProj p c : T -> (Is_proof Σ Γ (tProj p c) -> False) -> Informative Σ p.(proj_ind). @@ -676,56 +651,55 @@ Section no_prop_leq_type. Context `{cf : checker_flags}. Variable Hcf : prop_sub_type = false. -Variable Hcf' : check_univs. Lemma leq_term_prop_sorted_l {Σ Γ v v' u u'} : - wf_ext Σ -> + wf Σ -> PCUICEquality.leq_term Σ (global_ext_constraints Σ) v v' -> Σ;;; Γ |- v : tSort u -> Σ;;; Γ |- v' : tSort u' -> Universe.is_prop u -> leq_universe (global_ext_constraints Σ) u' u. -Proof using Hcf Hcf'. +Proof using Hcf. intros wfΣ leq hv hv' isp. eapply typing_leq_term_prop in leq; eauto. apply leq_prop_prop; intuition auto. eapply cumul_prop_sym in leq. - eapply cumul_prop_props in leq; eauto. auto. apply wfΣ. + eapply cumul_prop_props in leq; eauto. auto. Qed. Lemma leq_term_prop_sorted_r {Σ Γ v v' u u'} : - wf_ext Σ -> + wf Σ -> PCUICEquality.leq_term Σ (global_ext_constraints Σ) v v' -> Σ;;; Γ |- v : tSort u -> Σ;;; Γ |- v' : tSort u' -> Universe.is_prop u' -> leq_universe (global_ext_constraints Σ) u u'. -Proof using Hcf Hcf'. +Proof using Hcf. intros wfΣ leq hv hv' isp. eapply typing_leq_term_prop in leq; eauto. apply leq_prop_prop; intuition auto. - apply cumul_prop_props in leq; auto. apply wfΣ. + apply cumul_prop_props in leq; auto. Qed. Lemma leq_term_sprop_sorted_l {Σ Γ v v' u u'} : - wf_ext Σ -> + wf Σ -> PCUICEquality.leq_term Σ (global_ext_constraints Σ) v v' -> Σ;;; Γ |- v : tSort u -> Σ;;; Γ |- v' : tSort u' -> Universe.is_sprop u -> leq_universe (global_ext_constraints Σ) u' u. -Proof using Hcf Hcf'. +Proof using Hcf. intros wfΣ leq hv hv' isp. eapply typing_leq_term_prop in leq; eauto. apply leq_sprop_sprop; intuition auto. eapply cumul_prop_sym in leq. - eapply cumul_sprop_props in leq; auto. eauto. auto. apply wfΣ. + eapply cumul_sprop_props in leq; auto. eauto. Qed. Lemma leq_term_propositional_sorted_l {Σ Γ v v' u u'} : - wf_ext Σ -> + wf Σ -> PCUICEquality.leq_term Σ (global_ext_constraints Σ) v v' -> Σ;;; Γ |- v : tSort u -> Σ;;; Γ |- v' : tSort u' -> is_propositional u -> leq_universe (global_ext_constraints Σ) u' u. -Proof using Hcf Hcf'. +Proof using Hcf. intros wfΣ leq hv hv' isp. eapply orb_true_iff in isp as [isp | isp]. - eapply leq_term_prop_sorted_l; eauto. @@ -733,26 +707,26 @@ Proof using Hcf Hcf'. Qed. Lemma leq_term_sprop_sorted_r {Σ Γ v v' u u'} : - wf_ext Σ -> + wf Σ -> PCUICEquality.leq_term Σ (global_ext_constraints Σ) v v' -> Σ;;; Γ |- v : tSort u -> Σ;;; Γ |- v' : tSort u' -> Universe.is_sprop u' -> leq_universe (global_ext_constraints Σ) u u'. -Proof using Hcf Hcf'. +Proof using Hcf. intros wfΣ leq hv hv' isp. eapply typing_leq_term_prop in leq; eauto. apply leq_sprop_sprop; intuition auto. - apply cumul_sprop_props in leq; auto. apply wfΣ. + apply cumul_sprop_props in leq; auto. Qed. Lemma cumul_prop_inv (Σ : global_env_ext) Γ A B u u' : - wf_ext Σ -> + wf Σ -> Universe.is_prop u -> (((Σ ;;; Γ |- A : tSort u) * (Σ ;;; Γ |- B : tSort u')) + ((Σ ;;; Γ |- B : tSort u) * (Σ ;;; Γ |- A : tSort u')))%type -> Σ ;;; Γ |- A <= B -> ((Σ ;;; Γ |- A : tSort u) * (Σ ;;; Γ |- B : tSort u))%type. -Proof using Hcf Hcf'. +Proof using Hcf. intros wfΣ propu. intros [[HA HB]|[HB HA]] cum; split; auto; apply cumul_alt in cum as [v [v' [[redv redv'] leq]]]. @@ -773,13 +747,13 @@ Proof using Hcf Hcf'. Qed. Lemma cumul_sprop_inv (Σ : global_env_ext) Γ A B u u' : - wf_ext Σ -> + wf Σ -> Universe.is_sprop u -> (((Σ ;;; Γ |- A : tSort u) * (Σ ;;; Γ |- B : tSort u')) + ((Σ ;;; Γ |- B : tSort u) * (Σ ;;; Γ |- A : tSort u')))%type -> Σ ;;; Γ |- A <= B -> ((Σ ;;; Γ |- A : tSort u) * (Σ ;;; Γ |- B : tSort u))%type. -Proof using Hcf Hcf'. +Proof using Hcf. intros wfΣ propu. intros [[HA HB]|[HB HA]] cum; split; auto; apply cumul_alt in cum as [v [v' [[redv redv'] leq]]]. @@ -802,12 +776,12 @@ Proof using Hcf Hcf'. Qed. Lemma unique_sorting_equality_prop_l {pb} {Σ : global_env_ext} {Γ T U s s'} : - wf_ext Σ -> + wf Σ -> Σ ;;; Γ |- T : tSort s -> Σ ;;; Γ |- U : tSort s' -> Σ ;;; Γ ⊢ T ≤[pb] U -> Universe.is_prop s -> Universe.is_prop s'. -Proof using Hcf Hcf'. +Proof using Hcf. intros wfΣ HT HU cum isp. eapply PCUICSpine.ws_cumul_pb_le_le in cum. eapply ws_cumul_pb_alt_closed in cum as [v [v' [eqv]]]. @@ -818,12 +792,12 @@ Proof using Hcf Hcf'. Qed. Lemma unique_sorting_equality_prop_r {pb} {Σ : global_env_ext} {Γ T U s s'} : - wf_ext Σ -> + wf Σ -> Σ ;;; Γ |- T : tSort s -> Σ ;;; Γ |- U : tSort s' -> Σ ;;; Γ ⊢ T ≤[pb] U -> Universe.is_prop s' -> Universe.is_prop s. -Proof using Hcf Hcf'. +Proof using Hcf. intros wfΣ HT HU cum isp. eapply PCUICSpine.ws_cumul_pb_le_le in cum. eapply ws_cumul_pb_alt_closed in cum as [v [v' [eqv]]]. @@ -834,12 +808,12 @@ Proof using Hcf Hcf'. Qed. Lemma unique_sorting_equality_prop {pb} {Σ : global_env_ext} {Γ T U s s'} : - wf_ext Σ -> + wf Σ -> Σ ;;; Γ |- T : tSort s -> Σ ;;; Γ |- U : tSort s' -> Σ ;;; Γ ⊢ T ≤[pb] U -> Universe.is_prop s = Universe.is_prop s'. -Proof using Hcf Hcf'. +Proof using Hcf. intros wfΣ HT HU cum. apply iff_is_true_eq_bool. split. @@ -848,12 +822,12 @@ Proof using Hcf Hcf'. Qed. Lemma unique_sorting_equality_sprop_l {pb} {Σ : global_env_ext} {Γ T U s s'} : - wf_ext Σ -> + wf Σ -> Σ ;;; Γ |- T : tSort s -> Σ ;;; Γ |- U : tSort s' -> Σ ;;; Γ ⊢ T ≤[pb] U -> Universe.is_sprop s -> Universe.is_sprop s'. -Proof using Hcf Hcf'. +Proof using Hcf. intros wfΣ HT HU cum isp. eapply PCUICSpine.ws_cumul_pb_le_le in cum. eapply ws_cumul_pb_alt_closed in cum as [v [v' [eqv]]]. @@ -864,12 +838,12 @@ Proof using Hcf Hcf'. Qed. Lemma unique_sorting_equality_sprop_r {pb} {Σ : global_env_ext} {Γ T U s s'} : - wf_ext Σ -> + wf Σ -> Σ ;;; Γ |- T : tSort s -> Σ ;;; Γ |- U : tSort s' -> Σ ;;; Γ ⊢ T ≤[pb] U -> Universe.is_sprop s' -> Universe.is_sprop s. -Proof using Hcf Hcf'. +Proof using Hcf. intros wfΣ HT HU cum isp. eapply PCUICSpine.ws_cumul_pb_le_le in cum. eapply ws_cumul_pb_alt_closed in cum as [v [v' [eqv]]]. @@ -880,12 +854,12 @@ Proof using Hcf Hcf'. Qed. Lemma unique_sorting_equality_sprop {pb} {Σ : global_env_ext} {Γ T U s s'} : - wf_ext Σ -> + wf Σ -> Σ ;;; Γ |- T : tSort s -> Σ ;;; Γ |- U : tSort s' -> Σ ;;; Γ ⊢ T ≤[pb] U -> Universe.is_sprop s = Universe.is_sprop s'. -Proof using Hcf Hcf'. +Proof using Hcf. intros wfΣ HT HU cum. apply iff_is_true_eq_bool. split. @@ -894,12 +868,12 @@ Proof using Hcf Hcf'. Qed. Lemma unique_sorting_equality_propositional {pb} {Σ : global_env_ext} {Γ T U s s'} : - wf_ext Σ -> + wf Σ -> Σ ;;; Γ |- T : tSort s -> Σ ;;; Γ |- U : tSort s' -> Σ ;;; Γ ⊢ T ≤[pb] U -> is_propositional s = is_propositional s'. -Proof using Hcf Hcf'. +Proof using Hcf. intros wfΣ HT HU cum. unfold is_propositional. destruct (Universe.is_prop s) eqn:isp => /=. symmetry. @@ -914,13 +888,13 @@ Proof using Hcf Hcf'. Qed. Lemma cumul_prop1 (Σ : global_env_ext) Γ A B u : - wf_ext Σ -> + wf Σ -> Universe.is_prop u -> isType Σ Γ A -> Σ ;;; Γ |- B : tSort u -> Σ ;;; Γ |- A <= B -> Σ ;;; Γ |- A : tSort u. -Proof using Hcf Hcf'. +Proof using Hcf. intros. destruct X0 as (s & e & Hs). eapply cumul_prop_inv in H. 4:eauto. pcuicfo. auto. @@ -928,13 +902,13 @@ Proof using Hcf Hcf'. Qed. Lemma cumul_prop2 (Σ : global_env_ext) Γ A B u : - wf_ext Σ -> + wf Σ -> Universe.is_prop u -> isType Σ Γ B -> Σ ;;; Γ |- A <= B -> Σ ;;; Γ |- A : tSort u -> Σ ;;; Γ |- B : tSort u. -Proof using Hcf Hcf'. +Proof using Hcf. intros. destruct X0 as (s & e & Hs). eapply cumul_prop_inv in H. 4:eauto. pcuicfo. auto. @@ -942,13 +916,13 @@ Proof using Hcf Hcf'. Qed. Lemma cumul_sprop1 (Σ : global_env_ext) Γ A B u : - wf_ext Σ -> + wf Σ -> Universe.is_sprop u -> isType Σ Γ A -> Σ ;;; Γ |- B : tSort u -> Σ ;;; Γ |- A <= B -> Σ ;;; Γ |- A : tSort u. -Proof using Hcf Hcf'. +Proof using Hcf. intros. destruct X0 as (s & e & Hs). eapply cumul_sprop_inv in H. 4:eauto. pcuicfo. auto. @@ -956,13 +930,13 @@ Proof using Hcf Hcf'. Qed. Lemma cumul_sprop2 (Σ : global_env_ext) Γ A B u : - wf_ext Σ -> + wf Σ -> Universe.is_sprop u -> isType Σ Γ B -> Σ ;;; Γ |- A <= B -> Σ ;;; Γ |- A : tSort u -> Σ ;;; Γ |- B : tSort u. -Proof using Hcf Hcf'. +Proof using Hcf. intros. destruct X0 as (s & e & Hs). eapply cumul_sprop_inv in H. 4:eauto. pcuicfo. auto. diff --git a/pcuic/theories/PCUICExpandLetsCorrectness.v b/pcuic/theories/PCUICExpandLetsCorrectness.v index c7de784cc..5598517cf 100644 --- a/pcuic/theories/PCUICExpandLetsCorrectness.v +++ b/pcuic/theories/PCUICExpandLetsCorrectness.v @@ -3578,7 +3578,7 @@ Proof. destruct onc. red in on_ctype. destruct on_ctype as (s & e & Hs). rewrite /type_of_constructor. forward Hs. eauto. - exists (s@[u]); split; [apply e|]. + exists (s@[u]); split => //. 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. @@ -3621,6 +3621,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. } @@ -3675,7 +3676,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. @@ -3706,7 +3707,7 @@ Proof. rewrite !trans_bcontext. now eapply eq_context_gen_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. @@ -4697,11 +4698,11 @@ Proof. constructor; auto; try apply IHX. { now apply (fresh_global_map (Σ := Σ0)). } destruct d; cbn in *. - * cbn. red. move: ond; rewrite /on_constant_decl. - destruct c as [type [body|] univs] => /=. - intros Hty; eapply (pcuic_expand_lets (Σ0, univs) [] _ _ X Hty IHX). - intros (s & e & Hty). exists s. split; [apply e|]. - exact (pcuic_expand_lets (Σ0, univs) [] _ _ X Hty IHX). + * destruct ond as (onty & onbo); split. + + destruct onty as (s & e & Hty). exists s. split; [apply e|]. + eapply (pcuic_expand_lets (Σ0, _) [] _ _ X Hty IHX). + + destruct c as [type [body|] univs] => //=. + exact (pcuic_expand_lets (Σ0, _) [] _ _ X onbo IHX). * generalize ond. intros []; econstructor; eauto. + cbn. eapply (Alli_mapi _ _ _).1, Alli_impl; tea. diff --git a/pcuic/theories/PCUICInversion.v b/pcuic/theories/PCUICInversion.v index ebaab7f8e..15a229787 100644 --- a/pcuic/theories/PCUICInversion.v +++ b/pcuic/theories/PCUICInversion.v @@ -258,6 +258,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) + (sortrel : 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)))) diff --git a/pcuic/theories/PCUICParallelReductionConfluence.v b/pcuic/theories/PCUICParallelReductionConfluence.v index 36c827a0f..116ee1698 100644 --- a/pcuic/theories/PCUICParallelReductionConfluence.v +++ b/pcuic/theories/PCUICParallelReductionConfluence.v @@ -1697,9 +1697,9 @@ Section Rho. case e: lookup_env => [[decl|decl]|] //; simp rho. case eb: cst_body => [b|] //; simp rho. rewrite rename_inst inst_closed0 //. - apply declared_decl_closed in e => //. - hnf in e. rewrite eb in e. rewrite closedn_subst_instance; auto. - now move/andP: e => [-> _]. + apply declared_decl_closed in e as (_ & Hbo) => //. + rewrite eb in Hbo. rewrite closedn_subst_instance; auto. + now move/andP: Hbo => [-> _]. - (* construct/cofix iota reduction *) simpl; simp rho. diff --git a/pcuic/theories/PCUICPrincipality.v b/pcuic/theories/PCUICPrincipality.v index 7a380e2c3..f1b49cdcb 100644 --- a/pcuic/theories/PCUICPrincipality.v +++ b/pcuic/theories/PCUICPrincipality.v @@ -22,7 +22,7 @@ Set Equations With UIP. Section Principality. Context {cf : checker_flags}. Context (Σ : global_env_ext). - Context (wfΣ : wf_ext Σ). + Context (wfΣ : wf Σ). Ltac pih := lazymatch goal with @@ -68,7 +68,6 @@ Section Principality. Hint Extern 10 (isWfArity _ _ (tSort _)) => apply isWfArity_sort : pcuic. Ltac int inv := intros B hB; eapply inv in hB; auto; split; [|econstructor; eauto]. - Hint Resolve wf_ext_wf : core. Theorem principal_type {Γ u A} : Σ ;;; Γ |- u : A -> ∑ C, (forall B, Σ ;;; Γ |- u : B -> Σ ;;; Γ ⊢ C ≤ B × Σ ;;; Γ |- u : C). @@ -294,7 +293,7 @@ Section Principality. eapply type_reduction in t0. 2:eapply redr. eapply validity; eauto. * split. { eapply PCUICWeakeningTyp.weaken_wf_local; tea. pcuic. pcuic. - eapply (wf_projection_context _ (p:=p)); tea. pcuic. } + eapply (wf_projection_context _ (p:=p)); tea. } eapply (projection_subslet _ _ _ _ _ _ p); eauto. simpl. eapply validity; eauto. * constructor; auto. now eapply wt_cumul_pb_refl. now apply All2_rev. @@ -350,7 +349,7 @@ Section Principality. End Principality. -Lemma principal_type_ind {cf:checker_flags} {Σ Γ c ind u u' args args'} {wfΣ: wf_ext Σ} : +Lemma principal_type_ind {cf:checker_flags} {Σ Γ c ind u u' args args'} {wfΣ: wf Σ.1} : Σ ;;; Γ |- c : mkApps (tInd ind u) args -> Σ ;;; Γ |- c : mkApps (tInd ind u') args' -> (∑ ui', @@ -454,7 +453,6 @@ Qed. Lemma typing_leq_term {cf:checker_flags} (Σ : global_env_ext) Γ t t' T T' : wf Σ.1 -> - on_udecl Σ.1 Σ.2 -> Σ ;;; Γ |- t : T -> Σ ;;; Γ |- t' : T' -> leq_term empty_global_env Σ t' t -> @@ -462,11 +460,10 @@ Lemma typing_leq_term {cf:checker_flags} (Σ : global_env_ext) Γ t t' T T' : inductives in different sorts. *) Σ ;;; Γ |- t' : T. Proof. - intros wfΣ onu Ht. - revert Σ wfΣ Γ t T Ht onu t' T'. + intros wfΣ Ht. + revert Σ wfΣ Γ t T Ht t' T'. eapply (typing_ind_env (fun Σ Γ t T => - forall (onu : on_udecl Σ.1 Σ.2), forall t' T' : term, Σ ;;; Γ |- t' : T' -> leq_term empty_global_env Σ t' t -> Σ;;; Γ |- t' : T) (fun Σ Γ => wf_local Σ Γ)); auto;intros Σ wfΣ Γ wfΓ; intros. 1-13:match goal with @@ -482,14 +479,14 @@ Proof. apply cumul_Sort. now apply leq_universe_super. - eapply inversion_Prod in X4 as (s1' & s2' & e' & Ha & Hb & Hs); auto. - specialize (X1 onu _ _ Ha). + specialize (X1 _ _ Ha). specialize (X1 (eq_term_empty_leq_term X5_1)). apply eq_term_empty_eq_term in X5_1. eapply context_conversion in Hb. 3:{ constructor. apply conv_ctx_refl. constructor. eassumption. constructor. eauto. } all:eauto. 2:{ constructor; eauto. now exists s1. } - specialize (X3 onu _ _ Hb X5_2). + specialize (X3 _ _ Hb X5_2). econstructor; eauto. rewrite e; auto. apply leq_term_empty_leq_term in X5_2. @@ -500,11 +497,11 @@ Proof. constructor. now symmetry. - eapply inversion_Lambda in X4 as (s & B & e' & dom & codom & cum); auto. - specialize (X1 onu _ _ dom (eq_term_empty_leq_term X5_1)). + specialize (X1 _ _ dom (eq_term_empty_leq_term X5_1)). apply eq_term_empty_eq_term in X5_1. assert(conv_context cumulAlgo_gen Σ (Γ ,, vass na0 ty) (Γ ,, vass na t)). { repeat constructor; pcuic. } - specialize (X3 onu t0 B). + specialize (X3 t0 B). forward X3 by eapply context_conversion; eauto; pcuic. eapply (type_ws_cumul_pb (pb:=Conv)). * econstructor. 1,2: eauto. instantiate (1 := bty). @@ -519,14 +516,14 @@ Proof. now len in X2; len. - eapply inversion_LetIn in X6 as (s1' & A & e' & dom & bod & codom & cum); auto. - specialize (X1 onu _ _ dom (eq_term_empty_leq_term X7_2)). - specialize (X3 onu _ _ bod (eq_term_empty_leq_term X7_1)). + specialize (X1 _ _ dom (eq_term_empty_leq_term X7_2)). + specialize (X3 _ _ bod (eq_term_empty_leq_term X7_1)). apply eq_term_empty_eq_term in X7_1. apply eq_term_empty_eq_term in X7_2. assert(Σ ⊢ Γ ,, vdef na0 t ty = Γ ,, vdef na b b_ty). { constructor. eapply ws_cumul_ctx_pb_refl. fvs. constructor => //. constructor; fvs. constructor; fvs. } - specialize (X5 onu u A). + specialize (X5 u A). forward X5 by eapply closed_context_conversion; eauto; pcuic. specialize (X5 X7_3). eapply leq_term_empty_leq_term in X7_3. @@ -546,8 +543,8 @@ Proof. - eapply inversion_App in X6 as (na' & A' & B' & hf & ha & cum); auto. unfold leq_term in X1. eapply eq_term_upto_univ_empty_impl in X7_1. - specialize (X3 onu _ _ hf X7_1). all:try typeclasses eauto. - specialize (X5 onu _ _ ha (eq_term_empty_leq_term X7_2)). + specialize (X3 _ _ hf X7_1). all:try typeclasses eauto. + specialize (X5 _ _ ha (eq_term_empty_leq_term X7_2)). eapply leq_term_empty_leq_term in X7_1. eapply eq_term_empty_eq_term in X7_2. eapply type_ws_cumul_pb. @@ -626,7 +623,7 @@ Proof. eapply inversion_Case in X9 as (mdecl' & idecl' & decli' & indices' & data & cum); auto. destruct (declared_inductive_inj isdecl decli'). subst mdecl' idecl'. destruct data. - unshelve epose proof (X7 _ _ _ scrut_ty (eq_term_empty_leq_term X10)); tea. + unshelve epose proof (X7 _ _ scrut_ty (eq_term_empty_leq_term X10)); tea. pose proof (eq_term_empty_eq_term X10). destruct e as [eqpars [eqinst [eqpctx eqpret]]]. eapply eq_term_empty_eq_term in eqpret. @@ -646,19 +643,17 @@ Proof. rewrite /pre_case_predicate_context_gen. eapply eq_context_upto_inst_case_context => //. eapply All2_app. 2:constructor; pcuic. - specialize (X3 _ _ scrut_ty (eq_term_empty_leq_term X10)). - unshelve epose proof (principal_type_ind scrut_ty X3) as [_ indconv]; tea. - split; auto. + specialize (X7 _ _ scrut_ty (eq_term_empty_leq_term X10)). + unshelve epose proof (principal_type_ind scrut_ty X7) as [_ indconv]; tea. eapply All2_app_inv in indconv as [convpars convinds] => //. exact (All2_length eqpars). constructor => //; fvs. - eapply inversion_Proj in X3 as (u' & mdecl' & idecl' & cdecl' & pdecl' & args' & inv); auto. intuition auto. - specialize (X3 _ _ a0 (eq_term_empty_leq_term X4)). + specialize (X2 _ _ a0 (eq_term_empty_leq_term X4)). eapply eq_term_empty_eq_term in X4. - assert (wf_ext Σ) by (split; assumption). - pose proof (principal_type_ind X3 a0) as [Ruu' X3']. + pose proof (principal_type_ind X2 a0) as [Ruu' X3']. eapply (type_ws_cumul_pb (pb:=Conv)). * clear a0. econstructor; eauto. @@ -678,14 +673,14 @@ Proof. move=> i //. } eapply (substitution_ws_cumul_pb_subst_conv (Γ0 := ctx) (Γ1 := ctx) (Δ := [])); eauto. + eapply PCUICInductives.projection_subslet; eauto. - eapply validity in X3; auto. + eapply validity in X2; auto. + split. eapply PCUICWeakeningTyp.weaken_wf_local; tea. eapply wf_projection_context; tea. - eapply validity in X3. - now eapply (isType_mkApps_Ind_inv _ a) in X3 as [? [? []]]. + eapply validity in X2. + now eapply (isType_mkApps_Ind_inv _ a) in X2 as [? [? []]]. eapply PCUICInductives.projection_subslet; eauto. - eapply validity in X3; auto. + eapply validity in X2; auto. + constructor. constructor; fvs. eapply All2_rev. eapply ws_cumul_pb_terms_refl => //; fvs. + rewrite /ctx; eapply ws_cumul_pb_refl => //. @@ -722,14 +717,14 @@ Qed. Lemma typing_eq_term {cf:checker_flags} (Σ : global_env_ext) Γ t t' T T' : - wf_ext Σ -> + wf Σ -> Σ ;;; Γ |- t : T -> Σ ;;; Γ |- t' : T' -> eq_term empty_global_env Σ t t' -> Σ ;;; Γ |- t' : T. Proof. intros wfΣ ht ht' eq. - eapply typing_leq_term; eauto. apply wfΣ. + eapply typing_leq_term; eauto. now eapply eq_term_empty_leq_term. Qed. diff --git a/pcuic/theories/PCUICProgress.v b/pcuic/theories/PCUICProgress.v index 12f3e4ecb..20030f689 100644 --- a/pcuic/theories/PCUICProgress.v +++ b/pcuic/theories/PCUICProgress.v @@ -214,6 +214,7 @@ forall (P : global_env_ext -> context -> term -> term -> Type) wf_local Σ (Γ ,,, predctx) -> PΓ Σ (Γ ,,, predctx) -> is_allowed_elimination Σ idecl.(ind_kelim) ps -> + isSortRel ps ci.(ci_relevance) -> PCUICTyping.ctx_inst (Prop_conj typing P Σ) Γ (p.(pparams) ++ indices) (List.rev (subst_instance p.(puinst) (mdecl.(ind_params) ,,, idecl.(ind_indices)))) -> Σ ;;; Γ |- c : mkApps (tInd ci.(ci_ind) p.(puinst)) (p.(pparams) ++ indices) -> @@ -388,6 +389,7 @@ Lemma typing_ind_env `{cf : checker_flags} : wf_local Σ (Γ ,,, predctx) -> PΓ Σ (Γ ,,, predctx) -> is_allowed_elimination Σ idecl.(ind_kelim) ps -> + isSortRel ps ci.(ci_relevance) -> PCUICTyping.ctx_inst (Prop_conj typing P Σ) Γ (p.(pparams) ++ indices) (List.rev (subst_instance p.(puinst) (mdecl.(ind_params) ,,, idecl.(ind_indices)))) -> Σ ;;; Γ |- c : mkApps (tInd ci.(ci_ind) p.(puinst)) (p.(pparams) ++ indices) -> @@ -651,7 +653,7 @@ Proof with eauto with wcbv; try congruence. + eauto with wcbv. + red in Hax. eapply Hax in E; eauto. - intros Σ wfΣ Γ _ ci p c brs indices ps mdecl idecl Hidecl Hforall _ Heq Heq_context predctx Hwfpred Hcon Hreturn IHreturn Hwfl _. - intros Helim Hctxinst Hc IHc Hcof ptm Hwfbranches Hall Hax -> H. + intros Helim Hsortrel Hctxinst Hc IHc Hcof ptm Hwfbranches Hall Hax -> H. specialize (IHc Hax eq_refl) as [[t' IH] | IH]; eauto with wcbv. pose proof IH as IHv. eapply PCUICCanonicity.value_canonical in IH; eauto. diff --git a/pcuic/theories/PCUICRelevance.v b/pcuic/theories/PCUICRelevance.v index 8c9da0cf5..d9431d37e 100644 --- a/pcuic/theories/PCUICRelevance.v +++ b/pcuic/theories/PCUICRelevance.v @@ -10,9 +10,9 @@ Definition relevance_of_constant Σ kn := | None => Relevant end. -Definition relevance_of_constructor Σ ind (k : nat) := - match lookup_inductive Σ ind with - | Some (_, idecl) => idecl.(ind_relevance) +Definition relevance_of_constructor Σ ind k := + match lookup_constructor Σ ind k with + | Some (_, idecl, _) => idecl.(ind_relevance) | None => Relevant end. diff --git a/pcuic/theories/PCUICRelevanceTerm.v b/pcuic/theories/PCUICRelevanceTerm.v index 7e89970de..e18404f6f 100644 --- a/pcuic/theories/PCUICRelevanceTerm.v +++ b/pcuic/theories/PCUICRelevanceTerm.v @@ -31,8 +31,8 @@ Proof. destruct lookup_constant eqn:H => //. erewrite extends_lookup_constant; eauto. - unfold relevance_of_constructor. - destruct lookup_inductive eqn:H => //. - erewrite extends_lookup_inductive; eauto. + destruct lookup_constructor eqn:H => //. + erewrite extends_lookup_constructor; eauto. - unfold relevance_of_projection. destruct lookup_projection eqn:H => //. erewrite extends_lookup_projection; eauto. diff --git a/pcuic/theories/PCUICRelevanceTyp.v b/pcuic/theories/PCUICRelevanceTyp.v new file mode 100644 index 000000000..9cd1446b5 --- /dev/null +++ b/pcuic/theories/PCUICRelevanceTyp.v @@ -0,0 +1,344 @@ +From MetaCoq.Template Require Import config utils. +From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICRelevance + PCUICTyping PCUICInversion PCUICConversion PCUICCumulProp PCUICWeakeningTyp PCUICValidity PCUICWellScopedCumulativity BDUnique. + +Require Import ssreflect. + +Definition relevance_from_term_from_type {cf : checker_flags} (Σ : global_env_ext) Γ t T := + forall rel, isTermRel Σ (marks_of_context Γ) t rel <~> isTypeRel Σ Γ T rel. + +Lemma cumul_sort_relevance {cf : checker_flags} (Σ : global_env_ext) Γ T s1 s2 : + wf Σ.1 -> + Σ ;;; Γ |- T : tSort s1 -> Σ ;;; Γ |- T : tSort s2 -> relevance_of_sort s1 = relevance_of_sort s2. +Proof. + intros wfΣ Hs1 Hs2. + destruct (principal_type _ Hs1) as (ss & Hs & Hty); tea. + apply Hs in Hs1, Hs2. + eapply cumul_sort_confluence with (1 := Hs1) in Hs2 as (s & _ & le1 & le2). + transitivity (relevance_of_sort s). + * eapply leq_relevance => //. tea. + * eapply geq_relevance => //. tea. +Qed. + +Lemma isTypeRel2_relevance {cf : checker_flags} Σ Γ T rel1 rel2 : + wf Σ.1 -> + isTypeRel Σ Γ T rel1 -> isTypeRel Σ Γ T rel2 -> rel1 = rel2. +Proof. + intros wfΣ (s1 & <- & Hs1) (s2 & <- & Hs2). + now eapply cumul_sort_relevance. +Qed. + +Lemma ind_arity_relevant {cf : checker_flags} (Σ : global_env) mdecl idecl : + wf Σ -> + let ind_type := it_mkProd_or_LetIn (ind_params mdecl) (it_mkProd_or_LetIn (ind_indices idecl) (tSort (ind_sort idecl))) in + isType (Σ, ind_universes mdecl) [] ind_type -> + isTypeRel (Σ, ind_universes mdecl) [] ind_type Relevant. +Proof. + intros wfΣ ind_type (s & e & Hs). + eexists; split => //; tea. + rewrite /ind_type in Hs. + rewrite -it_mkProd_or_LetIn_app in Hs. + assert (wfΓ : wf_local (Σ, ind_universes mdecl) []) by pcuic. + revert wfΓ Hs. generalize ((ind_indices idecl ++ ind_params mdecl : context)) as l, ([]: context) as Γ. + induction l using rev_ind. + * cbn. intros Γ wfΓ Hty; apply inversion_Sort in Hty as (_ & _ & le); tea. + apply ws_cumul_pb_Sort_inv in le. + eapply leq_relevance; tea. + now destruct ind_sort. + * rewrite it_mkProd_or_LetIn_app; cbn. + destruct x, decl_body; cbn. + - intros Γ wfΓ Hty; pose proof Hty; apply inversion_LetIn in Hty as (? & ? & ? & ? & ? & ? & le); tea. + assert (wfΓ' : wf_local (Σ, ind_universes mdecl) (Γ,, vdef decl_name t decl_type)) by auto with pcuic. + eapply IHl. apply wfΓ'. + econstructor; eauto with pcuic. constructor; eauto with pcuic. + apply ws_cumul_pb_LetIn_l_inv in le. + epose proof (PCUICSpine.all_rels_subst_lift (Γ := Γ) (Δ := [vdef decl_name t decl_type]) (t := x0) (Δ' := []) _ _ _) => //=. + simpl in X0. rewrite PCUICLiftSubst.lift0_id in X0. rewrite PCUICLiftSubst.subst_empty in X0. + change (subst0 _ _) with ((lift 1 1 x0) {0 := lift0 1 t}) in X0. + rewrite -PCUICLiftSubst.distr_lift_subst10 in X0. + eapply PCUICContextConversion.ws_cumul_pb_eq_le in X0. + eapply cumulAlgo_cumulSpec. + etransitivity. 1: apply X0. + change (tSort _) with (lift0 1 (tSort s)). + eapply (weakening_ws_cumul_pb (Γ' := []) (Γ'' := [vdef _ _ _]) le). fvs. + - intros Γ wfΓ Hty; pose proof Hty; apply inversion_Prod in Hty as (? & s' & ? & ? & ? & le); tea. + assert (wfΓ' : wf_local (Σ, ind_universes mdecl) (Γ,, vass decl_name decl_type)) by auto with pcuic. + eapply IHl. apply wfΓ'. + econstructor; eauto with pcuic. constructor; eauto with pcuic. + eapply cumulAlgo_cumulSpec. + etransitivity. 2: eapply (weakening_ws_cumul_pb (Γ' := []) (Γ'' := [vass _ _]) le); fvs. + cbn. + constructor. 1-3: fvs. + constructor. apply leq_universe_product. + Unshelve. all: eauto. fvs. unfold app_context, app; fold (snoc Γ (vdef decl_name t decl_type)). fvs. +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 ((? & []) & ? & ? & []). + eapply infer_typing_sort_impl with _ on_ctype; [apply relevance_subst_opt|]; intros Hty. + instantiate (1 := u). + replace Γ with ([] ,,, Γ) by apply app_context_nil_l. + replace (type_of_constructor _ _ _ _) with (lift0 #|Γ| (type_of_constructor mdecl cdecl cstr u)). + 2: { rewrite PCUICLiftSubst.lift_closed //. eapply (PCUICClosedTyp.declared_constructor_closed_type isdecl). } + change (tSort _) with (lift0 #|Γ| (tSort on_ctype.π1)@[u]). + eapply @weakening with (Γ := []) => //. + rewrite app_context_nil_l => //. + destruct Σ as (Σ & φ). + eapply PCUICUnivSubstitutionTyp.typing_subst_instance' in Hty => //=; tea. + change ((tSort _)@[u]) with (subst0 (inds (inductive_mind cstr.1) u (ind_bodies mdecl)) (tSort on_ctype.π1)@[u]). + eapply @PCUICSubstitution.substitution with (Δ := []); tea. + 2: rewrite app_context_nil_l; apply Hty. + 2: { epose proof (PCUICWeakeningEnv.weaken_lookup_on_global_env' _ _ (InductiveDecl mdecl) wfΣ _); eauto. now split. Unshelve. 2: apply isdecl.p1.p1. } + rewrite /inds /arities_context. + clear -h_cuu isdecl wfΣ wfΓ. + pose proof (PCUICClosedTyp.declared_minductive_closed isdecl.p1.p1). + rewrite /PCUICClosed.closed_inductive_decl in H. move/andb_and: H => [] _ H. + apply forallb_All in H. + assert (Alli (fun i x => declared_inductive Σ {| inductive_mind := inductive_mind cstr.1; inductive_ind := i |} mdecl x) 0 (ind_bodies mdecl)). + { + apply forall_nth_error_Alli; intros. + unfold declared_inductive. + split => //=. + apply isdecl.p1.p1. + } + induction (ind_bodies mdecl) using rev_ind => //=. + - constructor. + - rewrite rev_map_app app_length Nat.add_1_r => //=. + apply All_app in H as (H & H1). + apply Alli_app in X as (X & X1). + constructor => //=; auto. + clear H X IHl. + inv H1. inv X1. clear X X0. + rewrite /PCUICClosed.closed_inductive_body in H. rtoProp. + rewrite PCUICClosed.subst_closedn. + rewrite PCUICClosed.closedn_subst_instance => //. + eapply type_Ind; eauto with pcuic. + now rewrite Nat.add_0_r in H0. +Qed. + +Lemma isType_of_projection {cf : checker_flags} Σ mdecl idecl cdecl pdecl p c args u Γ : + wf Σ.1 -> declared_projection Σ.1 p mdecl idecl cdecl pdecl -> + Σ;;; Γ |- c : mkApps (tInd (proj_ind p) u) args -> #|args| = ind_npars mdecl -> + isTypeRel Σ Γ (subst0 (c :: List.rev args) (proj_type pdecl)@[u]) pdecl.(proj_relevance). +Proof. +Admitted. + + +Lemma relevance_term_to_type {cf : checker_flags} : + env_prop relevance_from_term_from_type + (fun Σ Γ => All_local_env (lift_typing relevance_from_term_from_type Σ) Γ). +Proof using Type. + apply typing_ind_env; intros Σ wfΣ Γ wfΓ; intros **; rename_all_hyps; auto. + - eapply All_local_env_impl. + eapply All_local_env_over_2; tea. + intros ??? H; apply lift_typing_impl with (1 := H); now intros ? []. + - split; intro H. + + rewrite /relevance_of_term nth_error_map heq_nth_error /option_map /option_default /id in H. + rewrite -H. + destruct (PCUICTyping.nth_error_All_local_env heq_nth_error wfΓ); eauto. + apply isTypeRelOpt_lift => //. + pose proof (nth_error_Some_length heq_nth_error); lia. + + rewrite /relevance_of_term nth_error_map heq_nth_error /option_map /option_default /id. + pose proof (PCUICTyping.nth_error_All_local_env heq_nth_error wfΓ).1; eauto. + apply isTypeRelOpt_lift in X0 => //. 2: { pose proof (nth_error_Some_length heq_nth_error); lia. } + eapply isTypeRel2_relevance; tea. + - split; unfold relevance_of_term; intro Hr => //. + + eexists; split. + 2: { constructor; tea. eauto with pcuic. } + destruct u => //. + + destruct Hr as (s & <- & Hty) => //. + apply inversion_Sort in Hty as (_ & _ & le) => //. + eapply ws_cumul_pb_Sort_inv, leq_relevance in le => //. + destruct u => //. + - split; unfold relevance_of_term; intro Hr => //. + + rewrite -Hr. + eexists; split. + 2: { constructor; tea. eauto with pcuic. } + destruct s2, s1 => //. + + destruct Hr as (s & <- & Hty) => //. + apply inversion_Sort in Hty as (_ & _ & le) => //. + eapply ws_cumul_pb_Sort_inv, leq_relevance in le => //. + destruct s2, s1 => //. + - split; intro Hr => //. + + eapply X3 in Hr as (s & e & Hs). + eexists; split. + 2: { constructor; tea. } + destruct s, s1 => //. + + destruct Hr as (s & <- & Hs). + edestruct X3 as (_ & IH); eapply IH. + apply inversion_Prod in Hs as (s1' & s2 & e1 & _ & Hty & le) => //. + eexists; split; tea. + eapply ws_cumul_pb_Sort_inv, leq_relevance in le => //. + destruct s1', s2 => //. + - split; intro Hr => //. + + eapply X5 in Hr as (s & e & Hs). + exists s; split => //. + econstructor; tea. + eapply type_LetIn; eauto. constructor; eauto with pcuic. do 2 constructor. apply cumul_zeta. + + destruct Hr as (s & <- & Hs). + edestruct X5 as (_ & IH); eapply IH. + assert (wf_universe Σ s) by eauto with pcuic. + apply inversion_LetIn in Hs as (s1' & s2 & e1 & ? & ? & Hty & le) => //. + eapply ws_cumul_pb_Sort_r_inv in le as (ss & red & le) => //. + eexists; split; tea. + reflexivity. + econstructor; eauto with pcuic. constructor; eauto with pcuic. + eapply cumul_Trans. 4: apply cumul_Sort, le. 1,2: fvs. + apply cumulAlgo_cumulSpec. + apply invert_red_letin in red as [(? & ? & ? & []) | ?]; try discriminate. + epose proof (PCUICSpine.all_rels_subst_lift (Γ := Γ) (Δ := [vdef na b b_ty]) (t := s2) (Δ' := []) _ _ _) => //=. + simpl in X0. rewrite PCUICLiftSubst.lift0_id in X0. rewrite PCUICLiftSubst.subst_empty in X0. + change (subst0 _ _) with ((lift 1 1 s2) {0 := lift0 1 b}) in X0. + rewrite -PCUICLiftSubst.distr_lift_subst10 in X0. + apply PCUICContextConversion.ws_cumul_pb_eq_le in X0. + etransitivity. 1: apply X0. + change (tSort ss) with (lift0 1 (tSort ss)). + eapply red_conv in c. + eapply (weakening_ws_cumul_pb (Γ' := []) (Γ'' := [vdef _ _ _]) c). fvs. + - split; intro Hr => //. + + eapply X3 in Hr as (s' & e & Hs). + assert (wf_universe Σ s') by auto with pcuic. + apply inversion_Prod in Hs as (? & ss & ? & ? & ? & le); tea. + exists s'; split => //. + econstructor; tea. instantiate (1 := tSort ss). 2: constructor; eauto with pcuic. + 2: { apply cumul_Sort. apply ws_cumul_pb_Sort_inv in le. etransitivity. apply leq_universe_product. tea. } + change (tSort ss) with ((tSort ss) { 0 := u }). + eapply PCUICSubstitution.substitution0; tea. + + destruct Hr as (s' & <- & Hs). + edestruct X3 as (_ & IH); eapply IH; clear IH. + assert (wf_universe Σ s') by eauto with pcuic. + exists s; split => //. + apply inversion_Prod in IHB as (s1 & s2 & ? & ? & Hty & le) => //. + eapply (PCUICSubstitution.substitution0 Hty) in typeu; tea. cbn in typeu. + transitivity (relevance_of_sort s2). + 2: eapply cumul_sort_relevance; tea. + apply ws_cumul_pb_Sort_inv in le; destruct s2, s1, s => //. + - apply PCUICWeakeningEnvTyp.on_declared_constant in H as H'; tea. unfold on_constant_decl in H'. + split; intro Hr => //. + + unfold relevance_of_term, relevance_of_constant in Hr. + rewrite (declared_constant_lookup H) in Hr; rewrite -Hr. + destruct H' as (H' & _). + eapply infer_typing_sort_impl with _ H' => //; [apply relevance_subst_opt|]. intro Hty. + apply (weaken_ctx (Γ := [])); tea. + destruct Σ as (Σ & φ); unfold fst in *. + instantiate (1 := u). + eapply (PCUICUnivSubstitutionTyp.typing_subst_instance' _ _ [] _ (tSort H'.π1) u _); tea. + epose proof (PCUICWeakeningEnv.weaken_lookup_on_global_env' _ _ (ConstantDecl decl) wf _); eauto. + now split. + + unfold relevance_of_term, relevance_of_constant. + rewrite (declared_constant_lookup H). + destruct H' as (H' & _). + destruct Hr as (s1 & <- & Hs1), H' as (s2 & <- & Hs2). + erewrite <- relevance_subst_eq. + eapply cumul_sort_relevance; tea. + apply (weaken_ctx (Γ := [])); tea. + destruct Σ as (Σ & φ); unfold fst in *. instantiate (1 := u). + eapply (PCUICUnivSubstitutionTyp.typing_subst_instance' _ _ [] _ (tSort _) u _); tea. + epose proof (PCUICWeakeningEnv.weaken_lookup_on_global_env' _ _ (ConstantDecl decl) wf _); eauto. + now split. + - pose proof (onArity (PCUICWeakeningEnvTyp.on_declared_inductive isdecl).2). + destruct Σ as (Σ & φ); unfold fst in *. + split; intro Hr => //. + + unfold relevance_of_term in Hr; subst rel. + eapply infer_typing_sort_impl with _ X1; [apply relevance_subst_opt|]; intros Hty. + eapply PCUICUnivSubstitutionTyp.typing_subst_instance' in Hty. + eapply (weaken_ctx (Γ := [])) in Hty. + apply Hty. + all: tea. + epose proof (PCUICWeakeningEnv.weaken_lookup_on_global_env' _ _ (InductiveDecl mdecl) wf _); eauto. now split. + + unfold relevance_of_term. + destruct Hr as (s1 & <- & Hs1), X1 as (s2 & <- & Hs2). + eapply PCUICUnivSubstitutionTyp.typing_subst_instance' in Hs2. + eapply (weaken_ctx (Γ := [])) in Hs2. + all: tea. + 2: { epose proof (PCUICWeakeningEnv.weaken_lookup_on_global_env' _ _ (InductiveDecl mdecl) wf _); eauto. now split. } + erewrite <- relevance_subst_eq. + eapply cumul_sort_relevance; tea. apply wf. + - pose proof (declared_constructor_lookup isdecl). + destruct (PCUICWeakeningEnvTyp.on_declared_constructor isdecl) as ([[] []] & cu & ? & []); tea. + split; intro Hr => //. + + unfold relevance_of_term, relevance_of_constructor in Hr; subst rel. + destruct Σ as (Σ & φ); unfold fst, snd in *. + rewrite H. + apply isType_of_constructor; tea. + + unfold relevance_of_term, relevance_of_constructor. destruct Σ as (Σ & φ); unfold fst, snd in *. + rewrite H. + eassert (isTypeRel (Σ, φ) _ (type_of_constructor mdecl cdecl (ind, i) u) (idecl.(ind_relevance))) by (apply isType_of_constructor; tea). + eapply isTypeRel2_relevance; tea. apply wf. + - assert (Σ ;;; Γ |- mkApps ptm (indices ++ [c]) : tSort ps) by admit. + split; intro Hr => //. + + unfold relevance_of_term in Hr; subst rel. + exists ps; split => //. + + unfold relevance_of_term. + eapply isTypeRel2_relevance; tea. + exists ps; split => //. + - pose proof (declared_projection_lookup isdecl). + destruct (PCUICWeakeningEnvTyp.on_declared_projection isdecl) as [[[] e] [[[? []] ?] ?]]; tea. + unfold on_projection in o. + destruct nth_error eqn:EE in o => //. + destruct o. + split; intro Hr => //. + + unfold relevance_of_term, relevance_of_projection in Hr; subst rel. + destruct Σ as (Σ & φ); unfold fst, snd in *. + rewrite H. + eapply isType_of_projection; tea. + + unfold relevance_of_term, relevance_of_projection. destruct Σ as (Σ & φ); unfold fst, snd in *. + rewrite H. + eassert (isTypeRel (Σ, φ) _ (subst0 (c :: List.rev args) (proj_type pdecl)@[u]) (pdecl.(proj_relevance))) by (eapply isType_of_projection; tea). + eapply isTypeRel2_relevance; tea. apply wf. + - eapply All_nth_error in X0, X1; tea. + rewrite /on_def_type /on_def_body /lift_typing2 in X0, X1. + split; intro Hr => //. + + rewrite /relevance_of_term heq_nth_error /option_default in Hr; subst rel. + apply lift_typing_impl with (1 := X0); now intros ? []. + + rewrite /relevance_of_term heq_nth_error /option_default. + eapply isTypeRel2_relevance; tea. + apply infer_typing_sort_impl with id X0 => //. + now intros []. + - eapply All_nth_error in X0, X1; tea. + rewrite /on_def_type /on_def_body /lift_typing2 in X0, X1. + split; intro Hr => //. + + rewrite /relevance_of_term heq_nth_error /option_default in Hr; subst rel. + apply lift_typing_impl with (1 := X0); now intros ? []. + + rewrite /relevance_of_term heq_nth_error /option_default. + eapply isTypeRel2_relevance; tea. + apply infer_typing_sort_impl with id X0 => //. + now intros []. + - split; intro Hr => //. + + eexists; split; tea. + apply X1 in Hr as (s' & <- & Hs'). + enough (cumul_prop Σ Γ (tSort s) (tSort s')). + { apply PCUICCumulProp.cumul_sort_inv in X0. now destruct s, s'. } + pose proof (cumulSpec_typed_cumulAlgo typet typeB cumulA). + apply ws_cumul_pb_forget, PCUICCumulativity.cumul_alt in X0 as (A' & B' & [[red1 red2] leq]). + eapply PCUICSR.subject_reduction in red1, red2; tea. + eapply typing_leq_term_prop; tea. + + apply X1. + pose proof typet. + apply validity in X0 as (s' & _ & Hty). + eexists; split; tea. + enough (relevance_of_sort s = relevance_of_sort s'). + { cbn; rewrite -H. destruct Hr as (ss & <- & Hss). eapply cumul_sort_relevance; tea. } + enough (cumul_prop Σ Γ (tSort s) (tSort s')). + { apply PCUICCumulProp.cumul_sort_inv in X0. now destruct s, s'. } + pose proof (cumulSpec_typed_cumulAlgo typet typeB cumulA). + apply ws_cumul_pb_forget, PCUICCumulativity.cumul_alt in X0 as (A' & B' & [[red1 red2] leq]). + eapply PCUICSR.subject_reduction in red1, red2; tea. + eapply typing_leq_term_prop; tea. +Admitted. + + +Theorem relevance_from_type {cf : checker_flags} (Σ : global_env_ext) Γ t T rel : + wf Σ -> Σ ;;; Γ |- t : T -> + isTermRel Σ (marks_of_context Γ) t rel <~> isTypeRel Σ Γ T rel. +Proof. + intros wfΣ Hty. + pose proof relevance_term_to_type. + eapply env_prop_typing in X; tea. + apply X. +Qed. diff --git a/pcuic/theories/PCUICSR.v b/pcuic/theories/PCUICSR.v index fc46b5f18..05d0e0e47 100644 --- a/pcuic/theories/PCUICSR.v +++ b/pcuic/theories/PCUICSR.v @@ -1694,10 +1694,10 @@ Proof. - (* Constant unfolding *) unshelve epose proof (declared_constant_inj decl decl0 _ _); tea; subst decl. destruct decl0 as [ty body' univs]; simpl in *; subst body'. - eapply on_declared_constant in H; tas; cbn in H. + eapply on_declared_constant in H as (_ & Hbo); tas; cbn in Hbo. rewrite <- (app_context_nil_l Γ). - apply subject_closed in H as clb; tas. - apply type_closed in H as clty; tas. + apply subject_closed in Hbo as clb; tas. + apply type_closed in Hbo as clty; tas. replace (subst_instance u body) with (lift0 #|Γ| (subst_instance u body)). replace (subst_instance u ty) @@ -1758,7 +1758,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. @@ -2339,8 +2339,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 *. @@ -2407,14 +2407,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. diff --git a/pcuic/theories/PCUICSpine.v b/pcuic/theories/PCUICSpine.v index f5676bcef..d27c56a6c 100644 --- a/pcuic/theories/PCUICSpine.v +++ b/pcuic/theories/PCUICSpine.v @@ -873,8 +873,8 @@ Qed.*) now apply subslet_lift. Qed. - Lemma ctx_inst_length {Γ args Δ} : - ctx_inst Σ Γ args Δ -> + Lemma ctx_inst_length {P Γ args Δ} : + PCUICTyping.ctx_inst P Γ args Δ -> #|args| = context_assumptions Δ. Proof using Type. induction 1; simpl; auto. diff --git a/pcuic/theories/PCUICSubstitution.v b/pcuic/theories/PCUICSubstitution.v index 9cf3ae74a..55a6a2e25 100644 --- a/pcuic/theories/PCUICSubstitution.v +++ b/pcuic/theories/PCUICSubstitution.v @@ -362,19 +362,15 @@ Lemma subst_declared_constant `{H:checker_flags} Σ cst decl n k u : map_constant_body (subst_instance u) decl. Proof. intros. - eapply declared_decl_closed in H0; eauto. + eapply declared_decl_closed in H0 as (Hty & Hbo); eauto. unfold map_constant_body. - do 2 red in H0. destruct decl as [ty [body|] univs]; simpl in *. - - rewrite -> andb_and in H0. intuition auto. - rewrite <- (closedn_subst_instance 0 body u) in H1. - rewrite <- (closedn_subst_instance 0 ty u) in H2. - f_equal. - + apply subst_closedn; eauto using closed_upwards with arith wf. - + f_equal. apply subst_closedn; eauto using closed_upwards with arith wf. - - red in H0. f_equal. - intuition. simpl in *. - rewrite <- (closedn_subst_instance 0 ty u) in H0. - rewrite andb_true_r in H0. + simpl in *. f_equal. + - hnf in Hty. toProp Hty. destruct Hty as (Hty & _). + rewrite <- (closedn_subst_instance 0 _ u) in Hty. + apply subst_closedn; eauto using closed_upwards with arith wf. + - destruct cst_body => //=. f_equal. + hnf in Hbo. toProp Hbo. destruct Hbo as (Hbo & _). + rewrite <- (closedn_subst_instance 0 _ u) in Hbo. eapply subst_closedn; eauto using closed_upwards with arith wf. Qed. diff --git a/pcuic/theories/PCUICToTemplateCorrectness.v b/pcuic/theories/PCUICToTemplateCorrectness.v index b917afc94..bc97dc7a9 100644 --- a/pcuic/theories/PCUICToTemplateCorrectness.v +++ b/pcuic/theories/PCUICToTemplateCorrectness.v @@ -2290,10 +2290,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. @@ -2303,10 +2304,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. diff --git a/pcuic/theories/PCUICTyping.v b/pcuic/theories/PCUICTyping.v index ba3ee229f..7d8247b53 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)))) @@ -681,6 +682,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) -> @@ -778,8 +780,9 @@ Proof. * simpl. simpl in *. destruct d. - + destruct c; simpl in *. - destruct cst_body0; apply lift_typing_impl with (1 := Xg); intros ? Hs. + + destruct Xg; split. 2: destruct c, cst_body0 => //. + 1: rename o1 into Xg. 2: rename o2 into Xg. + all: apply lift_typing_impl with (1 := Xg); intros ? Hs. all: specialize (IH ((Σ', udecl); wfΣ; _; _; _; Hs)). all: forward IH; [constructor 1; simpl; subst Σ' Σg; cbn; lia|]. all: apply IH. @@ -1102,6 +1105,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) -> @@ -1183,18 +1187,6 @@ Section All_local_env. Context {cf: checker_flags}. - Lemma nth_error_All_local_env {P Γ n} (isdecl : n < #|Γ|) : - All_local_env P Γ -> - on_some (on_local_decl P (skipn (S n) Γ)) (nth_error Γ n). - Proof using Type. - induction 1 in n, isdecl |- *. red; simpl. - - destruct n; simpl; inv isdecl. - - destruct n. red; simpl. red. simpl. apply t0. - simpl. apply IHX. simpl in isdecl. lia. - - destruct n. auto. - apply IHX. simpl in *. lia. - Qed. - Lemma All_local_env_app (P : context -> term -> typ_or_sort -> Type) l l' : All_local_env P l * All_local_env (fun Γ t T => P (l ,,, Γ) t T) l' -> All_local_env P (l ,,, l'). diff --git a/pcuic/theories/PCUICValidity.v b/pcuic/theories/PCUICValidity.v index 175db0ef9..9d52667a5 100644 --- a/pcuic/theories/PCUICValidity.v +++ b/pcuic/theories/PCUICValidity.v @@ -116,9 +116,8 @@ Section Validity. on_local_decl P (skipn (S n) Γ) d. Proof using Type. intros heq hΓ. - epose proof (nth_error_Some_length heq). - eapply (nth_error_All_local_env) in H; tea. - now rewrite heq in H. + eapply (nth_error_All_local_env) in heq; tea. + destruct heq, d, decl_body => //. Qed. Notation type_ctx := (type_local_ctx (lift_typing typing)). @@ -280,17 +279,10 @@ Section Validity. do 2 constructor. apply leq_universe_product. - - destruct decl as [ty [b|] univs]; simpl in *. - * eapply declared_constant_inv in X; eauto. - red in X. simpl in X. - eapply isType_weakening; eauto. - eapply (isType_subst_instance_decl (Γ:=[])); eauto. simpl. - eapply weaken_env_prop_isType. - * have ond := on_declared_constant wf H. - do 2 red in ond. simpl in ond. - simpl in ond. - eapply isType_weakening; eauto. - eapply (isType_subst_instance_decl (Γ:=[])); eauto. + - destruct (on_declared_constant wf H) as (onty & _). + eapply isType_of_isTypeRel in onty. + eapply isType_weakening; eauto. + eapply (isType_subst_instance_decl (Γ:=[])); eauto. - (* Inductive type *) destruct (on_declared_inductive isdecl); pcuic. @@ -303,6 +295,7 @@ Section Validity. - (* Constructor type *) destruct (on_declared_constructor isdecl) as [[oni oib] [cs [declc onc]]]. unfold type_of_constructor. + eapply isType_of_isTypeRel. eapply infer_typing_sort_impl with _ (on_ctype onc); [apply relevance_subst_opt|]; intros Hs. eapply instantiate_minductive in Hs; eauto. 2:(destruct isdecl as [[] ?]; eauto). diff --git a/pcuic/theories/PCUICWcbvEval.v b/pcuic/theories/PCUICWcbvEval.v index af953cb78..d67bdca09 100644 --- a/pcuic/theories/PCUICWcbvEval.v +++ b/pcuic/theories/PCUICWcbvEval.v @@ -625,8 +625,8 @@ Section Wcbv. Proof using Type. move=> wfΣ Hc Hb. rewrite PCUICClosed.closedn_subst_instance. - apply declared_decl_closed in Hc => //. simpl in Hc. red in Hc. - rewrite Hb in Hc. simpl in Hc. now move/andP: Hc. + apply declared_decl_closed in Hc as (Hty & Hbo) => //. + rewrite Hb in Hbo. now move/andP: Hbo. Qed. Lemma closed_iota ci ind p c u args brs br : diff --git a/pcuic/theories/PCUICWfUniverses.v b/pcuic/theories/PCUICWfUniverses.v index 454ae5428..b76615d97 100644 --- a/pcuic/theories/PCUICWfUniverses.v +++ b/pcuic/theories/PCUICWfUniverses.v @@ -1038,10 +1038,10 @@ Qed. destruct X as [X _]. pose proof (nth_error_Some_length heq_nth_error). eapply nth_error_All_local_env in X; tea. - rewrite heq_nth_error /= in X. red in X. + destruct X as [Xty Xbod]. destruct decl as [na [b|] ty]; cbn -[skipn] in *. + now to_prop. - + now move: X => [s [] _ /andP[Hc _]]. + + now move: Xty => [s [] _ /andP[Hc _]]. - apply/andP; split; to_wfu; cbn ; eauto with pcuic. @@ -1053,15 +1053,8 @@ Qed. { rewrite wf_universeb_instance_forall. apply/wf_universe_instanceP. eapply consistent_instance_ext_wf; eauto. } - pose proof (declared_constant_inv _ _ _ _ wf_universes_weaken wf X H). - red in X1; cbn in X1. - destruct (cst_body decl). - * to_prop. - epose proof (weaken_lookup_on_global_env' Σ.1 _ _ wf H). - eapply wf_universes_inst. 2:eauto. all:eauto. - simpl in H2. - now eapply consistent_instance_ext_wf. - * move: X1 => [s [] _ /andP[Hc _]]. + pose proof (declared_constant_inv _ _ _ _ wf_universes_weaken wf X H) as (onty & onbo). + * move: onty => [_ [] _ /andP[Hc _]]. to_prop. eapply wf_universes_inst; eauto. exact (weaken_lookup_on_global_env' Σ.1 _ _ wf H). @@ -1093,8 +1086,8 @@ Qed. exact (weaken_lookup_on_global_env' Σ.1 _ _ wf (proj1 (proj1 isdecl))). now eapply consistent_instance_ext_wf. - - rewrite wf_universes_mkApps in H5. - move/andP: H5 => /= [] wfu; rewrite forallb_app. + - rewrite wf_universes_mkApps in H4. + move/andP: H4 => /= [] wfu; rewrite forallb_app. move/andP=> [] wfpars wfinds. cbn in wfu. rewrite wfu /= wfpars wf_universes_mkApps /= @@ -1132,7 +1125,7 @@ Qed. rewrite a1. now rewrite closedu_inds. * rewrite /ptm. - rewrite wf_universes_it_mkLambda_or_LetIn H4 andb_true_r. + rewrite wf_universes_it_mkLambda_or_LetIn H5 andb_true_r. rewrite /predctx. destruct X3 as [_ hctx]. move: hctx. now rewrite test_context_app => /andP[]. diff --git a/pcuic/theories/TemplateToPCUICCorrectness.v b/pcuic/theories/TemplateToPCUICCorrectness.v index 5c63d5ad4..adf0352aa 100644 --- a/pcuic/theories/TemplateToPCUICCorrectness.v +++ b/pcuic/theories/TemplateToPCUICCorrectness.v @@ -2089,6 +2089,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 @@ -2317,7 +2318,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|. @@ -2332,9 +2333,9 @@ Proof. eapply eq_binder_annots_eq. now eapply trans_ind_predicate_context. + cbn. split. cbn. - { epose proof (declared_minductive_ind_npars (proj1 H4)). - cbn in H4. len. rewrite -H0. - now rewrite context_assumptions_map in H5. } + { epose proof (declared_minductive_ind_npars (proj1 H5)). + cbn in H5. len. rewrite -H0. + 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. @@ -3009,11 +3010,9 @@ Proof. { split; rewrite trans_env_env_universes //. } have wfdecl := on_global_decl_wf (Σ := (Σg, udecl)) X0 o0. destruct d eqn:eqd. - * destruct c; simpl. destruct cst_body0; simpl in *. - red in o |- *. simpl in *. - eapply (X (Σg, cst_universes0) [] t (Typ cst_type0)); auto. - red in o0 |- *. simpl in *. - now apply (X (Σg, cst_universes0) [] cst_type0 (Sort None)). + * destruct o0, c; split. 2: destruct cst_body0 => //. + + now apply (X (Σg, cst_universes0) [] cst_type0 (SortRel cst_relevance0)). + + now apply (X (Σg, cst_universes0) [] t (Typ cst_type0)); auto. * destruct o0 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 *. @@ -3026,14 +3025,14 @@ Proof. eapply Alli_All; tea; cbv beta. move=> n x /Typing.onArity => o. eapply isType_of_isTypeRel. - apply (X (Σg, Ast.Env.ind_universes m) [] (Ast.Env.ind_type x) (SortRel (Ast.Env.ind_relevance x)) X0 o) => //. } + apply (X (Σg, Ast.Env.ind_universes m) [] (Ast.Env.ind_type x) (SortRel _) 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_rel (PCUICEnvTyping.lift_typing typing) (trans_global (Σg, Ast.Env.ind_universes m)) [] (ind_type (trans_one_ind_body (trans_global_env Σg) idecl)) - (ind_relevance (trans_one_ind_body (trans_global_env Σg) idecl)). - { apply ST.onArity in oni. unfold on_type in *; simpl in *. + Relevant. + { apply ST.onArity in oni. now apply (X (Σg, Ast.Env.ind_universes m) [] (Ast.Env.ind_type idecl) (SortRel _) _ oni). } unshelve refine {| ind_cunivs := oni.(ST.ind_cunivs) |}; tea. --- cbn -[trans_global_env]. rewrite oni.(ST.ind_arity_eq). @@ -3045,14 +3044,15 @@ 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)) - (Ast.Env.cstr_type x) (Sort None)). + (Ast.Env.cstr_type x) (SortRel (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) = diff --git a/pcuic/theories/TemplateToPCUICExpanded.v b/pcuic/theories/TemplateToPCUICExpanded.v index 4f9cb1871..e1ead9e17 100644 --- a/pcuic/theories/TemplateToPCUICExpanded.v +++ b/pcuic/theories/TemplateToPCUICExpanded.v @@ -333,12 +333,11 @@ Proof. cbn -[add_global_decl trans_global_env]. destruct decl as [kn []]; cbn in *; depelim H => //. * unfold trans_constant_body; cbn. - constructor. cbn. destruct p. - red in o. + constructor. cbn. destruct p as ((_ & Hbo) & _). destruct (Ast.Env.cst_body c); cbn => //. cbn in expanded_body. eapply trans_expanded in expanded_body; eauto. rewrite -/Σ' in expanded_body. now rewrite -eta_global_env. - red in o. now destruct o. + now destruct Hbo. * rewrite -eta_global_env. constructor. + cbn. eapply expanded_trans_local in expanded_params => //. diff --git a/pcuic/theories/Typing/PCUICClosedTyp.v b/pcuic/theories/Typing/PCUICClosedTyp.v index 05c889ece..f0ce359d7 100644 --- a/pcuic/theories/Typing/PCUICClosedTyp.v +++ b/pcuic/theories/Typing/PCUICClosedTyp.v @@ -190,19 +190,14 @@ Proof. - rewrite closedn_subst_instance. eapply lookup_on_global_env in X0; eauto. - destruct X0 as [Σ' [hext [onu HΣ'] IH]]. - repeat red in IH. destruct decl, cst_body0. simpl in *. - rewrite -> andb_and in IH. intuition auto. - eauto using closed_upwards with arith. - simpl in *. - repeat red in IH. destruct IH as (s & e & Hs). - rewrite -> andb_and in Hs. intuition auto. + destruct X0 as [Σ' [hext [onu HΣ'] (IHty & IHbo)]]. + destruct IHty as (_ & _ & (IH & _)%andb_and). eauto using closed_upwards with arith. - rewrite closedn_subst_instance. eapply declared_inductive_inv in X0; eauto. apply onArity in X0. repeat red in X0. - destruct X0 as (s & e & Hs). rewrite -> andb_and in Hs. + destruct X0 as (_ & _ & (X0 & _)%andb_and). intuition eauto using closed_upwards with arith. - destruct isdecl as [Hidecl Hcdecl]. @@ -215,17 +210,16 @@ Proof. pose proof X0.(onConstructors) as XX. eapply All2_nth_error_Some in Hcdecl; eauto. destruct Hcdecl as [? [? ?]]. cbn in *. - destruct o as [? ? (s & e' & Hs) _]. rewrite -> andb_and in Hs. - apply proj1 in Hs. + destruct o as [? ? (_ & _ & (Hs & _)%andb_and) _]. rewrite arities_context_length in Hs. 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). @@ -240,16 +234,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. - rewrite case_branch_type_fst in H8. *) - rewrite case_branch_context_length_args in H8 => //. + len in H9. + (*unfold case_branch_context_gen in H9. simpl in H9. + rewrite case_branch_type_fst in H9. *) + rewrite case_branch_context_length_args in H9 => //. now rewrite cstr_branch_context_length. + rewrite closedn_mkApps; auto. rewrite closedn_it_mkLambda_or_LetIn //. @@ -277,19 +271,17 @@ Proof. subst types. now rewrite app_context_length fix_context_length in H. eapply nth_error_all in X0; eauto. simpl in X0. intuition auto. rtoProp. - destruct X0 as (s & e & Hs & cl). now rewrite andb_true_r in cl. + now destruct X0 as (_ & _ & _ & (cl & _)%andb_and). - split. solve_all. destruct b. destruct x; simpl in *. unfold test_def. simpl. rtoProp. split. - destruct a as (s & e & Hs & cl). - now rewrite andb_true_r in cl. + now destruct a as (_ & _ & _ & (cl & _)%andb_and). rewrite -> app_context_length in H3, H4. rewrite -> Nat.add_comm in *. subst types. now rewrite fix_context_length in H3. eapply nth_error_all in X0; eauto. - destruct X0 as (s & e & Hs & cl). - now rewrite andb_true_r in cl. + now destruct X0 as (_ & _ & _ & (cl & _)%andb_and). Qed. Lemma declared_minductive_closed {cf:checker_flags} {Σ : global_env} {wfΣ : wf Σ} {mdecl mind} : @@ -495,13 +487,9 @@ Proof. intros h. unfold declared_constant in h. eapply lookup_on_global_env in h. 2: eauto. - destruct h as [Σ' [ext wfΣ' decl']]. - red in decl'. red in decl'. - destruct decl as [ty bo un]. simpl in *. - destruct bo as [t|]. - - now eapply type_closed in decl'. - - cbn in decl'. destruct decl' as (s & e & Hs). - now eapply subject_closed in Hs. Unshelve. all:tea. + destruct h as [Σ' [ext wfΣ' (declty' & declbo')]]. + destruct declty' as (? & _ & Hs). + now eapply subject_closed in Hs. Unshelve. all:tea. Qed. @@ -515,11 +503,9 @@ Proof. intros Σ cst decl body hΣ h e. unfold declared_constant in h. eapply lookup_on_global_env in h. 2: eauto. - destruct h as [Σ' [ext wfΣ' decl']]. - red in decl'. red in decl'. - destruct decl as [ty bo un]. simpl in *. - rewrite e in decl'. - now eapply subject_closed in decl'. + destruct h as [Σ' [ext wfΣ' (declty' & declbo')]]. + rewrite e in declbo'. + now eapply subject_closed in declbo'. Unshelve. all:tea. Qed. diff --git a/pcuic/theories/Typing/PCUICInstTyp.v b/pcuic/theories/Typing/PCUICInstTyp.v index b21a7717c..6dd45c770 100644 --- a/pcuic/theories/Typing/PCUICInstTyp.v +++ b/pcuic/theories/Typing/PCUICInstTyp.v @@ -438,7 +438,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 sortrel. intros IHctxi Hc IHc iscof ptm wfbrs Hbrs Δ f HΔ Hf. autorewrite with sigma. simpl. rewrite map_app. simpl. diff --git a/pcuic/theories/Typing/PCUICRenameTyp.v b/pcuic/theories/Typing/PCUICRenameTyp.v index 57818c4f8..b159f0886 100644 --- a/pcuic/theories/Typing/PCUICRenameTyp.v +++ b/pcuic/theories/Typing/PCUICRenameTyp.v @@ -886,7 +886,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 sortrel. intros IHctxi Hc IHc iscof ptm wfbrs Hbrs P Δ f Hf. simpl. rewrite rename_mkApps. diff --git a/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v b/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v index 464838090..65342ccd7 100644 --- a/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v +++ b/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v @@ -275,7 +275,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 sortrel IHctxi Hc IHc notCoFinite wfbrs hbrs i univs wfext cu. rewrite subst_instance_mkApps subst_instance_it_mkLambda_or_LetIn map_app. cbn. @@ -294,6 +294,7 @@ Proof using Type. + now rewrite -subst_instance_case_predicate_context -subst_instance_app_ctx. + cbn in *. eapply is_allowed_elimination_subst_instance; aa. + + apply relevance_subst => //. + move: IHctxi. simpl. rewrite -subst_instance_app. rewrite -subst_instance_two_context. diff --git a/pcuic/theories/Typing/PCUICWeakeningEnvTyp.v b/pcuic/theories/Typing/PCUICWeakeningEnvTyp.v index 5879403e4..8b60e43fe 100644 --- a/pcuic/theories/Typing/PCUICWeakeningEnvTyp.v +++ b/pcuic/theories/Typing/PCUICWeakeningEnvTyp.v @@ -125,11 +125,10 @@ Proof. intros HPΣ wfΣ wfΣ' Hext Hdecl. destruct decl. 1:{ - destruct c. destruct cst_body0. - - simpl in *. - red in Hdecl |- *. simpl in *. - eapply (HPΣ Σ Σ'); eauto. + destruct Hdecl; split. - eapply (HPΣ Σ Σ'); eauto. + - destruct cst_body => //. + eapply (HPΣ Σ Σ'); eauto. } simpl in *. destruct Hdecl as [onI onP onnP]; constructor; eauto. @@ -139,7 +138,7 @@ Proof. + unfold on_constructors in *. eapply All2_impl; eauto. intros. destruct X as [? ? ? ?]. unshelve econstructor; eauto. - * unfold on_type in *; eauto. + * unfold on_type_rel in *; eauto. * clear on_cindices cstr_eq cstr_args_length. revert on_cargs. induction (cstr_args x0) in y |- *; destruct y; simpl in *; eauto. @@ -202,11 +201,10 @@ Proof. pose proof (wfΣ := extends_decls_wf _ _ wfΣ' Hext). destruct decl. 1:{ - destruct c. destruct cst_body0. - - simpl in *. - red in Hdecl |- *. simpl in *. - eapply (HPΣ Σ Σ'); eauto. + destruct Hdecl; split. - eapply (HPΣ Σ Σ'); eauto. + - destruct cst_body => //. + eapply (HPΣ Σ Σ'); eauto. } simpl in *. destruct Hdecl as [onI onP onnP]; constructor; eauto. @@ -216,7 +214,7 @@ Proof. + unfold on_constructors in *. eapply All2_impl; eauto. intros. destruct X as [? ? ? ?]. unshelve econstructor; eauto. - * unfold on_type in *; eauto. + * unfold on_type_rel in *; eauto. * clear on_cindices cstr_eq cstr_args_length. revert on_cargs. induction (cstr_args x0) in y |- *; destruct y; simpl in *; eauto. diff --git a/pcuic/theories/Typing/PCUICWeakeningTyp.v b/pcuic/theories/Typing/PCUICWeakeningTyp.v index 634ca702d..ba9fb921d 100644 --- a/pcuic/theories/Typing/PCUICWeakeningTyp.v +++ b/pcuic/theories/Typing/PCUICWeakeningTyp.v @@ -206,11 +206,11 @@ Proof. rewrite app_context_assoc. apply X. Qed. -Lemma isType_lift {cf:checker_flags} {Σ : global_env_ext} {n Γ ty} +Lemma isTypeRelOpt_lift {cf:checker_flags} {Σ : global_env_ext} {n Γ ty relopt} (isdecl : n <= #|Γ|): wf Σ -> wf_local Σ Γ -> - isType Σ (skipn n Γ) ty -> - isType Σ Γ (lift0 n ty). + isTypeRelOpt Σ (skipn n Γ) ty relopt -> + isTypeRelOpt Σ Γ (lift0 n ty) relopt. Proof. intros wfΣ wfΓ wfty. rewrite <- (firstn_skipn n Γ) in wfΓ |- *. assert (n = #|firstn n Γ|). @@ -220,3 +220,12 @@ Proof. eapply (weakening_typing (Γ := skipn n Γ) (Γ' := []) (Γ'' := firstn n Γ) (T := tSort _)); eauto with wf. Qed. + +Lemma isType_lift {cf:checker_flags} {Σ : global_env_ext} {n Γ ty} + (isdecl : n <= #|Γ|): + wf Σ -> wf_local Σ Γ -> + isType Σ (skipn n Γ) ty -> + isType Σ Γ (lift0 n ty). +Proof. + apply isTypeRelOpt_lift => //. +Qed. diff --git a/template-coq/theories/Environment.v b/template-coq/theories/Environment.v index eef08fe8f..89a22ac03 100644 --- a/template-coq/theories/Environment.v +++ b/template-coq/theories/Environment.v @@ -487,7 +487,7 @@ 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. + Relevant) ind.(ind_type)) l. Lemma arities_context_length l : #|arities_context l| = #|l|. Proof. unfold arities_context. now rewrite rev_map_length. Qed. diff --git a/template-coq/theories/EnvironmentTyping.v b/template-coq/theories/EnvironmentTyping.v index bd05e444f..507fcea58 100644 --- a/template-coq/theories/EnvironmentTyping.v +++ b/template-coq/theories/EnvironmentTyping.v @@ -404,6 +404,26 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). inv X; intuition; red; simpl; eauto. Qed. + Definition on_decl (P : context -> term -> typ_or_sort -> Type) Γ d := + P Γ d.(decl_type) (SortRel d.(decl_name).(binder_relevance)) × + option_default (fun body => P Γ body (Typ d.(decl_type))) d.(decl_body) True. + + Lemma nth_error_All_local_env {P Γ n d} : + nth_error Γ n = Some d -> + All_local_env P Γ -> + on_decl P (skipn (S n) Γ) d. + Proof using Type. + intro Heq. + induction 1 in n, Heq |- *. + - destruct n; simpl in Heq; discriminate. + - destruct n. + + inv Heq. simpl. split => //. + + simpl in Heq. simpl. apply IHX => //. + - destruct n. + + inv Heq. simpl. split => //. + + simpl in Heq. simpl. apply IHX => //. + Qed. + Definition on_def_type (P : context -> term -> typ_or_sort -> Type) Γ d := P Γ d.(dtype) (SortRel d.(dname).(binder_relevance)). @@ -512,12 +532,37 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). End TypeLocalOver. Derive Signature for All_local_env_over_gen. + Notation Proploc_conj P Q := (fun Γ t T => P Γ t T × Q Γ t T). + + Lemma All_local_env_over_gen_2 checking sorting cproperty sproperty Γ wfΓ : + let cproperty_full Γ wfΓ t T check := cproperty Γ t T in + let sproperty_full Γ wfΓ t r check := sproperty Γ t r check in + All_local_env_over_gen checking sorting cproperty_full sproperty_full Γ wfΓ -> + All_local_env (lift_judgment (Proploc_conj checking cproperty) (fun Γ t r => { s : sorting Γ t r & sproperty Γ t r s })) Γ. + Proof. + intros cfull sfull. + induction 1; constructor => //. + all: cbn in tu |- *; eauto. + Qed. + Definition All_local_env_over typing property Σ := (All_local_env_over_gen (typing Σ) (infer_sort (typing_sort typing) Σ) (property Σ) (fun Γ H t relopt tu => property _ _ H _ _ tu.π2.2)). Definition All_local_env_over_sorting checking sorting cproperty (sproperty : forall Σ Γ _ t s, sorting Σ Γ t s -> Type) Σ := (All_local_env_over_gen (checking Σ) (infer_sort sorting Σ) (cproperty Σ) (fun Γ H t relopt tu => sproperty _ Γ H t tu.π1 tu.π2.2)). + Lemma All_local_env_over_2 typing property (Σ : global_env_ext) (Γ : context) (wfΓ : All_local_env (lift_typing typing Σ) Γ) : + let property_full Σ Γ wfΓ t T Hty := property Σ Γ t T in + All_local_env_over typing property_full Σ Γ wfΓ -> + All_local_env (lift_typing2 typing property Σ) Γ. + Proof. + intros full; unfold full, All_local_env_over. + intro H; eapply All_local_env_over_gen_2 in H. + apply All_local_env_impl with (1 := H); intros. + apply lift_judgment_impl with (1 := X); intros => //. + destruct X0 as ((s & e & Hs) & Hs'); exists s; now repeat split. + Qed. + Section TypeCtxInst. Context (typing : forall (Γ : context), term -> term -> Type). @@ -945,7 +990,8 @@ 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); + (* The ind_relevance part is redundant *) + 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; @@ -1080,7 +1126,8 @@ 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_rel Σ [] idecl.(ind_type) idecl.(ind_relevance); + (** The relevant part is redundant *) + onArity : on_type_rel Σ [] idecl.(ind_type) Relevant; (** The sorts of the arguments contexts of each constructor *) ind_cunivs : list constructor_univs; @@ -1145,10 +1192,8 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT (** *** Typing of constant declarations *) Definition on_constant_decl Σ d := - match d.(cst_body) with - | Some trm => P Σ [] trm (Typ d.(cst_type)) - | None => on_type Σ [] d.(cst_type) - end. + on_type_rel Σ [] d.(cst_type) d.(cst_relevance) × + option_default (fun trm => P Σ [] trm (Typ d.(cst_type))) d.(cst_body) True. Definition on_global_decl Σ kn decl := match decl with @@ -1190,6 +1235,23 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT Definition on_global_env_ext (Σ : global_env_ext) := on_global_env Σ.1 × on_udecl Σ.(universes) Σ.2. + Lemma on_global_env_ext_empty_ext g : + on_global_env g -> on_global_env_ext (empty_ext g). + Proof. + intro H; split => //. + unfold empty_ext, snd. repeat split. + - unfold levels_of_udecl. intros x e. lsets. + - unfold constraints_of_udecl. intros x e. csets. + - unfold satisfiable_udecl, univs_ext_constraints, constraints_of_udecl, fst_ctx, fst => //. + destruct H as ((cstrs & _ & consistent) & decls). + destruct consistent; eexists. + intros v e. specialize (H v e); tea. + - unfold valid_on_mono_udecl, constraints_of_udecl, consistent_extension_on. + intros v sat; exists v; split. + + intros x e. csets. + + intros x e => //. + Qed. + End GlobalMaps. Arguments cstr_args_length {_ Pcmp P Σ mdecl i idecl ind_indices cdecl cunivs}. @@ -1245,7 +1307,9 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT Proof. intros X X0. destruct d; simpl. - - destruct c; simpl. destruct cst_body0; cbn in *; now eapply X. + - destruct 1; split. + * eapply X => //. + * destruct cst_body => //. now eapply X. - intros [onI onP onNP]. constructor; auto. -- eapply Alli_impl; tea. intros. @@ -1386,7 +1450,9 @@ Module DeclarationTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E) (cu, wfΣ) (cu, X0) (cu, IHX0)); clear X. rename X' into X. clear IHX0. destruct d; simpl. - - destruct c; simpl. destruct cst_body0; simpl in *; now eapply X. + - destruct o0; split. + * eapply X => //. + * destruct cst_body => //. now eapply X. - red in o. simpl in *. destruct o0 as [onI onP onNP]. constructor; auto. diff --git a/template-coq/theories/EtaExpand.v b/template-coq/theories/EtaExpand.v index c9853d590..1426f085c 100644 --- a/template-coq/theories/EtaExpand.v +++ b/template-coq/theories/EtaExpand.v @@ -1110,7 +1110,7 @@ Proof. rewrite nth_error_app1. now rewrite nth_error_repeat. rewrite repeat_length. lia. - cbn. econstructor; eauto. * unfold map_branches. solve_all. - clear -X1 H8. + clear -X1 H9. set (Γ'' := map _ Γ'). cbn. enough (All (expanded Σ0 Γ'') (map (eta_expand (declarations Σ0) Γ') (pparams p ++ indices))). now rewrite map_app in X; eapply All_app in X as []. @@ -1132,7 +1132,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. 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 *. @@ -1383,11 +1383,11 @@ Lemma eta_expand_global_decl_expanded {cf : checker_flags} g kn d : Proof. intros wf ond. destruct d; cbn in *. - - unfold on_constant_decl in ond. + - destruct ond as (onty & onbo). destruct c as [na body ty rel]; cbn in *. destruct body. constructor => //; cbn. - apply (eta_expand_expanded (Σ := g) [] [] t na wf ond). constructor. - destruct ond as (s & e & Hs). constructor => //. + apply (eta_expand_expanded (Σ := g) [] [] t na wf onbo). constructor. + destruct onty as (s & e & Hs). constructor => //. - destruct ond as [onI onP onN onV]. constructor. cbn. eapply eta_expand_context => //. diff --git a/template-coq/theories/Typing.v b/template-coq/theories/Typing.v index 8b9b1850f..598fa92b7 100644 --- a/template-coq/theories/Typing.v +++ b/template-coq/theories/Typing.v @@ -809,6 +809,7 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> (List.rev (ind_params mdecl ,,, ind_indices idecl)@[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 @@ -1189,6 +1190,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 -> @@ -1281,8 +1283,9 @@ Proof. apply IH'; auto. - simpl. simpl in *. destruct d. - + destruct c; simpl in *. - destruct cst_body0; apply lift_typing_impl with (1 := Xg); intros ? Hs. + + destruct Xg; split. 2: destruct c, cst_body0 => //. + 1: rename o1 into Xg. 2: rename o2 into Xg. + all: apply lift_typing_impl with (1 := Xg); intros ? Hs. all: specialize (IH ((Σ', udecl); wfΣ; _; _; _; Hs)). all: forward IH; [constructor 1; simpl; subst Σ' Σg; cbn; lia|]. all: apply IH. @@ -1504,18 +1507,6 @@ Qed. (** * Lemmas about All_local_env *) -Lemma nth_error_All_local_env {P Γ n} (isdecl : n < #|Γ|) : - All_local_env P Γ -> - on_some (on_local_decl P (skipn (S n) Γ)) (nth_error Γ n). -Proof. - induction 1 in n, isdecl |- *. red; simpl. - - destruct n; simpl; inv isdecl. - - destruct n. red; simpl. red. simpl. apply t0. - simpl. apply IHX. simpl in isdecl. lia. - - destruct n. auto. - apply IHX. simpl in *. lia. -Qed. - Arguments on_global_env {cf} Pcmp P !g. Lemma lookup_on_global_env `{checker_flags} {Pcmp P} {Σ : global_env} {c decl} : diff --git a/template-coq/theories/TypingWf.v b/template-coq/theories/TypingWf.v index 24ea31e0c..3696d2dc6 100644 --- a/template-coq/theories/TypingWf.v +++ b/template-coq/theories/TypingWf.v @@ -793,11 +793,12 @@ Section WfRed. Proof using Type. intros wΣ h. unfold declared_constant in h. - destruct (lookup_on_global_env wΣ h) as [Σ' [wΣ' [ext h']]]. - simpl in h'. - destruct decl as [ty [bo|]]. all: cbn in *. - - destruct h'. intuition eauto using wf_extends. - - destruct h'. intuition eauto using wf_extends. + destruct (lookup_on_global_env wΣ h) as [Σ' [wΣ' [ext (hty & hbo)]]]. + split. + - rewrite /on_type_rel /wf_decl_pred in hty. destruct hty. eauto using wf_extends. + - destruct cst_body => //. + rewrite /option_default /wf_decl_pred in hbo. + rewrite /on_option. destruct hbo. eauto using wf_extends. Qed. Lemma wf_it_mkProd_or_LetIn_inv (Σ' : global_env_ext) Γ (wfΓ : wf_local Σ' Γ) @@ -879,9 +880,9 @@ Section TypingWf. induction X1; auto. apply IHX1. apply wf_subst. now destruct p0. destruct p. now inv w. - split. wf. apply wf_subst_instance. wf. - destruct (lookup_on_global_env X H) as [Σ' [wfΣ' [ext prf]]]; eauto. - red in prf. destruct decl; destruct cst_body0; red in prf; simpl in *; wf. - destruct prf as [s []]. wf. + destruct (lookup_on_global_env X H) as [Σ' [wfΣ' [ext (prfty & prfbo)]]]; eauto. + rewrite /on_type_rel //= in prfty. + destruct prfty as (? & _ & []). wf. - split. wf. apply wf_subst_instance. eapply declared_inductive_wf; eauto. diff --git a/template-coq/theories/Universes.v b/template-coq/theories/Universes.v index 431e6db08..2073ea719 100644 --- a/template-coq/theories/Universes.v +++ b/template-coq/theories/Universes.v @@ -1713,15 +1713,17 @@ Section Univ. Definition eq_levelalg φ (u u' : LevelAlgExpr.t) := if check_univs then eq0_levelalg φ u u' else True. - Definition eq_universe_ {CS} eq_levelalg (φ: CS) s s' := + Definition eq_universe_ {CS} prop_conv_type eq_levelalg (φ: CS) s s' := match s, s' with | Universe.lProp, Universe.lProp | Universe.lSProp, Universe.lSProp => True | Universe.lType u, Universe.lType u' => eq_levelalg φ u u' + | Universe.lType _, Universe.lProp + | Universe.lProp, Universe.lType _ => prop_conv_type | _, _ => False end. - Definition eq_universe := eq_universe_ eq_levelalg. + Definition eq_universe := eq_universe_ False eq_levelalg. Definition lt_levelalg := leq_levelalg_n 1. Definition leq_levelalg := leq_levelalg_n 0. @@ -2231,8 +2233,6 @@ Section no_prop_leq_type. Qed. Lemma leq_universe_props s1 s2 : - check_univs -> - consistent ϕ -> leq_universe ϕ s1 s2 -> match s1, s2 with | Universe.lProp, Universe.lProp => True @@ -2249,47 +2249,36 @@ Section no_prop_leq_type. Qed. Lemma leq_universe_prop_r s1 s2 : - check_univs -> - consistent ϕ -> leq_universe ϕ s1 s2 -> Universe.is_prop s2 -> Universe.is_prop s1. Proof using Type. - intros Hcf cu. destruct s2; cbn; [ | absurd | absurd]. destruct s1; cbn; [ auto | absurd | absurd]. Qed. Lemma leq_universe_sprop_r s1 s2 : - check_univs -> - consistent ϕ -> leq_universe ϕ s1 s2 -> Universe.is_sprop s2 -> Universe.is_sprop s1. Proof using Type. - intros Hcf cu. destruct s2; cbn; [ absurd | | absurd]. destruct s1; cbn; [ absurd | auto | absurd]. Qed. Lemma leq_universe_prop_no_prop_sub_type s1 s2 : - check_univs -> prop_sub_type = false -> - consistent ϕ -> leq_universe ϕ s1 s2 -> Universe.is_prop s1 -> Universe.is_prop s2. Proof using Type. - intros Hcf ps cu. + intros ps. destruct s1; cbn; [ | absurd | absurd]. rewrite ps. destruct s2; cbn; [ auto | absurd | absurd]. Qed. Lemma leq_universe_sprop_l s1 s2 : - check_univs -> - consistent ϕ -> leq_universe ϕ s1 s2 -> Universe.is_sprop s1 -> Universe.is_sprop s2. Proof using Type. - intros Hcf cu. destruct s1; cbn; [ absurd | | absurd]. destruct s2; cbn; [ absurd | auto | absurd]. Qed. @@ -2314,37 +2303,53 @@ Section no_prop_leq_type. Qed. Lemma leq_prop_is_prop_sprop {s1 s2} : - check_univs -> prop_sub_type = false -> - consistent ϕ -> leq_universe ϕ s1 s2 -> is_propositional s1 <-> is_propositional s2. Proof using Type. - intros Hcf ps cu. + intros ps. destruct s1, s2; cbn; try absurd; intros H; split; trivial. now rewrite ps in H. Qed. Lemma is_prop_gt s1 s2 : - check_univs -> - consistent ϕ -> leq_universe ϕ (Universe.super s1) s2 -> Universe.is_prop s2 -> False. Proof using Type. - intros Hcf cu Hleq Hprop. + intros Hleq Hprop. apply leq_universe_prop_r in Hleq; tas. now destruct s1. Qed. Lemma is_sprop_gt s1 s2 : - check_univs -> - consistent ϕ -> leq_universe ϕ (Universe.super s1) s2 -> Universe.is_sprop s2 -> False. Proof using Type. - intros Hcf cu Hleq Hprop. + intros Hleq Hprop. apply leq_universe_sprop_r in Hleq; tas. now destruct s1. Qed. + Lemma is_prop_superE s : Universe.is_prop (Universe.super s) -> False. + Proof using Type. + destruct s => //. + Qed. + + Lemma is_sprop_superE s : Universe.is_sprop (Universe.super s) -> False. + Proof using Type. + destruct s => //. + Qed. + + Lemma is_prop_prod {s s'} : Universe.is_prop s' -> Universe.is_prop (Universe.sort_of_product s s'). + Proof using Type. + intros isp. + unfold Universe.sort_of_product. rewrite isp. auto. + Qed. + + Lemma is_sprop_prod {s s'} : Universe.is_sprop s' -> Universe.is_sprop (Universe.sort_of_product s s'). + Proof using Type. + intros isp. + unfold Universe.sort_of_product. rewrite isp orb_true_r. auto. + Qed. + End no_prop_leq_type. diff --git a/template-coq/theories/common/uGraph.v b/template-coq/theories/common/uGraph.v index 53e566145..64558e06b 100644 --- a/template-coq/theories/common/uGraph.v +++ b/template-coq/theories/common/uGraph.v @@ -1992,7 +1992,7 @@ Section CheckLeq. leq_universe_n_ (fun n φ u u' => if check_univs then gc_leq0_levelalg_n n φ u u' else True) 0. Definition gc_eq_universe := - eq_universe_ (fun φ u u' => if check_univs then gc_eq0_levelalg φ u u' else True). + eq_universe_ False (fun φ u u' => if check_univs then gc_eq0_levelalg φ u u' else True). Let levels_declared_univ (u : Universe.t) := match u with From 120be9db727b5e6dd5faf46424220f8bd00d92a4 Mon Sep 17 00:00:00 2001 From: Yann Leray Date: Thu, 19 May 2022 18:57:23 +0200 Subject: [PATCH 06/17] Factor out admissions, safechecker --- pcuic/theories/PCUICRelevanceTyp.v | 59 ++- pcuic/theories/PCUICTyping.v | 9 +- .../theories/PCUICRetypingEnvIrrelevance.v | 15 +- safechecker/theories/PCUICSafeChecker.v | 293 +++++++------ safechecker/theories/PCUICSafeRetyping.v | 50 ++- safechecker/theories/PCUICTypeChecker.v | 393 +++++++++++++++--- 6 files changed, 598 insertions(+), 221 deletions(-) diff --git a/pcuic/theories/PCUICRelevanceTyp.v b/pcuic/theories/PCUICRelevanceTyp.v index 9cd1446b5..4fa297117 100644 --- a/pcuic/theories/PCUICRelevanceTyp.v +++ b/pcuic/theories/PCUICRelevanceTyp.v @@ -28,17 +28,17 @@ Proof. now eapply cumul_sort_relevance. Qed. -Lemma ind_arity_relevant {cf : checker_flags} (Σ : global_env) mdecl idecl : - wf Σ -> +Lemma ind_arity_relevant {cf : checker_flags} (Σ : global_env_ext) mdecl idecl : let ind_type := it_mkProd_or_LetIn (ind_params mdecl) (it_mkProd_or_LetIn (ind_indices idecl) (tSort (ind_sort idecl))) in - isType (Σ, ind_universes mdecl) [] ind_type -> - isTypeRel (Σ, ind_universes mdecl) [] ind_type Relevant. + isType Σ [] ind_type -> + isTypeRel Σ [] ind_type Relevant. Proof. - intros wfΣ ind_type (s & e & Hs). + assert (wfΣ : wf Σ) by admit. + intros ind_type (s & e & Hs). eexists; split => //; tea. rewrite /ind_type in Hs. rewrite -it_mkProd_or_LetIn_app in Hs. - assert (wfΓ : wf_local (Σ, ind_universes mdecl) []) by pcuic. + assert (wfΓ : wf_local Σ []) by pcuic. revert wfΓ Hs. generalize ((ind_indices idecl ++ ind_params mdecl : context)) as l, ([]: context) as Γ. induction l using rev_ind. * cbn. intros Γ wfΓ Hty; apply inversion_Sort in Hty as (_ & _ & le); tea. @@ -48,7 +48,7 @@ Proof. * rewrite it_mkProd_or_LetIn_app; cbn. destruct x, decl_body; cbn. - intros Γ wfΓ Hty; pose proof Hty; apply inversion_LetIn in Hty as (? & ? & ? & ? & ? & ? & le); tea. - assert (wfΓ' : wf_local (Σ, ind_universes mdecl) (Γ,, vdef decl_name t decl_type)) by auto with pcuic. + assert (wfΓ' : wf_local Σ (Γ,, vdef decl_name t decl_type)) by auto with pcuic. eapply IHl. apply wfΓ'. econstructor; eauto with pcuic. constructor; eauto with pcuic. apply ws_cumul_pb_LetIn_l_inv in le. @@ -62,7 +62,7 @@ Proof. change (tSort _) with (lift0 1 (tSort s)). eapply (weakening_ws_cumul_pb (Γ' := []) (Γ'' := [vdef _ _ _]) le). fvs. - intros Γ wfΓ Hty; pose proof Hty; apply inversion_Prod in Hty as (? & s' & ? & ? & ? & le); tea. - assert (wfΓ' : wf_local (Σ, ind_universes mdecl) (Γ,, vass decl_name decl_type)) by auto with pcuic. + assert (wfΓ' : wf_local Σ (Γ,, vass decl_name decl_type)) by auto with pcuic. eapply IHl. apply wfΓ'. econstructor; eauto with pcuic. constructor; eauto with pcuic. eapply cumulAlgo_cumulSpec. @@ -71,7 +71,32 @@ Proof. constructor. 1-3: fvs. constructor. apply leq_universe_product. Unshelve. all: eauto. fvs. unfold app_context, app; fold (snoc Γ (vdef decl_name t decl_type)). fvs. -Qed. +Admitted. + +Lemma isTypeRel_cstr_type {cf : checker_flags} Σ mdecl idecl cdecl n : + 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 (typing Σ) (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 H -> e1 <- c (s & _ & Hs); exists s; split => //=. + (* epose proof (nth_error_arities_context mdecl (#|ind_bodies mdecl| - S cstr.1.(inductive_ind)) idecl _). + Unshelve. 2: { rewrite nth_error_rev. len. apply nth_error_Some_length in di. lia. rewrite List.rev_involutive. replace _ with (inductive_ind cstr.1); tea. len. apply nth_error_Some_length in di. lia. } + rewrite e1 in H; clear e1. *) + admit. +Admitted. Lemma isType_of_constructor {cf : checker_flags} Σ mdecl idecl cdecl cstr u Γ : wf Σ.1 -> wf_local Σ Γ -> declared_constructor Σ.1 cstr mdecl idecl cdecl -> @@ -126,6 +151,19 @@ Lemma isType_of_projection {cf : checker_flags} Σ mdecl idecl cdecl pdecl p c a Σ;;; Γ |- c : mkApps (tInd (proj_ind p) u) args -> #|args| = ind_npars mdecl -> isTypeRel Σ Γ (subst0 (c :: List.rev args) (proj_type pdecl)@[u]) pdecl.(proj_relevance). Proof. + admit. +Admitted. + +Lemma apply_predctx {cf : checker_flags} Σ Γ ci p indices c ps mdecl idecl {wfΣ : wf Σ.1} : + wf_predicate mdecl idecl p -> + All2 (PCUICEquality.compare_decls eq eq) (pcontext p) (ind_predicate_context ci mdecl idecl) -> + ctx_inst (typing Σ) Γ (pparams p ++ indices) (List.rev (ind_params mdecl,,, ind_indices idecl)@[puinst p]) -> + Σ;;; Γ |- c : mkApps (tInd ci (puinst p)) (pparams p ++ indices) -> + let predctx := case_predicate_context ci mdecl idecl p in + let ptm := it_mkLambda_or_LetIn predctx (preturn p) in + Σ;;; Γ,,, predctx |- preturn p : tSort ps -> Σ ;;; Γ |- mkApps ptm (indices ++ [c]) : tSort ps. +Proof. + admit. Admitted. @@ -270,7 +308,8 @@ Proof using Type. rewrite H. eassert (isTypeRel (Σ, φ) _ (type_of_constructor mdecl cdecl (ind, i) u) (idecl.(ind_relevance))) by (apply isType_of_constructor; tea). eapply isTypeRel2_relevance; tea. apply wf. - - assert (Σ ;;; Γ |- mkApps ptm (indices ++ [c]) : tSort ps) by admit. + - assert (Σ ;;; Γ |- mkApps ptm (indices ++ [c]) : tSort ps). + { apply apply_predctx => //. apply ctx_inst_impl with (1 := X5) => ??[] //. } split; intro Hr => //. + unfold relevance_of_term in Hr; subst rel. exists ps; split => //. diff --git a/pcuic/theories/PCUICTyping.v b/pcuic/theories/PCUICTyping.v index 7d8247b53..2da8507f8 100644 --- a/pcuic/theories/PCUICTyping.v +++ b/pcuic/theories/PCUICTyping.v @@ -337,7 +337,14 @@ Definition isType_welltyped {cf Σ} {Γ T} Proof. intros []. now econstructor. Qed. -Global Hint Resolve isType_welltyped : pcuic. + +Definition isTypeRel_welltyped {cf Σ} {Γ T rel} + : isTypeRel Σ Γ T rel -> welltyped Σ Γ T. +Proof. + intros (s & e & Hs). now econstructor. +Qed. + +Global Hint Resolve isType_welltyped isTypeRel_welltyped : pcuic. Definition isWfArity {cf:checker_flags} Σ (Γ : context) T := (isType Σ Γ T × { ctx & { s & (destArity [] T = Some (ctx, s)) } }). diff --git a/safechecker/theories/PCUICRetypingEnvIrrelevance.v b/safechecker/theories/PCUICRetypingEnvIrrelevance.v index ef9ab5500..9364a1843 100644 --- a/safechecker/theories/PCUICRetypingEnvIrrelevance.v +++ b/safechecker/theories/PCUICRetypingEnvIrrelevance.v @@ -60,7 +60,7 @@ Qed. Section infer_irrel. Context {cf} {nor : normalizing_flags} {X_type : abstract_env_ext_impl} {X : X_type.π1} - {X_type' : abstract_env_ext_impl} {X' : X_type'.π1}. + {X_type' : abstract_env_ext_impl} {X' : X_type'.π1} {relopt : option relevance}. Context (hl : Hlookup X_type X X_type' X'). Definition same_prod (Γ : context) {T} @@ -371,14 +371,14 @@ Section infer_irrel. Lemma infer_as_sort_irrel {Γ t} (wf : forall Σ : global_env_ext, abstract_env_ext_rel X Σ -> ∥ wf_local Σ Γ ∥) (wf' : forall Σ : global_env_ext, abstract_env_ext_rel X' Σ -> ∥ wf_local Σ Γ ∥) - (wi : forall Σ : global_env_ext, abstract_env_ext_rel X Σ -> well_sorted Σ Γ t) - (wi' : forall Σ : global_env_ext, abstract_env_ext_rel X' Σ -> well_sorted Σ Γ t) - hp hp' : hp.π1 = hp'.π1 -> (infer_as_sort X_type X wf wi hp).π1 = (infer_as_sort X_type' X' wf' wi' hp').π1. + (wi : forall Σ : global_env_ext, abstract_env_ext_rel X Σ -> well_sorted Σ Γ t relopt) + (wi' : forall Σ : global_env_ext, abstract_env_ext_rel X' Σ -> well_sorted Σ Γ t relopt) + hp hp' : hp.π1 = hp'.π1 -> (infer_as_sort X_type X relopt wf wi hp).π1 = (infer_as_sort X_type' X' relopt wf' wi' hp').π1. Proof using hl. unfold infer_as_sort. unfold PCUICSafeRetyping.principal_type_type. - set (obl := fun Σ wfΣ => PCUICSafeRetyping.infer_as_sort_obligation_1 _ _ _ _ _ _ _ Σ wfΣ). - set (obl' := fun Σ wfΣ => PCUICSafeRetyping.infer_as_sort_obligation_1 _ _ _ _ _ _ _ Σ wfΣ). + set (obl := fun Σ wfΣ => PCUICSafeRetyping.infer_as_sort_obligation_1 _ _ _ _ relopt _ _ _ Σ wfΣ). + set (obl' := fun Σ wfΣ => PCUICSafeRetyping.infer_as_sort_obligation_1 _ _ _ _ relopt _ _ _ Σ wfΣ). clearbody obl obl'. destruct reduce_to_sort as [[s hr]|] eqn:h. 2:bang. destruct (reduce_to_sort _ hp'.π1 _) as [[s' hr']|] eqn:h'. 2:bang. @@ -471,11 +471,12 @@ Proof. - now cbn. - simpl. cbn. f_equal. f_equal; unfold PCUICSafeRetyping.principal_sort_sort. + set (relopt := Some n.(binder_relevance)). set (obl := fun a b => _). set (obl' := fun a b => _). set (obl'' := fun a b => _). set (obl''' := fun a b => _). - cbn. eapply infer_as_sort_irrel => //. now eapply H. + cbn. eapply (infer_as_sort_irrel (relopt := relopt)) => //. now eapply H. eapply infer_as_sort_irrel => //. now eapply H0. - cbn. f_equal. unfold PCUICSafeRetyping.principal_type_type. apply H; eauto. diff --git a/safechecker/theories/PCUICSafeChecker.v b/safechecker/theories/PCUICSafeChecker.v index bb9a556ab..f119eb6f0 100644 --- a/safechecker/theories/PCUICSafeChecker.v +++ b/safechecker/theories/PCUICSafeChecker.v @@ -8,7 +8,7 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICTactics PCUICPosition PCUICCumulativity PCUICSafeLemmata PCUICSN PCUICPretty PCUICArities PCUICConfluence PCUICSize PCUICContextConversion PCUICConversion PCUICWfUniverses - PCUICGlobalEnv + PCUICGlobalEnv PCUICRelevance (* Used for support lemmas *) PCUICInductives PCUICWfUniverses PCUICContexts PCUICSubstitution PCUICSpine PCUICInductiveInversion @@ -293,6 +293,11 @@ Section CheckEnv. Definition check_wf_type (kn : kername) X_ext t : EnvCheck X_env_ext_type (forall Σ : global_env_ext, abstract_env_ext_rel X_ext Σ -> ∥ isType Σ [] t ∥) := wrap_error _ X_ext (string_of_kername kn) (check_isType X_ext_impl X_ext [] (fun _ _ => sq_wfl_nil _) t). + + Definition check_wf_type_rel (kn : kername) X_ext t rel : + EnvCheck X_env_ext_type (forall Σ : global_env_ext, abstract_env_ext_rel X_ext Σ -> ∥ isTypeRel Σ [] t rel ∥) := + wrap_error _ X_ext (string_of_kername kn) (check_isTypeRel X_ext_impl X_ext [] (fun _ _ => sq_wfl_nil _) t rel). + Definition check_wf_judgement kn X_ext t ty : EnvCheck X_env_ext_type (forall Σ : global_env_ext, abstract_env_ext_rel X_ext Σ -> ∥ Σ ;;; [] |- t : ty ∥) @@ -470,17 +475,17 @@ Section CheckEnv. unfold isType. unfold PCUICTypingDef.typing. intros Γ Δ s A h. revert Γ s A h. induction Δ using rev_ind; intros. - simpl in h. eapply inversion_Sort in h as (?&?&?). - eexists; constructor; eauto. apply wfΣ. + eexists; split => //. constructor; eauto. apply wfΣ. - destruct x as [na [b|] ty]; simpl in *; rewrite it_mkProd_or_LetIn_app /= /mkProd_or_LetIn /= in h *. - * eapply inversion_LetIn in h as [s' [? [? [? [? ?]]]]]; auto. - specialize (IHΔ _ _ _ t1) as [s'' vdefty]. - exists s''. + * eapply inversion_LetIn in h as [s' (? & ? & ? & ? & ? & ?)]; auto. + specialize (IHΔ _ _ _ t1) as (s'' & ? & vdefty). + exists s''. split => //. eapply type_Cumul_alt. econstructor; eauto. pcuic. eapply red_cumul. repeat constructor. - * eapply inversion_Prod in h as [? [? [? [? ]]]]; auto. - specialize (IHΔ _ _ _ t0) as [s'' Hs'']. - eexists (Universe.sort_of_product x s''). + * eapply inversion_Prod in h as (? & ? & ? & ? & ? & ?); auto. + specialize (IHΔ _ _ _ t0) as (s'' & e'' & Hs''). + eexists (Universe.sort_of_product x s''); split => //. eapply type_Cumul'; eauto. econstructor; pcuic. pcuic. reflexivity. Qed. @@ -490,13 +495,16 @@ 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_universeb X_ext s with true => ret _ | false => raise (Msg "Ill-formed universe") end - | {| decl_body := None; decl_type := ty |} :: Δ => + | {| decl_name := na; decl_body := None; decl_type := ty |} :: Δ => checkΔ <- check_type_local_ctx X_ext Γ Δ s wfΓ ;; checkty <- check_type_wf_env X_ext (Γ ,,, Δ) _ ty (tSort s) ;; + check <- check_eq_true (relevance_of_sort s == na.(binder_relevance)) (Msg "Wrong relevance") ;; ret _ - | {| decl_body := Some b; decl_type := ty |} :: Δ => + | {| decl_name := na; decl_body := Some b; decl_type := ty |} :: Δ => checkΔ <- check_type_local_ctx X_ext Γ Δ s wfΓ ;; checkty <- check_type_wf_env X_ext (Γ ,,, Δ) _ b ty ;; + let '(tmrel; eqtmrel) := relevance_of_term_abstract_env X_ext_impl X_ext (marks_of_context (Γ ,,, Δ)) b in + check <- check_eq_true (tmrel == na.(binder_relevance)) (Msg "Wrong relevance") ;; ret _ end. Next Obligation. @@ -508,28 +516,36 @@ Section CheckEnv. Qed. Next Obligation. specialize_Σ H. - sq. split; auto. + sq. splits; auto. + apply eqb_eq, check. Qed. Next Obligation. specialize_Σ H. sq. now eapply PCUICContexts.type_local_ctx_wf_local in checkΔ. Qed. Next Obligation. + clear Heq_anonymous. pose proof (abstract_env_ext_wf _ H). specialize_Σ H. sq. split; auto. split; auto. - eapply PCUICValidity.validity in checkty; auto. + eapply PCUICRelevanceTyp.relevance_from_type in checkty. 2: apply H0. + apply checkty in eqtmrel; clear checkty. + apply eqb_eq in check; rewrite -check; clear check. + apply eqtmrel. Qed. Program Fixpoint infer_sorts_local_ctx X_ext Γ Δ (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_name := na; decl_body := None; decl_type := ty |} :: Δ => '(Δs; Δinfer) <- infer_sorts_local_ctx X_ext Γ Δ wfΓ ;; '(tys; tyinfer) <- infer_type_wf_env X_ext (Γ ,,, Δ) _ ty ;; + check <- check_eq_true (relevance_of_sort tys == na.(binder_relevance)) (Msg "Wrong relevance") ;; ret ((tys :: Δs); _) - | {| decl_body := Some b; decl_type := ty |} :: Δ => + | {| decl_name := na; decl_body := Some b; decl_type := ty |} :: Δ => '(Δs; Δinfer) <- infer_sorts_local_ctx X_ext Γ Δ wfΓ ;; checkty <- check_type_wf_env X_ext (Γ ,,, Δ) _ b ty ;; + let '(tmrel; eqtmrel) := relevance_of_term_abstract_env X_ext_impl X_ext (marks_of_context (Γ ,,, Δ)) b in + check <- check_eq_true (tmrel == na.(binder_relevance)) (Msg "Wrong relevance") ;; ret (Δs; _) end. Next Obligation. exact tt. Qed. @@ -537,14 +553,18 @@ Section CheckEnv. specialize_Σ H. sq. now eapply PCUICContexts.sorts_local_ctx_wf_local in Δinfer. Qed. Next Obligation. - specialize_Σ H. sq. split; auto. + specialize_Σ H. sq. splits; auto. apply eqb_eq, check. Qed. Next Obligation. specialize_Σ H. sq. now eapply PCUICContexts.sorts_local_ctx_wf_local in Δinfer. Qed. Next Obligation. - pose proof (abstract_env_ext_wf _ H). specialize_Σ H. sq. split; auto. split; auto. - eapply PCUICValidity.validity in checkty; auto. + clear Heq_anonymous. + pose proof (abstract_env_ext_wf _ H). specialize_Σ H. sq. splits; auto. + eapply PCUICRelevanceTyp.relevance_from_type in checkty. 2: apply H0. + apply checkty in eqtmrel; clear checkty. + apply eqb_eq in check; rewrite -check; clear check. + apply eqtmrel. Qed. Definition cumul_decl Pcmp Σ Γ (d d' : context_decl) : Type := cumul_decls Pcmp Σ Γ Γ d d'. @@ -799,8 +819,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 -> @@ -809,13 +834,15 @@ 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 PCUICRelevanceTyp.ind_arity_relevant => //). rewrite rev_map_cons /=. eapply All_local_env_app; split. constructor; pcuic. eapply All_local_env_impl; eauto. move=> Γ t [] /=. * move=> ty ht. eapply weaken_ctx; eauto. constructor; pcuic. - * move=> [s Hs]; exists s. + * intros relopt H; apply infer_sort_impl with id H => Hty //. eapply weaken_ctx; eauto. constructor; pcuic. Qed. @@ -874,7 +901,7 @@ Section CheckEnv. eapply eqb_eq in eqbindices. eapply eqb_eq in eqbpars. sq; red; cbn. intuition auto. - eexists; eauto. + eexists; split => //; tea. 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. @@ -916,8 +943,7 @@ Section CheckEnv. isType Σ Γ (it_mkProd_or_LetIn Δ T) -> isType Σ (Γ ,,, Δ) T. Proof using Type. - move=> wfX [u Hu]. - exists u. unfold PCUICTypingDef.typing in *. + intros wfΣ H; apply infer_typing_sort_impl with id H => Hu //. now eapply PCUICSpine.inversion_it_mkProd_or_LetIn in Hu. Qed. @@ -926,7 +952,7 @@ Section CheckEnv. ∑ fty s, (Σ ;;; Γ |- f : fty) * typing_spine Σ Γ fty args (tSort s). Proof using Type. - intros [s Hs]. + intros (s & e & Hs). eapply inversion_mkApps in Hs as [fty [Hf Hargs]]; auto. now exists fty, s. Qed. @@ -934,7 +960,7 @@ Section CheckEnv. 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) |}; + Some {| decl_name := {| binder_name := nNamed idecl.(ind_name); binder_relevance := Relevant |}; decl_body := None; decl_type := idecl.(ind_type) |}. Proof using Type. @@ -995,8 +1021,8 @@ Section CheckEnv. eapply type_ws_cumul_pb; pcuic. eapply ws_cumul_pb_eq_le. now symmetry. rewrite subst_telescope_subst_context. cbn in *. have tyhd : Σ ;;; Γ |- hd0 : ty'. - { eapply type_ws_cumul_pb; tea. eapply isType_tProd in isty as []. - pcuic. eapply ws_cumul_pb_eq_le. now symmetry. } + { eapply type_ws_cumul_pb; tea. eapply isType_tProd in isty as []. eapply isType_of_isTypeRel, i. + eapply ws_cumul_pb_eq_le. now symmetry. } eapply X. now len. pcuic. eapply substitution_wf_local; eauto. eapply subslet_ass_tip; tea. @@ -1081,7 +1107,7 @@ Section CheckEnv. This is a quite common operation in tactics. Making this work up-to unfolding of let-ins is the tricky part. *) - Lemma typing_spine_it_mkProd_or_LetIn_ext_list_inv_gen {Σ : global_env_ext} {wfX : wf Σ} {Γ Δ Δ' Δ'' s args s'} : + Lemma typing_spine_it_mkProd_or_LetIn_ext_list_inv_gen {Σ : global_env_ext} {wfX : wf Σ} {Γ Δ Δ' Δ'' : context} {s args s'} : wf_local Σ (Γ ,,, Δ ,,, Δ'') -> typing_spine Σ (Γ ,,, Δ ,,, Δ') (it_mkProd_or_LetIn (lift_context #|Δ ,,, Δ'| 0 (Δ ,,, Δ'')) (tSort s)) (to_extended_list_k Δ #|Δ'| ++ args) (tSort s') -> @@ -1097,7 +1123,7 @@ Section CheckEnv. assert (wf_universe Σ s). { eapply typing_spine_isType_dom in sp. eapply isType_it_mkProd_or_LetIn_inv in sp; auto. - destruct sp as [? Hs]. now eapply inversion_Sort in Hs as [? []]. } + destruct sp as (? & _ & Hs). now eapply inversion_Sort in Hs as [? []]. } destruct d as [na [b|] ty]. simpl in sp; unfold mkProd_or_LetIn in sp; simpl in *. - len in sp. rewrite lift_context_app /= it_mkProd_or_LetIn_app lift_context_app it_mkProd_or_LetIn_app /= @@ -1193,7 +1219,7 @@ Section CheckEnv. welltyped Σ Γ (tProd na ty ty') -> welltyped Σ Γ ty * welltyped Σ (Γ ,, vass na ty) ty'. Proof using Type. - intros [s [s1 [s2 [hty [hty' hcum]]]]%inversion_Prod]; auto. + intros (s & (s1 & s2 & e1 & hty & hty' & hcum)%inversion_Prod); auto. split; eexists; eauto. Qed. @@ -1203,15 +1229,15 @@ Section CheckEnv. welltyped Σ Γ b * welltyped Σ (Γ ,, vdef na b ty) t. Proof using Type. - intros [s [s1 [s2 [hty [hdef [hty' hcum]]]]]%inversion_LetIn]; auto. - split; [split|]; eexists; eauto. + intros (s & (s1 & s2 & e1 & hty & hdef & hty' & hcum)%inversion_LetIn); auto. + splits; eexists; eauto. Qed. Lemma welltyped_letin_red {Σ : global_env_ext} {Γ na b ty t} {wfX : wf Σ} : welltyped Σ Γ (tLetIn na b ty t) -> welltyped Σ Γ (subst0 [b] t). Proof using Type. - intros [s [s1 [s2 [hty [hdef [hty' hcum]]]]]%inversion_LetIn]; auto. + intros (s & (s1 & s2 & e1 & hty & hdef & hty' & hcum)%inversion_LetIn); auto. exists (subst0 [b] s2). now eapply substitution_let in hty'. Qed. @@ -1383,23 +1409,29 @@ Section CheckEnv. 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 |} :: Γ => + | {| decl_name := na; decl_body := Some b; decl_type := ty |} :: Γ => wfΓ <- check_wf_local X_ext Γ ;; wfty <- check_type_wf_env X_ext Γ wfΓ b ty ;; + let '(tmrel; eqtmrel) := relevance_of_term_abstract_env X_ext_impl X_ext (marks_of_context Γ) b in + check <- check_eq_true (tmrel == na.(binder_relevance)) (Msg "Wrong relevance") ;; ret _ - | {| decl_body := None; decl_type := ty |} :: Γ => + | {| decl_name := na; decl_body := None; decl_type := ty |} :: Γ => wfΓ <- check_wf_local X_ext Γ ;; - wfty <- infer_type_wf_env X_ext Γ wfΓ ty ;; + '(s; wfty) <- infer_type_wf_env X_ext Γ wfΓ ty ;; + check <- check_eq_true (relevance_of_sort s == na.(binder_relevance)) (Msg "Wrong relevance") ;; ret _ end. Next Obligation. + clear Heq_anonymous. pose proof (abstract_env_ext_wf _ H); specialize_Σ H. sq. constructor; eauto. - eapply validity in wfty. apply wfty. + apply eqb_eq in check; subst tmrel. + eapply PCUICRelevanceTyp.relevance_from_type; tea. apply H0. Qed. Next Obligation. pose proof (abstract_env_ext_wf _ H); specialize_Σ H. - sq. constructor; auto. now exists wfty. + apply eqb_eq in check; subst. + sq. constructor; auto. now exists x. Qed. Definition wt_indices Σ mdecl indices cs := @@ -1416,7 +1448,7 @@ Section CheckEnv. Qed. (* Now in PCUIC *) - Lemma type_smash {Σ : global_env_ext} {wfΣ : wf Σ} {Γ Δ t T} : + Lemma type_smash {Σ : global_env_ext} {wfΣ : wf Σ} {Γ Δ : context} {t T} : Σ ;;; Γ ,,, Δ |- t : T -> Σ ;;; Γ ,,, smash_context [] Δ |- expand_lets Δ t : expand_lets Δ T. Proof using Type. @@ -1529,20 +1561,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 ∥. + 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. @@ -1560,7 +1593,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. @@ -1577,7 +1609,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. @@ -1692,7 +1723,7 @@ Section CheckEnv. Proof using Type. intros wfΣ. unfold wf_ctx_universes. induction 1 => //. cbn. rewrite IHX andb_true_r. - destruct t0 as [s Hs]. move/typing_wf_universes: Hs => /andP[] //. + destruct t0 as (s & e & Hs). move/typing_wf_universes: Hs => /andP[] //. cbn. move/typing_wf_universes: t1 => /=; rewrite /wf_decl_universes /on_decl_universes /= => -> /= //. Qed. @@ -1849,11 +1880,10 @@ 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));; @@ -1865,8 +1895,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 _ then true else is_assumption_context (cstr_args x)) @@ -1889,7 +1919,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). specialize_Σ H. unfold check_constructor_spec in Hcs; simpl in *. sq. solve_all. @@ -1900,6 +1930,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 PCUICRelevanceTyp.isTypeRel_cstr_type; eauto. - now destruct wtinds. - destruct lets_in_constructor_types; eauto. Qed. @@ -2090,8 +2123,8 @@ End monad_Alli_nth_forall. induction Δ in s |- *; destruct s; simpl; auto. destruct a as [na [b|] ty]. - intros []. eauto. - - intros []; eauto. constructor; eauto. - now eapply typing_wf_universe in t0. + - intros (Hs & e & Ht); eauto. constructor; eauto. + now eapply typing_wf_universe in Ht. Qed. Obligation Tactic := Program.Tactics.program_simplify. @@ -2178,17 +2211,53 @@ End monad_Alli_nth_forall. rewrite -eqvaru in e; discriminate. Qed. + Program Definition check_ind_type X_ext (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")) ;; + 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 (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 exists y. - 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 (mind : kername) (mdecl : mutual_inductive_body) @@ -2200,41 +2269,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 _) ;; - 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. + idecl.(ind_indices) cs _ idecl.(ind_sort) _) ;; + onindices <- check_indices X mdecl mind _ mdeclvar idecl.(ind_indices) _ ;; + ret _. Next Obligation. intros ? ?. pose proof (abstract_env_ext_wf _ H). @@ -2253,8 +2295,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. @@ -2269,36 +2311,26 @@ 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]. - 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. - refine - {| ind_arity_eq := _; onArity := _; - ind_cunivs := cs; - onConstructors := oncstrs; - onProjections := onprojs; - 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. - - red. red. - eapply nth_error_all in wfars; eauto; simpl in wfars. - destruct wfars as [s Hs]. now exists s. - - now apply eqb_eq in eqsort; subst. + rename x into cs. destruct Σ as [Σ ext]. + pose proof (abstract_env_exists X) as [[Σ0 wfΣ0]]. + 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 := e; onArity := _; + ind_cunivs := cs; + onConstructors := oncstrs; + onProjections := onprojs; + ind_sorts := onsorts; + ind_relevance_compat := eqrel; + onIndices := _; |}. + - rewrite e in isty |- *; apply PCUICRelevanceTyp.ind_arity_relevant => //. - erewrite (abstract_env_ext_irr _ _ pf); eauto. Unshelve. eauto. Qed. @@ -2307,15 +2339,19 @@ End monad_Alli_nth_forall. kn (d : global_decl) (pf : check_wf_env_ext_prop X X_ext (universes_decl_of_decl d)) : EnvCheck X_env_ext_type (forall Σ, abstract_env_ext_rel X_ext Σ -> ∥ on_global_decl cumulSpec0 (lift_typing typing) Σ kn d ∥) := + let id := string_of_kername kn in match d with | ConstantDecl cst => match cst.(cst_body) with | Some term => + let '(tmrel; eqtmrel) := relevance_of_term_abstract_env X_ext_impl X_ext [] term in + check <- wrap_error _ X_ext id (check_eq_true (tmrel == cst.(cst_relevance)) (Msg "Wrong relevance")) ;; check_wf_judgement kn X_ext term cst.(cst_type) ;; ret _ - | None => check_wf_type kn X_ext cst.(cst_type) ;; ret _ + | None => + s <- check_wf_type_rel kn X_ext cst.(cst_type) cst.(cst_relevance);; + ret _ end | InductiveDecl mdecl => - let id := string_of_kername kn in check_var <- @check_variance X kn (ind_universes mdecl) (ind_variance mdecl) _ ;; check_pars <- wrap_error _ X_ext id (check_context_wf_env X_ext (ind_params mdecl)) ;; check_npars <- wrap_error _ X_ext id @@ -2326,11 +2362,14 @@ End monad_Alli_nth_forall. ret (Build_on_inductive_sq check_bodies check_pars check_npars _) end. Next Obligation. - specialize_Σ H0. sq. unfold on_constant_decl; rewrite <- Heq_anonymous. - eassumption. + clear Heq_anonymous. + specialize_Σ H0. pose proof (hΣ _ _ H0). sq. unfold on_constant_decl; rewrite <- Heq_anonymous0. + split => //. + apply eqb_eq in check; subst tmrel. + eapply PCUICRelevanceTyp.relevance_from_type; tea. Qed. Next Obligation. - specialize_Σ H0. sq. unfold on_constant_decl. rewrite <- Heq_anonymous; tea. + specialize_Σ H. sq. unfold on_constant_decl. rewrite <- Heq_anonymous; split => //. Qed. Next Obligation. edestruct pf as [? ?]; specialize_Σ H. now pose proof (abstract_env_ext_wf _ H0). diff --git a/safechecker/theories/PCUICSafeRetyping.v b/safechecker/theories/PCUICSafeRetyping.v index d26c0c2d7..89e688d26 100644 --- a/safechecker/theories/PCUICSafeRetyping.v +++ b/safechecker/theories/PCUICSafeRetyping.v @@ -66,13 +66,13 @@ Qed. Inductive wellinferred {cf: checker_flags} Σ Γ t : Prop := | iswellinferred T : Σ ;;; Γ |- t ▹ T -> wellinferred Σ Γ t. -Definition well_sorted {cf:checker_flags} Σ Γ T := - ∥ ∑ u, Σ ;;; Γ |- T ▹□ u ∥. +Definition well_sorted {cf:checker_flags} Σ Γ T relopt := + ∥ infer_sort infering_sort Σ Γ T relopt ∥. -Lemma well_sorted_wellinferred {cf:checker_flags} {Σ Γ T} : - well_sorted Σ Γ T -> wellinferred Σ Γ T. +Lemma well_sorted_wellinferred {cf:checker_flags} {Σ Γ T relopt} : + well_sorted Σ Γ T relopt -> wellinferred Σ Γ T. Proof. - intros [[s []]]. + intros [[s [e []]]]. now econstructor. Qed. @@ -153,10 +153,10 @@ Context {cf : checker_flags} {nor : normalizing_flags}. Definition on_subterm P Pty Γ t : Type := match t with - | tProd na t b => Pty Γ t * Pty (Γ ,, vass na t) b + | tProd na t b => Pty Γ t (Some na.(binder_relevance)) * Pty (Γ ,, vass na t) b None | tLetIn na d t t' => - Pty Γ t * P Γ d * P (Γ ,, vdef na d t) t' - | tLambda na t b => Pty Γ t * P (Γ ,, vass na t) b + Pty Γ t (Some na.(binder_relevance)) * P Γ d * P (Γ ,, vdef na d t) t' + | tLambda na t b => Pty Γ t (Some na.(binder_relevance)) * P (Γ ,, vass na t) b | _ => True end. @@ -174,19 +174,19 @@ Qed. #[local] Definition principal_type Γ t := ∑ T : term, forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ Σ ;;; Γ |- t ▹ T ∥. - #[local] Definition principal_sort Γ T := - ∑ u, forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ Σ ;;; Γ |- T ▹□ u ∥. + #[local] Definition principal_sort Γ T relopt := + ∑ u, forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ isSortRelOpt u relopt × Σ ;;; Γ |- T ▹□ u ∥. #[local] Definition principal_type_type {Γ t} (wt : principal_type Γ t) : term := projT1 wt. - #[local] Definition principal_sort_sort {Γ T} (ps : principal_sort Γ T) : Universe.t + #[local] Definition principal_sort_sort {Γ T relopt} (ps : principal_sort Γ T relopt) : Universe.t := projT1 ps. #[local] Coercion principal_type_type : principal_type >-> term. #[local] Coercion principal_sort_sort : principal_sort >-> Universe.t. - Program Definition infer_as_sort {Γ T} + Program Definition infer_as_sort {Γ T} relopt (wfΓ : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ wf_local Σ Γ ∥) - (wf : forall Σ (wfΣ : abstract_env_ext_rel X Σ), well_sorted Σ Γ T) - (tx : principal_type Γ T) : principal_sort Γ T := + (wf : forall Σ (wfΣ : abstract_env_ext_rel X Σ), well_sorted Σ Γ T relopt) + (tx : principal_type Γ T) : principal_sort Γ T relopt := match @reduce_to_sort cf nor _ X Γ tx _ with | Checked_comp (u;_) => (u;_) | TypeError_comp e _ => ! @@ -205,8 +205,12 @@ Qed. pose proof (s Σ wfΣ) as s'. cbn in *. sq. - econstructor ; tea. - now eapply closed_red_red. + assert (Hty : Σ;;; Γ |- T ▹□ u). + { econstructor ; tea. now eapply closed_red_red. } + split => //. + inversion wf. destruct X0 as (ss & e & Hs). + assert (u = ss) by (destruct (hΣ _ wfΣ); eapply infering_sort_sort; tea). + now subst ss. Qed. Next Obligation. clear Heq_anonymous. @@ -215,7 +219,7 @@ Qed. pose proof (s Σ wfΣ) as s'. cbn in *. sq. - destruct wf as [[? i]], (hΣ _ wfΣ) as [wΣ]. + destruct wf as [[? [_ i]]], (hΣ _ wfΣ) as [wΣ]. eapply infering_sort_infering in i ; eauto. eapply wildcard'. exists x0. intros. erewrite(abstract_env_ext_irr _ _ wfΣ); eauto. @@ -302,9 +306,9 @@ Qed. infer Γ wfΓ (tProd n ty b) wt := let wfΓ' : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ wf_local Σ (Γ ,, vass n ty) ∥ := _ in let ty1 := infer Γ wfΓ ty (fun a b => (welltyped_subterm (wt a b)).1) in - let s1 := infer_as_sort wfΓ (fun a b => (welltyped_subterm (wt a b)).1) ty1 in + let s1 := infer_as_sort (Some n.(binder_relevance)) wfΓ (fun a b => (welltyped_subterm (wt a b)).1) ty1 in let ty2 := infer (Γ ,, vass n ty) wfΓ' b (fun a b => (welltyped_subterm (wt a b)).2) in - let s2 := infer_as_sort wfΓ' (fun a b => (welltyped_subterm (wt a b)).2) ty2 in + let s2 := infer_as_sort None wfΓ' (fun a b => (welltyped_subterm (wt a b)).2) ty2 in ret (tSort (Universe.sort_of_product s1 s2)); infer Γ wfΓ (tLambda n t b) wt := @@ -399,6 +403,7 @@ Qed. Next Obligation. cbn ; intros. destruct s1, s2. cbn. specialize_Σ wfΣ. sq. + destruct s, s0. now constructor. Defined. @@ -406,7 +411,8 @@ Qed. pose (hΣ _ wfΣ). specialize_Σ wfΣ. inversion wt. sq. inversion X0 ; subst. constructor ; tea. - now eapply infering_sort_isType. + eexists; split; tea. + now eapply infering_sort_typing. Defined. Next Obligation. case t2 as []. intros; cbn. specialize_Σ wfΣ. @@ -420,7 +426,7 @@ Qed. pose (hΣ _ wfΣ). specialize_Σ wfΣ. inversion wt. sq. inversion X0 ; subst. constructor ; tea. - 1: now eapply infering_sort_isType. + 1: eexists; split; tea; now eapply infering_sort_typing. apply checking_typing ; eauto. now eapply infering_sort_isType. Defined. @@ -822,7 +828,7 @@ Qed. cbn in ns. clear ns. specialize (wt _ wfΣ). destruct T as [T HT]. cbn in *. destruct (HT _ wfΣ) as [[hty hp]]. - eapply validity in hty. destruct wt as [[s Hs]]. + eapply validity in hty. destruct wt as [[s [e Hs]]]. red in hp. specialize (hp _ Hs). eapply ws_cumul_pb_Sort_r_inv in hp as [s' [hs' _]]. eapply (H s' hs'). diff --git a/safechecker/theories/PCUICTypeChecker.v b/safechecker/theories/PCUICTypeChecker.v index 67bee2124..20392df90 100644 --- a/safechecker/theories/PCUICTypeChecker.v +++ b/safechecker/theories/PCUICTypeChecker.v @@ -9,13 +9,13 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICTactics PCUICPretty PCUICArities PCUICConfluence PCUICSize PCUICContextConversion PCUICContextConversionTyp PCUICConversion PCUICWfUniverses - PCUICGlobalEnv PCUICSigmaCalculus + PCUICGlobalEnv PCUICSigmaCalculus PCUICRelevance (* Used for support lemmas *) PCUICInductives PCUICWfUniverses PCUICOnFreeVars PCUICWellScopedCumulativity PCUICContexts PCUICSubstitution PCUICSpine PCUICInductiveInversion PCUICClosed PCUICClosedTyp - PCUICUnivSubstitutionConv PCUICUnivSubstitutionTyp. + PCUICUnivSubstitutionConv PCUICUnivSubstitutionTyp PCUICRelevanceTyp. From MetaCoq.PCUIC Require Import BDTyping BDToPCUIC BDFromPCUIC BDUnique. @@ -417,6 +417,39 @@ Section Typecheck. Unshelve. eauto. 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 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_isTypeRel Γ HΓ T rel := + s <- infer_type Γ HΓ T ;; + check_eq_true ((relevance_of_sort s.π1) == rel) (Msg "Wrong relevance") ;; + ret _. + Next Obligation. + pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. + exists s; split => //=. + - apply eqb_eq in i => //. + - eapply infering_sort_typing; eauto. + Qed. + Next Obligation. + destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. + pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. + apply absurd. case: eqb_spec => a //. exfalso; apply a. + eapply infering_sort_typing in X0; eauto. + destruct H as (s' & e' & Hs'). + rewrite <- e'. + eapply cumul_sort_relevance; tea. + Qed. + Next Obligation. + destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. + pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. + apply absurd. + eapply isType_of_isTypeRel, 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 := @@ -475,18 +508,23 @@ Section Typecheck. 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 |} :: Γ) := + check_context ({| decl_name := na; decl_body := None; decl_type := A |} :: Γ) := HΓ <- check_context Γ ;; - infer_type Γ HΓ A ;; + infer_isTypeRel Γ HΓ A na.(binder_relevance) ;; ret _ ; - check_context ({| decl_body := Some t; decl_type := A |} :: Γ) := + check_context ({| decl_name := na; decl_body := Some t; decl_type := A |} :: Γ) := HΓ <- check_context Γ ;; - infer_isType Γ HΓ A ;; + infer_isTypeRel Γ HΓ A na.(binder_relevance) ;; (* Or do we check t's syntactic relevance ? *) bdcheck Γ HΓ t A _ ;; ret _. + Next Obligation. + specialize_Σ wfΣ. sq. + now eapply isType_of_isTypeRel. + Qed. Next Obligation. pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. econstructor ; tea. + apply isType_of_isTypeRel in s. now eapply checking_typing. Qed. Next Obligation. @@ -505,22 +543,19 @@ Section Typecheck. now inversion H. Qed. Next Obligation. - pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. econstructor; tas. - eexists. - now eapply infering_sort_typing. - 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. + pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. econstructor; tas. Qed. Next Obligation. eapply absurd. intros. specialize_Σ wfΣ. sq. now inversion H. Qed. + Next Obligation. + destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. + pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. eapply absurd. + inversion H ; subst. + intros. erewrite (abstract_env_ext_irr _ _ wfΣ); eauto. + Unshelve. eauto. + Qed. Lemma sq_wf_local_app {Γ Δ} : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ wf_local Σ Γ ∥ -> ∥ wf_local_rel Σ Γ Δ ∥ -> ∥ wf_local Σ (Γ ,,, Δ) ∥. @@ -533,22 +568,29 @@ Section Typecheck. check_context_rel Γ wfΓ [] := ret _ ; - check_context_rel Γ wfΓ ({| decl_body := None; decl_type := A |} :: Δ) := + check_context_rel Γ wfΓ ({| decl_name := na; 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 ;; + infer_isTypeRel (Γ ,,, Δ) (fun Σ wfΣ => sq_wf_local_app Σ wfΣ (wfΓ Σ wfΣ) (wfΔ Σ wfΣ)) A na.(binder_relevance) ;; ret _ ; - check_context_rel Γ wfΓ ({| decl_body := Some t; decl_type := A |} :: Δ) := + check_context_rel Γ wfΓ ({| decl_name := na; decl_body := Some t; decl_type := A |} :: Δ) := 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 ;; + infer_isTypeRel (Γ ,,, Δ) (fun Σ wfΣ => sq_wf_local_app Σ wfΣ (wfΓ Σ wfΣ) (wfΔ Σ wfΣ)) A na.(binder_relevance) ;; + (* Or do we check t's syntactic relevance ? *) + bdcheck (Γ ,,, Δ) (fun Σ wfΣ => sq_wf_local_app Σ wfΣ (wfΓ Σ wfΣ) (wfΔ Σ wfΣ)) t A _ ;; ret _. Next Obligation. sq. now constructor. Qed. + Next Obligation. + pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. + now eapply isType_of_isTypeRel. + Qed. Next Obligation. pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. constructor ; auto. - eapply checking_typing ; pcuic. + eapply checking_typing; tea. + - apply wf_local_app; tea. + - now eapply isType_of_isTypeRel. Qed. Next Obligation. apply absurd. intros. @@ -558,9 +600,7 @@ Section Typecheck. Qed. Next Obligation. apply absurd. intros. specialize_Σ wfΣ. sq. - inversion H ; subst ; cbn in *. - destruct X1 as [s ?]. - now exists s. + now inversion H. Qed. Next Obligation. apply absurd. intros. specialize_Σ wfΣ. sq. @@ -745,8 +785,8 @@ Section Typecheck. Next Obligation. specialize_Σ wfΣ. sq. depelim wfΔ; simpl. - destruct l; eexists; eauto. - destruct l; split; eexists; eauto. + - pcuic. + - split; pcuic. eexists; eauto. Qed. Next Obligation. pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. @@ -852,6 +892,7 @@ Section Typecheck. Qed. Next Obligation. specialize_Σ wfΣ; sq. eapply All_local_env_app_inv in wfΔ as [wt _]. + eapply isType_of_isTypeRel. now depelim wt. Qed. Next Obligation. @@ -863,6 +904,7 @@ Section Typecheck. rewrite subst_empty. apply checking_typing ; auto. apply wf_local_rel_app_inv in wfΔ as [wt _]. + eapply isType_of_isTypeRel. now depelim wt. Qed. Next Obligation. @@ -873,6 +915,7 @@ Section Typecheck. constructor ; tea. apply checking_typing ; auto. eapply All_local_env_app_l in wfΔ. + eapply isType_of_isTypeRel. now inversion wfΔ ; subst. Qed. Next Obligation. @@ -965,6 +1008,147 @@ Section Typecheck. congruence. Qed. + Definition abstract_lookup_constant kn := + match abstract_env_lookup X kn with + | Some (ConstantDecl d) => Some d + | _ => None + end. + + Definition abstract_lookup_minductive mind := + match abstract_env_lookup X mind with + | Some (InductiveDecl decl) => Some decl + | _ => None + end. + + Definition abstract_lookup_inductive ind := + match abstract_lookup_minductive (inductive_mind ind) with + | Some mdecl => + match nth_error mdecl.(ind_bodies) (inductive_ind ind) with + | Some idecl => Some (mdecl, idecl) + | None => None + end + | None => None + end. + + Definition abstract_lookup_constructor ind k := + match abstract_lookup_inductive ind with + | Some (mdecl, idecl) => + match nth_error idecl.(ind_ctors) k with + | Some cdecl => Some (mdecl, idecl, cdecl) + | None => None + end + | _ => None + end. + + Definition abstract_lookup_projection p := + match abstract_lookup_constructor p.(proj_ind) 0 with + | Some (mdecl, idecl, cdecl) => + match nth_error idecl.(ind_projs) p.(proj_arg) with + | Some pdecl => Some (mdecl, idecl, cdecl, pdecl) + | None => None + end + | _ => None + end. + + Equations relevance_of_term_abstract_env (Γ : mark_context) (t : term) : { rel & forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ isTermRel Σ Γ t rel ∥ } by struct t := + + relevance_of_term_abstract_env Γ (tRel n) + with inspect (nth_error Γ n) := { + | exist (Some c) e => (c; _) ; + | exist None e => (Relevant; _) + } ; + + relevance_of_term_abstract_env Γ (tLambda na A t) := + relevance_of_term_abstract_env (Γ ,, na.(binder_relevance)) t ; + + relevance_of_term_abstract_env Γ (tLetIn na b b_ty b') := + relevance_of_term_abstract_env (Γ ,, na.(binder_relevance)) b' ; + + relevance_of_term_abstract_env Γ (tApp t u) := + relevance_of_term_abstract_env Γ t ; + + relevance_of_term_abstract_env Γ (tConst cst u) + with inspect (abstract_lookup_constant cst) := { + | exist (Some d) e => (d.(cst_relevance); _) ; + | exist None e => (Relevant; _) ; + } ; + + relevance_of_term_abstract_env Γ (tConstruct ind k u) + with inspect (abstract_lookup_constructor ind k) := { + | exist (Some (_, idecl, _)) e := (idecl.(ind_relevance); _) ; + | exist None _ := (Relevant; _) ; + }; + + relevance_of_term_abstract_env Γ (tCase ci p c brs) := + (ci.(ci_relevance); _) ; + + relevance_of_term_abstract_env Γ (tProj p c) + with inspect (abstract_lookup_projection p) := { + | exist (Some (_, _, _, pdecl)) e := (pdecl.(proj_relevance); _) ; + | exist None _ := (Relevant; _) ; + }; + + relevance_of_term_abstract_env Γ (tFix mfix n) + with inspect (nth_error mfix n) := { + | exist None _ := (Relevant; _) ; + | exist (Some decl) e := (decl.(dname).(binder_relevance); _) ; + } ; + + relevance_of_term_abstract_env Γ (tCoFix mfix n) + with inspect (nth_error mfix n) := { + | exist None _ := (Relevant; _) ; + | exist (Some decl) e := (decl.(dname).(binder_relevance); _) ; + } ; + + relevance_of_term_abstract_env Γ (tVar _) := (Relevant; _) ; + relevance_of_term_abstract_env Γ (tEvar ev _) := (Relevant; _) ; + relevance_of_term_abstract_env Γ (tSort u) := (Relevant; _) ; + relevance_of_term_abstract_env Γ (tProd na A B) := (Relevant; _) ; + relevance_of_term_abstract_env Γ (tInd ind u) := (Relevant; _). + + Next Obligation. + unfold relevance_of_term, relevance_of_constant, relevance_of_constructor, relevance_of_projection; + unfold lookup_constant, lookup_projection, lookup_constructor, lookup_inductive, lookup_minductive; + unfold abstract_lookup_constant, abstract_lookup_projection, abstract_lookup_constructor, abstract_lookup_inductive, abstract_lookup_minductive in e; + erewrite <- abstract_env_lookup_correct in e; eauto; + now rewrite <- e. + Qed. + Next Obligation. + unfold relevance_of_term, relevance_of_constant, relevance_of_constructor, relevance_of_projection. + unfold lookup_constant, lookup_projection, lookup_constructor, lookup_inductive, lookup_minductive. + unfold abstract_lookup_constant, abstract_lookup_projection, abstract_lookup_constructor, abstract_lookup_inductive, abstract_lookup_minductive in e. + erewrite <- abstract_env_lookup_correct in e; eauto. + now rewrite <- e. + Qed. + Next Obligation. + unfold relevance_of_term, relevance_of_constant, relevance_of_constructor, relevance_of_projection. + unfold lookup_constant, lookup_projection, lookup_constructor, lookup_inductive, lookup_minductive. + unfold abstract_lookup_constant, abstract_lookup_projection, abstract_lookup_constructor, abstract_lookup_inductive, abstract_lookup_minductive in e. + erewrite <- abstract_env_lookup_correct in e; eauto. + now rewrite <- e. + Qed. + Next Obligation. + unfold relevance_of_term, relevance_of_constant, relevance_of_constructor, relevance_of_projection. + unfold lookup_constant, lookup_projection, lookup_constructor, lookup_inductive, lookup_minductive. + unfold abstract_lookup_constant, abstract_lookup_projection, abstract_lookup_constructor, abstract_lookup_inductive, abstract_lookup_minductive in e. + erewrite <- abstract_env_lookup_correct in e; eauto. + now rewrite <- e. + Qed. + Next Obligation. + unfold relevance_of_term, relevance_of_constant, relevance_of_constructor, relevance_of_projection. + unfold lookup_constant, lookup_projection, lookup_constructor, lookup_inductive, lookup_minductive. + unfold abstract_lookup_constant, abstract_lookup_projection, abstract_lookup_constructor, abstract_lookup_inductive, abstract_lookup_minductive in e. + erewrite <- abstract_env_lookup_correct in e; eauto. + now rewrite <- e. + Qed. + Next Obligation. + unfold relevance_of_term, relevance_of_constant, relevance_of_constructor, relevance_of_projection. + unfold lookup_constant, lookup_projection, lookup_constructor, lookup_inductive, lookup_minductive. + unfold abstract_lookup_constant, abstract_lookup_projection, abstract_lookup_constructor, abstract_lookup_inductive, abstract_lookup_minductive in e. + erewrite <- abstract_env_lookup_correct in e; eauto. + now rewrite <- e. + Qed. + Definition abstract_env_level_mem_forallb {Σ} (wfΣ : abstract_env_ext_rel X Σ) u : forallb (level_mem Σ) u = forallb (abstract_env_level_mem X) u. Proof using Type. @@ -1166,8 +1350,8 @@ Section Typecheck. Next Obligation. eapply branch_helper in i; tea. pose proof (heΣ _ wfΣ) as [heΣ]. specialize_Σ wfΣ; sq. - destruct i as [? []]. - exists ps. + destruct i as (wfb & wfl & isty); fold brctxty in wfl, isty. + exists ps; split => //. apply checking_typing ; eauto. eapply isType_Sort ; eauto. apply infering_sort_typing, validity, isType_Sort_inv in wfpret ; eauto. @@ -1231,8 +1415,11 @@ Section Typecheck. Context (infer : forall (Γ : context) (HΓ : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ wf_local Σ Γ ∥) (t : term), typing_result_comp ({ A : term & forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ Σ ;;; Γ |- t ▹ A ∥ })) (Γ : context) (wfΓ : forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ wf_local Σ Γ ∥). + 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_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_typing 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 ;; @@ -1240,24 +1427,36 @@ Section Typecheck. ret _. *) check_mfix_types (def :: mfix) := s <- infer_type infer Γ wfΓ (dtype def) ;; + check_eq_true (relevance_of_sort s.π1 == def.(dname).(binder_relevance)) (Msg "Wrong relevance") ;; check_mfix_types mfix ;; ret _. Next Obligation. pose proof (heΣ _ wfΣ) as [heΣ]. specialize_Σ wfΣ; sq. constructor ; tea. - exists s. - now apply infering_sort_typing. + exists s; split. + - apply eqb_eq, i. + - now apply infering_sort_typing. Qed. Next Obligation. apply absurd. intros; specialize_Σ wfΣ; sq. now depelim H. Qed. + Next Obligation. + destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. + pose (hΣ _ wfΣ). specialize_Σ wfΣ. sq. + apply absurd. case: eqb_spec => _i //; exfalso; apply _i; clear _i. + eapply infering_sort_typing in X0; eauto. + depelim H. + destruct o as (s' & e' & Hs'). + rewrite <- e'. + eapply cumul_sort_relevance; tea. + Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]; pose proof (heΣ _ wfΣ) as [heΣ]; specialize_Σ wfΣ. sq. depelim H. apply absurd. - apply isType_infering_sort in i as [u ?]; tea. + apply isType_of_isTypeRel, isType_infering_sort in o as [u ?]; tea. exists u. intros. erewrite (abstract_env_ext_irr _ _ wfΣ); eauto. Unshelve. eauto. Qed. @@ -1339,17 +1538,20 @@ Section Typecheck. } ; infer Γ HΓ (tProd na A B) := - s1 <- infer_type infer Γ HΓ A ;; - s2 <- infer_type infer (Γ,,vass 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 (Universe.sort_of_product s1.π1 s2.π1);_) ; infer Γ HΓ (tLambda na A t) := - infer_type infer Γ HΓ A ;; + s <- infer_type infer Γ HΓ A ;; + check_eq_true (relevance_of_sort s.π1 == na.(binder_relevance)) (Msg "Wrong 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 ;; + s <- infer_type infer Γ HΓ b_ty ;; + check_eq_true (relevance_of_sort s.π1 == n.(binder_relevance)) (Msg "Wrong 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; _) ; @@ -1417,6 +1619,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 _ ;; @@ -1492,11 +1695,15 @@ 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. + eexists; split. + - apply eqb_eq in i; apply i. + - now apply infering_sort_typing. Qed. Next Obligation. (* intros Γ HΓ t na A B Heq_t [s1 ?] [s2 ?]; *) - cbn. specialize_Σ wfΣ. sq. econstructor; try eassumption. + cbn. specialize_Σ wfΣ. sq. + econstructor; tea. + apply eqb_eq, i. Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. @@ -1504,6 +1711,15 @@ 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. + rewrite -H3. + eapply cumul_sort_relevance; tea. + all: apply infering_sort_typing; eauto with pcuic. + Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. specialize_Σ wfΣ; sq. inversion X1; subst. apply absurd. @@ -1515,12 +1731,15 @@ 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. + eexists; split. + - apply eqb_eq in i; apply i. + - now apply infering_sort_typing. Qed. Next Obligation. (* intros Γ HΓ t0 na A t Heq_t [s ?] [B ?]; *) cbn; pose proof (heΣ _ wfΣ) as [heΣ]. specialize_Σ wfΣ. - sq; econstructor; eassumption. + sq; econstructor; tea. + apply eqb_eq, i. Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. @@ -1528,6 +1747,15 @@ 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. + rewrite -H3. + eapply cumul_sort_relevance; tea. + all: apply infering_sort_typing; eauto with pcuic. + Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. specialize_Σ wfΣ; sq. inversion X1; subst. apply absurd. @@ -1537,18 +1765,22 @@ Section Typecheck. (* tLetIn *) Next Obligation. pose proof (heΣ _ wfΣ) as [heΣ]. specialize_Σ wfΣ. - sq; econstructor; tea. + sq; econstructor; tea. split => //. eapply infering_sort_typing ; eauto. Qed. Next Obligation. (* intros Γ HΓ t n b b_ty b' Heq_t [? ?] H0; *) pose proof (heΣ _ wfΣ) as [heΣ]. specialize_Σ wfΣ ; sq. econstructor ; tea. - 2: apply checking_typing ; eauto. - all: now eapply infering_sort_isType. + eexists; split. + - apply eqb_eq in i; apply i. + - now apply infering_sort_typing. + - apply checking_typing ; eauto. + now eapply infering_sort_isType. Qed. Next Obligation. cbn; specialize_Σ wfΣ; sq; econstructor; eauto. + apply eqb_eq, i. Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. @@ -1562,6 +1794,15 @@ 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. + rewrite -H4. + eapply cumul_sort_relevance; tea. + all: apply infering_sort_typing; eauto with pcuic. + Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. specialize_Σ wfΣ; sq. inversion X1; subst. apply absurd. @@ -1581,10 +1822,10 @@ Section Typecheck. specialize_Σ wfΣ ; sq. eapply infering_typing, type_reduction_closed, validity in X3. 2-4: eauto. - destruct X3 as [s HH]. + destruct X3 as (s & e & HH). eapply inversion_Prod in HH ; auto. - destruct HH as [s1 [_ [HH _]]]. - eexists. eassumption. + destruct HH as (s1 & _ & _ & HH & _ & _). + eexists; split => //. eassumption. Qed. Next Obligation. cbn in *; specialize_Σ wfΣ ; sq. @@ -1734,7 +1975,7 @@ Section Typecheck. (* tCase *) Next Obligation. cbn in *. pose proof (heΣ _ wfΣ) as [heΣ]. specialize_Σ wfΣ ; sq. - eapply infering_typing, validity in X0 as []; eauto. + eapply infering_typing, validity in X0 as (s & e & Hs); eauto. eexists; eauto using validity_wf. Qed. Next Obligation. @@ -1783,7 +2024,7 @@ Section Typecheck. eapply ctx_inst_smash in wt_params. unshelve epose proof (ctx_inst_spine_subst _ wt_params). { eapply weaken_wf_local; eauto. eapply on_minductive_wf_params; eauto. } - eexists; eapply isType_mkApps_Ind_smash; tea. + eexists; split => //. eapply isType_mkApps_Ind_smash; tea. rewrite subst_instance_app List.rev_app_distr smash_context_app_expand. have wf_ctx : wf_local Σ (Γ,,, smash_context [] (ind_params d)@[puinst p],,, @@ -1885,6 +2126,14 @@ Section Typecheck. cbn in *. specialize_Σ wfΣ ; now sq. Qed. + Next Obligation. + intros. + destruct ps as [ps ?]. + cbn in *. + specialize_Σ wfΣ ; sq. + assumption. + Qed. + Next Obligation. intros. cbn in *. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. @@ -1929,6 +2178,7 @@ Section Typecheck. now apply closed_red_red. - now eapply All2_fold_All2 in check_wfpctx_conv. - now eapply wf_local_rel_wf_local_bd, wf_local_app_inv, wf_case_predicate_context. + - apply eqb_eq, i3. - eapply ctx_inst_typing_bd ; eauto. eapply ctx_inst_smash. now rewrite subst_instance_smash /= in wt_params. @@ -1989,6 +2239,26 @@ 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Σ]. + cbn in *. specialize_Σ wfΣ. + destruct X0 as [? [ty]]; eauto. + inversion ty ; subst. + eapply declared_inductive_inj in isdecl as []; tea. + subst. sq. + eapply infering_sort_sort in s as <- ; eauto. + now eapply wf_case_predicate_context. + Qed. + Next Obligation. intros; clearbody isty wfp. destruct cty as [A cty]. @@ -2409,6 +2679,12 @@ Section Typecheck. Qed. (* tFix *) + Next Obligation. + pose proof (heΣ _ wfΣ) as [heΣ]. + cbn in *. specialize_Σ wfΣ ; sq. + apply All_impl with (1 := wf_types); intro d. + apply isType_of_isTypeRel. + Qed. Next Obligation. pose proof (heΣ _ wfΣ) as [heΣ]. cbn in *. specialize_Σ wfΣ ; sq. @@ -2421,7 +2697,7 @@ Section Typecheck. constructor; auto. eapply All_impl ; tea. intros. - now apply isType_infering_sort. + now apply isTypeRel_infering_sort. erewrite abstract_wf_fixpoint in wffix; eauto. Qed. Next Obligation. @@ -2451,7 +2727,8 @@ Section Typecheck. eapply All_impl. 1: eexact X2. intros. - now eapply einfering_sort_isType. + apply infer_sort_impl with id X0 => //; intro Hty. + apply infering_sort_typing; eauto. Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. @@ -2461,6 +2738,12 @@ Section Typecheck. Qed. (* tCoFix *) + Next Obligation. + pose proof (heΣ _ wfΣ) as [heΣ]. + cbn in *. specialize_Σ wfΣ ; sq. + apply All_impl with (1 := wf_types); intro d. + apply isType_of_isTypeRel. + Qed. Next Obligation. pose proof (heΣ _ wfΣ) as [heΣ]. cbn in *. specialize_Σ wfΣ ; sq. @@ -2473,7 +2756,7 @@ Section Typecheck. constructor; auto. eapply All_impl ; tea. intros. - now apply isType_infering_sort. + now apply isTypeRel_infering_sort. erewrite abstract_wf_cofixpoint in wfcofix; eauto. Qed. Next Obligation. @@ -2503,7 +2786,8 @@ Section Typecheck. eapply All_impl. 1: eexact X2. intros. - now eapply einfering_sort_isType. + apply infer_sort_impl with id X0 => //; intro Hty. + apply infering_sort_typing; eauto. Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. @@ -2526,6 +2810,7 @@ Section Typecheck. Defined. *) 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 ∥) := From a6d7c5f8712a0b3c6dd2844b98e98248564a7eb9 Mon Sep 17 00:00:00 2001 From: Yann Leray Date: Fri, 20 May 2022 18:50:06 +0200 Subject: [PATCH 07/17] Erasure --- erasure/theories/EArities.v | 170 ++++++++---------- erasure/theories/EDeps.v | 14 +- erasure/theories/EOptimizePropDiscr.v | 6 +- erasure/theories/ESubstitution.v | 71 ++++---- erasure/theories/ErasureCorrectness.v | 35 ++-- erasure/theories/ErasureFunction.v | 38 ++-- erasure/theories/ErasureProperties.v | 59 +++--- erasure/theories/Prelim.v | 10 +- pcuic/theories/PCUICRelevanceTyp.v | 28 ++- .../src/metacoq_safechecker_plugin.mlpack | 1 + 10 files changed, 204 insertions(+), 228 deletions(-) diff --git a/erasure/theories/EArities.v b/erasure/theories/EArities.v index 492c0d74c..4ad9da358 100644 --- a/erasure/theories/EArities.v +++ b/erasure/theories/EArities.v @@ -31,16 +31,17 @@ Qed. Lemma isType_isErasable Σ Γ T : isType Σ Γ T -> isErasable Σ Γ T. Proof. - intros [s Hs]. - exists (tSort s). intuition auto. left; simpl; auto. + intros (s & e & Hs). + exists (tSort s). intuition auto. Qed. Lemma isType_red: forall (Σ : global_env_ext) (Γ : context) (T : term), wf Σ -> wf_local Σ Γ -> isType Σ Γ T -> forall x5 : term, red Σ Γ T x5 -> isType Σ Γ x5. Proof. - intros. destruct X1 as []. - eexists. eapply subject_reduction ; eauto. + intros. + apply infer_typing_sort_impl with id X1 => // Hs. + eapply subject_reduction; eauto. Qed. Lemma it_mkProd_isArity: @@ -72,7 +73,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. - split; auto. + split; eauto using isType_of_isTypeRel. split; auto. clear dom codom. eapply destArity_app_Some in e as (? & ? & ?); subst. @@ -209,7 +210,8 @@ 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. - now eapply isType_tProd in i as []. + eapply isType_tProd in i as []. + now apply isType_of_isTypeRel in i. Qed. Lemma typing_spine_Is_conv_to_Arity {Σ : global_env_ext} {wfΣ : wf Σ} {Γ Δ ind u args args' T'} : @@ -258,7 +260,7 @@ Lemma nIs_conv_to_Arity_nArity {Σ : global_env_ext} {wfΣ : wf Σ} {Γ T} : Proof. intros isty nisc isa. apply nisc. exists T. split => //. sq. - destruct isty as [s Hs]. + destruct isty as (s & e & Hs). eapply wt_closed_red_refl; tea. Qed. @@ -295,7 +297,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 in a; eauto; apply isType_of_isTypeRel 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. @@ -322,9 +324,8 @@ Qed. Section Elim'. Context `{cf : checker_flags}. -Context {Σ : global_env_ext} {wfΣ : wf_ext Σ}. +Context {Σ : global_env_ext} {wfΣ : wf Σ}. Variable Hcf : prop_sub_type = false. -Variable Hcf' : check_univs. Lemma cumul_prop1 Γ A B u : Universe.is_prop u -> @@ -332,7 +333,7 @@ Lemma cumul_prop1 Γ A B u : Σ ;;; Γ |- B : tSort u -> Σ ;;; Γ ⊢ A ≤ B -> Σ ;;; Γ |- A : tSort u. -Proof using Hcf Hcf' wfΣ. +Proof using Hcf wfΣ. intros; eapply cumul_prop1; tea. now apply ws_cumul_pb_forget in X1. Qed. @@ -343,7 +344,7 @@ Lemma cumul_prop2 Γ A B u : Σ ;;; Γ ⊢ A ≤ B -> Σ ;;; Γ |- A : tSort u -> Σ ;;; Γ |- B : tSort u. -Proof using Hcf Hcf' wfΣ. +Proof using Hcf wfΣ. intros. eapply cumul_prop2; tea. now apply ws_cumul_pb_forget in X0. Qed. @@ -354,7 +355,7 @@ Lemma cumul_sprop1 Γ A B u : Σ ;;; Γ |- B : tSort u -> Σ ;;; Γ ⊢ A ≤ B -> Σ ;;; Γ |- A : tSort u. -Proof using Hcf Hcf' wfΣ. +Proof using Hcf wfΣ. intros. eapply cumul_sprop1; tea. now apply ws_cumul_pb_forget in X1. Qed. @@ -365,14 +366,14 @@ Lemma cumul_sprop2 Γ A B u : Σ ;;; Γ ⊢ A ≤ B -> Σ ;;; Γ |- A : tSort u -> Σ ;;; Γ |- B : tSort u. -Proof using Hcf Hcf' wfΣ. +Proof using Hcf wfΣ. intros. eapply cumul_sprop2; tea. now apply ws_cumul_pb_forget in X0. Qed. End Elim'. Lemma cumul_propositional (Σ : global_env_ext) Γ A B u : - wf_ext Σ -> + wf Σ -> is_propositional u -> isType Σ Γ B -> Σ ;;; Γ ⊢ A ≤ B -> @@ -389,7 +390,7 @@ Qed. Lemma sort_typing_spine: forall (Σ : global_env_ext) (Γ : context) (L : list term) (u : Universe.t) (x x0 : term), - wf_ext Σ -> + wf Σ -> is_propositional u -> typing_spine Σ Γ x L x0 -> Σ;;; Γ |- x : tSort u -> @@ -402,7 +403,7 @@ Proof. - destruct u => //. eapply cumul_prop2 in c0; eauto. eapply cumul_sprop2 in c0; eauto. - eapply cumul_propositional in c0; auto. 2-3: tea. - eapply inversion_Prod in c0 as (? & ? & ? & ? & e0) ; auto. + eapply inversion_Prod in c0 as (? & ? & ? & ? & ? & e0) ; auto. eapply ws_cumul_pb_Sort_inv in e0. unfold is_propositional in H. destruct (Universe.is_prop u) eqn:isp => //. @@ -416,7 +417,7 @@ Proof. eapply substitution0; eauto. Qed. -Lemma arity_type_inv (Σ : global_env_ext) Γ t T1 T2 : wf_ext Σ -> wf_local Σ Γ -> +Lemma arity_type_inv (Σ : global_env_ext) Γ t T1 T2 : wf Σ -> wf_local Σ Γ -> Σ ;;; Γ |- t : T1 -> isArity T1 -> Σ ;;; Γ |- t : T2 -> Is_conv_to_Arity Σ Γ T2. Proof. intros wfΣ wfΓ. intros. @@ -424,13 +425,12 @@ Proof. eapply invert_cumul_arity_l_gen; tea. eapply invert_cumul_arity_r_gen. 2:exact e. exists T1. split; auto. sq. - eapply PCUICValidity.validity in X as [s Hs]. + eapply PCUICValidity.validity in X as (s & _ & Hs). eapply wt_closed_red_refl; eauto. Qed. Lemma cumul_prop1' (Σ : global_env_ext) Γ A B u : - check_univs -> - wf_ext Σ -> + wf Σ -> isType Σ Γ A -> is_propositional u -> Σ ;;; Γ |- B : tSort u -> @@ -445,8 +445,7 @@ Proof. Qed. Lemma cumul_prop2' (Σ : global_env_ext) Γ A B u : - check_univs -> - wf_ext Σ -> + wf Σ -> isType Σ Γ A -> is_propositional u -> Σ ;;; Γ |- B : tSort u -> @@ -461,7 +460,7 @@ Proof. Qed. Lemma leq_term_propositional_sorted_l {Σ Γ v v' u u'} : - wf_ext Σ -> + wf Σ -> PCUICEquality.leq_term Σ (global_ext_constraints Σ) v v' -> Σ;;; Γ |- v : tSort u -> Σ;;; Γ |- v' : tSort u' -> is_propositional u -> @@ -475,7 +474,7 @@ Proof. Qed. Lemma leq_term_propopositional_sorted_r {Σ Γ v v' u u'} : - wf_ext Σ -> + wf Σ -> PCUICEquality.leq_term Σ (global_ext_constraints Σ) v v' -> Σ;;; Γ |- v : tSort u -> Σ;;; Γ |- v' : tSort u' -> is_propositional u' -> @@ -489,7 +488,7 @@ Proof. Qed. Lemma Is_type_app (Σ : global_env_ext) Γ t L T : - wf_ext Σ -> + wf Σ -> wf_local Σ Γ -> Σ ;;; Γ |- mkApps t L : T -> isErasable Σ Γ t -> @@ -525,32 +524,20 @@ Proof. Qed. Lemma leq_universe_propositional_r {cf : checker_flags} (ϕ : ConstraintSet.t) (u1 u2 : Universe.t_) : - check_univs -> - consistent ϕ -> leq_universe ϕ u1 u2 -> is_propositional u2 -> is_propositional u1. Proof. - intros cu cons leq; unfold is_propositional. - destruct u2 => //. - apply leq_universe_prop_r in leq => //. - now rewrite leq. - intros _. - apply leq_universe_sprop_r in leq => //. - now rewrite leq orb_true_r. + intros leq; unfold is_propositional. + destruct u2, u1 => //. Qed. Lemma leq_universe_propositional_l {cf : checker_flags} (ϕ : ConstraintSet.t) (u1 u2 : Universe.t_) : - check_univs -> prop_sub_type = false -> - consistent ϕ -> leq_universe ϕ u1 u2 -> is_propositional u1 -> is_propositional u2. Proof. - intros cu ps cons leq; unfold is_propositional. - destruct u1 => //. - eapply leq_universe_prop_no_prop_sub_type in leq => //. - now rewrite leq. - intros _. - apply leq_universe_sprop_l in leq => //. - now rewrite leq orb_true_r. + intros ps leq; unfold is_propositional. + destruct u1, u2 => //. + unfold leq_universe, leq_universe_n, leq_universe_n_ in leq. + now rewrite ps in leq. Qed. Lemma is_propositional_sort_prod x2 x3 : @@ -566,20 +553,20 @@ Proof. Qed. Lemma Is_type_lambda (Σ : global_env_ext) Γ na T1 t : - wf_ext Σ -> + wf Σ -> wf_local Σ Γ -> isErasable Σ Γ (tLambda na T1 t) -> ∥isErasable Σ (vass na T1 :: Γ) t∥. Proof. intros ? ? (T & ? & ?). - eapply inversion_Lambda in t0 as (? & ? & ? & ? & e); auto. + eapply inversion_Lambda in t0 as (? & ? & ? & ? & ? & e); auto. destruct s as [ | (u & ? & ?)]. - eapply invert_cumul_arity_r in e; eauto. destruct e as (? & [] & ?). eapply invert_red_prod in X1 as (? & ? & []); eauto; subst. cbn in H. econstructor. exists x3. econstructor. eapply type_reduction_closed; eauto. econstructor; eauto. - sq. eapply cumul_prop1' in e; eauto. - eapply inversion_Prod in e as (? & ? & ? & ? & e) ; auto. + eapply inversion_Prod in e as (? & ? & ? & ? & ? & e) ; auto. eapply ws_cumul_pb_Sort_inv in e. eapply leq_universe_propositional_r in e as H0; cbn; eauto. eexists. split. eassumption. right. eexists. split. eassumption. @@ -615,7 +602,7 @@ Qed. on well-typed terms. *) Lemma Is_type_eval_inv (Σ : global_env_ext) t v: - wf_ext Σ -> + wf Σ -> welltyped Σ [] t -> PCUICWcbvEval.eval Σ t v -> isErasable Σ [] v -> @@ -623,7 +610,7 @@ Lemma Is_type_eval_inv (Σ : global_env_ext) t v: Proof. intros wfΣ [T HT] ev [vt [Ht Hp]]. eapply wcbeval_red in ev; eauto. - pose proof (subject_reduction _ _ _ _ _ wfΣ.1 HT ev). + pose proof (subject_reduction _ _ _ _ _ wfΣ HT ev). pose proof (common_typing _ wfΣ Ht X) as [P [Pvt [Pt vP]]]. destruct Hp. eapply arity_type_inv in X. 5:eauto. all:eauto. @@ -640,7 +627,7 @@ Qed. Lemma isType_closed_red_refl {Σ} {wfΣ : wf Σ} {Γ T} : isType Σ Γ T -> Σ ;;; Γ ⊢ T ⇝ T. Proof. - intros [s hs]; eapply wt_closed_red_refl; tea. + intros (s & e & hs); eapply wt_closed_red_refl; tea. Qed. Lemma nIs_conv_to_Arity_isWfArity_elim {Σ} {wfΣ : wf Σ} {Γ x} : @@ -660,7 +647,7 @@ Definition isErasable_Type (Σ : global_env_ext) Γ T := (Is_conv_to_Arity Σ Γ T + (∑ u : Universe.t, Σ;;; Γ |- T : tSort u × is_propositional u))%type. -Lemma isErasable_any_type {Σ} {wfΣ : wf_ext Σ} {Γ t T} : +Lemma isErasable_any_type {Σ} {wfΣ : wf Σ} {Γ t T} : isErasable Σ Γ t -> Σ ;;; Γ |- t : T -> isErasable_Type Σ Γ T. @@ -678,7 +665,7 @@ Proof. Qed. Lemma Is_proof_ty Σ Γ t : - wf_ext Σ -> + wf Σ -> Is_proof Σ Γ t -> forall t' ty, Σ ;;; Γ |- t : ty -> @@ -699,16 +686,15 @@ Qed. Lemma is_propositional_bottom {Σ Γ T s s'} : - wf_ext Σ -> - check_univs -> + wf Σ -> prop_sub_type = false -> Σ ;;; Γ ⊢ T ≤ tSort s -> Σ ;;; Γ ⊢ T ≤ tSort s' -> PCUICCumulProp.eq_univ_prop s s'. Proof. - intros wf cu pst h h'; rewrite /PCUICCumulProp.eq_univ_prop. - split. split; eapply PCUICCumulProp.is_prop_bottom; tea. - split; eapply PCUICCumulProp.is_sprop_bottom; tea. + intros wf pst h h'. + destruct (PCUICCumulProp.cumul_sort_confluence h h') as [x' [conv [leq leq']]]. + destruct x' => //; destruct s => //; destruct s' => //. Qed. Import PCUICGlobalEnv PCUICUnivSubst PCUICValidity PCUICCumulProp. @@ -716,84 +702,70 @@ Import PCUICGlobalEnv PCUICUnivSubst PCUICValidity PCUICCumulProp. Notation " Σ ;;; Γ |- t ~~ u " := (cumul_prop Σ Γ t u) (at level 50, Γ, t, u at next level) : type_scope. Lemma is_propositional_bottom' {Σ Γ T s s'} : - wf_ext Σ -> - check_univs -> + wf Σ -> prop_sub_type = false -> Σ ;;; Γ |- T ~~ tSort s -> Σ ;;; Γ |- T ~~ tSort s' -> PCUICCumulProp.eq_univ_prop s s'. Proof. - intros wf cu pst h h'; rewrite /PCUICCumulProp.eq_univ_prop. - pose proof (cumul_prop_trans _ _ _ _ _ _ (cumul_prop_sym _ _ _ _ _ h') h). - split. split; intros; eapply PCUICCumulProp.cumul_prop_props; tea. now symmetry. - split; intros; eapply PCUICCumulProp.cumul_sprop_props; tea. now symmetry. + intros wf pst h h'. + pose proof (cumul_prop_trans _ _ _ _ _ _ (cumul_prop_sym _ _ _ _ _ h) h'). + now eapply PCUICCumulProp.cumul_sort_inv. Qed. Lemma is_propositional_lower {Σ s u u'} : - consistent Σ -> leq_universe Σ s u -> leq_universe Σ s u' -> PCUICCumulProp.eq_univ_prop u u'. Proof. - intros wf leu leu'. - unfold eq_univ_prop; split. - - split. intros pu. eapply leq_universe_prop_r in leu; tea => //. - eapply leq_universe_prop_no_prop_sub_type in leu'; trea => //. - intros pu'. eapply leq_universe_prop_r in leu'; tea => //. - eapply leq_universe_prop_no_prop_sub_type in leu; tea => //. - - split. intros pu. eapply leq_universe_sprop_r in leu; tea => //. - eapply leq_universe_sprop_l in leu'; tea => //. - intros pu'. eapply leq_universe_sprop_r in leu'; tea => //. - eapply leq_universe_sprop_l in leu; tea => //. + intros leu leu'. + destruct s => //; destruct u => //; destruct u' => //. Qed. Lemma typing_spine_inj {Σ Γ Δ s args args' u u'} : - wf_ext Σ -> - check_univs -> + wf Σ -> prop_sub_type = false -> let T := it_mkProd_or_LetIn Δ (tSort s) in typing_spine Σ Γ T args (tSort u) -> typing_spine Σ Γ T args' (tSort u') -> PCUICCumulProp.eq_univ_prop u u'. Proof. - intros wf cu ips T. + intros wf ips T. move/typing_spine_it_mkProd_or_LetIn_full_inv => su. move/typing_spine_it_mkProd_or_LetIn_full_inv => su'. - eapply is_propositional_lower; tea. apply wf. + eapply is_propositional_lower; tea. Qed. Lemma Is_proof_ind Σ Γ t : - wf_ext Σ -> + wf Σ -> Is_proof Σ Γ t -> forall t' ind u args args', Σ ;;; Γ |- t : mkApps (tInd ind u) args -> Σ ;;; Γ |- t' : mkApps (tInd ind u) args' -> Is_proof Σ Γ t'. Proof. - intros wfΣ [ty [u [Hty isp]]]. + intros wfΣ (ty & u & Hty & Hu & isp). intros t' ind u' args args' Hty' Hty''. epose proof (PCUICPrincipality.common_typing _ wfΣ Hty Hty') as [C [Cty [Cty' Ht'']]]. - destruct isp. assert (Σ ;;; Γ |- C : tSort u). eapply cumul_prop1'; tea => //. now eapply validity. assert (Σ ;;; Γ |- mkApps (tInd ind u') args : tSort u). - eapply cumul_prop2'; tea => //. now eapply validity. - eapply inversion_mkApps in X0 as x1. destruct x1 as [? []]. - eapply inversion_Ind in t1 as [mdecl [idecl [wf [decli ?]]]]; eauto. - destruct (validity Hty'') as [u'' tyargs']. - eapply inversion_mkApps in X0 as x1. destruct x1 as [? []]. - eapply invert_type_mkApps_ind in X0 as [sp cum]; eauto. - eapply invert_type_mkApps_ind in tyargs' as f; tea. destruct f as [sp' cum']; eauto. + eapply cumul_prop2'; tea => //. now eapply validity. pose proof (X1 := X0). + eapply inversion_mkApps in X0 as _x; destruct _x as (A & HA & HA'). + eapply inversion_Ind in HA as _x; tea; destruct _x as (mdecl & idecl & wfΓ & decli & cu & leq). + destruct (validity Hty'') as (u'' & _ & tyargs'). + eapply invert_type_mkApps_ind in X0 as _x; tea; destruct _x as [sp cum]. + eapply invert_type_mkApps_ind in tyargs' as _x; tea; destruct _x as [sp' cum']. do 2 eexists. split => //. tea. instantiate (1 := u''). split => //. rewrite (declared_inductive_type decli) in sp, sp'. rewrite subst_instance_it_mkProd_or_LetIn /= in sp, sp'. - eapply typing_spine_inj in sp. 5:exact sp'. all:eauto. - destruct sp as [H H0]. apply/orP. rewrite H H0. now apply/orP. + eapply typing_spine_inj in sp. 4:exact sp'. all:eauto. + move/orP: isp => [] isp; destruct u => //; destruct u'' => //. Qed. -Lemma red_case_isproof {Σ : global_env_ext} {Γ ip p discr discr' brs T} {wfΣ : wf_ext Σ} : +Lemma red_case_isproof {Σ : global_env_ext} {Γ ip p discr discr' brs T} {wfΣ : wf Σ} : PCUICReduction.red Σ Γ (tCase ip p discr brs) (tCase ip p discr' brs) -> Σ ;;; Γ |- tCase ip p discr brs : T -> Is_proof Σ Γ discr -> Is_proof Σ Γ discr'. @@ -809,7 +781,7 @@ Proof. exact (X _ _ _ _ _ scrut_ty scrut_ty0). Qed. -Lemma Is_proof_app {Σ Γ t args ty} {wfΣ : wf_ext Σ} : +Lemma Is_proof_app {Σ Γ t args ty} {wfΣ : wf Σ} : Is_proof Σ Γ t -> Σ ;;; Γ |- mkApps t args : ty -> Is_proof Σ Γ (mkApps t args). @@ -821,7 +793,7 @@ Proof. epose proof (PCUICPrincipality.common_typing _ wfΣ Hty Ht) as [C [Cty [Cty' Ht'']]]. eapply PCUICSpine.typing_spine_strengthen in sp. 3:tea. edestruct (sort_typing_spine _ _ _ u _ _ _ pu sp) as [u' [Hty' isp']]. - eapply cumul_prop1'. 5:tea. all:eauto. + eapply cumul_prop1'. 4:tea. all:eauto. eapply validity; eauto. exists ty, u'; split; auto. eapply PCUICSpine.type_mkApps; tea; eauto. @@ -829,7 +801,7 @@ Proof. Qed. Lemma isErasable_Propositional {Σ : global_env_ext} {Γ ind n u args} : - wf_ext Σ -> + wf Σ -> isErasable Σ Γ (mkApps (tConstruct ind n u) args) -> isPropositional Σ ind true. Proof. intros wfΣ ise. @@ -861,7 +833,7 @@ Proof. Qed. Lemma nisErasable_Propositional {Σ : global_env_ext} {Γ ind n u} : - wf_ext Σ -> + wf Σ -> welltyped Σ Γ (tConstruct ind n u) -> (isErasable Σ Γ (tConstruct ind n u) -> False) -> isPropositional Σ ind false. Proof. @@ -893,12 +865,12 @@ Proof. eapply cumul_propositional; eauto. rewrite is_propositional_subst_instance => //. eapply PCUICValidity.validity; eauto. - destruct X as [cty ty]. + destruct X as (cty & _ & ty). eapply type_Cumul_alt; eauto. eapply isType_Sort. 2:eauto. destruct (ind_sort x0) => //. eapply PCUICSpine.inversion_it_mkProd_or_LetIn in ty; eauto. - epose proof (typing_spine_proofs _ _ [] _ _ _ [] _ _ eq_refl wfΣ ty). + epose proof (typing_spine_proofs _ _ [] _ _ _ [] _ _ wfΣ ty). forward H0 by constructor. eexists; eauto. simpl. now exists cty. eapply PCUICConversion.ws_cumul_pb_eq_le_gen, PCUICSR.wt_cumul_pb_refl; eauto. destruct H0 as [_ sorts]. @@ -976,7 +948,7 @@ Proof. Qed. Lemma Informative_cofix v ci p brs T (Σ : global_env_ext) : - wf_ext Σ -> + wf Σ -> forall (mdecl : mutual_inductive_body) (idecl : one_inductive_body) mfix idx, declared_inductive Σ.1 ci.(ci_ind) mdecl idecl -> forall (args : list term), Informative Σ ci.(ci_ind) -> @@ -992,7 +964,7 @@ Proof. eapply inversion_Case in X0 as (? & ? & ? & ? & [] & ?); eauto. destruct (declared_inductive_inj x8 x4); subst. destruct (declared_inductive_inj x8 H); subst. - eapply H0; eauto. reflexivity. + eapply H0; eauto. 2: reflexivity. tea. eapply Is_proof_ind; tea. Qed. diff --git a/erasure/theories/EDeps.v b/erasure/theories/EDeps.v index 1b72d5c12..17862cb0a 100644 --- a/erasure/theories/EDeps.v +++ b/erasure/theories/EDeps.v @@ -574,8 +574,8 @@ Proof. repeat eapply conj; try eassumption. cbn in *. now rewrite H8, H9. Qed. -Lemma erases_deps_single Σ Σ' Γ t T et : - wf_ext Σ -> +Lemma erases_deps_single (Σ : global_env_ext) Σ' Γ t T et : + wf Σ -> Σ;;; Γ |- t : T -> Σ;;; Γ |- t ⇝ℇ et -> globals_erased_with_deps Σ Σ' -> @@ -631,6 +631,7 @@ Proof. - constructor. apply inversion_Fix in wt as (?&?&?&?&?&?&?); eauto. clear -wf a0 X H Σer. + unfold on_def_body in *. revert a0 X H Σer. generalize mfix at 1 2 4 6. intros mfix_gen. @@ -645,6 +646,7 @@ Proof. - constructor. apply inversion_CoFix in wt as (?&?&?&?&?&?&?); eauto. clear -wf a0 X H Σer. + unfold on_def_body in *. revert a0 X H Σer. generalize mfix at 1 2 4 6. intros mfix_gen. @@ -696,7 +698,7 @@ Proof. depelim wf. depelim o0. cbn in *. eapply (erases_extends ({| universes := univs; declarations := Σ |}, cst_universes cst')); eauto. cbn. 4:{ split; eauto; cbn; try reflexivity. eexists [_]; cbn; reflexivity. } - constructor; auto. cbn. red in o2. rewrite E in o2. exact o2. + constructor; auto. cbn. red in o2. rewrite E in o2. exact o2.2. split; auto. * intros. eapply (erases_deps_cons {| universes := univs; declarations := Σ |} _ kn (PCUICEnvironment.ConstantDecl cst')); auto. @@ -734,7 +736,7 @@ Proof. destruct ?; [|easy]. eapply (erases_extends (Σu, cst_universes cst')). 4:{ split; cbn; auto. eexists [_]; cbn; reflexivity. } - all: cbn; eauto. + all: cbn; eauto. apply decl_ext. * intros. apply H0 in H1. eapply (erases_deps_cons Σu); eauto. apply wfΣu. apply wf. @@ -774,8 +776,8 @@ Proof. apply neqb in n. destruct eqb; cbn in n; try congruence. Qed. -Lemma erases_global_erases_deps Σ Γ t T et Σ' : - wf_ext Σ -> +Lemma erases_global_erases_deps (Σ : global_env_ext) Γ t T et Σ' : + wf Σ -> Σ;;; Γ |- t : T -> Σ;;; Γ |- t ⇝ℇ et -> erases_global Σ Σ' -> diff --git a/erasure/theories/EOptimizePropDiscr.v b/erasure/theories/EOptimizePropDiscr.v index 64a4ef6f7..fc9290fb4 100644 --- a/erasure/theories/EOptimizePropDiscr.v +++ b/erasure/theories/EOptimizePropDiscr.v @@ -295,7 +295,7 @@ Lemma isType_tSort {cf:checker_flags} {Σ : global_env_ext} {Γ l A} {wfΣ : wf Proof. intros HT. eapply inversion_Sort in HT as [l' [wfΓ Hs]]; auto. - eexists; econstructor; eauto. + eexists; split => //; econstructor; eauto. Qed. Lemma isType_it_mkProd {cf:checker_flags} {Σ : global_env_ext} {Γ na dom codom A} {wfΣ : wf Σ} : @@ -303,8 +303,8 @@ Lemma isType_it_mkProd {cf:checker_flags} {Σ : global_env_ext} {Γ na dom codom isType Σ Γ (tProd na dom codom). Proof. intros HT. - eapply inversion_Prod in HT as (? & ? & ? & ? & ?); auto. - eexists; econstructor; eauto. + eapply inversion_Prod in HT as (? & ? & ? & ? & ? & ?); auto. + eexists; split => //; econstructor; eauto. Qed. Definition optimize_constant_decl Σ cb := diff --git a/erasure/theories/ESubstitution.v b/erasure/theories/ESubstitution.v index 48933e81c..97ca925c4 100644 --- a/erasure/theories/ESubstitution.v +++ b/erasure/theories/ESubstitution.v @@ -61,7 +61,6 @@ Proof. unfold PCUICAst.declared_minductive in *. eapply PCUICWeakeningEnv.extends_lookup in H1; eauto; tc. - 2:{ cbn. apply extends_refl. } rewrite H1 in H. inversion H. subst. clear H. rewrite H3 in H4. inversion H4. subst. clear H4. split. eauto. econstructor. eauto. @@ -148,7 +147,7 @@ Proof. intros hctx. rewrite /inst_case_branch_context /= /id. rewrite -rename_context_lift_context. - rewrite PCUICRenameConv.rename_inst_case_context_wf //. + rewrite PCUICRenameTerm.rename_inst_case_context_wf //. f_equal. apply map_ext => x. now setoid_rewrite <- PCUICSigmaCalculus.lift_rename. Qed. @@ -166,7 +165,7 @@ Lemma erases_weakening' (Σ : global_env_ext) (Γ Γ' Γ'' : context) (t T : ter Σ ;;; Γ ,,, Γ'' ,,, lift_context #|Γ''| 0 Γ' |- (lift #|Γ''| #|Γ'| t) ⇝ℇ (ELiftSubst.lift #|Γ''| #|Γ'| t'). Proof. intros HΣ HΓ'' * H He. - generalize_eqs H. intros eqw. rewrite <- eqw in *. + remember (Γ ,,, Γ') as Γ0 eqn:eqw. revert Γ Γ' Γ'' HΓ'' eqw t' He. revert Σ HΣ Γ0 t T H . apply (typing_ind_env (fun Σ Γ0 t T => forall Γ Γ' Γ'', @@ -185,27 +184,27 @@ Proof. - destruct ?; econstructor. - econstructor. unfold app_context, snoc in *. - pose proof (H0 Γ (vass n t :: Γ') Γ''). - rewrite lift_context_snoc0 - plus_n_O in H1. - eapply H1; eauto. cbn. econstructor. - eauto. cbn. exists s1. eapply (weakening_typing (T := tSort s1)); eauto. + pose proof (H1 Γ (vass na t :: Γ') Γ''). + rewrite lift_context_snoc0 - plus_n_O in H2. + eapply H2; eauto. cbn. econstructor. + eauto. cbn. exists s1. split => //. eapply (weakening_typing (T := tSort s1)); eauto. now apply All_local_env_app_inv in X2. - econstructor. - + eapply H0; eauto. - + pose proof (H1 Γ (vdef n b b_ty :: Γ') Γ''). - rewrite lift_context_snoc0 -plus_n_O in H2. - eapply H2; eauto. cbn. econstructor. + + eapply H1; eauto. + + pose proof (H2 Γ (vdef na b b_ty :: Γ') Γ''). + rewrite lift_context_snoc0 -plus_n_O in H3. + eapply H3; eauto. cbn. econstructor. eauto. hnf. 2: cbn; eapply weakening_typing; eauto. eapply weakening_typing in X0; eauto. now apply All_local_env_app_inv in X3. 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. + clear H7. eapply All2i_nth_hyp in X6. eapply All2_map. eapply All2i_All2_All2; tea; cbv beta. @@ -221,8 +220,8 @@ Proof. rewrite -app_context_assoc -{1}(Nat.add_0_r #|Γ'|) -(lift_context_app _ 0). assert (#|inst_case_branch_context p br| = #|bcontext br|). { rewrite /inst_case_branch_context. now len. } - rewrite /map_branch_k /= -H6 -app_length. - rewrite -e2 map_length -H6 -app_length. + rewrite /map_branch_k /= -H7 -app_length. + rewrite -e2 map_length -H7 -app_length. rewrite -(PCUICCasesContexts.inst_case_branch_context_eq a). eapply e. eapply weakening_wf_local => //. @@ -378,7 +377,7 @@ Lemma erases_subst (Σ : global_env_ext) Γ Γ' Δ t s t' s' T : Σ ;;; (Γ ,,, subst_context s 0 Δ) |- (subst s #|Δ| t) ⇝ℇ ELiftSubst.subst s' #|Δ| t'. Proof. intros HΣ HΔ Hs Ht He. - generalize_eqs Ht. intros eqw. + remember (Γ ,,, Γ' ,,, Δ) as Γ0 eqn:eqw in Ht. revert Γ Γ' Δ t' s Hs HΔ He eqw. revert Σ HΣ Γ0 t T Ht. eapply (typing_ind_env (fun Σ Γ0 t T => @@ -408,24 +407,24 @@ Proof. + econstructor. eapply is_type_subst; eauto. - inv H0. econstructor. eapply is_type_subst; eauto. - - inv H1. econstructor. + - inv H2. econstructor. eapply is_type_subst; eauto. - - inv H1. + - inv H2. + cbn. econstructor. - specialize (H0 Γ Γ' (vass n t :: Δ) t'0 s). + specialize (H1 Γ Γ' (vass na t :: Δ) t'0 s). (* unfold app_context, snoc in *. *) - rewrite subst_context_snoc0 in H0. - eapply H0; eauto. + rewrite subst_context_snoc0 in H1. + eapply H1; eauto. cbn. econstructor. eauto. - cbn. exists s1. eapply (substitution (T := tSort s1)); eauto. + cbn. exists s1. split => //. eapply (substitution (T := tSort s1)); eauto. + econstructor. eapply is_type_subst; eauto. - - inv H2. + - inv H3. + cbn. econstructor. eauto. - specialize (H1 Γ Γ' (vdef n b b_ty :: Δ) t2' s). - rewrite subst_context_snoc0 in H1. - eapply H1; eauto. + specialize (H2 Γ Γ' (vdef na b b_ty :: Δ) t2' s). + rewrite subst_context_snoc0 in H2. + eapply H2; eauto. cbn. econstructor. eauto. hnf. eapply substitution in X0; eauto. @@ -448,15 +447,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 := _ |}). @@ -473,10 +472,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 f559b1aa2..d17d9cdae 100644 --- a/erasure/theories/ErasureCorrectness.v +++ b/erasure/theories/ErasureCorrectness.v @@ -39,8 +39,8 @@ Notation "Σ ⊢ s ▷ t" := (EWcbvEval.eval Σ s t) (at level 50, s, t at next Import ssrbool. -Lemma erases_correct (wfl := default_wcbv_flags) Σ t T t' v Σ' : - wf_ext Σ -> +Lemma erases_correct (wfl := default_wcbv_flags) (Σ : global_env_ext) t T t' v Σ' : + wf Σ -> Σ;;; [] |- t : T -> Σ;;; [] |- t ⇝ℇ t' -> erases_deps Σ Σ' t' -> @@ -59,7 +59,7 @@ Proof. eapply IHeval1 in H4 as (vf' & Hvf' & [He_vf']); eauto. eapply IHeval2 in H6 as (vu' & Hvu' & [He_vu']); eauto. pose proof (subject_reduction_eval t0 H). - eapply inversion_Lambda in X0 as (? & ? & ? & ? & e0). + eapply inversion_Lambda in X0 as (? & ? & ? & ? & ? & e0). assert (Σ ;;; [] |- a' : t). { eapply subject_reduction_eval; eauto. eapply PCUICConversion.ws_cumul_pb_Prod_Prod_inv in e0 as [? e1 e2]. @@ -89,7 +89,7 @@ Proof. eapply (is_type_subst Σ [] [vass na _] [] _ [a']) in X1 ; auto. cbn in X1. eapply Is_type_eval; try assumption. - eauto. eapply H1. rewrite <-eqs. eassumption. + eapply H1. rewrite <-eqs. eassumption. all: eauto. econstructor. econstructor. rewrite PCUICLiftSubst.subst_empty. eauto. constructor. econstructor. eauto. eauto. * auto. @@ -99,7 +99,7 @@ Proof. + auto. - assert (Hty' := Hty). assert (Σ |-p tLetIn na b0 t b1 ▷ res) by eauto. - eapply inversion_LetIn in Hty' as (? & ? & ? & ? & ? & ?); auto. + eapply inversion_LetIn in Hty' as (? & ? & ? & ? & ? & ? & ?); auto. invs He. + depelim Hed. eapply IHeval1 in H6 as (vt1' & Hvt2' & [He_vt1']); eauto. @@ -149,14 +149,14 @@ Proof. edestruct IHeval as (? & ? & [?]). * cbn in *. assert (isdecl'' := isdecl'). - eapply declared_constant_inv in isdecl'; [| |now eauto|now apply wfΣ]. + eapply declared_constant_inv in isdecl' as (onty & onbo); [| |now eauto|now apply wfΣ]. 2:eapply weaken_env_prop_typing. - unfold on_constant_decl in isdecl'. rewrite e in isdecl'. red in isdecl'. + rewrite e in onbo. red in onbo. unfold declared_constant in isdecl''. now eapply @typing_subst_instance_decl with (Σ := Σ) (Γ := []); eauto. * assert (isdecl'' := isdecl'). - eapply declared_constant_inv in isdecl'; [| |now eauto|now apply wfΣ]. - unfold on_constant_decl in isdecl'. rewrite e in isdecl'. cbn in *. + eapply declared_constant_inv in isdecl' as (onty & onbo); [| |now eauto|now apply wfΣ]. + rewrite e in onbo. cbn in *. 2:eapply weaken_env_prop_typing. now eapply erases_subst_instance_decl with (Σ := Σ) (Γ := []); eauto. * now eauto. @@ -177,7 +177,7 @@ Proof. assert (Σ ;;; [] |- mkApps (tConstruct ind c u) args : mkApps (tInd ind (puinst p)) (pparams p ++ indices)). eapply subject_reduction_eval; eauto. assert (Hcstr := X0). - eapply Construct_Ind_ind_eq in X0; tea. 2:pcuic. + eapply Construct_Ind_ind_eq in X0; tea. destruct X0 as (((([_ Ru] & cu) & cpu) & ?) & (parsubst & argsubst & parsubst' & argsubst' & [])). invs He. + depelim Hed. @@ -656,7 +656,7 @@ Proof. -- cbn. destruct a2 as [eqb eqrar isl isl' e5]. eapply (erases_subst Σ [] (fix_context mfix) [] dbody (fix_subst mfix)) in e5; cbn; eauto. - ++ eapply subslet_fix_subst. now eapply wf_ext_wf. all: eassumption. + ++ eapply subslet_fix_subst. all: eassumption. ++ eapply nth_error_all in a1; eauto. cbn in a1. eapply a1. ++ eapply All2_from_nth_error. @@ -775,7 +775,7 @@ Proof. specialize (IHeval2 _ htunfcof). eapply inversion_CoFix in t; destruct_sigma t; auto. eapply PCUICSpine.typing_spine_strengthen in t0; eauto. - 2:{ now eapply nth_error_all in a; tea. } + 2:{ now eapply nth_error_all, isType_of_isTypeRel in a; tea. } invs He. * edestruct IHeval1 as (? & ? & ?); eauto. now depelim Hed. depelim Hed. @@ -821,7 +821,8 @@ Proof. unfold unfold_cofix. intros hnth; rewrite hnth. intros [=]. subst fn narg. eapply All_nth_error in a0 as a'; tea. - eapply erases_subst0. eauto. 2:eauto. pcuic. all:tea. + rewrite /on_def_body //= in a'. + eapply erases_subst0. 3:eauto. 2:pcuic. all:tea. { rewrite app_context_nil_l. eapply subslet_cofix_subst; eauto. econstructor; eauto. } {eapply All2_from_nth_error. @@ -935,7 +936,8 @@ Proof. unfold unfold_cofix. intros hnth'; rewrite hnth'. intros [=]. subst fn narg. rewrite hnth' in e2. noconf e2. eapply All_nth_error in a0 as a'; tea. - eapply erases_subst0. eauto. 2:eauto. pcuic. all:tea. + rewrite /on_def_body //= in a'. + eapply erases_subst0. 3:eauto. 2:pcuic. all:tea. { rewrite app_context_nil_l. eapply subslet_cofix_subst; eauto. econstructor; eauto. } {eapply All2_from_nth_error. @@ -1128,9 +1130,10 @@ Proof. - constructor. - cbn. destruct cb' as [[]]. cbn in *. depelim wf. + destruct o0 as (onty & onbo). rewrite [forallb _ _](IHer wf) andb_true_r. red in H. destruct cb as [ty []]; cbn in *. - unshelve eapply PCUICClosedTyp.subject_closed in o0. cbn. split; auto. + unshelve eapply PCUICClosedTyp.subject_closed in onbo. cbn. split; auto. eapply erases_closed in H; tea. elim H. cbn. apply IHer. now depelim wf. - depelim wf. @@ -1305,4 +1308,4 @@ Proof. + constructor => //. eapply erases_deps_mkApps_inv in etaΣ as []. solve_all. -Qed. \ No newline at end of file +Qed. diff --git a/erasure/theories/ErasureFunction.v b/erasure/theories/ErasureFunction.v index 4b6fc0280..cd476f471 100644 --- a/erasure/theories/ErasureFunction.v +++ b/erasure/theories/ErasureFunction.v @@ -149,9 +149,9 @@ Section fix_sigma. - intros. destruct H' as []. rewrite <- (abstract_env_ext_irr _ H2) in X0; eauto. rewrite <- (abstract_env_ext_irr _ H2) in wf; eauto. - eapply inversion_Prod in X0 as (? & ? & ? & ? & ?) ; auto. + eapply inversion_Prod in X0 as (? & ? & ? & ? & ? & ?) ; auto. eapply cored_red in H0 as []. - econstructor. econstructor. econstructor. eauto. + econstructor. econstructor. tea. econstructor. eauto. 2:reflexivity. econstructor; pcuic. rewrite <- (abstract_env_ext_irr _ H2) in X0; eauto. eapply subject_reduction; eauto. @@ -191,7 +191,7 @@ Section fix_sigma. pose proof (abstract_env_ext_wf _ wfΣ') as [wf]. sq. eapply subject_reduction_closed in HT; tea. - eapply inversion_Prod in HT as [? [? [? []]]]. + eapply inversion_Prod in HT as (? & ? & ? & ? & ? & ?). now eapply typing_wf_local in t1. pcuic. pcuic. - clear rprod is_arity rsort a0. intros Σ' wfΣ'; specialize (H Σ' wfΣ'). @@ -200,7 +200,7 @@ Section fix_sigma. pose proof (abstract_env_ext_wf _ wfΣ') as [wf]. sq. eapply subject_reduction_closed in HT; tea. - eapply inversion_Prod in HT as [? [? [? []]]]. + eapply inversion_Prod in HT as (? & ? & ? & ? & ? & ?). now eexists. all:pcuic. - cbn. clear rsort is_arity rprod. intros Σ' wfΣ'; specialize (H Σ' wfΣ'). @@ -336,7 +336,7 @@ Proof. { apply nisa. intros. rewrite (abstract_env_ext_irr _ H wfΣ). eapply invert_cumul_arity_r; tea. } { destruct s as [Hs]. - unshelve epose proof (H := unique_sorting_equality_propositional _ _ wf Hs Hu' p) => //. reflexivity. reflexivity. congruence. } + unshelve epose proof (H := unique_sorting_equality_propositional _ wf Hs Hu' p) => //. reflexivity. congruence. } Qed. Equations? is_erasable {X_type X} (Γ : context) (t : PCUICAst.term) @@ -478,12 +478,12 @@ Section Erase. - now eapply inversion_Evar in Ht. - discriminate. - discriminate. - - eapply inversion_Lambda in Ht as (? & ? & ? & ? & ?); auto. + - eapply inversion_Lambda in Ht as (? & ? & ? & ? & ? & ?); auto. eexists; eauto. - - eapply inversion_LetIn in Ht as (? & ? & ? & ? & ? & ?); auto. + - eapply inversion_LetIn in Ht as (? & ? & ? & ? & ? & ? & ?); auto. eexists; eauto. - simpl in *. - eapply inversion_LetIn in Ht as (? & ? & ? & ? & ? & ?); auto. + eapply inversion_LetIn in Ht as (? & ? & ? & ? & ? & ? & ?); auto. eexists; eauto. - eapply inversion_App in Ht as (? & ? & ? & ? & ? & ?); auto. eexists; eauto. @@ -671,7 +671,7 @@ Proof. split => //. cbn. rewrite eq. now constructor. split => //. cbn. rewrite eq. destruct H2. - eapply Is_type_lambda in X2; tea. 2:pcuic. destruct X2. + eapply Is_type_lambda in X2; tea. 2: apply wf. 2:pcuic. destruct X2. now constructor. Qed. @@ -1024,9 +1024,9 @@ Proof. | [ H : KernameSet.In _ (KernameSet.union _ _) |- _ ] => apply KernameSet.union_spec in hin as [?|?] end. - - apply inversion_Lambda in wt as (? & ? & ? & ? & ?); eauto. - - apply inversion_LetIn in wt as (? & ? & ? & ? & ? & ?); eauto. - - apply inversion_LetIn in wt as (? & ? & ? & ? & ? & ?); eauto. + - apply inversion_Lambda in wt as (? & ? & ? & ? & ? & ?); eauto. + - apply inversion_LetIn in wt as (? & ? & ? & ? & ? & ? & ?); eauto. + - apply inversion_LetIn in wt as (? & ? & ? & ? & ? & ? & ?); eauto. - apply inversion_App in wt as (? & ? & ? & ? & ? & ?); eauto. - apply inversion_App in wt as (? & ? & ? & ? & ? & ?); eauto. - apply inversion_Const in wt as (? & ? & ? & ? & ?); eauto. @@ -1103,8 +1103,8 @@ Proof. end. - now apply inversion_Evar in wt. - constructor. - now apply inversion_Lambda in wt as (? & ? & ? & ? & ?); eauto. - - apply inversion_LetIn in wt as (? & ? & ? & ? & ? & ?); eauto. + now apply inversion_Lambda in wt as (? & ? & ? & ? & ? & ?); eauto. + - apply inversion_LetIn in wt as (? & ? & ? & ? & ? & ? & ?); eauto. constructor; eauto. - apply inversion_App in wt as (? & ? & ? & ? & ? & ?); eauto. now constructor; eauto. @@ -1224,9 +1224,9 @@ Proof. destruct (cst_body cb) eqn:cbe; destruct (E.cst_body cb') eqn:cbe'; auto. specialize (H3 _ eq_refl). - eapply on_declared_constant in H; auto. + eapply on_declared_constant in H as (onty & onbo); auto. 2:{ inv wfΣ. now inv X. } - red in H. rewrite cbe in H. simpl in H. + rewrite cbe in onbo. simpl in onbo. eapply (erases_weakeninv_env (Σ := (Σ, cst_universes cb)) (Σ' := (add_global_decl Σ (kn, d), cst_universes cb))); eauto. simpl. @@ -1356,7 +1356,7 @@ Lemma erase_constant_body_correct X_type X cb Proof. red. destruct cb as [name [bod|] univs]; simpl; eauto. intros. set (ecbo := erase_constant_body_obligation_1 X_type X _ _ _ _). clearbody ecbo. - cbn in *. specialize_Σ H. sq. + cbn in *. specialize_Σ H. sq. destruct onc as (_ & onbo); simpl in onbo. eapply (erases_weakeninv_env (Σ := Σ) (Σ' := (Σ', univs))); simpl; eauto. now eapply erases_erase. Qed. @@ -1576,7 +1576,7 @@ Proof. epose proof (abstract_env_exists X) as [[Σ wfΣX]]. now rewrite (abstract_make_wf_env_ext_correct X univs wfext Σ _ wfΣX wfΣex). epose proof (abstract_env_exists X) as [[Σ wfΣX]]. - eapply erases_correct; tea. + eapply erases_correct; tea. apply wfΣ. rewrite (abstract_make_wf_env_ext_correct X univs wfext _ _ wfΣX wfΣex); eauto. Qed. @@ -1886,7 +1886,7 @@ Proof. eapply expanded_erases. apply wf. eapply erases_erase; eauto. assumption. pose proof (wt _ wfΣ). destruct H as [T ht]. - eapply erases_global_erases_deps; tea. + eapply erases_global_erases_deps; tea. apply wf. eapply erases_erase; eauto. Qed. diff --git a/erasure/theories/ErasureProperties.v b/erasure/theories/ErasureProperties.v index 8450c1c80..20888ecf4 100644 --- a/erasure/theories/ErasureProperties.v +++ b/erasure/theories/ErasureProperties.v @@ -42,12 +42,14 @@ Proof. intros Σ wfΣ Γ Γ' X1 Γ0 ? w0. induction w0. - econstructor. - econstructor; eauto. cbn in *. - destruct t0. exists x. eapply context_conversion with (Γ ,,, Γ0); eauto. + apply infer_typing_sort_impl with id t0 => Ht //. + eapply context_conversion with (Γ ,,, Γ0); eauto. * eapply wf_local_app; eauto. * eapply conv_context_app_same; eauto. - econstructor; eauto. + cbn in *. - destruct t0. exists x. eapply context_conversion with (Γ ,,, Γ0); eauto. + apply infer_typing_sort_impl with id t0 => Ht //. + eapply context_conversion with (Γ ,,, Γ0); eauto. * eapply wf_local_app; eauto. * eapply conv_context_app_same; eauto. + cbn in *. eapply context_conversion with (Γ ,,, Γ0); eauto. @@ -65,7 +67,7 @@ Proof. Qed. -Lemma type_closed_subst {Σ t T} u : wf_ext Σ -> +Lemma type_closed_subst {Σ t T} u : wf Σ.1 -> Σ ;;; [] |- t : T -> subst1 t 0 u = PCUICCSubst.csubst t 0 u. Proof. @@ -199,11 +201,11 @@ Proof. - econstructor. eapply h_forall_Γ'0. econstructor. eauto. now constructor. constructor; auto. - exists s1. + exists s1; split => //. eapply context_conversion. 3:eauto. all:eauto. - econstructor. eauto. eapply h_forall_Γ'1. econstructor. eauto. now constructor. - constructor; auto. exists s1. + constructor; auto. exists s1; split => //. eapply context_conversion with Γ; eauto. eapply context_conversion with Γ; eauto. eassumption. @@ -300,16 +302,17 @@ Proof. all: match goal with [ H : erases _ _ ?a _ |- ?G ] => tryif is_var a then idtac else invs H end. all: try now (econstructor; eauto 2 using isErasable_subst_instance). - cbn. econstructor. - eapply H0 in X2; eauto. apply X2. - cbn. econstructor. eauto. cbn. econstructor. + eapply H1 in X2; eauto. apply X2. + cbn. econstructor. eauto. cbn. eexists; split. apply relevance_subst, H. eapply typing_subst_instance in X0; eauto. apply snd in X0. cbn in X0. destruct X0. refine (t0 _ _ _ _); eauto. - cbn. econstructor. - eapply H0 in X3; eauto. - eapply H1 in X3; eauto. exact X3. + eapply H1 in X3; eauto. + eapply H2 in X3; eauto. exact X3. cbn. econstructor. eauto. cbn. econstructor. eapply typing_subst_instance in X0; eauto. apply snd in X0. cbn in X0. + split. apply relevance_subst, H. eapply X0; eauto. cbn. eapply typing_subst_instance in X1; eauto. apply snd in X1. cbn in X1. eapply X1; eauto. @@ -330,7 +333,7 @@ Proof. now rewrite inst_case_branch_context_subst_instance. - assert (Hw : wf_local (Σ.1, univs) (subst_instance u (Γ ,,, types))). { (* rewrite subst_instance_app. *) - assert(All (fun d => isType Σ Γ (dtype d)) mfix). + assert(All (on_def_type (lift_typing typing Σ) Γ) mfix). eapply (All_impl X0); pcuicfo. now destruct X5 as [s [Hs ?]]; exists s. eapply All_mfix_wf in X5; auto. subst types. @@ -339,13 +342,13 @@ Proof. induction 1. - eauto. - cbn. econstructor; eauto. cbn in *. - destruct t0 as (? & ?). eexists. - cbn. eapply typing_subst_instance in t0; eauto. apply snd in t0. cbn in t0. - eapply t0; eauto. + eapply infer_typing_sort_impl with _ t0; [apply relevance_subst_opt|]. intros Hs. + eapply typing_subst_instance in Hs; eauto. apply snd in Hs. cbn in Hs. + eapply Hs; eauto. - cbn. econstructor; eauto. cbn in *. - destruct t0 as (? & ?). eexists. - cbn. eapply typing_subst_instance in t0; eauto. apply snd in t0. - eapply t0; eauto. + eapply infer_typing_sort_impl with _ t0; [apply relevance_subst_opt|]. intros Hs. + eapply typing_subst_instance in Hs; eauto. apply snd in Hs. cbn in Hs. + eapply Hs; eauto. cbn in *. eapply typing_subst_instance in t1; eauto. apply snd in t1. eapply t1. all:eauto. } @@ -366,22 +369,22 @@ Proof. - assert (Hw : wf_local (Σ.1, univs) (subst_instance u (Γ ,,, types))). { (* rewrite subst_instance_app. *) - assert(All (fun d => isType Σ Γ (dtype d)) mfix). + assert(All (on_def_type (lift_typing typing Σ) Γ) mfix). eapply (All_impl X0); pcuicfo. - destruct X5 as [s [Hs ?]]; now exists s. + now apply infer_typing_sort_impl with id X5 => // []. eapply All_mfix_wf in X5; auto. subst types. revert X5. clear - wfΣ wfΓ H2 X2 X3. induction 1. - eauto. - cbn. econstructor; eauto. cbn in *. - destruct t0 as (? & ?). eexists. - cbn. eapply typing_subst_instance in t0; eauto. apply snd in t0. cbn in t0. - eapply t0; eauto. + eapply infer_typing_sort_impl with _ t0; [apply relevance_subst_opt|]. intros Hs. + eapply typing_subst_instance in Hs; eauto. apply snd in Hs. cbn in Hs. + eapply Hs; eauto. - cbn. econstructor; eauto. cbn in *. - destruct t0 as (? & ?). eexists. - cbn. eapply typing_subst_instance in t0; eauto. apply snd in t0. - eapply t0; eauto. + eapply infer_typing_sort_impl with _ t0; [apply relevance_subst_opt|]. intros Hs. + eapply typing_subst_instance in Hs; eauto. apply snd in Hs. cbn in Hs. + eapply Hs; eauto. cbn in *. eapply typing_subst_instance in t1; eauto. apply snd in t1. eapply t1. all:eauto. } @@ -528,17 +531,17 @@ Section wellscoped. - now rewrite (declared_inductive_lookup isdecl). - rewrite (declared_constructor_lookup isdecl) //. - now rewrite (declared_inductive_lookup 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. - now rewrite (declared_projection_lookup isdecl). - now eapply nth_error_Some_length, Nat.ltb_lt in H0. - move/andb_and: H2 => [] hb _. - solve_all. destruct a as [s []], a0. + solve_all. destruct a as [s (_ & ? & ?)], a0. unfold test_def. len in b0. rewrite b0. now rewrite i i0. - now eapply nth_error_Some_length, Nat.ltb_lt in H0. - - solve_all. destruct a as [s []], b. + - solve_all. destruct a as [s (_ & ? & ?)], b. unfold test_def. len in i0. now rewrite i i0. Qed. diff --git a/erasure/theories/Prelim.v b/erasure/theories/Prelim.v index 7166b4341..9aca75f48 100644 --- a/erasure/theories/Prelim.v +++ b/erasure/theories/Prelim.v @@ -52,7 +52,7 @@ Proof. Qed. Lemma cumul_Sort_Prod_discr {Σ Γ T s na A B} : - wf_ext Σ -> + wf Σ.1 -> Σ ;;; Γ ⊢ T ≤ tSort s -> Σ ;;; Γ ⊢ T ≤ tProd na A B -> False. Proof. @@ -214,7 +214,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, isType_of_isTypeRel in a; tea. Qed. (** Assumption contexts: constructor arguments/case branches contexts contain only assumptions, no local definitions *) @@ -237,7 +237,7 @@ Proof. destruct nas; cbn; auto; constructor. auto. Qed. -Lemma declared_constructor_assumption_context (wfl := default_wcbv_flags) {Σ c mdecl idecl cdecl} {wfΣ : wf_ext Σ} : +Lemma declared_constructor_assumption_context (wfl := default_wcbv_flags) {Σ c mdecl idecl cdecl} {wfΣ : wf Σ} : declared_constructor Σ c mdecl idecl cdecl -> assumption_context (cstr_args cdecl). Proof. @@ -247,7 +247,7 @@ Proof. now eapply is_assumption_context_spec. Qed. -Lemma assumption_context_cstr_branch_context (wfl := default_wcbv_flags) {Σ} {wfΣ : wf_ext Σ} {c mdecl idecl cdecl} : +Lemma assumption_context_cstr_branch_context (wfl := default_wcbv_flags) {Σ} {wfΣ : wf Σ} {c mdecl idecl cdecl} : declared_constructor Σ c mdecl idecl cdecl -> assumption_context (cstr_branch_context c.1 mdecl cdecl). Proof. @@ -256,7 +256,7 @@ Proof. rewrite /cstr_branch_context. pcuic. Qed. -Lemma expand_lets_erasure (wfl := default_wcbv_flags) {Σ mdecl idecl cdecl c brs p} {wfΣ : wf_ext Σ} : +Lemma expand_lets_erasure (wfl := default_wcbv_flags) {Σ mdecl idecl cdecl c brs p} {wfΣ : wf Σ} : declared_constructor Σ c mdecl idecl cdecl -> wf_branches idecl brs -> All2i (fun i cdecl br => diff --git a/pcuic/theories/PCUICRelevanceTyp.v b/pcuic/theories/PCUICRelevanceTyp.v index 4fa297117..0b319852b 100644 --- a/pcuic/theories/PCUICRelevanceTyp.v +++ b/pcuic/theories/PCUICRelevanceTyp.v @@ -1,5 +1,6 @@ From MetaCoq.Template Require Import config utils. From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICRelevance + PCUICClosedTyp PCUICGlobalEnv PCUICTyping PCUICInversion PCUICConversion PCUICCumulProp PCUICWeakeningTyp PCUICValidity PCUICWellScopedCumulativity BDUnique. Require Import ssreflect. @@ -229,7 +230,8 @@ Proof using Type. eapply cumul_Trans. 4: apply cumul_Sort, le. 1,2: fvs. apply cumulAlgo_cumulSpec. apply invert_red_letin in red as [(? & ? & ? & []) | ?]; try discriminate. - epose proof (PCUICSpine.all_rels_subst_lift (Γ := Γ) (Δ := [vdef na b b_ty]) (t := s2) (Δ' := []) _ _ _) => //=. + unshelve epose proof (PCUICSpine.all_rels_subst_lift (Γ := Γ) (Δ := [vdef na b b_ty]) (t := s2) (Δ' := []) _ _ _) => //=. + all: change (Γ ,,, [_]) with (Γ ,, vdef na b b_ty). pcuic. fvs. fvs. simpl in X0. rewrite PCUICLiftSubst.lift0_id in X0. rewrite PCUICLiftSubst.subst_empty in X0. change (subst0 _ _) with ((lift 1 1 s2) {0 := lift0 1 b}) in X0. rewrite -PCUICLiftSubst.distr_lift_subst10 in X0. @@ -266,7 +268,7 @@ Proof using Type. destruct Σ as (Σ & φ); unfold fst in *. instantiate (1 := u). eapply (PCUICUnivSubstitutionTyp.typing_subst_instance' _ _ [] _ (tSort H'.π1) u _); tea. - epose proof (PCUICWeakeningEnv.weaken_lookup_on_global_env' _ _ (ConstantDecl decl) wf _); eauto. + unshelve epose proof (PCUICWeakeningEnv.weaken_lookup_on_global_env' _ _ (ConstantDecl decl) wf _); eauto. now split. + unfold relevance_of_term, relevance_of_constant. rewrite (declared_constant_lookup H). @@ -277,7 +279,7 @@ Proof using Type. apply (weaken_ctx (Γ := [])); tea. destruct Σ as (Σ & φ); unfold fst in *. instantiate (1 := u). eapply (PCUICUnivSubstitutionTyp.typing_subst_instance' _ _ [] _ (tSort _) u _); tea. - epose proof (PCUICWeakeningEnv.weaken_lookup_on_global_env' _ _ (ConstantDecl decl) wf _); eauto. + unshelve epose proof (PCUICWeakeningEnv.weaken_lookup_on_global_env' _ _ (ConstantDecl decl) wf _); eauto. now split. - pose proof (onArity (PCUICWeakeningEnvTyp.on_declared_inductive isdecl).2). destruct Σ as (Σ & φ); unfold fst in *. @@ -288,13 +290,13 @@ Proof using Type. eapply (weaken_ctx (Γ := [])) in Hty. apply Hty. all: tea. - epose proof (PCUICWeakeningEnv.weaken_lookup_on_global_env' _ _ (InductiveDecl mdecl) wf _); eauto. now split. + unshelve epose proof (PCUICWeakeningEnv.weaken_lookup_on_global_env' _ _ (InductiveDecl mdecl) wf isdecl.p1); eauto. now split. + unfold relevance_of_term. destruct Hr as (s1 & <- & Hs1), X1 as (s2 & <- & Hs2). eapply PCUICUnivSubstitutionTyp.typing_subst_instance' in Hs2. eapply (weaken_ctx (Γ := [])) in Hs2. all: tea. - 2: { epose proof (PCUICWeakeningEnv.weaken_lookup_on_global_env' _ _ (InductiveDecl mdecl) wf _); eauto. now split. } + 2: { epose proof (PCUICWeakeningEnv.weaken_lookup_on_global_env' _ _ (InductiveDecl mdecl) wf isdecl.p1); eauto. now split. } erewrite <- relevance_subst_eq. eapply cumul_sort_relevance; tea. apply wf. - pose proof (declared_constructor_lookup isdecl). @@ -317,19 +319,13 @@ Proof using Type. eapply isTypeRel2_relevance; tea. exists ps; split => //. - pose proof (declared_projection_lookup isdecl). - destruct (PCUICWeakeningEnvTyp.on_declared_projection isdecl) as [[[] e] [[[? []] ?] ?]]; tea. - unfold on_projection in o. - destruct nth_error eqn:EE in o => //. - destruct o. + assert (isTypeRel Σ Γ (subst0 (c :: List.rev args) (proj_type pdecl)@[u]) (pdecl.(proj_relevance))) by (eapply isType_of_projection; tea). split; intro Hr => //. + unfold relevance_of_term, relevance_of_projection in Hr; subst rel. - destruct Σ as (Σ & φ); unfold fst, snd in *. - rewrite H. - eapply isType_of_projection; tea. - + unfold relevance_of_term, relevance_of_projection. destruct Σ as (Σ & φ); unfold fst, snd in *. + now rewrite H. + + unfold relevance_of_term, relevance_of_projection. rewrite H. - eassert (isTypeRel (Σ, φ) _ (subst0 (c :: List.rev args) (proj_type pdecl)@[u]) (pdecl.(proj_relevance))) by (eapply isType_of_projection; tea). - eapply isTypeRel2_relevance; tea. apply wf. + eapply isTypeRel2_relevance; tea. - eapply All_nth_error in X0, X1; tea. rewrite /on_def_type /on_def_body /lift_typing2 in X0, X1. split; intro Hr => //. @@ -369,7 +365,7 @@ Proof using Type. apply ws_cumul_pb_forget, PCUICCumulativity.cumul_alt in X0 as (A' & B' & [[red1 red2] leq]). eapply PCUICSR.subject_reduction in red1, red2; tea. eapply typing_leq_term_prop; tea. -Admitted. +Qed. Theorem relevance_from_type {cf : checker_flags} (Σ : global_env_ext) Γ t T rel : diff --git a/safechecker/src/metacoq_safechecker_plugin.mlpack b/safechecker/src/metacoq_safechecker_plugin.mlpack index 1c0e60a37..167205fcd 100644 --- a/safechecker/src/metacoq_safechecker_plugin.mlpack +++ b/safechecker/src/metacoq_safechecker_plugin.mlpack @@ -19,6 +19,7 @@ PCUICTyping PCUICWfUniverses PCUICNormal PCUICPosition +PCUICRelevance PCUICPretty PCUICProgram TemplateToPCUIC From 477ea5df1e33b4a9c384643e65c9224019704f5b Mon Sep 17 00:00:00 2001 From: Yann Leray Date: Mon, 23 May 2022 14:59:25 +0200 Subject: [PATCH 08/17] Add the missing file to PluginProject --- safechecker/_PluginProject.in | 2 ++ 1 file changed, 2 insertions(+) diff --git a/safechecker/_PluginProject.in b/safechecker/_PluginProject.in index 788756c48..96bc42b84 100644 --- a/safechecker/_PluginProject.in +++ b/safechecker/_PluginProject.in @@ -72,6 +72,8 @@ src/pCUICPretty.mli src/pCUICPretty.ml src/pCUICProgram.mli src/pCUICProgram.ml +src/pCUICRelevance.ml +src/pCUICRelevance.mli # From SafeChecker # src/string2pos.mli From eeb701438cba04c0e915b6a8da9c1f5d1fe5afd1 Mon Sep 17 00:00:00 2001 From: Yann Leray Date: Mon, 23 May 2022 15:15:46 +0200 Subject: [PATCH 09/17] Add wf hypothesis to admitted lemmas --- pcuic/theories/PCUICRelevanceTyp.v | 16 +++++++++------- safechecker/theories/PCUICSafeChecker.v | 4 +++- 2 files changed, 12 insertions(+), 8 deletions(-) diff --git a/pcuic/theories/PCUICRelevanceTyp.v b/pcuic/theories/PCUICRelevanceTyp.v index 0b319852b..9d3701449 100644 --- a/pcuic/theories/PCUICRelevanceTyp.v +++ b/pcuic/theories/PCUICRelevanceTyp.v @@ -29,13 +29,13 @@ Proof. now eapply cumul_sort_relevance. Qed. -Lemma ind_arity_relevant {cf : checker_flags} (Σ : global_env_ext) mdecl idecl : +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 Relevant. Proof. - assert (wfΣ : wf Σ) by admit. - intros ind_type (s & e & Hs). + intros wfΣ ind_type (s & e & Hs). eexists; split => //; tea. rewrite /ind_type in Hs. rewrite -it_mkProd_or_LetIn_app in Hs. @@ -53,7 +53,9 @@ Proof. eapply IHl. apply wfΓ'. econstructor; eauto with pcuic. constructor; eauto with pcuic. apply ws_cumul_pb_LetIn_l_inv in le. - epose proof (PCUICSpine.all_rels_subst_lift (Γ := Γ) (Δ := [vdef decl_name t decl_type]) (t := x0) (Δ' := []) _ _ _) => //=. + unshelve epose proof (PCUICSpine.all_rels_subst_lift (Γ := Γ) (Δ := [vdef decl_name t decl_type]) (t := x0) (Δ' := []) _ _ _) => //=. + all: change (Γ,,, [vdef decl_name t decl_type]) with (Γ ,, vdef decl_name t decl_type); tea. + 1,2: fvs. simpl in X0. rewrite PCUICLiftSubst.lift0_id in X0. rewrite PCUICLiftSubst.subst_empty in X0. change (subst0 _ _) with ((lift 1 1 x0) {0 := lift0 1 t}) in X0. rewrite -PCUICLiftSubst.distr_lift_subst10 in X0. @@ -71,10 +73,10 @@ Proof. cbn. constructor. 1-3: fvs. constructor. apply leq_universe_product. - Unshelve. all: eauto. fvs. unfold app_context, app; fold (snoc Γ (vdef decl_name t decl_type)). fvs. -Admitted. +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 @@ -92,7 +94,7 @@ Lemma isTypeRel_cstr_type {cf : checker_flags} Σ mdecl idecl cdecl n : isType Σ (arities_context mdecl.(ind_bodies)) (cstr_type cdecl) -> isTypeRel Σ (arities_context mdecl.(ind_bodies)) (cstr_type cdecl) idecl.(ind_relevance). Proof. - intros H -> e1 <- c (s & _ & Hs); exists s; split => //=. + intros wfΣ H -> e1 <- c (s & _ & Hs); exists s; split => //=. (* epose proof (nth_error_arities_context mdecl (#|ind_bodies mdecl| - S cstr.1.(inductive_ind)) idecl _). Unshelve. 2: { rewrite nth_error_rev. len. apply nth_error_Some_length in di. lia. rewrite List.rev_involutive. replace _ with (inductive_ind cstr.1); tea. len. apply nth_error_Some_length in di. lia. } rewrite e1 in H; clear e1. *) diff --git a/safechecker/theories/PCUICSafeChecker.v b/safechecker/theories/PCUICSafeChecker.v index f119eb6f0..3b03beb44 100644 --- a/safechecker/theories/PCUICSafeChecker.v +++ b/safechecker/theories/PCUICSafeChecker.v @@ -1920,6 +1920,7 @@ Section CheckEnv. Qed. Next Obligation. epose proof (get_wt_indices X_ext wfar wfpars _ _ hnth Hcs). + pose proof (hΣ _ _ H). specialize_Σ H. unfold check_constructor_spec in Hcs; simpl in *. sq. solve_all. @@ -2319,6 +2320,7 @@ End monad_Alli_nth_forall. Next Obligation. 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. sq. eapply All_nth_error in wfars as (isty & e & eqrel); tea. @@ -2330,7 +2332,7 @@ End monad_Alli_nth_forall. ind_sorts := onsorts; ind_relevance_compat := eqrel; onIndices := _; |}. - - rewrite e in isty |- *; apply PCUICRelevanceTyp.ind_arity_relevant => //. + - rewrite e in isty |- *; apply PCUICRelevanceTyp.ind_arity_relevant => //. apply H0. - erewrite (abstract_env_ext_irr _ _ pf); eauto. Unshelve. eauto. Qed. From 761890ad66a30f3a82362efb8815dc5bcce142bb Mon Sep 17 00:00:00 2001 From: Yann Leray Date: Tue, 24 May 2022 15:38:34 +0200 Subject: [PATCH 10/17] Make isTermRel an inductive predicate --- pcuic/theories/Conversion/PCUICInstConv.v | 39 +- pcuic/theories/Conversion/PCUICRenameTerm.v | 28 +- .../Conversion/PCUICUnivSubstitutionConv.v | 13 +- pcuic/theories/PCUICRelevance.v | 75 ++-- pcuic/theories/PCUICRelevanceTerm.v | 13 +- pcuic/theories/PCUICRelevanceTyp.v | 78 ++-- safechecker/theories/PCUICSafeChecker.v | 25 +- safechecker/theories/PCUICTypeChecker.v | 352 +++++++++++++----- template-coq/theories/BasicAst.v | 4 + template-coq/theories/Universes.v | 5 + 10 files changed, 400 insertions(+), 232 deletions(-) diff --git a/pcuic/theories/Conversion/PCUICInstConv.v b/pcuic/theories/Conversion/PCUICInstConv.v index c69be650d..d32ed9457 100644 --- a/pcuic/theories/Conversion/PCUICInstConv.v +++ b/pcuic/theories/Conversion/PCUICInstConv.v @@ -1413,7 +1413,7 @@ Proof. intros [HΔ h] HAσ; split. - apply closed_subst_Up; assumption. - intros [|n] decl e. - * now inversion e. + * constructor. now inversion e. * cbn -[rshiftk marks_of_context] in *. rewrite /subst_compose. specialize (h _ _ e). @@ -1503,7 +1503,7 @@ Proof. intros [HΔ h] HAσ; split. - apply closed_subst_Up'; assumption. - intros [|n] decl e. - * now inversion e. + * constructor. now inversion e. * cbn -[rshiftk marks_of_context] in *. rewrite /subst_compose. specialize (h _ _ e). @@ -1601,8 +1601,8 @@ Proof. unfold Upn, subst_consn, subst_compose. destruct (nth_error Δ' x) eqn:E. * assert (x < #|Δ'|) by (apply nth_error_Some; intros e; rewrite e in E; discriminate). - rewrite nth_error_idsn_Some => //. - cbn. rewrite nth_error_map nth_error_app1. + rewrite nth_error_idsn_Some => //. + constructor. rewrite nth_error_map nth_error_app1. { len. } rewrite -nth_error_map /inst_context -map_map map_decl_name_fold_context_k map_map nth_error_map. erewrite <- nth_error_app1, H => //. @@ -1775,37 +1775,34 @@ Lemma relevance_of_term_inst Σ Γ Δ t rel σ : isTermRel Σ (marks_of_context Δ) t.[σ] rel. Proof. intros hσ HΓ Ht h. - induction t using term_forall_list_ind in Ht, σ, Γ, Δ, hσ, HΓ, h |- *. - all: cbn in Ht |- *; auto. + induction t using term_forall_list_ind in Ht, σ, Γ, Δ, hσ, HΓ, h |- *; depelim h. + all: try solve [ try rewrite H; econstructor => //; eauto ]. - destruct hσ as (hσ & hrel). - rewrite /relevance_of_term /marks_of_context nth_error_map /option_map in h. - destruct (nth_error Γ n) eqn:E. - 2: { unfold shiftnP in *. rewrite orb_false_r in Ht. apply Nat.ltb_lt in Ht. apply nth_error_None in E; lia. } - rewrite -h; now apply hrel. + rewrite nth_error_map /option_map in e. + destruct (nth_error Γ n) eqn:E => //. inv e. + now apply hrel. - assert (is_open_term Γ t1). - { rewrite shiftnP_add in Ht; toProp; apply Ht. } - eapply IHt2 with (Δ := Δ ,, vass n t1.[σ]). + { simpl in Ht. toProp; apply Ht. } + constructor; eapply IHt2 with (Δ := Δ ,, vass n t1.[σ]). 4: { instantiate (1 := Γ ,, vass n t1). cbn in h; apply h. } 2: cbn; rewrite alli_app; toProp; tas; cbn; rewrite andb_true_r; len. - 2: rewrite shiftnP_add in Ht; toProp; apply Ht. + 2: rewrite //= shiftnP_add in Ht; toProp; apply Ht. eapply valid_subst_Up in hσ; tas. * eapply valid_subst_ext. 1: apply hσ. now sigma. * eapply inst_is_open_term; tea. apply hσ. - - rewrite shiftnP_add in Ht; move/andP: Ht => [] H0 /andP[] H1 H2. - eapply IHt3 with (Δ := Δ ,, vdef n t1.[σ] t2.[σ]). + - cbn in Ht. move/andP: Ht => [] H1 /andP[] H2 H3. rewrite -> shiftnP_add in H3. + constructor. eapply IHt3 with (Δ := Δ ,, vdef n t1.[σ] t2.[σ]). 4: { instantiate (1 := Γ ,, vdef n t1 t2). cbn in h; apply h. } 2: cbn; rewrite alli_app; toProp; tas; cbn; rewrite andb_true_r; len. - 3: apply H2. + 3: apply H3. 2: { rewrite /on_free_vars_decl /test_decl /decl_body /decl_type /vdef /option_default. toProp => //. } eapply valid_subst_Up' in hσ; tas. * eapply valid_subst_ext. 1: apply hσ. now sigma. * eapply inst_is_open_term; tea. apply hσ. * eapply inst_is_open_term; tea. apply hσ. - - toProp; eapply IHt1; tea. apply Ht. - - unfold option_default. rewrite nth_error_map. unfold relevance_of_term in h. - now destruct (nth_error m). - - unfold option_default. rewrite nth_error_map. unfold relevance_of_term in h. - now destruct (nth_error m). + - constructor. eapply IHt1; tea. simpl in Ht. toProp Ht. apply Ht. + - erewrite map_dname. constructor. rewrite nth_error_map e //. + - erewrite map_dname. constructor. rewrite nth_error_map e //. Qed. diff --git a/pcuic/theories/Conversion/PCUICRenameTerm.v b/pcuic/theories/Conversion/PCUICRenameTerm.v index 7ac7bef32..be3b0c783 100644 --- a/pcuic/theories/Conversion/PCUICRenameTerm.v +++ b/pcuic/theories/Conversion/PCUICRenameTerm.v @@ -1114,28 +1114,28 @@ Lemma rename_relevance_of_term P f Σ Δ Γ t rel : isTermRel Σ (marks_of_context Γ) t rel -> isTermRel Σ (marks_of_context Δ) (rename f t) rel. Proof. - revert f Γ Δ. - induction t using term_forall_list_ind; cbnr; intros f Γ Δ Hur (* wfΓ wfΔ *) wft H; auto. - - unfold relevance_of_term, option_default in *. - rewrite nth_error_map /option_map in H. + intros sP Hur wft H. + induction t using term_forall_list_ind in f, Γ, Δ, sP, Hur, wft, H |- *; depelim H. + all: try solve [ try rewrite H; econstructor => //; eauto ]. + - constructor. + rewrite nth_error_map /option_map in e. rewrite nth_error_map /option_map. - destruct (nth_error Γ n) eqn:E. - 2: { unfold shiftnP in *. rewrite orb_false_r in wft. apply Nat.ltb_lt in wft. apply nth_error_None in E; lia. } + destruct nth_error eqn:E => //. eapply Hur in E as (decl & -> & <- & _) => //. - clear -wft. unfold shiftnP in *. toProp. rewrite orb_false_r in wft. now left. - - cbn in wft; rtoProp. + clear -wft. unfold shiftnP in *. unfold sP. toProp. rewrite //= orb_false_r in wft. now left. + - constructor. + cbn in wft; rtoProp. eapply IHt2 with (Δ := Δ,, vass n (rename _ t1)). 3: instantiate (1 := Γ,, vass n t1); apply H. * eapply urenaming_impl. 1: intro; rewrite shiftnP_S; eauto. apply urenaming_vass; eauto. * eapply on_free_vars_impl; tea. 1: intro; rewrite -shiftnP_S; eauto. - - cbn in wft; rtoProp. + - constructor. + cbn in wft; rtoProp. eapply IHt3 with (Δ := Δ,, vdef n (rename _ t1) (rename _ t2)). 3: instantiate (1 := Γ,, vdef n t1 t2); apply H. * eapply urenaming_impl. 1: intro; rewrite shiftnP_S; eauto. apply urenaming_vdef; eauto. * eapply on_free_vars_impl; tea. 1: intro; rewrite -shiftnP_S; eauto. - - eapply IHt1; tea. now toProp wft. - - unfold relevance_of_term, option_default in *. - rewrite nth_error_map. unfold option_map; cbn. destruct (nth_error m) eqn:E; tas. - - unfold relevance_of_term, option_default in *. - rewrite nth_error_map. unfold option_map; cbn. destruct (nth_error m) eqn:E; tas. + - constructor. eapply IHt1; tea. apply andb_and in wft; now destruct wft. + - erewrite map_dname. econstructor. rewrite nth_error_map e => //. + - erewrite map_dname. econstructor. rewrite nth_error_map e => //. Qed. diff --git a/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v b/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v index 800525136..10fccbb3b 100644 --- a/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v +++ b/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v @@ -1694,14 +1694,11 @@ Definition wf_global_ext {cf : checker_flags} Σ ext := wf_ext_wk (Σ, ext). Lemma relevance_of_term_subst_instance Σ Γ u t rel : isTermRel Σ Γ t rel -> isTermRel Σ Γ t@[u] rel. Proof. - revert Γ. - induction t using term_forall_list_ind; intros Γ Hirrel; cbnr; auto. - - unfold subst_instance, relevance_of_term, option_default in *. - destruct (nth_error m) eqn:E; - rewrite nth_error_map E; now cbn. - - unfold subst_instance, relevance_of_term, option_default in *. - destruct (nth_error m) eqn:E; - rewrite nth_error_map E; now cbn. + intro h. + induction t using term_forall_list_ind in Γ, h |- *; depelim h. + all: try solve [ try rewrite H; econstructor => //; eauto ]. + - erewrite map_dname. econstructor. rewrite nth_error_map e => //. + - erewrite map_dname. econstructor. rewrite nth_error_map e => //. Qed. Require Import Morphisms. diff --git a/pcuic/theories/PCUICRelevance.v b/pcuic/theories/PCUICRelevance.v index d9431d37e..4ee26fd9b 100644 --- a/pcuic/theories/PCUICRelevance.v +++ b/pcuic/theories/PCUICRelevance.v @@ -4,43 +4,50 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils. Require Import ssreflect. -Definition relevance_of_constant Σ kn := - match lookup_constant Σ kn with - | Some decl => decl.(cst_relevance) - | None => Relevant - end. - -Definition relevance_of_constructor Σ ind k := - match lookup_constructor Σ ind k with - | Some (_, idecl, _) => idecl.(ind_relevance) - | None => Relevant - end. - -Definition relevance_of_projection Σ (p: projection) := - match lookup_projection Σ p with - | Some (_, _, _, pdecl) => pdecl.(proj_relevance) - | None => Relevant - end. - Definition mark_context := list relevance. Definition marks_of_context (Γ: context) := List.map (binder_relevance ∘ decl_name) Γ. -Fixpoint relevance_of_term (Σ : global_env) (Γ : mark_context) (t: term) : relevance := - match t with - | tRel n => - option_default id (nth_error Γ n) Relevant - | tLambda na A t => relevance_of_term Σ (Γ ,, na.(binder_relevance)) t - | tLetIn na b B t => relevance_of_term Σ (Γ ,, na.(binder_relevance)) t - | tApp u _ => relevance_of_term Σ Γ u - | tConst k _ => relevance_of_constant Σ k - | tConstruct ind i _ => relevance_of_constructor Σ ind i - | tCase ci _ _ _ => ci.(ci_relevance) - | tProj p _ => relevance_of_projection Σ p - | tFix mfix n | tCoFix mfix n => option_default (binder_relevance ∘ dname) (nth_error mfix n) Relevant - | tVar _ | tEvar _ _ | tSort _ | tProd _ _ _ | tInd _ _ => Relevant - end. - -Notation isTermRel Σ Γ t rel := (relevance_of_term Σ Γ t = rel). +Inductive isTermRel (Σ : global_env) (Γ : mark_context) : term -> relevance -> Type := + | rel_Rel n rel : + nth_error Γ n = Some rel -> isTermRel Σ Γ (tRel n) rel + + | rel_Lambda na A t rel : + isTermRel Σ (Γ ,, na.(binder_relevance)) t rel -> isTermRel Σ Γ (tLambda na A t) rel + + | rel_LetIn na b B t rel : + isTermRel Σ (Γ ,, na.(binder_relevance)) t rel -> isTermRel Σ Γ (tLetIn na b B t) rel + + | rel_App t u rel : + isTermRel Σ Γ t rel -> isTermRel Σ Γ (tApp t u) rel + + | rel_Const kn u decl : + declared_constant Σ kn decl -> isTermRel Σ Γ (tConst kn u) decl.(cst_relevance) + + | rel_Construct ind i u mdecl idecl cdecl : + declared_constructor Σ (ind, i) mdecl idecl cdecl -> isTermRel Σ Γ (tConstruct ind i u) idecl.(ind_relevance) + + | rel_Case ci p c brs : + isTermRel Σ Γ (tCase ci p c brs) ci.(ci_relevance) + + | rel_Proj p u mdecl idecl cdecl pdecl : + declared_projection Σ p mdecl idecl cdecl pdecl -> isTermRel Σ Γ (tProj p u) pdecl.(proj_relevance) + + | rel_Fix mfix n def : + nth_error mfix n = Some def -> isTermRel Σ Γ (tFix mfix n) def.(dname).(binder_relevance) + + | rel_CoFix mfix n def : + nth_error mfix n = Some def -> isTermRel Σ Γ (tCoFix mfix n) def.(dname).(binder_relevance) + + | rel_Sort s : + isTermRel Σ Γ (tSort s) Relevant + + | rel_Prod na A B : + isTermRel Σ Γ (tProd na A B) Relevant + + | rel_Ind ind u : + isTermRel Σ Γ (tInd ind u) Relevant. + +Derive Signature for isTermRel. Lemma mark_of_context_app Γ Γ' Δ Δ' : marks_of_context Γ = marks_of_context Γ' -> marks_of_context Δ = marks_of_context Δ' -> marks_of_context (Γ,,, Δ) = marks_of_context (Γ',,, Δ'). diff --git a/pcuic/theories/PCUICRelevanceTerm.v b/pcuic/theories/PCUICRelevanceTerm.v index e18404f6f..a7e5bbb04 100644 --- a/pcuic/theories/PCUICRelevanceTerm.v +++ b/pcuic/theories/PCUICRelevanceTerm.v @@ -2,6 +2,7 @@ From MetaCoq.Template Require Import config utils. From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICRelevance PCUICInduction PCUICCases PCUICSigmaCalculus PCUICLiftSubst PCUICWeakeningEnv. Require Import ssreflect. +From Equations Require Import Equations. Lemma mark_inst_case_context params puinst (pctx : context) : @@ -26,15 +27,7 @@ Lemma extends_irrelevant {cf : checker_flags} {Pcmp P} Σ Σ' Γ t : isTermRel Σ' Γ t Irrelevant. Proof. induction t in Γ |- * using term_forall_list_ind; - cbn; intros ext Hirr; auto. - - unfold relevance_of_constant. - destruct lookup_constant eqn:H => //. - erewrite extends_lookup_constant; eauto. - - unfold relevance_of_constructor. - destruct lookup_constructor eqn:H => //. - erewrite extends_lookup_constructor; eauto. - - unfold relevance_of_projection. - destruct lookup_projection eqn:H => //. - erewrite extends_lookup_projection; eauto. + intros wfΣ' ext Hirr; depelim Hirr; try econstructor; eauto. + all: solve [ rewrite H; econstructor => //; eauto with extends ]. Qed. diff --git a/pcuic/theories/PCUICRelevanceTyp.v b/pcuic/theories/PCUICRelevanceTyp.v index 9d3701449..65b187cca 100644 --- a/pcuic/theories/PCUICRelevanceTyp.v +++ b/pcuic/theories/PCUICRelevanceTyp.v @@ -4,6 +4,7 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICRelevance PCUICTyping PCUICInversion PCUICConversion PCUICCumulProp PCUICWeakeningTyp PCUICValidity PCUICWellScopedCumulativity BDUnique. Require Import ssreflect. +From Equations Require Import Equations. Definition relevance_from_term_from_type {cf : checker_flags} (Σ : global_env_ext) Γ t T := forall rel, isTermRel Σ (marks_of_context Γ) t rel <~> isTypeRel Σ Γ T rel. @@ -179,50 +180,52 @@ Proof using Type. eapply All_local_env_over_2; tea. intros ??? H; apply lift_typing_impl with (1 := H); now intros ? []. - split; intro H. - + rewrite /relevance_of_term nth_error_map heq_nth_error /option_map /option_default /id in H. - rewrite -H. + + depelim H. rewrite nth_error_map heq_nth_error in e. inv e. destruct (PCUICTyping.nth_error_All_local_env heq_nth_error wfΓ); eauto. apply isTypeRelOpt_lift => //. pose proof (nth_error_Some_length heq_nth_error); lia. - + rewrite /relevance_of_term nth_error_map heq_nth_error /option_map /option_default /id. + + constructor. rewrite nth_error_map heq_nth_error /option_map. pose proof (PCUICTyping.nth_error_All_local_env heq_nth_error wfΓ).1; eauto. apply isTypeRelOpt_lift in X0 => //. 2: { pose proof (nth_error_Some_length heq_nth_error); lia. } - eapply isTypeRel2_relevance; tea. - - split; unfold relevance_of_term; intro Hr => //. - + eexists; split. + eapply f_equal, isTypeRel2_relevance; tea. + - split; intro Hr => //. + + inv Hr. + eexists; split. 2: { constructor; tea. eauto with pcuic. } - destruct u => //. + apply relevance_super. + destruct Hr as (s & <- & Hty) => //. apply inversion_Sort in Hty as (_ & _ & le) => //. eapply ws_cumul_pb_Sort_inv, leq_relevance in le => //. - destruct u => //. - - split; unfold relevance_of_term; intro Hr => //. - + rewrite -Hr. + rewrite le relevance_super. constructor. + - split; intro Hr. + + inv Hr. eexists; split. 2: { constructor; tea. eauto with pcuic. } - destruct s2, s1 => //. + apply relevance_super. + destruct Hr as (s & <- & Hty) => //. apply inversion_Sort in Hty as (_ & _ & le) => //. eapply ws_cumul_pb_Sort_inv, leq_relevance in le => //. - destruct s2, s1 => //. + rewrite le relevance_super. constructor. - split; intro Hr => //. - + eapply X3 in Hr as (s & e & Hs). + + inv Hr. + eapply X3 in X0 as (s & e & Hs). eexists; split. 2: { constructor; tea. } destruct s, s1 => //. + destruct Hr as (s & <- & Hs). - edestruct X3 as (_ & IH); eapply IH. + constructor. edestruct X3 as (_ & IH); eapply IH. apply inversion_Prod in Hs as (s1' & s2 & e1 & _ & Hty & le) => //. eexists; split; tea. eapply ws_cumul_pb_Sort_inv, leq_relevance in le => //. destruct s1', s2 => //. - split; intro Hr => //. - + eapply X5 in Hr as (s & e & Hs). + + inv Hr. + eapply X5 in X0 as (s & e & Hs). exists s; split => //. econstructor; tea. eapply type_LetIn; eauto. constructor; eauto with pcuic. do 2 constructor. apply cumul_zeta. + destruct Hr as (s & <- & Hs). - edestruct X5 as (_ & IH); eapply IH. + constructor. edestruct X5 as (_ & IH); eapply IH. assert (wf_universe Σ s) by eauto with pcuic. apply inversion_LetIn in Hs as (s1' & s2 & e1 & ? & ? & Hty & le) => //. eapply ws_cumul_pb_Sort_r_inv in le as (ss & red & le) => //. @@ -243,7 +246,7 @@ Proof using Type. eapply red_conv in c. eapply (weakening_ws_cumul_pb (Γ' := []) (Γ'' := [vdef _ _ _]) c). fvs. - split; intro Hr => //. - + eapply X3 in Hr as (s' & e & Hs). + + inv Hr. eapply X3 in X0 as (s' & e & Hs). assert (wf_universe Σ s') by auto with pcuic. apply inversion_Prod in Hs as (? & ss & ? & ? & ? & le); tea. exists s'; split => //. @@ -252,7 +255,7 @@ Proof using Type. change (tSort ss) with ((tSort ss) { 0 := u }). eapply PCUICSubstitution.substitution0; tea. + destruct Hr as (s' & <- & Hs). - edestruct X3 as (_ & IH); eapply IH; clear IH. + constructor. edestruct X3 as (_ & IH); eapply IH; clear IH. assert (wf_universe Σ s') by eauto with pcuic. exists s; split => //. apply inversion_Prod in IHB as (s1 & s2 & ? & ? & Hty & le) => //. @@ -262,8 +265,8 @@ Proof using Type. apply ws_cumul_pb_Sort_inv in le; destruct s2, s1, s => //. - apply PCUICWeakeningEnvTyp.on_declared_constant in H as H'; tea. unfold on_constant_decl in H'. split; intro Hr => //. - + unfold relevance_of_term, relevance_of_constant in Hr. - rewrite (declared_constant_lookup H) in Hr; rewrite -Hr. + + inv Hr. + assert (decl = decl0). { unfold declared_constant in H, H0. rewrite H in H0. now do 2 inversion H. } subst decl0. destruct H' as (H' & _). eapply infer_typing_sort_impl with _ H' => //; [apply relevance_subst_opt|]. intro Hty. apply (weaken_ctx (Γ := [])); tea. @@ -272,8 +275,7 @@ Proof using Type. eapply (PCUICUnivSubstitutionTyp.typing_subst_instance' _ _ [] _ (tSort H'.π1) u _); tea. unshelve epose proof (PCUICWeakeningEnv.weaken_lookup_on_global_env' _ _ (ConstantDecl decl) wf _); eauto. now split. - + unfold relevance_of_term, relevance_of_constant. - rewrite (declared_constant_lookup H). + + enough (cst_relevance decl = rel) by (subst rel; constructor => //). destruct H' as (H' & _). destruct Hr as (s1 & <- & Hs1), H' as (s2 & <- & Hs2). erewrite <- relevance_subst_eq. @@ -286,14 +288,14 @@ Proof using Type. - pose proof (onArity (PCUICWeakeningEnvTyp.on_declared_inductive isdecl).2). destruct Σ as (Σ & φ); unfold fst in *. split; intro Hr => //. - + unfold relevance_of_term in Hr; subst rel. + + inv Hr. eapply infer_typing_sort_impl with _ X1; [apply relevance_subst_opt|]; intros Hty. eapply PCUICUnivSubstitutionTyp.typing_subst_instance' in Hty. eapply (weaken_ctx (Γ := [])) in Hty. apply Hty. all: tea. unshelve epose proof (PCUICWeakeningEnv.weaken_lookup_on_global_env' _ _ (InductiveDecl mdecl) wf isdecl.p1); eauto. now split. - + unfold relevance_of_term. + + enough (Relevant = rel) by (subst rel; constructor => //). destruct Hr as (s1 & <- & Hs1), X1 as (s2 & <- & Hs2). eapply PCUICUnivSubstitutionTyp.typing_subst_instance' in Hs2. eapply (weaken_ctx (Γ := [])) in Hs2. @@ -304,45 +306,45 @@ Proof using Type. - pose proof (declared_constructor_lookup isdecl). destruct (PCUICWeakeningEnvTyp.on_declared_constructor isdecl) as ([[] []] & cu & ? & []); tea. split; intro Hr => //. - + unfold relevance_of_term, relevance_of_constructor in Hr; subst rel. + + inv Hr. + assert (idecl = idecl0). { clear -H0 isdecl. destruct H0 as ((H0m & H0i) & H0c), isdecl as ((Hm & Hi) & Hc). unfold declared_minductive in Hm, H0m. rewrite H0m in Hm. do 2 inversion Hm. subst mdecl0. rewrite H0i in Hi. now inversion Hi. } subst idecl0. destruct Σ as (Σ & φ); unfold fst, snd in *. - rewrite H. apply isType_of_constructor; tea. - + unfold relevance_of_term, relevance_of_constructor. destruct Σ as (Σ & φ); unfold fst, snd in *. - rewrite H. + + enough (idecl.(ind_relevance) = rel) by (subst rel; econstructor; apply isdecl). + destruct Σ as (Σ & φ); unfold fst, snd in *. eassert (isTypeRel (Σ, φ) _ (type_of_constructor mdecl cdecl (ind, i) u) (idecl.(ind_relevance))) by (apply isType_of_constructor; tea). eapply isTypeRel2_relevance; tea. apply wf. - assert (Σ ;;; Γ |- mkApps ptm (indices ++ [c]) : tSort ps). { apply apply_predctx => //. apply ctx_inst_impl with (1 := X5) => ??[] //. } split; intro Hr => //. - + unfold relevance_of_term in Hr; subst rel. + + inv Hr. exists ps; split => //. - + unfold relevance_of_term. + + enough (rel = ci.(ci_relevance)) by (subst rel; constructor). eapply isTypeRel2_relevance; tea. exists ps; split => //. - pose proof (declared_projection_lookup isdecl). assert (isTypeRel Σ Γ (subst0 (c :: List.rev args) (proj_type pdecl)@[u]) (pdecl.(proj_relevance))) by (eapply isType_of_projection; tea). split; intro Hr => //. - + unfold relevance_of_term, relevance_of_projection in Hr; subst rel. - now rewrite H. - + unfold relevance_of_term, relevance_of_projection. - rewrite H. + + inv Hr. + assert (pdecl = pdecl0). { clear -H0 isdecl. destruct H0 as (((H0m & H0i) & H0c) & H0p & _), isdecl as (((Hm & Hi) & Hc) & Hp & _). unfold declared_minductive in Hm, H0m. rewrite H0m in Hm. do 2 inversion Hm. subst mdecl0. rewrite H0i in Hi. inversion Hi. subst idecl0. rewrite H0c in Hc. inversion Hc. subst cdecl0. rewrite H0p in Hp. now inversion Hp. } subst pdecl0. + assumption. + + enough (pdecl.(proj_relevance) = rel) by (subst rel; econstructor; apply isdecl). eapply isTypeRel2_relevance; tea. - eapply All_nth_error in X0, X1; tea. rewrite /on_def_type /on_def_body /lift_typing2 in X0, X1. split; intro Hr => //. - + rewrite /relevance_of_term heq_nth_error /option_default in Hr; subst rel. + + inv Hr. rewrite H0 in heq_nth_error; inversion heq_nth_error; subst def. apply lift_typing_impl with (1 := X0); now intros ? []. - + rewrite /relevance_of_term heq_nth_error /option_default. + + eenough (rel = _) by (erewrite H0; constructor; tea). eapply isTypeRel2_relevance; tea. apply infer_typing_sort_impl with id X0 => //. now intros []. - eapply All_nth_error in X0, X1; tea. rewrite /on_def_type /on_def_body /lift_typing2 in X0, X1. split; intro Hr => //. - + rewrite /relevance_of_term heq_nth_error /option_default in Hr; subst rel. + + inv Hr. rewrite H0 in heq_nth_error; inversion heq_nth_error; subst def. apply lift_typing_impl with (1 := X0); now intros ? []. - + rewrite /relevance_of_term heq_nth_error /option_default. + + eenough (rel = _) by (erewrite H0; constructor; tea). eapply isTypeRel2_relevance; tea. apply infer_typing_sort_impl with id X0 => //. now intros []. diff --git a/safechecker/theories/PCUICSafeChecker.v b/safechecker/theories/PCUICSafeChecker.v index 3b03beb44..2059ba736 100644 --- a/safechecker/theories/PCUICSafeChecker.v +++ b/safechecker/theories/PCUICSafeChecker.v @@ -503,7 +503,7 @@ Section CheckEnv. | {| decl_name := na; decl_body := Some b; decl_type := ty |} :: Δ => checkΔ <- check_type_local_ctx X_ext Γ Δ s wfΓ ;; checkty <- check_type_wf_env X_ext (Γ ,,, Δ) _ b ty ;; - let '(tmrel; eqtmrel) := relevance_of_term_abstract_env X_ext_impl X_ext (marks_of_context (Γ ,,, Δ)) b in + let '(tmrel; eqtmrel) := relevance_of_term_abstract_env X_ext_impl X_ext (Γ ,,, Δ) b _ in check <- check_eq_true (tmrel == na.(binder_relevance)) (Msg "Wrong relevance") ;; ret _ end. @@ -522,6 +522,9 @@ Section CheckEnv. Next Obligation. specialize_Σ H. sq. now eapply PCUICContexts.type_local_ctx_wf_local in checkΔ. Qed. + Next Obligation. + specialize_Σ wfΣ; sq; now eexists. + Qed. Next Obligation. clear Heq_anonymous. pose proof (abstract_env_ext_wf _ H). specialize_Σ H. sq. @@ -544,7 +547,7 @@ Section CheckEnv. | {| decl_name := na; decl_body := Some b; decl_type := ty |} :: Δ => '(Δs; Δinfer) <- infer_sorts_local_ctx X_ext Γ Δ wfΓ ;; checkty <- check_type_wf_env X_ext (Γ ,,, Δ) _ b ty ;; - let '(tmrel; eqtmrel) := relevance_of_term_abstract_env X_ext_impl X_ext (marks_of_context (Γ ,,, Δ)) b in + let '(tmrel; eqtmrel) := relevance_of_term_abstract_env X_ext_impl X_ext (Γ ,,, Δ) b _ in check <- check_eq_true (tmrel == na.(binder_relevance)) (Msg "Wrong relevance") ;; ret (Δs; _) end. @@ -559,7 +562,10 @@ Section CheckEnv. specialize_Σ H. sq. now eapply PCUICContexts.sorts_local_ctx_wf_local in Δinfer. Qed. Next Obligation. - clear Heq_anonymous. + specialize_Σ wfΣ; sq; now eexists. + Qed. + Next Obligation. + clear Heq_anonymous Heq_x. pose proof (abstract_env_ext_wf _ H). specialize_Σ H. sq. splits; auto. eapply PCUICRelevanceTyp.relevance_from_type in checkty. 2: apply H0. apply checkty in eqtmrel; clear checkty. @@ -1412,7 +1418,7 @@ Section CheckEnv. | {| decl_name := na; decl_body := Some b; decl_type := ty |} :: Γ => wfΓ <- check_wf_local X_ext Γ ;; wfty <- check_type_wf_env X_ext Γ wfΓ b ty ;; - let '(tmrel; eqtmrel) := relevance_of_term_abstract_env X_ext_impl X_ext (marks_of_context Γ) b in + let '(tmrel; eqtmrel) := relevance_of_term_abstract_env X_ext_impl X_ext Γ b _ in check <- check_eq_true (tmrel == na.(binder_relevance)) (Msg "Wrong relevance") ;; ret _ | {| decl_name := na; decl_body := None; decl_type := ty |} :: Γ => @@ -1421,6 +1427,9 @@ Section CheckEnv. check <- check_eq_true (relevance_of_sort s == na.(binder_relevance)) (Msg "Wrong relevance") ;; ret _ end. + Next Obligation. + specialize_Σ wfΣ; sq; now eexists. + Qed. Next Obligation. clear Heq_anonymous. pose proof (abstract_env_ext_wf _ H); specialize_Σ H. @@ -2346,9 +2355,10 @@ End monad_Alli_nth_forall. | ConstantDecl cst => match cst.(cst_body) with | Some term => - let '(tmrel; eqtmrel) := relevance_of_term_abstract_env X_ext_impl X_ext [] term in + check_wf_judgement kn X_ext term cst.(cst_type) ;; + let '(tmrel; eqtmrel) := relevance_of_term_abstract_env X_ext_impl X_ext [] term _ in check <- wrap_error _ X_ext id (check_eq_true (tmrel == cst.(cst_relevance)) (Msg "Wrong relevance")) ;; - check_wf_judgement kn X_ext term cst.(cst_type) ;; ret _ + ret _ | None => s <- check_wf_type_rel kn X_ext cst.(cst_type) cst.(cst_relevance);; ret _ @@ -2363,6 +2373,9 @@ End monad_Alli_nth_forall. check_bodies <- monad_Alli_nth_forall mdecl.(ind_bodies) (fun i oib Hoib => check_one_ind_body X X_ext kn mdecl _ check_pars onarities check_var i oib Hoib);; ret (Build_on_inductive_sq check_bodies check_pars check_npars _) end. + Next Obligation. + specialize_Σ wfΣ; sq; now eexists. + Qed. Next Obligation. clear Heq_anonymous. specialize_Σ H0. pose proof (hΣ _ _ H0). sq. unfold on_constant_decl; rewrite <- Heq_anonymous0. diff --git a/safechecker/theories/PCUICTypeChecker.v b/safechecker/theories/PCUICTypeChecker.v index 20392df90..be4b1da55 100644 --- a/safechecker/theories/PCUICTypeChecker.v +++ b/safechecker/theories/PCUICTypeChecker.v @@ -1008,145 +1008,295 @@ Section Typecheck. congruence. Qed. - Definition abstract_lookup_constant kn := - match abstract_env_lookup X kn with - | Some (ConstantDecl d) => Some d - | _ => None - end. - - Definition abstract_lookup_minductive mind := - match abstract_env_lookup X mind with - | Some (InductiveDecl decl) => Some decl - | _ => None - end. - - Definition abstract_lookup_inductive ind := - match abstract_lookup_minductive (inductive_mind ind) with - | Some mdecl => - match nth_error mdecl.(ind_bodies) (inductive_ind ind) with - | Some idecl => Some (mdecl, idecl) - | None => None - end - | None => None - end. + Equations abstract_lookup_constant kn : { decl & forall Σ (wfΣ : abstract_env_ext_rel X Σ), lookup_constant Σ kn = decl } := + abstract_lookup_constant kn with inspect (abstract_env_lookup X kn) := { + | @exist (Some (ConstantDecl d)) e => (Some d; _) ; + | @exist _ e => (None; _) + }. + Next Obligation. + assert (lookup_env Σ kn = abstract_env_lookup X kn) by (eapply abstract_env_lookup_correct; eauto). + rewrite /lookup_constant H -e //. + Qed. + Next Obligation. + assert (lookup_env Σ kn = abstract_env_lookup X kn) by (eapply abstract_env_lookup_correct; eauto). + rewrite /lookup_constant H -e //. + Qed. + Next Obligation. + assert (lookup_env Σ kn = abstract_env_lookup X kn) by (eapply abstract_env_lookup_correct; eauto). + rewrite /lookup_constant H -e //. + Qed. + + + Equations abstract_lookup_minductive mind : { mdecl & forall Σ (wfΣ : abstract_env_ext_rel X Σ), lookup_minductive Σ mind = mdecl } := + abstract_lookup_minductive mind with inspect (abstract_env_lookup X mind) := { + | @exist (Some (InductiveDecl mdecl)) e => (Some mdecl; _) ; + | @exist _ e => (None; _) + }. + Next Obligation. + assert (lookup_env Σ mind = abstract_env_lookup X mind) by (eapply abstract_env_lookup_correct; eauto). + rewrite /lookup_minductive H -e //. + Qed. + Next Obligation. + assert (lookup_env Σ mind = abstract_env_lookup X mind) by (eapply abstract_env_lookup_correct; eauto). + rewrite /lookup_minductive H -e //. + Qed. + Next Obligation. + assert (lookup_env Σ mind = abstract_env_lookup X mind) by (eapply abstract_env_lookup_correct; eauto). + rewrite /lookup_minductive H -e //. + Qed. - Definition abstract_lookup_constructor ind k := - match abstract_lookup_inductive ind with - | Some (mdecl, idecl) => - match nth_error idecl.(ind_ctors) k with - | Some cdecl => Some (mdecl, idecl, cdecl) - | None => None - end - | _ => None - end. + Equations abstract_lookup_inductive ind : { midecl & forall Σ (wfΣ : abstract_env_ext_rel X Σ), lookup_inductive Σ ind = midecl } := + abstract_lookup_inductive ind with abstract_lookup_minductive ind.(inductive_mind) := { + | (None; _) => (None; _) ; + | (Some mdecl; _) with inspect (nth_error mdecl.(ind_bodies) ind.(inductive_ind)) := { + | @exist None ee => (None; _) ; + | @exist (Some idecl) ee => (Some (mdecl, idecl); _) ; + } + }. + Next Obligation. + specialize_Σ wfΣ. + rewrite /lookup_inductive e -ee //. + Qed. + Next Obligation. + specialize_Σ wfΣ. + rewrite /lookup_inductive e -ee //. + Qed. + Next Obligation. + specialize_Σ wfΣ. + rewrite /lookup_inductive e //. + Qed. - Definition abstract_lookup_projection p := - match abstract_lookup_constructor p.(proj_ind) 0 with - | Some (mdecl, idecl, cdecl) => - match nth_error idecl.(ind_projs) p.(proj_arg) with - | Some pdecl => Some (mdecl, idecl, cdecl, pdecl) - | None => None - end - | _ => None - end. + Equations abstract_lookup_constructor ind k : { micdecl & forall Σ (wfΣ : abstract_env_ext_rel X Σ), lookup_constructor Σ ind k = micdecl } := + abstract_lookup_constructor ind k with abstract_lookup_inductive ind := { + | (None; _) => (None; _) ; + | (Some (mdecl, idecl); _) with inspect (nth_error idecl.(ind_ctors) k) := { + | @exist None ee => (None; _) ; + | @exist (Some cdecl) ee => (Some (mdecl, idecl, cdecl); _) ; + } + }. + Next Obligation. + specialize_Σ wfΣ. + rewrite /lookup_constructor e -ee //. + Qed. + Next Obligation. + specialize_Σ wfΣ. + rewrite /lookup_constructor e -ee //. + Qed. + Next Obligation. + specialize_Σ wfΣ. + rewrite /lookup_constructor e //. + Qed. - Equations relevance_of_term_abstract_env (Γ : mark_context) (t : term) : { rel & forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ isTermRel Σ Γ t rel ∥ } by struct t := + Equations abstract_lookup_projection proj : { micpdecl & forall Σ (wfΣ : abstract_env_ext_rel X Σ), lookup_projection Σ proj = micpdecl } := + abstract_lookup_projection p with abstract_lookup_constructor p.(proj_ind) 0 := { + | (None; _) => (None; _) ; + | (Some (mdecl, idecl, cdecl); _) with inspect (nth_error idecl.(ind_projs) p.(proj_arg)) := { + | @exist None ee => (None; _) ; + | @exist (Some pdecl) ee => (Some (mdecl, idecl, cdecl, pdecl); _) ; + } + }. + Next Obligation. + specialize_Σ wfΣ. + rewrite /lookup_projection e -ee //. + Qed. + Next Obligation. + specialize_Σ wfΣ. + rewrite /lookup_projection e -ee //. + Qed. + Next Obligation. + specialize_Σ wfΣ. + rewrite /lookup_projection e //. + Qed. - relevance_of_term_abstract_env Γ (tRel n) + Equations relevance_of_term_abstract_env (Γ : context) (t : term) (H : forall Σ (wfΣ : abstract_env_ext_rel X Σ), welltyped Σ Γ t ) : + { rel & forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ isTermRel Σ (marks_of_context Γ) t rel ∥ } by struct t := + + relevance_of_term_abstract_env Γ (tRel n) H with inspect (nth_error Γ n) := { - | exist (Some c) e => (c; _) ; - | exist None e => (Relevant; _) + | exist (Some c) e => (c.(decl_name).(binder_relevance); _) ; + | exist None e => ! } ; - relevance_of_term_abstract_env Γ (tLambda na A t) := - relevance_of_term_abstract_env (Γ ,, na.(binder_relevance)) t ; + relevance_of_term_abstract_env Γ (tLambda na A t) H := + let '(rel; _) := relevance_of_term_abstract_env (Γ ,, vass na A) t _ in + (rel; _) ; - relevance_of_term_abstract_env Γ (tLetIn na b b_ty b') := - relevance_of_term_abstract_env (Γ ,, na.(binder_relevance)) b' ; + relevance_of_term_abstract_env Γ (tLetIn na b b_ty b') H := + let '(rel; _) := relevance_of_term_abstract_env (Γ ,, vdef na b b_ty) b' _ in + (rel; _) ; - relevance_of_term_abstract_env Γ (tApp t u) := - relevance_of_term_abstract_env Γ t ; + relevance_of_term_abstract_env Γ (tApp t u) H := + let '(rel; _) := relevance_of_term_abstract_env Γ t _ in + (rel; _) ; - relevance_of_term_abstract_env Γ (tConst cst u) - with inspect (abstract_lookup_constant cst) := { - | exist (Some d) e => (d.(cst_relevance); _) ; - | exist None e => (Relevant; _) ; + relevance_of_term_abstract_env Γ (tConst cst u) H + with (abstract_lookup_constant cst) := { + | (Some d; _) => (d.(cst_relevance); _) ; + | (None; _) => ! ; } ; - relevance_of_term_abstract_env Γ (tConstruct ind k u) - with inspect (abstract_lookup_constructor ind k) := { - | exist (Some (_, idecl, _)) e := (idecl.(ind_relevance); _) ; - | exist None _ := (Relevant; _) ; + relevance_of_term_abstract_env Γ (tConstruct ind k u) H + with abstract_lookup_constructor ind k := { + | (Some (_, idecl, _); _) := (idecl.(ind_relevance); _) ; + | (None; _) := ! ; }; - relevance_of_term_abstract_env Γ (tCase ci p c brs) := + relevance_of_term_abstract_env Γ (tCase ci p c brs) H := (ci.(ci_relevance); _) ; - relevance_of_term_abstract_env Γ (tProj p c) - with inspect (abstract_lookup_projection p) := { - | exist (Some (_, _, _, pdecl)) e := (pdecl.(proj_relevance); _) ; - | exist None _ := (Relevant; _) ; + relevance_of_term_abstract_env Γ (tProj p c) H + with (abstract_lookup_projection p) := { + | (Some (_, _, _, pdecl); _) := (pdecl.(proj_relevance); _) ; + | (None; _) := ! ; }; - relevance_of_term_abstract_env Γ (tFix mfix n) + relevance_of_term_abstract_env Γ (tFix mfix n) H with inspect (nth_error mfix n) := { - | exist None _ := (Relevant; _) ; | exist (Some decl) e := (decl.(dname).(binder_relevance); _) ; + | exist None _ := ! ; } ; - relevance_of_term_abstract_env Γ (tCoFix mfix n) + relevance_of_term_abstract_env Γ (tCoFix mfix n) H with inspect (nth_error mfix n) := { - | exist None _ := (Relevant; _) ; | exist (Some decl) e := (decl.(dname).(binder_relevance); _) ; + | exist None _ := ! ; } ; - relevance_of_term_abstract_env Γ (tVar _) := (Relevant; _) ; - relevance_of_term_abstract_env Γ (tEvar ev _) := (Relevant; _) ; - relevance_of_term_abstract_env Γ (tSort u) := (Relevant; _) ; - relevance_of_term_abstract_env Γ (tProd na A B) := (Relevant; _) ; - relevance_of_term_abstract_env Γ (tInd ind u) := (Relevant; _). + relevance_of_term_abstract_env Γ (tVar _) H := ! ; + relevance_of_term_abstract_env Γ (tEvar ev _) H := ! ; + relevance_of_term_abstract_env Γ (tSort u) H := (Relevant; _) ; + relevance_of_term_abstract_env Γ (tProd na A B) H := (Relevant; _) ; + relevance_of_term_abstract_env Γ (tInd ind u) H := (Relevant; _). Next Obligation. - unfold relevance_of_term, relevance_of_constant, relevance_of_constructor, relevance_of_projection; - unfold lookup_constant, lookup_projection, lookup_constructor, lookup_inductive, lookup_minductive; - unfold abstract_lookup_constant, abstract_lookup_projection, abstract_lookup_constructor, abstract_lookup_inductive, abstract_lookup_minductive in e; - erewrite <- abstract_env_lookup_correct in e; eauto; - now rewrite <- e. + sq. constructor => //. rewrite nth_error_map -e //. + Qed. + Next Obligation. + destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. + pose proof hΣ. + specialize_Σ wfΣ. destruct H. sq. + apply inversion_Rel in X0 as (? & _ & ? & _); tea. + congruence. + Qed. + Next Obligation. + destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. + pose proof hΣ. + specialize_Σ wfΣ. destruct H. sq. + apply inversion_Var in X0 => //. + Qed. + Next Obligation. + destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. + pose proof hΣ. + specialize_Σ wfΣ. destruct H. sq. + apply inversion_Evar in X0 => //. + Qed. + Next Obligation. + sq. constructor => //. Qed. Next Obligation. - unfold relevance_of_term, relevance_of_constant, relevance_of_constructor, relevance_of_projection. - unfold lookup_constant, lookup_projection, lookup_constructor, lookup_inductive, lookup_minductive. - unfold abstract_lookup_constant, abstract_lookup_projection, abstract_lookup_constructor, abstract_lookup_inductive, abstract_lookup_minductive in e. - erewrite <- abstract_env_lookup_correct in e; eauto. - now rewrite <- e. + sq. constructor => //. Qed. Next Obligation. - unfold relevance_of_term, relevance_of_constant, relevance_of_constructor, relevance_of_projection. - unfold lookup_constant, lookup_projection, lookup_constructor, lookup_inductive, lookup_minductive. - unfold abstract_lookup_constant, abstract_lookup_projection, abstract_lookup_constructor, abstract_lookup_inductive, abstract_lookup_minductive in e. - erewrite <- abstract_env_lookup_correct in e; eauto. - now rewrite <- e. + pose proof hΣ. + specialize_Σ wfΣ. destruct H. sq. + apply inversion_Lambda in X0 as (_ & B & _ & _ & ? & _) => //. + now eexists. Qed. Next Obligation. - unfold relevance_of_term, relevance_of_constant, relevance_of_constructor, relevance_of_projection. - unfold lookup_constant, lookup_projection, lookup_constructor, lookup_inductive, lookup_minductive. - unfold abstract_lookup_constant, abstract_lookup_projection, abstract_lookup_constructor, abstract_lookup_inductive, abstract_lookup_minductive in e. - erewrite <- abstract_env_lookup_correct in e; eauto. - now rewrite <- e. + specialize_Σ wfΣ. sq. + constructor => //. Qed. Next Obligation. - unfold relevance_of_term, relevance_of_constant, relevance_of_constructor, relevance_of_projection. - unfold lookup_constant, lookup_projection, lookup_constructor, lookup_inductive, lookup_minductive. - unfold abstract_lookup_constant, abstract_lookup_projection, abstract_lookup_constructor, abstract_lookup_inductive, abstract_lookup_minductive in e. - erewrite <- abstract_env_lookup_correct in e; eauto. - now rewrite <- e. + pose proof hΣ. + specialize_Σ wfΣ. destruct H. sq. + apply inversion_LetIn in X0 as (_ & B & _ & _ & _ & ? & _) => //. + now eexists. Qed. Next Obligation. - unfold relevance_of_term, relevance_of_constant, relevance_of_constructor, relevance_of_projection. - unfold lookup_constant, lookup_projection, lookup_constructor, lookup_inductive, lookup_minductive. - unfold abstract_lookup_constant, abstract_lookup_projection, abstract_lookup_constructor, abstract_lookup_inductive, abstract_lookup_minductive in e. - erewrite <- abstract_env_lookup_correct in e; eauto. - now rewrite <- e. + specialize_Σ wfΣ. sq. + constructor => //. + Qed. + Next Obligation. + pose proof hΣ. + specialize_Σ wfΣ. destruct H. sq. + apply inversion_App in X0 as (? & ? & ? & ? & _) => //. + now eexists. + Qed. + Next Obligation. + specialize_Σ wfΣ. sq. + constructor => //. + Qed. + Next Obligation. + pose proof hΣ. + specialize_Σ wfΣ. destruct H. sq. + constructor => //. + apply lookup_constant_declared => //. + Qed. + Next Obligation. + destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. + pose proof hΣ. + specialize_Σ wfΣ. destruct H. sq. + apply inversion_Const in X0 as (? & _ & H & _); tea. + apply declared_constant_lookup in H. congruence. + Qed. + Next Obligation. + specialize_Σ wfΣ. sq. + constructor => //. + Qed. + Next Obligation. + pose proof hΣ. + specialize_Σ wfΣ. destruct H. sq. + econstructor => //. + apply lookup_constructor_declared => //; tea. + Qed. + Next Obligation. + destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. + pose proof hΣ. + specialize_Σ wfΣ. destruct H. sq. + apply inversion_Construct in X0 as (m & i & c & _ & H & _); tea. + apply declared_constructor_lookup in H. rewrite e in H => //. + Qed. + Next Obligation. + specialize_Σ wfΣ. sq. + constructor => //. + Qed. + Next Obligation. + pose proof hΣ. + specialize_Σ wfΣ. destruct H. sq. + econstructor => //. + apply lookup_projection_declared => //; tea. + apply inversion_Proj in X0 as (_ & m' & i' & c' & p' & _ & H & _); tea. + apply declared_projection_lookup in H as ee. + rewrite e in ee; inversion ee. + destruct H as (_ & _ & ?). + congruence. + Qed. + Next Obligation. + destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. + pose proof hΣ. + specialize_Σ wfΣ. destruct H. sq. + apply inversion_Proj in X0 as (_ & m' & i' & c' & p' & _ & H & _); tea. + apply declared_projection_lookup in H. rewrite e in H => //. + Qed. + Next Obligation. + specialize_Σ wfΣ. sq. + constructor => //. + Qed. + Next Obligation. + destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. + pose proof hΣ. + specialize_Σ wfΣ. destruct H. sq. + apply inversion_Fix in X0 as (decl & _ & H & _); tea. congruence. + Qed. + Next Obligation. + specialize_Σ wfΣ. sq. + constructor => //. + Qed. + Next Obligation. + destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. + pose proof hΣ. + specialize_Σ wfΣ. destruct H. sq. + apply inversion_CoFix in X0 as (decl & _ & H & _); tea. congruence. Qed. Definition abstract_env_level_mem_forallb {Σ} (wfΣ : abstract_env_ext_rel X Σ) u : diff --git a/template-coq/theories/BasicAst.v b/template-coq/theories/BasicAst.v index 9c8717cfe..275432f95 100644 --- a/template-coq/theories/BasicAst.v +++ b/template-coq/theories/BasicAst.v @@ -133,6 +133,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) := diff --git a/template-coq/theories/Universes.v b/template-coq/theories/Universes.v index 2073ea719..8027a07a2 100644 --- a/template-coq/theories/Universes.v +++ b/template-coq/theories/Universes.v @@ -2312,6 +2312,11 @@ Section no_prop_leq_type. now rewrite ps in H. Qed. + Lemma relevance_super s : relevance_of_sort (Universe.super s) = Relevant. + Proof using Type. + now destruct s. + Qed. + Lemma is_prop_gt s1 s2 : leq_universe ϕ (Universe.super s1) s2 -> Universe.is_prop s2 -> False. Proof using Type. From dfaf09497a24047473e2381b698827b43eef71d2 Mon Sep 17 00:00:00 2001 From: Yann Leray Date: Tue, 24 May 2022 17:45:23 +0200 Subject: [PATCH 11/17] Plugin fix --- safechecker/_PluginProject.in | 2 -- safechecker/src/metacoq_safechecker_plugin.mlpack | 1 - 2 files changed, 3 deletions(-) diff --git a/safechecker/_PluginProject.in b/safechecker/_PluginProject.in index 96bc42b84..788756c48 100644 --- a/safechecker/_PluginProject.in +++ b/safechecker/_PluginProject.in @@ -72,8 +72,6 @@ src/pCUICPretty.mli src/pCUICPretty.ml src/pCUICProgram.mli src/pCUICProgram.ml -src/pCUICRelevance.ml -src/pCUICRelevance.mli # From SafeChecker # src/string2pos.mli diff --git a/safechecker/src/metacoq_safechecker_plugin.mlpack b/safechecker/src/metacoq_safechecker_plugin.mlpack index 167205fcd..1c0e60a37 100644 --- a/safechecker/src/metacoq_safechecker_plugin.mlpack +++ b/safechecker/src/metacoq_safechecker_plugin.mlpack @@ -19,7 +19,6 @@ PCUICTyping PCUICWfUniverses PCUICNormal PCUICPosition -PCUICRelevance PCUICPretty PCUICProgram TemplateToPCUIC From 4c382919dc14800dd33e0ba4f292e7df5a81f812 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Wed, 25 May 2022 01:35:50 +0200 Subject: [PATCH 12/17] Fill the auxilliary lemmas about the relevance of constructors/inductives --- pcuic/theories/PCUICRelevanceTyp.v | 187 +++++++++++++++++++++++++++-- 1 file changed, 177 insertions(+), 10 deletions(-) diff --git a/pcuic/theories/PCUICRelevanceTyp.v b/pcuic/theories/PCUICRelevanceTyp.v index 65b187cca..952da1408 100644 --- a/pcuic/theories/PCUICRelevanceTyp.v +++ b/pcuic/theories/PCUICRelevanceTyp.v @@ -76,6 +76,45 @@ Proof. constructor. apply leq_universe_product. 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 := Relevant |}; + 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. + +Lemma inversion_Rel_wt {cf : checker_flags} : +forall {Σ : global_env_ext} {Γ n T} {wfΣ : wf Σ}, + Σ ;;; Γ |- tRel n : T -> + ∑ decl, + [× wf_local Σ Γ, + nth_error Γ n = Some decl, + isTypeRel Σ Γ (lift0 (S n) (decl_type decl)) decl.(decl_name).(binder_relevance) & + Σ ;;; Γ ⊢ lift0 (S n) (decl_type decl) ≤ T]. +Proof. + intros Σ Γ n T wfΣ h. depind h. + - exists decl. + assert (isTypeRel Σ Γ (lift0 (S n) (decl_type decl)) (binder_relevance (decl_name decl))). + { destruct (PCUICTyping.nth_error_All_local_env e a); eauto. cbn in o. + apply isTypeRelOpt_lift in l => //. 1:{ pose proof (nth_error_Some_length e). lia. } } + split => //. + apply isType_ws_cumul_pb_refl. destruct X as [s [_ Hs]]. now exists s. + - destruct (IHh1 wfΣ) as [decl []]. exists decl; split => //. + etransitivity; tea. eapply cumulSpec_typed_cumulAlgo; tea => //. +Qed. + Lemma isTypeRel_cstr_type {cf : checker_flags} Σ mdecl idecl cdecl n : wf Σ.1 -> nth_error (ind_bodies mdecl) n = Some idecl -> @@ -96,11 +135,19 @@ Lemma isTypeRel_cstr_type {cf : checker_flags} Σ mdecl idecl cdecl n : isTypeRel Σ (arities_context mdecl.(ind_bodies)) (cstr_type cdecl) idecl.(ind_relevance). Proof. intros wfΣ H -> e1 <- c (s & _ & Hs); exists s; split => //=. - (* epose proof (nth_error_arities_context mdecl (#|ind_bodies mdecl| - S cstr.1.(inductive_ind)) idecl _). - Unshelve. 2: { rewrite nth_error_rev. len. apply nth_error_Some_length in di. lia. rewrite List.rev_involutive. replace _ with (inductive_ind cstr.1); tea. len. apply nth_error_Some_length in di. lia. } - rewrite e1 in H; clear e1. *) - admit. -Admitted. + eapply PCUICSpine.inversion_it_mkProd_or_LetIn in Hs. + eapply inversion_mkApps in Hs as [tyapp [Happ Hsp]]. + eapply inversion_Rel_wt in Happ as [cdecl' [wfl hnth ht cum]] => //. + 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. + destruct ht as [s' [_ hs]]. now exists s'. +Qed. Lemma isType_of_constructor {cf : checker_flags} Σ mdecl idecl cdecl cstr u Γ : wf Σ.1 -> wf_local Σ Γ -> declared_constructor Σ.1 cstr mdecl idecl cdecl -> @@ -158,7 +205,115 @@ Proof. admit. Admitted. -Lemma apply_predctx {cf : checker_flags} Σ Γ ci p indices c ps mdecl idecl {wfΣ : wf Σ.1} : + +Section OnInductives. + Context {cf : checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ}. + + Context {mdecl ind idecl} + (decli : declared_inductive Σ ind mdecl idecl). + +Lemma on_minductive_wf_params_indices_inst_weaken {Γ} (u : Instance.t) : +consistent_instance_ext Σ (ind_universes mdecl) u -> +wf_local Σ Γ -> +wf_local Σ (Γ ,,, (ind_params mdecl ,,, ind_indices idecl)@[u]). +Proof. +intros. eapply weaken_wf_local; tea. +eapply PCUICInductives.on_minductive_wf_params_indices_inst; tea. +Qed. +End OnInductives. + +Lemma spine_subst_alpha {cf : checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} Γ args inst Δ Δ' : + PCUICSpine.spine_subst Σ Γ args inst Δ -> + All2 (PCUICEquality.compare_decls eq eq) Δ Δ' -> + PCUICSpine.spine_subst Σ Γ args inst Δ'. +Proof. + intros [] eqd. + split => //. + eapply PCUICSpine.wf_local_alpha; tea. eapply All2_app; trea. + 2:{ eapply PCUICSpine.subslet_eq_context_alpha; tea. } + move: inst_ctx_subst. clear -eqd. + induction 1 in Δ', eqd |- *; depelim eqd; try constructor; depelim c; subst. + - constructor. now apply IHinst_ctx_subst. + - constructor. eauto. +Qed. + +Import PCUICAlpha. + +Lemma arity_spine_alpha {cf : checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} Γ args T T' T'' : + isType Σ Γ T'' -> + PCUICSpine.arity_spine Σ Γ T args T'' -> + T ≡α T' -> + PCUICSpine.arity_spine Σ Γ T' args T''. +Proof. + intros isty sp. revert T'. induction sp; intros. + - intros. constructor. auto. + pose proof (isType_alpha _ _ _ isty X). + eapply PCUICContextConversion.ws_cumul_pb_eq_le. + destruct isty as [? []]. + constructor; fvs. red. apply PCUICEquality.upto_names_impl_eq_term. now symmetry. + - intros. constructor => //. + transitivity ty => //. + eapply PCUICContextConversion.ws_cumul_pb_eq_le. + constructor; fvs. + eapply PCUICConfluence.eq_term_upto_univ_napp_on_free_vars; tea. fvs. + apply PCUICEquality.upto_names_impl_eq_term. now symmetry. + - depelim X. + constructor. eapply IHsp => //. + eapply PCUICEquality.eq_term_upto_univ_subst; tc; auto. + - depelim X. + constructor. + pose proof (validity t). + eapply isType_alpha in X; tea. destruct X as [s [_ Hs]]. + econstructor; tea. eapply convSpec_cumulSpec. now apply eq_term_upto_univ_cumulSpec, PCUICEquality.upto_names_impl_eq_term. + eapply IHsp => //. eapply PCUICEquality.eq_term_upto_univ_subst; tc; auto. reflexivity. +Qed. + +Lemma eq_term_upto_univ_it_mkProd_or_LetIn (Γ Γ' : context) T : + All2 (PCUICEquality.compare_decls eq eq) Γ Γ' -> + it_mkProd_or_LetIn Γ T ≡α it_mkProd_or_LetIn Γ' T. +Proof. + induction Γ in Γ' |- * using PCUICInduction.ctx_length_rev_ind. + - intros a; depelim a. reflexivity. + - intros a. + eapply All2_app_inv_l in a as [? [? [? []]]]. subst. + depelim a0. depelim a0. + rewrite !it_mkProd_or_LetIn_app. + depelim c; subst; constructor; cbn; auto; try reflexivity. + apply X => //. + apply X => //. +Qed. + +Lemma arity_spine_mkProd_alpha {cf : checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} {Γ args} {Δ Δ' : context} {T T'} : + isType Σ Γ T' -> + PCUICSpine.arity_spine Σ Γ (it_mkProd_or_LetIn Δ T) args T' -> + All2 (PCUICEquality.compare_decls eq eq) Δ Δ' -> + PCUICSpine.arity_spine Σ Γ (it_mkProd_or_LetIn Δ' T) args T'. +Proof. + intros isty sp eq. + eapply arity_spine_alpha => //. exact sp. + now apply eq_term_upto_univ_it_mkProd_or_LetIn. +Qed. + +Lemma eq_annots_expand_lets_ctx (Γ : list aname) (Δ Δ' : context) : + PCUICEquality.eq_annots Γ (expand_lets_ctx Δ Δ') <-> PCUICEquality.eq_annots Γ Δ'. +Proof. + rewrite /expand_lets_ctx /expand_lets_k_ctx. + etransitivity. eapply PCUICEquality.eq_annots_subst_context. + eapply PCUICEquality.eq_annots_lift_context. +Qed. + +Lemma eq_annots_ind_predicate_context ind mdecl idecl (pctx : list aname) : + PCUICEquality.eq_annots pctx (ind_predicate_context ind mdecl idecl) -> + PCUICEquality.eq_annots pctx (idecl_binder idecl :: ind_indices idecl). +Proof. + rewrite /ind_predicate_context. + intros eq. depelim eq; cbn in *. + constructor => //. + now eapply eq_annots_expand_lets_ctx. +Qed. + +Lemma apply_predctx {cf : checker_flags} {Σ : global_env_ext} Γ (ci : case_info) p indices c ps mdecl idecl {wfΣ : wf Σ.1} : + declared_inductive Σ ci mdecl idecl -> wf_predicate mdecl idecl p -> All2 (PCUICEquality.compare_decls eq eq) (pcontext p) (ind_predicate_context ci mdecl idecl) -> ctx_inst (typing Σ) Γ (pparams p ++ indices) (List.rev (ind_params mdecl,,, ind_indices idecl)@[puinst p]) -> @@ -167,9 +322,21 @@ Lemma apply_predctx {cf : checker_flags} Σ Γ ci p indices c ps mdecl idecl {wf let ptm := it_mkLambda_or_LetIn predctx (preturn p) in Σ;;; Γ,,, predctx |- preturn p : tSort ps -> Σ ;;; Γ |- mkApps ptm (indices ++ [c]) : tSort ps. Proof. - admit. -Admitted. - + intros decli wfp hpctx ctxi hd predctx ptm hret. + pose proof (validity hd) as ist. + pose proof (PCUICInductiveInversion.isType_mkApps_Ind_smash decli ist). forward X. apply wfp. + destruct X as [sppars [spinds cu]]. + eapply type_mkApps_arity. subst ptm. eapply PCUICGeneration.type_it_mkLambda_or_LetIn; tea. + assert (wfΓ : wf_local Σ Γ) by pcuic. + assert (wfu : wf_universe Σ ps). now eapply PCUICWfUniverses.typing_wf_universe in hret. + unshelve epose proof (PCUICInductives.arity_spine_case_predicate (ci:=ci) (indices:=indices) wfΓ decli cu wfu sppars _ hd). + now rewrite PCUICSR.smash_context_subst_context_let_expand in spinds. + unshelve eapply (arity_spine_mkProd_alpha _ X). + { now eapply PCUICArities.isType_Sort. } + unfold predctx. unfold case_predicate_context, case_predicate_context_gen. + symmetry; eapply PCUICCasesContexts.eq_binder_annots_eq. + now eapply wf_pre_case_predicate_context_gen. +Qed. Lemma relevance_term_to_type {cf : checker_flags} : env_prop relevance_from_term_from_type @@ -315,7 +482,7 @@ Proof using Type. eassert (isTypeRel (Σ, φ) _ (type_of_constructor mdecl cdecl (ind, i) u) (idecl.(ind_relevance))) by (apply isType_of_constructor; tea). eapply isTypeRel2_relevance; tea. apply wf. - assert (Σ ;;; Γ |- mkApps ptm (indices ++ [c]) : tSort ps). - { apply apply_predctx => //. apply ctx_inst_impl with (1 := X5) => ??[] //. } + { apply apply_predctx => //. eapply ctx_inst_impl with (2 := X5) => ???[] //. } split; intro Hr => //. + inv Hr. exists ps; split => //. From 5162310b0c9e53340b5adf448ce19875ba8dae60 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Wed, 25 May 2022 11:29:28 +0200 Subject: [PATCH 13/17] Fill the proof for projections as well --- pcuic/theories/PCUICInductives.v | 27 +++++++------- pcuic/theories/PCUICRelevanceTyp.v | 56 ++++++++++++++++++++++++++---- 2 files changed, 63 insertions(+), 20 deletions(-) diff --git a/pcuic/theories/PCUICInductives.v b/pcuic/theories/PCUICInductives.v index 8011fa280..bf7cb8aa3 100644 --- a/pcuic/theories/PCUICInductives.v +++ b/pcuic/theories/PCUICInductives.v @@ -718,7 +718,7 @@ Proof. destruct onps. destruct (ind_indices idecl). cbn. eapply typing_spine_refl. eapply isType_Sort; tea. now eapply on_inductive_sort. - now cbn in *. + cbn in on_projs_noidx. congruence. eapply isType_weaken; auto. now rewrite (oib.(ind_arity_eq)) in X. Qed. @@ -988,11 +988,12 @@ Lemma declared_projections_subslet_ind {cf:checker_flags} {Σ : global_env_ext} i0 < i -> forall x, nth_error (ind_projs idecl) i0 = Some x -> - isType (Σ.1, ind_universes mdecl) + isTypeRel (Σ.1, ind_universes mdecl) (smash_context [] (ind_params mdecl),, vass indb (mkApps (tInd ind (abstract_instance (ind_universes mdecl))) (to_extended_list (smash_context [] (ind_params mdecl))))) x.(proj_type) + x.(proj_relevance) × (∑ decl : context_decl, [× nth_error (smash_context [] (cstr_args c)) (context_assumptions (cstr_args c) - S i0) = Some decl, @@ -1079,10 +1080,11 @@ Proof. reflexivity. autorewrite with len. simpl. rewrite context_assumptions_smash_context /= //. + destruct IH as [isTy [decl [nthdecl _ _ eqpdecl ptyeq]]]. + eapply PCUICDeclarationTyping.isType_of_isTypeRel in isTy. assert(subst_instance (abstract_instance (ind_universes mdecl)) pdecl.(proj_type) = pdecl.(proj_type)) as ->. { eapply (isType_subst_instance_id (Σ.1, ind_universes mdecl)); eauto with pcuic. eapply declared_inductive_wf_ext_wk; eauto with pcuic. } - destruct IH as [isTy [decl [nthdecl _ _ eqpdecl ptyeq]]]. move ptyeq at bottom. rewrite nthdecl in Hnth. noconf Hnth. simpl in ptyeq. rewrite -eqpdecl. simpl. @@ -1111,7 +1113,7 @@ Proof. ++ constructor. apply smash_context_assumption_context; constructor. ++ unfold to_extended_list, to_extended_list_k. simpl. rewrite -reln_lift /= (reln_acc [_]) rev_app_distr /= //. - ++ now rewrite (Nat.add_1_r i). + ++ now rewrite (Nat.add_1_r i). Qed. Lemma wf_local_nth_isType {cf} {Σ} {Γ n d} : @@ -1141,7 +1143,8 @@ Lemma declared_projections {cf:checker_flags} {Σ : global_env_ext} {mdecl ind i on_projections mdecl (inductive_mind ind) (inductive_ind ind) idecl (ind_indices idecl) cs -> Alli (fun i pdecl => - isType (Σ.1, ind_universes mdecl) (projection_context_gen ind mdecl idecl) pdecl.(proj_type) * + isTypeRel (Σ.1, ind_universes mdecl) (projection_context_gen ind mdecl idecl) pdecl.(proj_type) + pdecl.(proj_relevance) * ∑ decl, [× nth_error (smash_context [] (cstr_args cs)) (context_assumptions (cstr_args cs) - S i) = Some decl, @@ -1410,6 +1413,7 @@ Proof. split. unfold projection_type in H0. rewrite H0. exists s; split; auto. + now rewrite on_proj_relevance. rewrite -/indb in Hs. rewrite /projection_type' -/indb -/indsubst -/projsubst. rewrite Nat.add_1_r in Hs. exact Hs. @@ -1467,10 +1471,8 @@ Lemma declared_projection_type {cf:checker_flags} {Σ : global_env_ext} {mdecl i wf Σ.1 -> declared_projection Σ p mdecl idecl cdecl pdecl -> let u := abstract_instance (ind_universes mdecl) in - isType (Σ.1, ind_universes mdecl) - ((vass {| binder_name := nNamed idecl.(ind_name); binder_relevance := idecl.(ind_relevance) |} (mkApps (tInd p.(proj_ind) u) - (to_extended_list (smash_context [] (ind_params mdecl))))):: - smash_context [] (ind_params mdecl)) pdecl.(proj_type). + isTypeRel (Σ.1, ind_universes mdecl) + (projection_context_gen p.(proj_ind) mdecl idecl) pdecl.(proj_type) pdecl.(proj_relevance). Proof. intros wfΣ declp. destruct (on_declared_projection declp) as [oni onp]. @@ -1493,11 +1495,8 @@ Lemma declared_projection_type_and_eq {cf:checker_flags} {Σ : global_env_ext} { forall (wfΣ : wf Σ.1) (Hdecl : declared_projection Σ p mdecl idecl cdecl pdecl), let u := abstract_instance (ind_universes mdecl) in (ind_ctors idecl = [cdecl]) * - isType (Σ.1, ind_universes mdecl) - ((vass {| binder_name := nNamed idecl.(ind_name); binder_relevance := idecl.(ind_relevance) |} - (mkApps (tInd p.(proj_ind) u) - (to_extended_list (smash_context [] (ind_params mdecl))))):: - smash_context [] (ind_params mdecl)) pdecl.(proj_type) * + isTypeRel (Σ.1, ind_universes mdecl) + (projection_context_gen p.(proj_ind) mdecl idecl) pdecl.(proj_type) pdecl.(proj_relevance) * ∑ decl, [× nth_error (smash_context [] (cstr_args cdecl)) (context_assumptions (cstr_args cdecl) - S p.(proj_arg)) = Some decl, diff --git a/pcuic/theories/PCUICRelevanceTyp.v b/pcuic/theories/PCUICRelevanceTyp.v index 952da1408..d81226a13 100644 --- a/pcuic/theories/PCUICRelevanceTyp.v +++ b/pcuic/theories/PCUICRelevanceTyp.v @@ -197,14 +197,58 @@ Proof. now rewrite Nat.add_0_r in H0. Qed. -Lemma isType_of_projection {cf : checker_flags} Σ mdecl idecl cdecl pdecl p c args u Γ : +Lemma isTypeRel_subst_instance_decl {cf : checker_flags} {Σ Γ 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. + intros wfΣ look isty cu. + eapply infer_typing_sort_impl with _ isty; [apply relevance_subst_opt|]. + intros Hs. + eapply (PCUICUnivSubstitutionTyp.typing_subst_instance_decl _ _ _ (tSort _)) in Hs; tea. +Qed. + +Lemma declared_projection_type_inst {cf:checker_flags} {Σ : global_env_ext} {mdecl idecl p cdecl pdecl Γ c u args} : + wf Σ.1 -> + declared_projection Σ p mdecl idecl cdecl pdecl -> + Σ;;; Γ |- c : mkApps (tInd (proj_ind p) u) args -> + isTypeRel Σ (PCUICInductives.projection_context p.(proj_ind) mdecl idecl u) + pdecl.(proj_type)@[u] pdecl.(proj_relevance). +Proof. + intros wfΣ declp hc. + eapply validity in hc. + assert (wfΓ : wf_local Σ Γ) by pcuic. + epose proof (PCUICInductives.isType_mkApps_Ind_inv wfΣ declp wfΓ hc) as [pars [argsub [? ? ? ? cu]]]. + epose proof (PCUICInductives.declared_projection_type wfΣ declp). + rewrite -(PCUICInductiveInversion.projection_context_gen_inst declp cu). + eapply isTypeRel_subst_instance_decl; tea. eapply declp. apply X. + apply cu. +Qed. + +Lemma isTypeRel_weaken {cf:checker_flags} {Σ : global_env_ext} {Γ Δ ty rel} : + wf Σ -> wf_local Σ Γ -> + isTypeRel Σ Δ ty rel -> + isTypeRel Σ (Γ ,,, Δ) ty rel. +Proof. + intros wfΣ wfΓ [s [hrel hty]]. + exists s; split => //. + now eapply weaken_ctx. +Qed. + +Lemma isTypeRel_projection {cf : checker_flags} Σ mdecl idecl cdecl pdecl p c args u Γ : wf Σ.1 -> declared_projection Σ.1 p mdecl idecl cdecl pdecl -> Σ;;; Γ |- c : mkApps (tInd (proj_ind p) u) args -> #|args| = ind_npars mdecl -> isTypeRel Σ Γ (subst0 (c :: List.rev args) (proj_type pdecl)@[u]) pdecl.(proj_relevance). Proof. - admit. -Admitted. - + intros wfΣ declp declc hargs. + epose proof (declared_projection_type_inst wfΣ declp declc). + eapply (PCUICSubstitution.isTypeRel_substitution (Δ := [])). + 2:{ cbn. eapply isTypeRel_weaken; tea. pcuic. } + eapply PCUICInductives.projection_subslet; tea. + now eapply validity. +Qed. Section OnInductives. Context {cf : checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ}. @@ -490,10 +534,10 @@ Proof using Type. eapply isTypeRel2_relevance; tea. exists ps; split => //. - pose proof (declared_projection_lookup isdecl). - assert (isTypeRel Σ Γ (subst0 (c :: List.rev args) (proj_type pdecl)@[u]) (pdecl.(proj_relevance))) by (eapply isType_of_projection; tea). + assert (isTypeRel Σ Γ (subst0 (c :: List.rev args) (proj_type pdecl)@[u]) (pdecl.(proj_relevance))) by (eapply isTypeRel_projection; tea). split; intro Hr => //. + inv Hr. - assert (pdecl = pdecl0). { clear -H0 isdecl. destruct H0 as (((H0m & H0i) & H0c) & H0p & _), isdecl as (((Hm & Hi) & Hc) & Hp & _). unfold declared_minductive in Hm, H0m. rewrite H0m in Hm. do 2 inversion Hm. subst mdecl0. rewrite H0i in Hi. inversion Hi. subst idecl0. rewrite H0c in Hc. inversion Hc. subst cdecl0. rewrite H0p in Hp. now inversion Hp. } subst pdecl0. + destruct (declared_projection_inj isdecl H0) as [? [? []]]; subst mdecl0 idecl0 cdecl0 pdecl0. assumption. + enough (pdecl.(proj_relevance) = rel) by (subst rel; econstructor; apply isdecl). eapply isTypeRel2_relevance; tea. From ea6493aa3a6477e75b5b87b8ee6a03e478e567e1 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Wed, 25 May 2022 11:45:10 +0200 Subject: [PATCH 14/17] Adapt to new projection_type(_and_eq) lemmas --- pcuic/theories/PCUICExpandLetsCorrectness.v | 5 +++-- pcuic/theories/PCUICInductiveInversion.v | 3 ++- pcuic/theories/PCUICSR.v | 3 +++ pcuic/theories/PCUICValidity.v | 2 +- 4 files changed, 9 insertions(+), 4 deletions(-) diff --git a/pcuic/theories/PCUICExpandLetsCorrectness.v b/pcuic/theories/PCUICExpandLetsCorrectness.v index 5598517cf..db2fec2a5 100644 --- a/pcuic/theories/PCUICExpandLetsCorrectness.v +++ b/pcuic/theories/PCUICExpandLetsCorrectness.v @@ -3777,8 +3777,9 @@ Proof. - rewrite (trans_subst (shiftnP #|projection_context p.(proj_ind) mdecl idecl u| xpred0) (shiftnP #|Γ| xpred0)). { rewrite /projection_context /=; len. cbn. - destruct (declared_projection_type_and_eq _ isdecl) as [[] ?]. - eapply isType_is_open_term in i. cbn in i; len in i. + destruct (declared_projection_type_and_eq _ isdecl) as [[? hty] ?]. + apply isType_of_isTypeRel in hty. + eapply isType_is_open_term in hty. cbn in hty; len in hty. rewrite on_free_vars_subst_instance //. } { generalize (subject_is_open_term X1). move/type_is_open_term: X1. now rewrite on_free_vars_mkApps /= forallb_rev => -> ->. } diff --git a/pcuic/theories/PCUICInductiveInversion.v b/pcuic/theories/PCUICInductiveInversion.v index e4acc8256..1470f1f94 100644 --- a/pcuic/theories/PCUICInductiveInversion.v +++ b/pcuic/theories/PCUICInductiveInversion.v @@ -3006,7 +3006,8 @@ Proof. epose proof (declared_projection_type_and_eq wfΣ declp). destruct (on_declared_projection declp). set (oib := declared_inductive_inv _ _ _ _) in *. simpl in p1, X. - destruct X as [[? ?] ?]. + destruct X as [[? hty] ?]. + eapply isType_of_isTypeRel in hty. destruct p1 as [[[H onps] onidx] onproj]. simpl in *. destruct ind_cunivs as [|? []] eqn:cseq => //. diff --git a/pcuic/theories/PCUICSR.v b/pcuic/theories/PCUICSR.v index 05d0e0e47..b71202d9a 100644 --- a/pcuic/theories/PCUICSR.v +++ b/pcuic/theories/PCUICSR.v @@ -2458,6 +2458,7 @@ Proof. destruct (ind_cunivs oib) as [|? []] eqn:hcunivs; try rewrite hcunivs // in tyargctx. specialize (projsubs onps). epose proof (declared_projection_type_and_eq wf isdecl) as [[hctors' hty] [decl'' [Hdecl Hty wfdecl peq peq']]]. + eapply isType_of_isTypeRel in hty. unfold projection_context_gen in hty. red in onp. unfold projection_type, projection_type' in *. set (indsubst := inds _ _ _) in *. @@ -2551,6 +2552,7 @@ Proof. move: (proj2 declc). destruct ind_ctors as [|? []] => //. intros [= ->]. destruct X2 as [[_ projty] projeq]. + eapply isType_of_isTypeRel in projty. unfold projection_context_gen in projty. destruct typec' as [[[[[_ equ] _] cu] eqargs] [cparsubst [cargsubst [iparsubst [iidxsubst ci]]]]]. destruct ci as [cparsubst0 iparsubst0 idxsubst0 subsidx [s [typectx [Hpars Hargs]]]]. destruct ind_cunivs as [|? ?] => //; noconf Hnth. @@ -2817,6 +2819,7 @@ Proof. move/validity/isType_open: typec => -> //. * have := (declared_projection_type_and_eq wf isdecl). move=> [[hctors isTy] Hdecl]. + eapply isType_of_isTypeRel in isTy. move/validity/(isType_mkApps_Ind_proj_inv _ isdecl): typec => [sppars hpars hargs cu]. move/(isType_subst_instance_decl _ _ _ _ (InductiveDecl mdecl) u _ _) : isTy. move/(_ _ isdecl.p1.p1.p1 cu). diff --git a/pcuic/theories/PCUICValidity.v b/pcuic/theories/PCUICValidity.v index 9d52667a5..cab145386 100644 --- a/pcuic/theories/PCUICValidity.v +++ b/pcuic/theories/PCUICValidity.v @@ -361,7 +361,7 @@ Section Validity. unshelve eapply isType_mkApps_Ind_inv in X2 as [parsubst [argsubst [sppar sparg lenpars lenargs cu]]]; eauto. 2:eapply isdecl.p1. - eapply infer_typing_sort_impl with _ isdecl'; [apply relevance_subst_opt|]; intros Hs. + eapply isType_of_isTypeRel, infer_typing_sort_impl with _ isdecl'; [apply relevance_subst_opt|]; intros Hs. eapply (typing_subst_instance_decl _ _ _ _ _ _ _ wf isdecl.p1.p1.p1) in Hs; eauto. simpl in Hs. eapply (weaken_ctx Γ) in Hs; eauto. From 35cd102141e61ffb575558fa6a04e108aad44caf Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Thu, 26 May 2022 02:32:37 +0200 Subject: [PATCH 15/17] Move lemmas to their right place --- erasure/_PluginProject.in | 2 - erasure/src/metacoq_erasure_plugin.mlpack | 13 - erasure/theories/ErasureFunction.v | 29 -- erasure/theories/Prelim.v | 4 +- pcuic/theories/PCUICAlpha.v | 83 +++- pcuic/theories/PCUICArities.v | 11 + pcuic/theories/PCUICContexts.v | 10 + pcuic/theories/PCUICEquality.v | 18 + pcuic/theories/PCUICInductiveInversion.v | 154 ++++++++ pcuic/theories/PCUICInductives.v | 25 ++ pcuic/theories/PCUICRelevanceTyp.v | 357 +----------------- pcuic/theories/PCUICSR.v | 15 - pcuic/theories/PCUICSpine.v | 15 + pcuic/theories/Syntax/PCUICClosed.v | 26 -- pcuic/theories/Syntax/PCUICLiftSubst.v | 25 ++ .../Typing/PCUICUnivSubstitutionTyp.v | 13 + pcuic/theories/Typing/PCUICWeakeningTyp.v | 10 + safechecker/theories/PCUICSafeChecker.v | 6 +- 18 files changed, 371 insertions(+), 445 deletions(-) diff --git a/erasure/_PluginProject.in b/erasure/_PluginProject.in index aa47c2482..9807b6804 100644 --- a/erasure/_PluginProject.in +++ b/erasure/_PluginProject.in @@ -106,8 +106,6 @@ src/eInduction.ml src/eInduction.mli src/eEtaExpanded.mli src/eEtaExpanded.ml -# src/eInduction.ml -# src/eInduction.mli src/eLiftSubst.ml src/eLiftSubst.mli src/eCSubst.mli diff --git a/erasure/src/metacoq_erasure_plugin.mlpack b/erasure/src/metacoq_erasure_plugin.mlpack index 520b59552..b98b63f00 100644 --- a/erasure/src/metacoq_erasure_plugin.mlpack +++ b/erasure/src/metacoq_erasure_plugin.mlpack @@ -1,30 +1,18 @@ -MSetWeakList -EqdepFacts Ssrbool Ssreflect -Utils WGraph UGraph0 EnvMap -WcbvEval -Classes0 -Logic1 -Relation -Relation_Properties PCUICAst PCUICCases PCUICAstUtils -EqDecInstances PCUICEquality PCUICTyping -PCUICInduction PCUICWfUniverses PCUICNormal PCUICPosition -PCUICPretty -PCUICWcbvEval PCUICProgram TemplateToPCUIC PCUICExpandLets @@ -34,7 +22,6 @@ PCUICErrors PCUICEqualityDec PCUICWfEnv PCUICWfEnvImpl -PCUICUnivSubst PCUICSafeReduce PCUICSafeRetyping diff --git a/erasure/theories/ErasureFunction.v b/erasure/theories/ErasureFunction.v index cd476f471..6a0b7d62f 100644 --- a/erasure/theories/ErasureFunction.v +++ b/erasure/theories/ErasureFunction.v @@ -41,35 +41,6 @@ Proof. eapply All2_app => //. reflexivity. Qed. -Section OnInductives. - Context {cf : checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ}. - - Lemma on_minductive_wf_params_weaken {ind mdecl Γ} {u} : - declared_minductive Σ.1 ind mdecl -> - consistent_instance_ext Σ (ind_universes mdecl) u -> - wf_local Σ Γ -> - wf_local Σ (Γ ,,, (ind_params mdecl)@[u]). - Proof. - intros. - eapply weaken_wf_local; tea. - eapply PCUICArities.on_minductive_wf_params; tea. - Qed. - - Context {mdecl ind idecl} - (decli : declared_inductive Σ ind mdecl idecl). - - Lemma on_minductive_wf_params_indices_inst_weaken {Γ} (u : Instance.t) : - consistent_instance_ext Σ (ind_universes mdecl) u -> - wf_local Σ Γ -> - wf_local Σ (Γ ,,, (ind_params mdecl ,,, ind_indices idecl)@[u]). - Proof. - intros. eapply weaken_wf_local; tea. - eapply PCUICInductives.on_minductive_wf_params_indices_inst; tea. - Qed. - - -End OnInductives. - Local Existing Instance extraction_checker_flags. (* Definition wf_ext_wf Σ : wf_ext Σ -> wf Σ := fst. #[global] diff --git a/erasure/theories/Prelim.v b/erasure/theories/Prelim.v index 9aca75f48..ce2af07a3 100644 --- a/erasure/theories/Prelim.v +++ b/erasure/theories/Prelim.v @@ -346,8 +346,8 @@ Proof. assert (smash_context [] (cstr_args cdecl)@[puinst p] = (cstr_args cdecl)@[puinst p]). { rewrite smash_assumption_context //. pcuic. } rewrite -H. - rewrite -PCUICClosed.smash_context_subst /= subst_context_nil. - rewrite -PCUICClosed.smash_context_subst /= subst_context_nil. apply eqctx. + rewrite -smash_context_subst /= subst_context_nil. + rewrite -smash_context_subst /= subst_context_nil. apply eqctx. Qed. diff --git a/pcuic/theories/PCUICAlpha.v b/pcuic/theories/PCUICAlpha.v index 94df50c65..02faaf966 100644 --- a/pcuic/theories/PCUICAlpha.v +++ b/pcuic/theories/PCUICAlpha.v @@ -7,7 +7,7 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICTactics PCUICConversion PCUICContextConversion PCUICContextConversionTyp PCUICValidity PCUICArities PCUICSpine PCUICInductives PCUICInductiveInversion PCUICOnFreeVars - PCUICWellScopedCumulativity PCUICGuardCondition. + PCUICWellScopedCumulativity PCUICGuardCondition PCUICContexts. From Equations Require Import Equations. @@ -33,6 +33,21 @@ Proof. eapply eq_term_upto_univ_empty_impl; tea; tc. Qed. +Lemma eq_term_upto_univ_it_mkProd_or_LetIn (Γ Γ' : context) T : + All2 (PCUICEquality.compare_decls eq eq) Γ Γ' -> + it_mkProd_or_LetIn Γ T ≡α it_mkProd_or_LetIn Γ' T. +Proof. + induction Γ in Γ' |- * using PCUICInduction.ctx_length_rev_ind. + - intros a; depelim a. reflexivity. + - intros a. + eapply All2_app_inv_l in a as [? [? [? []]]]. subst. + depelim a0. depelim a0. + rewrite !it_mkProd_or_LetIn_app. + depelim c; subst; constructor; cbn; auto; try reflexivity. + apply X => //. + apply X => //. +Qed. + Section Alpha. Context {cf:checker_flags}. @@ -1147,4 +1162,70 @@ Section Alpha. exists ctx', s; auto. Qed. + Lemma arity_spine_alpha {Σ : global_env_ext} {wfΣ : wf Σ} Γ args T T' T'' : + isType Σ Γ T'' -> + arity_spine Σ Γ T args T'' -> + T ≡α T' -> + arity_spine Σ Γ T' args T''. + Proof using Type. + intros isty sp. revert T'. induction sp; intros. + - intros. constructor. auto. + pose proof (isType_alpha _ _ _ isty X). + eapply PCUICContextConversion.ws_cumul_pb_eq_le. + destruct isty as [? []]. + constructor; fvs. red. apply PCUICEquality.upto_names_impl_eq_term. now symmetry. + - intros. constructor => //. + transitivity ty => //. + eapply PCUICContextConversion.ws_cumul_pb_eq_le. + constructor; fvs. + eapply PCUICConfluence.eq_term_upto_univ_napp_on_free_vars; tea. fvs. + apply PCUICEquality.upto_names_impl_eq_term. now symmetry. + - depelim X. + constructor. eapply IHsp => //. + eapply PCUICEquality.eq_term_upto_univ_subst; tc; auto. + - depelim X. + constructor. + pose proof (validity t). + eapply isType_alpha in X; tea. destruct X as [s [_ Hs]]. + econstructor; tea. eapply convSpec_cumulSpec. now apply eq_term_upto_univ_cumulSpec, PCUICEquality.upto_names_impl_eq_term. + eapply IHsp => //. eapply PCUICEquality.eq_term_upto_univ_subst; tc; auto. reflexivity. + Qed. + + Lemma arity_spine_mkProd_alpha {Σ : global_env_ext} {wfΣ : wf Σ} {Γ args} {Δ Δ' : context} {T T'} : + isType Σ Γ T' -> + arity_spine Σ Γ (it_mkProd_or_LetIn Δ T) args T' -> + All2 (PCUICEquality.compare_decls eq eq) Δ Δ' -> + arity_spine Σ Γ (it_mkProd_or_LetIn Δ' T) args T'. + Proof. + intros isty sp eq. + eapply arity_spine_alpha => //. exact sp. + now apply eq_term_upto_univ_it_mkProd_or_LetIn. + Qed. + + Lemma apply_predctx {Σ : global_env_ext} Γ (ci : case_info) p indices c ps mdecl idecl {wfΣ : wf Σ.1} : + declared_inductive Σ ci mdecl idecl -> + wf_predicate mdecl idecl p -> + All2 (PCUICEquality.compare_decls eq eq) (pcontext p) (ind_predicate_context ci mdecl idecl) -> + ctx_inst Σ Γ (pparams p ++ indices) (List.rev (ind_params mdecl,,, ind_indices idecl)@[puinst p]) -> + Σ;;; Γ |- c : mkApps (tInd ci (puinst p)) (pparams p ++ indices) -> + let predctx := case_predicate_context ci mdecl idecl p in + let ptm := it_mkLambda_or_LetIn predctx (preturn p) in + Σ;;; Γ,,, predctx |- preturn p : tSort ps -> Σ ;;; Γ |- mkApps ptm (indices ++ [c]) : tSort ps. + Proof. + intros decli wfp hpctx ctxi hd predctx ptm hret. + pose proof (validity hd) as ist. + pose proof (isType_mkApps_Ind_smash decli ist). forward X. apply wfp. + destruct X as [sppars [spinds cu]]. + eapply type_mkApps_arity. subst ptm. eapply PCUICGeneration.type_it_mkLambda_or_LetIn; tea. + assert (wfΓ : wf_local Σ Γ) by pcuic. + assert (wfu : wf_universe Σ ps). now eapply PCUICWfUniverses.typing_wf_universe in hret. + unshelve epose proof (PCUICInductives.arity_spine_case_predicate (ci:=ci) (indices:=indices) wfΓ decli cu wfu sppars _ hd). + now rewrite smash_context_subst_context_let_expand in spinds. + unshelve eapply (arity_spine_mkProd_alpha _ X). + { now eapply PCUICArities.isType_Sort. } + unfold predctx. unfold case_predicate_context, case_predicate_context_gen. + symmetry; eapply PCUICCasesContexts.eq_binder_annots_eq. + now eapply wf_pre_case_predicate_context_gen. + Qed. + End Alpha. diff --git a/pcuic/theories/PCUICArities.v b/pcuic/theories/PCUICArities.v index d5e2340f5..aa9e1f3c4 100644 --- a/pcuic/theories/PCUICArities.v +++ b/pcuic/theories/PCUICArities.v @@ -659,6 +659,17 @@ Section WfEnv. now apply onParams in H. Qed. + Lemma on_minductive_wf_params_weaken {ind mdecl Γ} {u} : + declared_minductive Σ.1 ind mdecl -> + consistent_instance_ext Σ (ind_universes mdecl) u -> + wf_local Σ Γ -> + wf_local Σ (Γ ,,, (ind_params mdecl)@[u]). + Proof using wfΣ. + intros. + eapply weaken_wf_local; tea. + eapply on_minductive_wf_params; tea. + Qed. + Lemma it_mkProd_or_LetIn_wf_local {Γ Δ T U} : Σ ;;; Γ |- it_mkProd_or_LetIn Δ T : U -> wf_local Σ (Γ ,,, Δ). Proof using wfΣ. diff --git a/pcuic/theories/PCUICContexts.v b/pcuic/theories/PCUICContexts.v index 747c76a1e..1787c6f47 100644 --- a/pcuic/theories/PCUICContexts.v +++ b/pcuic/theories/PCUICContexts.v @@ -746,6 +746,16 @@ Proof. now rewrite !to_extended_list_k_subst to_extended_list_k_lift_context. Qed. +Lemma smash_context_subst_context_let_expand s Γ Δ : + smash_context [] (subst_context_let_expand s Γ Δ) = + subst_context_let_expand s Γ (smash_context [] Δ). +Proof. + rewrite /subst_context_let_expand. + rewrite (smash_context_subst []). + now rewrite /expand_lets_ctx /expand_lets_k_ctx (smash_context_subst []) + (smash_context_lift []). +Qed. + Lemma map2_set_binder_name_alpha (nas : list aname) (Δ Δ' : context) : All2 (fun x y => eq_binder_annot x (decl_name y)) nas Δ -> eq_context_upto_names Δ Δ' -> diff --git a/pcuic/theories/PCUICEquality.v b/pcuic/theories/PCUICEquality.v index 5e6a614f4..fef77cd0f 100644 --- a/pcuic/theories/PCUICEquality.v +++ b/pcuic/theories/PCUICEquality.v @@ -1812,3 +1812,21 @@ Proof. symmetry; eapply (eq_annots_subst_context _ (List.rev pars) 0). reflexivity. Qed. + +Lemma eq_annots_expand_lets_ctx (Γ : list aname) (Δ Δ' : context) : + eq_annots Γ (expand_lets_ctx Δ Δ') <-> eq_annots Γ Δ'. +Proof. + rewrite /expand_lets_ctx /expand_lets_k_ctx. + etransitivity. eapply eq_annots_subst_context. + eapply eq_annots_lift_context. +Qed. + +Lemma eq_annots_ind_predicate_context ind mdecl idecl (pctx : list aname) : + eq_annots pctx (ind_predicate_context ind mdecl idecl) -> + eq_annots pctx (idecl_binder idecl :: ind_indices idecl). +Proof. + rewrite /ind_predicate_context. + intros eq. depelim eq; cbn in *. + constructor => //. + now eapply eq_annots_expand_lets_ctx. +Qed. diff --git a/pcuic/theories/PCUICInductiveInversion.v b/pcuic/theories/PCUICInductiveInversion.v index 1470f1f94..2ba49dce9 100644 --- a/pcuic/theories/PCUICInductiveInversion.v +++ b/pcuic/theories/PCUICInductiveInversion.v @@ -4506,3 +4506,157 @@ Proof. eapply wf_case_branch_type; tea. split; eauto. 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 Relevant. +Proof. + intros wfΣ ind_type (s & e & Hs). + eexists; split => //; tea. + rewrite /ind_type in Hs. + rewrite -it_mkProd_or_LetIn_app in Hs. + assert (wfΓ : wf_local Σ []) by pcuic. + revert wfΓ Hs. generalize ((ind_indices idecl ++ ind_params mdecl : context)) as l, ([]: context) as Γ. + induction l using rev_ind. + * cbn. intros Γ wfΓ Hty; apply inversion_Sort in Hty as (_ & _ & le); tea. + apply ws_cumul_pb_Sort_inv in le. + eapply leq_relevance; tea. + now destruct ind_sort. + * rewrite it_mkProd_or_LetIn_app; cbn. + destruct x, decl_body; cbn. + - intros Γ wfΓ Hty; pose proof Hty; apply inversion_LetIn in Hty as (? & ? & ? & ? & ? & ? & le); tea. + assert (wfΓ' : wf_local Σ (Γ,, vdef decl_name t decl_type)) by auto with pcuic. + eapply IHl. apply wfΓ'. + econstructor; eauto with pcuic. constructor; eauto with pcuic. + apply ws_cumul_pb_LetIn_l_inv in le. + unshelve epose proof (PCUICSpine.all_rels_subst_lift (Γ := Γ) (Δ := [vdef decl_name t decl_type]) (t := x0) (Δ' := []) _ _ _) => //=. + all: change (Γ,,, [vdef decl_name t decl_type]) with (Γ ,, vdef decl_name t decl_type); tea. + 1,2: fvs. + simpl in X0. rewrite PCUICLiftSubst.lift0_id in X0. rewrite PCUICLiftSubst.subst_empty in X0. + change (subst0 _ _) with ((lift 1 1 x0) {0 := lift0 1 t}) in X0. + rewrite -PCUICLiftSubst.distr_lift_subst10 in X0. + eapply PCUICContextConversion.ws_cumul_pb_eq_le in X0. + eapply cumulAlgo_cumulSpec. + etransitivity. 1: apply X0. + change (tSort _) with (lift0 1 (tSort s)). + eapply (weakening_ws_cumul_pb (Γ' := []) (Γ'' := [vdef _ _ _]) le). fvs. + - intros Γ wfΓ Hty; pose proof Hty; apply inversion_Prod in Hty as (? & s' & ? & ? & ? & le); tea. + assert (wfΓ' : wf_local Σ (Γ,, vass decl_name decl_type)) by auto with pcuic. + eapply IHl. apply wfΓ'. + econstructor; eauto with pcuic. constructor; eauto with pcuic. + eapply cumulAlgo_cumulSpec. + etransitivity. 2: eapply (weakening_ws_cumul_pb (Γ' := []) (Γ'' := [vass _ _]) le); fvs. + cbn. + constructor. 1-3: fvs. + constructor. apply leq_universe_product. +Qed. + + +Lemma inversion_Rel_wt {cf} {Σ : global_env_ext} {wfΣ : wf Σ}: + forall {Γ n T}, + Σ ;;; Γ |- tRel n : T -> + ∑ decl, + [× wf_local Σ Γ, + nth_error Γ n = Some decl, + isTypeRel Σ Γ (lift0 (S n) (decl_type decl)) decl.(decl_name).(binder_relevance) & + Σ ;;; Γ ⊢ lift0 (S n) (decl_type decl) ≤ T]. +Proof. + intros Γ n T h. depind h. + - exists decl. + assert (isTypeRel Σ Γ (lift0 (S n) (decl_type decl)) (binder_relevance (decl_name decl))). + { destruct (PCUICTyping.nth_error_All_local_env e a); eauto. cbn in o. + apply isTypeRelOpt_lift in l => //. 1:{ pose proof (nth_error_Some_length e). lia. } } + split => //. + apply isType_ws_cumul_pb_refl. destruct X as [s [_ Hs]]. now exists s. + - destruct IHh1 as [decl []]. exists decl; split => //. + etransitivity; tea. eapply PCUICConversion.cumulSpec_typed_cumulAlgo; tea => //. +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); exists s; split => //=. + eapply PCUICSpine.inversion_it_mkProd_or_LetIn in Hs. + eapply inversion_mkApps in Hs as [tyapp [Happ Hsp]]. + eapply inversion_Rel_wt in Happ as [cdecl' [wfl hnth ht cum]] => //. + 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. + destruct ht as [s' [_ hs]]. now exists s'. +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 ((? & []) & ? & ? & []). + eapply infer_typing_sort_impl with _ on_ctype; [apply relevance_subst_opt|]; intros Hty. + instantiate (1 := u). + replace Γ with ([] ,,, Γ) by apply app_context_nil_l. + replace (type_of_constructor _ _ _ _) with (lift0 #|Γ| (type_of_constructor mdecl cdecl cstr u)). + 2: { rewrite PCUICLiftSubst.lift_closed //. eapply (PCUICClosedTyp.declared_constructor_closed_type isdecl). } + change (tSort _) with (lift0 #|Γ| (tSort on_ctype.π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.π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 isdecl.p1.p1. } + eapply subslet_inds; tea. exact isdecl. +Qed. + +Lemma declared_projection_type_inst {cf:checker_flags} {Σ : global_env_ext} {mdecl idecl p cdecl pdecl Γ c u args} : + wf Σ.1 -> + declared_projection Σ p mdecl idecl cdecl pdecl -> + Σ;;; Γ |- c : mkApps (tInd (proj_ind p) u) args -> + isTypeRel Σ (PCUICInductives.projection_context p.(proj_ind) mdecl idecl u) + pdecl.(proj_type)@[u] pdecl.(proj_relevance). +Proof. + intros wfΣ declp hc. + eapply validity in hc. + assert (wfΓ : wf_local Σ Γ) by pcuic. + epose proof (isType_mkApps_Ind_inv wfΣ declp wfΓ hc) as [pars [argsub [? ? ? ? cu]]]. + epose proof (declared_projection_type wfΣ declp). + rewrite -(projection_context_gen_inst declp cu). + eapply isTypeRel_subst_instance_decl; tea. eapply declp. apply X. apply cu. +Qed. + +Lemma isTypeRel_projection {cf : checker_flags} Σ mdecl idecl cdecl pdecl p c args u Γ : + wf Σ.1 -> declared_projection Σ.1 p mdecl idecl cdecl pdecl -> + Σ;;; Γ |- c : mkApps (tInd (proj_ind p) u) args -> #|args| = ind_npars mdecl -> + isTypeRel Σ Γ (subst0 (c :: List.rev args) (proj_type pdecl)@[u]) pdecl.(proj_relevance). +Proof. + intros wfΣ declp declc hargs. + epose proof (declared_projection_type_inst wfΣ declp declc). + eapply (isTypeRel_substitution (Δ := [])). + 2:{ cbn. eapply isTypeRel_weaken; tea. pcuic. } + eapply projection_subslet; tea. now eapply validity. +Qed. diff --git a/pcuic/theories/PCUICInductives.v b/pcuic/theories/PCUICInductives.v index bf7cb8aa3..dafe54536 100644 --- a/pcuic/theories/PCUICInductives.v +++ b/pcuic/theories/PCUICInductives.v @@ -196,6 +196,22 @@ Proof. eapply wf_arities_context; tea. 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 := Relevant |}; + 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. + rewrite /arities_context rev_map_spec nth_error_rev; autorewrite with len; auto. + now rewrite List.rev_involutive nth_error_map hnth. +Qed. + Lemma instantiate_inds {cf:checker_flags} {Σ} {wfΣ : wf Σ.1} {u mind mdecl} : declared_minductive Σ.1 mind mdecl -> consistent_instance_ext Σ (ind_universes mdecl) u -> @@ -272,6 +288,15 @@ Section OnInductives. now eapply on_minductive_wf_params_indices. Qed. + Lemma on_minductive_wf_params_indices_inst_weaken {Γ} (u : Instance.t) : + consistent_instance_ext Σ (ind_universes mdecl) u -> + wf_local Σ Γ -> + wf_local Σ (Γ ,,, (ind_params mdecl ,,, ind_indices idecl)@[u]). + Proof. + intros. eapply weaken_wf_local; tea. + eapply on_minductive_wf_params_indices_inst; tea. + Qed. + Lemma on_inductive_inst Γ u : wf_local Σ Γ -> consistent_instance_ext Σ (ind_universes mdecl) u -> diff --git a/pcuic/theories/PCUICRelevanceTyp.v b/pcuic/theories/PCUICRelevanceTyp.v index d81226a13..917be751d 100644 --- a/pcuic/theories/PCUICRelevanceTyp.v +++ b/pcuic/theories/PCUICRelevanceTyp.v @@ -1,7 +1,8 @@ From MetaCoq.Template Require Import config utils. From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICRelevance PCUICClosedTyp PCUICGlobalEnv - PCUICTyping PCUICInversion PCUICConversion PCUICCumulProp PCUICWeakeningTyp PCUICValidity PCUICWellScopedCumulativity BDUnique. + PCUICTyping PCUICInversion PCUICConversion PCUICCumulProp PCUICWeakeningTyp PCUICValidity PCUICWellScopedCumulativity + PCUICInductives PCUICInductiveInversion BDUnique. Require Import ssreflect. From Equations Require Import Equations. @@ -30,358 +31,6 @@ Proof. now eapply cumul_sort_relevance. 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 Relevant. -Proof. - intros wfΣ ind_type (s & e & Hs). - eexists; split => //; tea. - rewrite /ind_type in Hs. - rewrite -it_mkProd_or_LetIn_app in Hs. - assert (wfΓ : wf_local Σ []) by pcuic. - revert wfΓ Hs. generalize ((ind_indices idecl ++ ind_params mdecl : context)) as l, ([]: context) as Γ. - induction l using rev_ind. - * cbn. intros Γ wfΓ Hty; apply inversion_Sort in Hty as (_ & _ & le); tea. - apply ws_cumul_pb_Sort_inv in le. - eapply leq_relevance; tea. - now destruct ind_sort. - * rewrite it_mkProd_or_LetIn_app; cbn. - destruct x, decl_body; cbn. - - intros Γ wfΓ Hty; pose proof Hty; apply inversion_LetIn in Hty as (? & ? & ? & ? & ? & ? & le); tea. - assert (wfΓ' : wf_local Σ (Γ,, vdef decl_name t decl_type)) by auto with pcuic. - eapply IHl. apply wfΓ'. - econstructor; eauto with pcuic. constructor; eauto with pcuic. - apply ws_cumul_pb_LetIn_l_inv in le. - unshelve epose proof (PCUICSpine.all_rels_subst_lift (Γ := Γ) (Δ := [vdef decl_name t decl_type]) (t := x0) (Δ' := []) _ _ _) => //=. - all: change (Γ,,, [vdef decl_name t decl_type]) with (Γ ,, vdef decl_name t decl_type); tea. - 1,2: fvs. - simpl in X0. rewrite PCUICLiftSubst.lift0_id in X0. rewrite PCUICLiftSubst.subst_empty in X0. - change (subst0 _ _) with ((lift 1 1 x0) {0 := lift0 1 t}) in X0. - rewrite -PCUICLiftSubst.distr_lift_subst10 in X0. - eapply PCUICContextConversion.ws_cumul_pb_eq_le in X0. - eapply cumulAlgo_cumulSpec. - etransitivity. 1: apply X0. - change (tSort _) with (lift0 1 (tSort s)). - eapply (weakening_ws_cumul_pb (Γ' := []) (Γ'' := [vdef _ _ _]) le). fvs. - - intros Γ wfΓ Hty; pose proof Hty; apply inversion_Prod in Hty as (? & s' & ? & ? & ? & le); tea. - assert (wfΓ' : wf_local Σ (Γ,, vass decl_name decl_type)) by auto with pcuic. - eapply IHl. apply wfΓ'. - econstructor; eauto with pcuic. constructor; eauto with pcuic. - eapply cumulAlgo_cumulSpec. - etransitivity. 2: eapply (weakening_ws_cumul_pb (Γ' := []) (Γ'' := [vass _ _]) le); fvs. - cbn. - constructor. 1-3: fvs. - constructor. apply leq_universe_product. -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 := Relevant |}; - 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. - -Lemma inversion_Rel_wt {cf : checker_flags} : -forall {Σ : global_env_ext} {Γ n T} {wfΣ : wf Σ}, - Σ ;;; Γ |- tRel n : T -> - ∑ decl, - [× wf_local Σ Γ, - nth_error Γ n = Some decl, - isTypeRel Σ Γ (lift0 (S n) (decl_type decl)) decl.(decl_name).(binder_relevance) & - Σ ;;; Γ ⊢ lift0 (S n) (decl_type decl) ≤ T]. -Proof. - intros Σ Γ n T wfΣ h. depind h. - - exists decl. - assert (isTypeRel Σ Γ (lift0 (S n) (decl_type decl)) (binder_relevance (decl_name decl))). - { destruct (PCUICTyping.nth_error_All_local_env e a); eauto. cbn in o. - apply isTypeRelOpt_lift in l => //. 1:{ pose proof (nth_error_Some_length e). lia. } } - split => //. - apply isType_ws_cumul_pb_refl. destruct X as [s [_ Hs]]. now exists s. - - destruct (IHh1 wfΣ) as [decl []]. exists decl; split => //. - etransitivity; tea. eapply cumulSpec_typed_cumulAlgo; tea => //. -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 (typing Σ) (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); exists s; split => //=. - eapply PCUICSpine.inversion_it_mkProd_or_LetIn in Hs. - eapply inversion_mkApps in Hs as [tyapp [Happ Hsp]]. - eapply inversion_Rel_wt in Happ as [cdecl' [wfl hnth ht cum]] => //. - 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. - destruct ht as [s' [_ hs]]. now exists s'. -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 ((? & []) & ? & ? & []). - eapply infer_typing_sort_impl with _ on_ctype; [apply relevance_subst_opt|]; intros Hty. - instantiate (1 := u). - replace Γ with ([] ,,, Γ) by apply app_context_nil_l. - replace (type_of_constructor _ _ _ _) with (lift0 #|Γ| (type_of_constructor mdecl cdecl cstr u)). - 2: { rewrite PCUICLiftSubst.lift_closed //. eapply (PCUICClosedTyp.declared_constructor_closed_type isdecl). } - change (tSort _) with (lift0 #|Γ| (tSort on_ctype.π1)@[u]). - eapply @weakening with (Γ := []) => //. - rewrite app_context_nil_l => //. - destruct Σ as (Σ & φ). - eapply PCUICUnivSubstitutionTyp.typing_subst_instance' in Hty => //=; tea. - change ((tSort _)@[u]) with (subst0 (inds (inductive_mind cstr.1) u (ind_bodies mdecl)) (tSort on_ctype.π1)@[u]). - eapply @PCUICSubstitution.substitution with (Δ := []); tea. - 2: rewrite app_context_nil_l; apply Hty. - 2: { epose proof (PCUICWeakeningEnv.weaken_lookup_on_global_env' _ _ (InductiveDecl mdecl) wfΣ _); eauto. now split. Unshelve. 2: apply isdecl.p1.p1. } - rewrite /inds /arities_context. - clear -h_cuu isdecl wfΣ wfΓ. - pose proof (PCUICClosedTyp.declared_minductive_closed isdecl.p1.p1). - rewrite /PCUICClosed.closed_inductive_decl in H. move/andb_and: H => [] _ H. - apply forallb_All in H. - assert (Alli (fun i x => declared_inductive Σ {| inductive_mind := inductive_mind cstr.1; inductive_ind := i |} mdecl x) 0 (ind_bodies mdecl)). - { - apply forall_nth_error_Alli; intros. - unfold declared_inductive. - split => //=. - apply isdecl.p1.p1. - } - induction (ind_bodies mdecl) using rev_ind => //=. - - constructor. - - rewrite rev_map_app app_length Nat.add_1_r => //=. - apply All_app in H as (H & H1). - apply Alli_app in X as (X & X1). - constructor => //=; auto. - clear H X IHl. - inv H1. inv X1. clear X X0. - rewrite /PCUICClosed.closed_inductive_body in H. rtoProp. - rewrite PCUICClosed.subst_closedn. - rewrite PCUICClosed.closedn_subst_instance => //. - eapply type_Ind; eauto with pcuic. - now rewrite Nat.add_0_r in H0. -Qed. - -Lemma isTypeRel_subst_instance_decl {cf : checker_flags} {Σ Γ 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. - intros wfΣ look isty cu. - eapply infer_typing_sort_impl with _ isty; [apply relevance_subst_opt|]. - intros Hs. - eapply (PCUICUnivSubstitutionTyp.typing_subst_instance_decl _ _ _ (tSort _)) in Hs; tea. -Qed. - -Lemma declared_projection_type_inst {cf:checker_flags} {Σ : global_env_ext} {mdecl idecl p cdecl pdecl Γ c u args} : - wf Σ.1 -> - declared_projection Σ p mdecl idecl cdecl pdecl -> - Σ;;; Γ |- c : mkApps (tInd (proj_ind p) u) args -> - isTypeRel Σ (PCUICInductives.projection_context p.(proj_ind) mdecl idecl u) - pdecl.(proj_type)@[u] pdecl.(proj_relevance). -Proof. - intros wfΣ declp hc. - eapply validity in hc. - assert (wfΓ : wf_local Σ Γ) by pcuic. - epose proof (PCUICInductives.isType_mkApps_Ind_inv wfΣ declp wfΓ hc) as [pars [argsub [? ? ? ? cu]]]. - epose proof (PCUICInductives.declared_projection_type wfΣ declp). - rewrite -(PCUICInductiveInversion.projection_context_gen_inst declp cu). - eapply isTypeRel_subst_instance_decl; tea. eapply declp. apply X. - apply cu. -Qed. - -Lemma isTypeRel_weaken {cf:checker_flags} {Σ : global_env_ext} {Γ Δ ty rel} : - wf Σ -> wf_local Σ Γ -> - isTypeRel Σ Δ ty rel -> - isTypeRel Σ (Γ ,,, Δ) ty rel. -Proof. - intros wfΣ wfΓ [s [hrel hty]]. - exists s; split => //. - now eapply weaken_ctx. -Qed. - -Lemma isTypeRel_projection {cf : checker_flags} Σ mdecl idecl cdecl pdecl p c args u Γ : - wf Σ.1 -> declared_projection Σ.1 p mdecl idecl cdecl pdecl -> - Σ;;; Γ |- c : mkApps (tInd (proj_ind p) u) args -> #|args| = ind_npars mdecl -> - isTypeRel Σ Γ (subst0 (c :: List.rev args) (proj_type pdecl)@[u]) pdecl.(proj_relevance). -Proof. - intros wfΣ declp declc hargs. - epose proof (declared_projection_type_inst wfΣ declp declc). - eapply (PCUICSubstitution.isTypeRel_substitution (Δ := [])). - 2:{ cbn. eapply isTypeRel_weaken; tea. pcuic. } - eapply PCUICInductives.projection_subslet; tea. - now eapply validity. -Qed. - -Section OnInductives. - Context {cf : checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ}. - - Context {mdecl ind idecl} - (decli : declared_inductive Σ ind mdecl idecl). - -Lemma on_minductive_wf_params_indices_inst_weaken {Γ} (u : Instance.t) : -consistent_instance_ext Σ (ind_universes mdecl) u -> -wf_local Σ Γ -> -wf_local Σ (Γ ,,, (ind_params mdecl ,,, ind_indices idecl)@[u]). -Proof. -intros. eapply weaken_wf_local; tea. -eapply PCUICInductives.on_minductive_wf_params_indices_inst; tea. -Qed. -End OnInductives. - -Lemma spine_subst_alpha {cf : checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} Γ args inst Δ Δ' : - PCUICSpine.spine_subst Σ Γ args inst Δ -> - All2 (PCUICEquality.compare_decls eq eq) Δ Δ' -> - PCUICSpine.spine_subst Σ Γ args inst Δ'. -Proof. - intros [] eqd. - split => //. - eapply PCUICSpine.wf_local_alpha; tea. eapply All2_app; trea. - 2:{ eapply PCUICSpine.subslet_eq_context_alpha; tea. } - move: inst_ctx_subst. clear -eqd. - induction 1 in Δ', eqd |- *; depelim eqd; try constructor; depelim c; subst. - - constructor. now apply IHinst_ctx_subst. - - constructor. eauto. -Qed. - -Import PCUICAlpha. - -Lemma arity_spine_alpha {cf : checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} Γ args T T' T'' : - isType Σ Γ T'' -> - PCUICSpine.arity_spine Σ Γ T args T'' -> - T ≡α T' -> - PCUICSpine.arity_spine Σ Γ T' args T''. -Proof. - intros isty sp. revert T'. induction sp; intros. - - intros. constructor. auto. - pose proof (isType_alpha _ _ _ isty X). - eapply PCUICContextConversion.ws_cumul_pb_eq_le. - destruct isty as [? []]. - constructor; fvs. red. apply PCUICEquality.upto_names_impl_eq_term. now symmetry. - - intros. constructor => //. - transitivity ty => //. - eapply PCUICContextConversion.ws_cumul_pb_eq_le. - constructor; fvs. - eapply PCUICConfluence.eq_term_upto_univ_napp_on_free_vars; tea. fvs. - apply PCUICEquality.upto_names_impl_eq_term. now symmetry. - - depelim X. - constructor. eapply IHsp => //. - eapply PCUICEquality.eq_term_upto_univ_subst; tc; auto. - - depelim X. - constructor. - pose proof (validity t). - eapply isType_alpha in X; tea. destruct X as [s [_ Hs]]. - econstructor; tea. eapply convSpec_cumulSpec. now apply eq_term_upto_univ_cumulSpec, PCUICEquality.upto_names_impl_eq_term. - eapply IHsp => //. eapply PCUICEquality.eq_term_upto_univ_subst; tc; auto. reflexivity. -Qed. - -Lemma eq_term_upto_univ_it_mkProd_or_LetIn (Γ Γ' : context) T : - All2 (PCUICEquality.compare_decls eq eq) Γ Γ' -> - it_mkProd_or_LetIn Γ T ≡α it_mkProd_or_LetIn Γ' T. -Proof. - induction Γ in Γ' |- * using PCUICInduction.ctx_length_rev_ind. - - intros a; depelim a. reflexivity. - - intros a. - eapply All2_app_inv_l in a as [? [? [? []]]]. subst. - depelim a0. depelim a0. - rewrite !it_mkProd_or_LetIn_app. - depelim c; subst; constructor; cbn; auto; try reflexivity. - apply X => //. - apply X => //. -Qed. - -Lemma arity_spine_mkProd_alpha {cf : checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} {Γ args} {Δ Δ' : context} {T T'} : - isType Σ Γ T' -> - PCUICSpine.arity_spine Σ Γ (it_mkProd_or_LetIn Δ T) args T' -> - All2 (PCUICEquality.compare_decls eq eq) Δ Δ' -> - PCUICSpine.arity_spine Σ Γ (it_mkProd_or_LetIn Δ' T) args T'. -Proof. - intros isty sp eq. - eapply arity_spine_alpha => //. exact sp. - now apply eq_term_upto_univ_it_mkProd_or_LetIn. -Qed. - -Lemma eq_annots_expand_lets_ctx (Γ : list aname) (Δ Δ' : context) : - PCUICEquality.eq_annots Γ (expand_lets_ctx Δ Δ') <-> PCUICEquality.eq_annots Γ Δ'. -Proof. - rewrite /expand_lets_ctx /expand_lets_k_ctx. - etransitivity. eapply PCUICEquality.eq_annots_subst_context. - eapply PCUICEquality.eq_annots_lift_context. -Qed. - -Lemma eq_annots_ind_predicate_context ind mdecl idecl (pctx : list aname) : - PCUICEquality.eq_annots pctx (ind_predicate_context ind mdecl idecl) -> - PCUICEquality.eq_annots pctx (idecl_binder idecl :: ind_indices idecl). -Proof. - rewrite /ind_predicate_context. - intros eq. depelim eq; cbn in *. - constructor => //. - now eapply eq_annots_expand_lets_ctx. -Qed. - -Lemma apply_predctx {cf : checker_flags} {Σ : global_env_ext} Γ (ci : case_info) p indices c ps mdecl idecl {wfΣ : wf Σ.1} : - declared_inductive Σ ci mdecl idecl -> - wf_predicate mdecl idecl p -> - All2 (PCUICEquality.compare_decls eq eq) (pcontext p) (ind_predicate_context ci mdecl idecl) -> - ctx_inst (typing Σ) Γ (pparams p ++ indices) (List.rev (ind_params mdecl,,, ind_indices idecl)@[puinst p]) -> - Σ;;; Γ |- c : mkApps (tInd ci (puinst p)) (pparams p ++ indices) -> - let predctx := case_predicate_context ci mdecl idecl p in - let ptm := it_mkLambda_or_LetIn predctx (preturn p) in - Σ;;; Γ,,, predctx |- preturn p : tSort ps -> Σ ;;; Γ |- mkApps ptm (indices ++ [c]) : tSort ps. -Proof. - intros decli wfp hpctx ctxi hd predctx ptm hret. - pose proof (validity hd) as ist. - pose proof (PCUICInductiveInversion.isType_mkApps_Ind_smash decli ist). forward X. apply wfp. - destruct X as [sppars [spinds cu]]. - eapply type_mkApps_arity. subst ptm. eapply PCUICGeneration.type_it_mkLambda_or_LetIn; tea. - assert (wfΓ : wf_local Σ Γ) by pcuic. - assert (wfu : wf_universe Σ ps). now eapply PCUICWfUniverses.typing_wf_universe in hret. - unshelve epose proof (PCUICInductives.arity_spine_case_predicate (ci:=ci) (indices:=indices) wfΓ decli cu wfu sppars _ hd). - now rewrite PCUICSR.smash_context_subst_context_let_expand in spinds. - unshelve eapply (arity_spine_mkProd_alpha _ X). - { now eapply PCUICArities.isType_Sort. } - unfold predctx. unfold case_predicate_context, case_predicate_context_gen. - symmetry; eapply PCUICCasesContexts.eq_binder_annots_eq. - now eapply wf_pre_case_predicate_context_gen. -Qed. - Lemma relevance_term_to_type {cf : checker_flags} : env_prop relevance_from_term_from_type (fun Σ Γ => All_local_env (lift_typing relevance_from_term_from_type Σ) Γ). @@ -526,7 +175,7 @@ Proof using Type. eassert (isTypeRel (Σ, φ) _ (type_of_constructor mdecl cdecl (ind, i) u) (idecl.(ind_relevance))) by (apply isType_of_constructor; tea). eapply isTypeRel2_relevance; tea. apply wf. - assert (Σ ;;; Γ |- mkApps ptm (indices ++ [c]) : tSort ps). - { apply apply_predctx => //. eapply ctx_inst_impl with (2 := X5) => ???[] //. } + { apply PCUICAlpha.apply_predctx => //. eapply ctx_inst_impl with (1 := X5) => ??[] //. } split; intro Hr => //. + inv Hr. exists ps; split => //. diff --git a/pcuic/theories/PCUICSR.v b/pcuic/theories/PCUICSR.v index b71202d9a..8d2fae0b8 100644 --- a/pcuic/theories/PCUICSR.v +++ b/pcuic/theories/PCUICSR.v @@ -491,16 +491,6 @@ Proof. now rewrite Nat.add_comm. Qed. -Lemma smash_context_subst_context_let_expand s Γ Δ : - smash_context [] (subst_context_let_expand s Γ Δ) = - subst_context_let_expand s Γ (smash_context [] Δ). -Proof. - rewrite /subst_context_let_expand. - rewrite (smash_context_subst []). - now rewrite /expand_lets_ctx /expand_lets_k_ctx (smash_context_subst []) - (smash_context_lift []). -Qed. - Lemma on_constructor_wf_args {cf} {Σ} {wfΣ : wf Σ} {ind c mdecl idecl cdecl u} : declared_constructor Σ (ind, c) mdecl idecl cdecl -> consistent_instance_ext Σ (ind_universes mdecl) u -> @@ -1088,11 +1078,6 @@ Definition closed_red1_ind' := ltac:(let T := type of closed_red1_ind in let T' := eval cbn in T in exact (closed_red1_ind : T')). -(* -Ltac revert_until x := - repeat lazymatch goal with - | [ H : _ |- _ ] => revert H - end. *) Ltac invert_closed_red H := generalize_eqs H; destruct H using closed_red1_ind'; simplify_dep_elim. diff --git a/pcuic/theories/PCUICSpine.v b/pcuic/theories/PCUICSpine.v index d27c56a6c..33a4f814e 100644 --- a/pcuic/theories/PCUICSpine.v +++ b/pcuic/theories/PCUICSpine.v @@ -453,6 +453,21 @@ Section WfEnv. apply IHsp. now eapply isType_apply in wf => //. Qed. + Lemma spine_subst_alpha {Γ args inst Δ Δ'} : + spine_subst Σ Γ args inst Δ -> + All2 (PCUICEquality.compare_decls eq eq) Δ Δ' -> + spine_subst Σ Γ args inst Δ'. + Proof. + intros [wf wf' cs sub] eqd. + split => //. + eapply wf_local_alpha; tea. eapply All2_app; trea. + 2:{ eapply subslet_eq_context_alpha; tea. } + move: cs. clear -eqd. + induction 1 in Δ', eqd |- *; depelim eqd; try constructor; depelim c; subst. + - constructor. now apply IHcs. + - constructor. eauto. + Qed. + Import PCUICConversion. Lemma arity_typing_spine {Γ Γ' s inst s'} : diff --git a/pcuic/theories/Syntax/PCUICClosed.v b/pcuic/theories/Syntax/PCUICClosed.v index 7f970a044..59de9a508 100644 --- a/pcuic/theories/Syntax/PCUICClosed.v +++ b/pcuic/theories/Syntax/PCUICClosed.v @@ -639,32 +639,6 @@ Qed. #[global] Hint Rewrite to_extended_list_k_length : len. -Lemma smash_context_subst Δ s n Γ : smash_context (subst_context s (n + #|Γ|) Δ) (subst_context s n Γ) = - subst_context s n (smash_context Δ Γ). -Proof. - revert Δ. induction Γ as [|[na [b|] ty]]; intros Δ; simpl; auto. - - now rewrite Nat.add_0_r. - - rewrite -IHΓ. - rewrite subst_context_snoc /=. f_equal. - rewrite !subst_context_alt !mapi_compose. - apply mapi_ext=> n' x. - destruct x as [na' [b'|] ty']; simpl. - * rewrite !mapi_length /subst_decl /= /map_decl /=; f_equal. - + rewrite Nat.add_0_r distr_subst_rec. simpl. lia_f_equal. - + rewrite Nat.add_0_r distr_subst_rec; simpl. lia_f_equal. - * rewrite !mapi_length /subst_decl /= /map_decl /=; f_equal. - rewrite Nat.add_0_r distr_subst_rec /=. lia_f_equal. - - rewrite -IHΓ. - rewrite subst_context_snoc /= // /subst_decl /map_decl /=. - f_equal. - rewrite subst_context_app. simpl. - rewrite /app_context. f_equal. - + lia_f_equal. - + rewrite /subst_context // /fold_context_k /= /map_decl /=. - lia_f_equal. -Qed. - - Lemma smash_context_app_def Γ na b ty : smash_context [] (Γ ++ [{| decl_name := na; decl_body := Some b; decl_type := ty |}]) = smash_context [] (subst_context [b] 0 Γ). diff --git a/pcuic/theories/Syntax/PCUICLiftSubst.v b/pcuic/theories/Syntax/PCUICLiftSubst.v index c1b1929b6..544f7a65f 100644 --- a/pcuic/theories/Syntax/PCUICLiftSubst.v +++ b/pcuic/theories/Syntax/PCUICLiftSubst.v @@ -855,5 +855,30 @@ Proof. apply nth_error_appP. Qed. +Lemma smash_context_subst Δ s n Γ : smash_context (subst_context s (n + #|Γ|) Δ) (subst_context s n Γ) = + subst_context s n (smash_context Δ Γ). +Proof. + revert Δ. induction Γ as [|[na [b|] ty]]; intros Δ; simpl; auto. + - now rewrite Nat.add_0_r. + - rewrite -IHΓ. + rewrite subst_context_snoc /=. f_equal. + rewrite !subst_context_alt !mapi_compose. + apply mapi_ext=> n' x. + destruct x as [na' [b'|] ty']; simpl. + * rewrite !mapi_length /subst_decl /= /map_decl /=; f_equal. + + rewrite Nat.add_0_r distr_subst_rec. simpl. lia_f_equal. + + rewrite Nat.add_0_r distr_subst_rec; simpl. lia_f_equal. + * rewrite !mapi_length /subst_decl /= /map_decl /=; f_equal. + rewrite Nat.add_0_r distr_subst_rec /=. lia_f_equal. + - rewrite -IHΓ. + rewrite subst_context_snoc /= // /subst_decl /map_decl /=. + f_equal. + rewrite subst_context_app. simpl. + rewrite /app_context. f_equal. + + lia_f_equal. + + rewrite /subst_context // /fold_context_k /= /map_decl /=. + lia_f_equal. +Qed. + diff --git a/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v b/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v index 65342ccd7..38ece128d 100644 --- a/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v +++ b/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v @@ -485,6 +485,19 @@ Proof using Type. intros Hs; now eapply (typing_subst_instance_decl _ _ _ (tSort _)). 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. + intros wfΣ look isty cu. + eapply infer_typing_sort_impl with _ isty; [apply relevance_subst_opt|]. + intros Hs. + eapply (typing_subst_instance_decl _ _ _ (tSort _)) in Hs; tea. +Qed. + Lemma isArity_subst_instance u T : isArity T -> isArity (subst_instance u T). diff --git a/pcuic/theories/Typing/PCUICWeakeningTyp.v b/pcuic/theories/Typing/PCUICWeakeningTyp.v index ba9fb921d..47ee85b31 100644 --- a/pcuic/theories/Typing/PCUICWeakeningTyp.v +++ b/pcuic/theories/Typing/PCUICWeakeningTyp.v @@ -229,3 +229,13 @@ Lemma isType_lift {cf:checker_flags} {Σ : global_env_ext} {n Γ ty} Proof. apply isTypeRelOpt_lift => //. Qed. + +Lemma isTypeRel_weaken {cf:checker_flags} {Σ : global_env_ext} {Γ Δ ty rel} : + wf Σ -> wf_local Σ Γ -> + isTypeRel Σ Δ ty rel -> + isTypeRel Σ (Γ ,,, Δ) ty rel. +Proof. + intros wfΣ wfΓ [s [hrel hty]]. + exists s; split => //. + now eapply weaken_ctx. +Qed. diff --git a/safechecker/theories/PCUICSafeChecker.v b/safechecker/theories/PCUICSafeChecker.v index 2059ba736..388d453d5 100644 --- a/safechecker/theories/PCUICSafeChecker.v +++ b/safechecker/theories/PCUICSafeChecker.v @@ -841,7 +841,7 @@ Section CheckEnv. 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 PCUICRelevanceTyp.ind_arity_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; split. constructor; pcuic. eapply All_local_env_impl; eauto. @@ -1942,7 +1942,7 @@ Section CheckEnv. 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 PCUICRelevanceTyp.isTypeRel_cstr_type; eauto. + eapply PCUICInductiveInversion.isTypeRel_cstr_type; eauto. - now destruct wtinds. - destruct lets_in_constructor_types; eauto. Qed. @@ -2341,7 +2341,7 @@ End monad_Alli_nth_forall. ind_sorts := onsorts; ind_relevance_compat := eqrel; onIndices := _; |}. - - rewrite e in isty |- *; apply PCUICRelevanceTyp.ind_arity_relevant => //. apply H0. + - rewrite e in isty |- *; apply PCUICInductiveInversion.ind_arity_relevant => //. apply H0. - erewrite (abstract_env_ext_irr _ _ pf); eauto. Unshelve. eauto. Qed. From c2699f99cf703c95062603adb8fede3171690099 Mon Sep 17 00:00:00 2001 From: Yann Leray Date: Fri, 10 Jun 2022 13:58:42 +0200 Subject: [PATCH 16/17] Replace typ_or_sort with a judgment type (binary or ternary with options) --- template-coq/theories/Ast.v | 34 +- template-coq/theories/BasicAst.v | 28 +- template-coq/theories/Environment.v | 5 +- template-coq/theories/EnvironmentTyping.v | 503 +++++++++++----------- template-coq/theories/EtaExpand.v | 21 +- template-coq/theories/Typing.v | 181 ++++---- template-coq/theories/TypingWf.v | 70 +-- template-coq/theories/WfAst.v | 5 +- template-coq/theories/utils.v | 2 + 9 files changed, 431 insertions(+), 418 deletions(-) diff --git a/template-coq/theories/Ast.v b/template-coq/theories/Ast.v index 03016987f..4776fcb8d 100644 --- a/template-coq/theories/Ast.v +++ b/template-coq/theories/Ast.v @@ -651,10 +651,38 @@ Definition inds ind u (l : list one_inductive_body) := end in aux (List.length l). +Module TemplateLookup := EnvironmentTyping.Lookup TemplateTerm Env. +Include TemplateLookup. + + +Inductive isTermRel (Σ : global_env) (Γ : mark_context) : term -> relevance -> Type := +| rel_Rel n rel : nth_error Γ n = Some rel -> isTermRel Σ Γ (tRel n) rel +| rel_Lambda na A t rel : isTermRel Σ (Γ ,, na.(binder_relevance)) t rel -> isTermRel Σ Γ (tLambda na A t) rel +| rel_LetIn na b B t rel : isTermRel Σ (Γ ,, na.(binder_relevance)) t rel -> isTermRel Σ Γ (tLetIn na b B t) rel +| rel_App t u rel : isTermRel Σ Γ t rel -> isTermRel Σ Γ (tApp t u) rel +| rel_Const kn u decl : declared_constant Σ kn decl -> isTermRel Σ Γ (tConst kn u) decl.(cst_relevance) +| rel_Construct ind i u mdecl idecl cdecl : + declared_constructor Σ (ind, i) mdecl idecl cdecl -> isTermRel Σ Γ (tConstruct ind i u) idecl.(ind_relevance) +| rel_Case ci p c brs : isTermRel Σ Γ (tCase ci p c brs) ci.(ci_relevance) +| rel_Proj p u mdecl idecl cdecl pdecl : + declared_projection Σ p mdecl idecl cdecl pdecl -> isTermRel Σ Γ (tProj p u) pdecl.(proj_relevance) +| rel_Fix mfix n def : + nth_error mfix n = Some def -> isTermRel Σ Γ (tFix mfix n) def.(dname).(binder_relevance) +| rel_CoFix mfix n def : + nth_error mfix n = Some def -> isTermRel Σ Γ (tCoFix mfix n) def.(dname).(binder_relevance) +| rel_Sort s : isTermRel Σ Γ (tSort s) Relevant +| rel_Prod na A B : isTermRel Σ Γ (tProd na A B) Relevant +| rel_Ind ind u : isTermRel Σ Γ (tInd ind u) Relevant. + +Derive Signature for isTermRel. +Definition isTermRelOpt Σ Γ t relopt := option_default (isTermRel Σ Γ t) relopt unit. + + Module TemplateTermUtils <: TermUtils TemplateTerm Env. Definition destArity := destArity. Definition inds := inds. +Definition isTermRelOpt := isTermRelOpt. End TemplateTermUtils. @@ -665,10 +693,8 @@ Ltac unf_term := unfold TemplateTerm.term in *; unfold TemplateTerm.tRel in *; unfold TemplateTerm.lift in *; unfold TemplateTerm.subst in *; unfold TemplateTerm.closedn in *; unfold TemplateTerm.noccur_between in *; unfold TemplateTerm.subst_instance_constr in *; - unfold TemplateTermUtils.destArity; unfold TemplateTermUtils.inds. - -Module TemplateLookup := EnvironmentTyping.Lookup TemplateTerm Env. -Include TemplateLookup. + unfold TemplateTermUtils.destArity; unfold TemplateTermUtils.inds; + unfold TemplateTermUtils.isTermRelOpt. Definition tDummy := tVar "". diff --git a/template-coq/theories/BasicAst.v b/template-coq/theories/BasicAst.v index 275432f95..c60379751 100644 --- a/template-coq/theories/BasicAst.v +++ b/template-coq/theories/BasicAst.v @@ -212,21 +212,21 @@ Proof. eapply map_def_spec; eauto. Qed. -Variant typ_or_sort_ {term} := Typ (T : term) | Sort (relopt : option relevance). -Arguments typ_or_sort_ : clear implicits. - -Notation SortRel rel := (Sort (Some rel)). - -Definition typ_or_sort_map {T T'} (f: T -> T') t := - match t with - | Typ T => Typ (f T) - | Sort relopt => Sort relopt - end. - -Definition typ_or_sort_default {T A} (f: T -> A) t d := +Variant judgment_ {term} := + | Typ (t T : term) + | TripleOpt (t : option term) (T : term) (relopt : option relevance). +Arguments judgment_ : clear implicits. + +Notation Triple t T := (TripleOpt (Some t) T None). +Notation TripleRel t T rel := (TripleOpt (Some t) T (Some rel)). +Notation TripleRelOpt t T rel := (TripleOpt t T (Some rel)). +Notation Sort T := (TripleOpt None T None). +Notation SortRel T rel := (TripleOpt None T (Some rel)). + +Definition judgment_map {T T'} (f: T -> T') t := match t with - | Typ T => f T - | Sort _ => d + | Typ t T => Typ (f t) (f T) + | TripleOpt t T relopt => TripleOpt (option_map f t) (f T) relopt end. Section Contexts. diff --git a/template-coq/theories/Environment.v b/template-coq/theories/Environment.v index 89a22ac03..f38e5c4cd 100644 --- a/template-coq/theories/Environment.v +++ b/template-coq/theories/Environment.v @@ -30,7 +30,7 @@ Module Environment (T : Term). Import T. #[global] Existing Instance subst_instance_constr. - Definition typ_or_sort := typ_or_sort_ term. + Definition judgment := judgment_ term. (** ** Declarations *) Notation context_decl := (context_decl term). @@ -48,6 +48,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 *) @@ -725,5 +727,6 @@ Module Type TermUtils (T: Term) (E: EnvironmentSig T). Parameter Inline destArity : context -> term -> option (context × Universe.t). Parameter Inline inds : kername -> Instance.t -> list one_inductive_body -> list term. + Parameter Inline isTermRelOpt : global_env -> mark_context -> term -> option relevance -> Type. End TermUtils. diff --git a/template-coq/theories/EnvironmentTyping.v b/template-coq/theories/EnvironmentTyping.v index 507fcea58..ee9e2d0d6 100644 --- a/template-coq/theories/EnvironmentTyping.v +++ b/template-coq/theories/EnvironmentTyping.v @@ -28,32 +28,32 @@ Module Lookup (T : Term) (E : EnvironmentSig T). declared_constructor Σ (proj.(proj_ind), 0) mdecl idecl cdecl /\ List.nth_error idecl.(ind_projs) proj.(proj_arg) = Some pdecl /\ mdecl.(ind_npars) = proj.(proj_npars). - - Definition lookup_constant Σ kn := + + Definition lookup_constant Σ kn := match lookup_env Σ kn with | Some (ConstantDecl d) => Some d | _ => None end. - + Definition lookup_minductive Σ mind := match lookup_env Σ mind with | Some (InductiveDecl decl) => Some decl | _ => None end. - + Definition lookup_inductive Σ ind := match lookup_minductive Σ (inductive_mind ind) with - | Some mdecl => + | Some mdecl => match nth_error mdecl.(ind_bodies) (inductive_ind ind) with | Some idecl => Some (mdecl, idecl) | None => None end | None => None end. - + Definition lookup_constructor Σ ind k := match lookup_inductive Σ ind with - | Some (mdecl, idecl) => + | Some (mdecl, idecl) => match nth_error idecl.(ind_ctors) k with | Some cdecl => Some (mdecl, idecl, cdecl) | None => None @@ -63,14 +63,14 @@ Module Lookup (T : Term) (E : EnvironmentSig T). Definition lookup_projection Σ p := match lookup_constructor Σ p.(proj_ind) 0 with - | Some (mdecl, idecl, cdecl) => + | Some (mdecl, idecl, cdecl) => match nth_error idecl.(ind_projs) p.(proj_arg) with | Some pdecl => Some (mdecl, idecl, cdecl, pdecl) | None => None end | _ => None end. - + Lemma declared_constant_lookup {Σ kn cdecl} : declared_constant Σ kn cdecl -> lookup_constant Σ kn = Some cdecl. @@ -79,13 +79,13 @@ Module Lookup (T : Term) (E : EnvironmentSig T). Qed. Lemma lookup_constant_declared {Σ kn cdecl} : - lookup_constant Σ kn = Some cdecl -> + lookup_constant Σ kn = Some cdecl -> declared_constant Σ kn cdecl. Proof. unfold declared_constant, lookup_constant. destruct lookup_env as [[]|] => //. congruence. Qed. - + Lemma declared_minductive_lookup {Σ ind mdecl} : declared_minductive Σ ind mdecl -> lookup_minductive Σ ind = Some mdecl. @@ -93,7 +93,7 @@ Module Lookup (T : Term) (E : EnvironmentSig T). rewrite /declared_minductive /lookup_minductive. now intros ->. Qed. - + Lemma lookup_minductive_declared {Σ ind mdecl} : lookup_minductive Σ ind = Some mdecl -> declared_minductive Σ ind mdecl. @@ -101,7 +101,7 @@ Module Lookup (T : Term) (E : EnvironmentSig T). rewrite /declared_minductive /lookup_minductive. destruct lookup_env as [[]|] => //. congruence. Qed. - + Lemma declared_inductive_lookup {Σ ind mdecl idecl} : declared_inductive Σ ind mdecl idecl -> lookup_inductive Σ ind = Some (mdecl, idecl). @@ -109,7 +109,7 @@ Module Lookup (T : Term) (E : EnvironmentSig T). rewrite /declared_inductive /lookup_inductive. intros []. now rewrite (declared_minductive_lookup H) H0. Qed. - + Lemma lookup_inductive_declared {Σ ind mdecl idecl} : lookup_inductive Σ ind = Some (mdecl, idecl) -> declared_inductive Σ ind mdecl idecl. @@ -122,7 +122,7 @@ Module Lookup (T : Term) (E : EnvironmentSig T). Qed. Lemma declared_constructor_lookup {Σ id mdecl idecl cdecl} : - declared_constructor Σ id mdecl idecl cdecl -> + declared_constructor Σ id mdecl idecl cdecl -> lookup_constructor Σ id.1 id.2 = Some (mdecl, idecl, cdecl). Proof. intros []. unfold lookup_constructor. @@ -141,13 +141,13 @@ Module Lookup (T : Term) (E : EnvironmentSig T). Qed. Lemma declared_projection_lookup {Σ p mdecl idecl cdecl pdecl} : - declared_projection Σ p mdecl idecl cdecl pdecl -> + declared_projection Σ p mdecl idecl cdecl pdecl -> lookup_projection Σ p = Some (mdecl, idecl, cdecl, pdecl). Proof. intros [? []]. unfold lookup_projection. rewrite (declared_constructor_lookup (Σ := Σ) H) /= H0 //. Qed. - + Lemma lookup_projection_declared {Σ p mdecl idecl cdecl pdecl} : ind_npars mdecl = p.(proj_npars) -> lookup_projection Σ p = Some (mdecl, idecl, cdecl, pdecl) -> @@ -165,7 +165,7 @@ Module Lookup (T : Term) (E : EnvironmentSig T). | ConstantDecl cb => F cb.(cst_universes) | InductiveDecl mb => F mb.(ind_universes) end. - + Definition universes_decl_of_decl := on_udecl_decl (fun x => x). (* Definition LevelSet_add_list l := LevelSet.union (LevelSetProp.of_list l). *) @@ -179,7 +179,7 @@ Module Lookup (T : Term) (E : EnvironmentSig T). apply LevelSet.union_spec; right. now apply LevelSet.singleton_spec. Qed. - + Lemma global_levels_memSet univs : LevelSet.mem Level.lzero (global_levels univs) = true. Proof. @@ -237,10 +237,10 @@ Module Lookup (T : Term) (E : EnvironmentSig T). Definition consistent_instance_ext `{checker_flags} Σ := consistent_instance (global_ext_levels Σ) (global_ext_constraints Σ). - + Lemma consistent_instance_length {cf : checker_flags} {Σ : global_env_ext} {univs u} : consistent_instance_ext Σ univs u -> - #|u| = #|abstract_instance univs|. + #|u| = #|abstract_instance univs|. Proof. unfold consistent_instance_ext, consistent_instance. destruct univs; simpl; auto. @@ -253,7 +253,7 @@ Module Lookup (T : Term) (E : EnvironmentSig T). Universe.on_sort (fun u => forall l, LevelExprSet.In l u -> LevelSet.In (LevelExpr.get_level l) (global_ext_levels Σ)) True s. - + End Lookup. Module Type LookupSig (T : Term) (E : EnvironmentSig T). @@ -265,7 +265,7 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). Import T E TU. Section TypeLocal. - Context (typing : forall (Γ : context), term -> typ_or_sort -> Type). + Context (typing : forall (Γ : context), judgment -> Type). Inductive All_local_env : context -> Type := | localenv_nil : @@ -273,13 +273,12 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). | localenv_cons_abs Γ na t : All_local_env Γ -> - typing Γ t (SortRel na.(binder_relevance)) -> + typing Γ (SortRel t na.(binder_relevance)) -> All_local_env (Γ ,, vass na t) | localenv_cons_def Γ na b t : All_local_env Γ -> - typing Γ t (SortRel na.(binder_relevance)) -> - typing Γ b (Typ t) -> + typing Γ (TripleRel b t na.(binder_relevance)) -> All_local_env (Γ ,, vdef na b t). Derive Signature NoConfusion for All_local_env. End TypeLocal. @@ -289,7 +288,7 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). Arguments localenv_cons_abs {_ _ _ _} _ _. Lemma All_local_env_fold P f Γ : - All_local_env (fun Γ t T => P (fold_context_k f Γ) (f #|Γ| t) (typ_or_sort_map (f #|Γ|) T)) Γ <~> + All_local_env (fun Γ T => P (fold_context_k f Γ) (judgment_map (f #|Γ|) T)) Γ <~> All_local_env P (fold_context_k f Γ). Proof. split. @@ -299,46 +298,45 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). * destruct a as [na [b|] ty]; depelim H; specialize (IHΓ H); constructor; simpl; auto. Qed. - Lemma All_local_env_impl (P Q : context -> term -> typ_or_sort -> Type) l : + Lemma All_local_env_impl (P Q : context -> judgment -> Type) l : All_local_env P l -> - (forall Γ t T, P Γ t T -> Q Γ t T) -> + (forall Γ T, P Γ T -> Q Γ T) -> All_local_env Q l. Proof. induction 1; intros; simpl; econstructor; eauto. Qed. - Lemma All_local_env_impl_ind {P Q : context -> term -> typ_or_sort -> Type} {l} : + Lemma All_local_env_impl_ind {P Q : context -> judgment -> Type} {l} : All_local_env P l -> - (forall Γ t T, All_local_env Q Γ -> P Γ t T -> Q Γ t T) -> + (forall Γ T, All_local_env Q Γ -> P Γ T -> Q Γ T) -> All_local_env Q l. Proof. induction 1; intros; simpl; econstructor; eauto. Qed. - + Lemma All_local_env_skipn P Γ : All_local_env P Γ -> forall n, All_local_env P (skipn n Γ). Proof. induction 1; simpl; intros; destruct n; simpl; try econstructor; eauto. Qed. - #[global] + #[global] Hint Resolve All_local_env_skipn : wf. Section All_local_env_rel. Definition All_local_rel P Γ Γ' - := (All_local_env (fun Γ0 t T => P (Γ ,,, Γ0) t T) Γ'). + := (All_local_env (fun Γ0 T => P (Γ ,,, Γ0) T) Γ'). Definition All_local_rel_nil {P Γ} : All_local_rel P Γ [] := localenv_nil. Definition All_local_rel_abs {P Γ Γ' A na} : - All_local_rel P Γ Γ' -> P (Γ ,,, Γ') A (SortRel na.(binder_relevance)) + All_local_rel P Γ Γ' -> P (Γ ,,, Γ') (SortRel A na.(binder_relevance)) -> All_local_rel P Γ (Γ',, vass na A) := localenv_cons_abs. Definition All_local_rel_def {P Γ Γ' t A na} : All_local_rel P Γ Γ' -> - P (Γ ,,, Γ') A (SortRel na.(binder_relevance)) -> - P (Γ ,,, Γ') t (Typ A) -> + P (Γ ,,, Γ') (TripleRel t A na.(binder_relevance)) -> All_local_rel P Γ (Γ',, vdef na t A) := localenv_cons_def. @@ -349,7 +347,7 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). Proof. intros P Γ h. eapply All_local_env_impl. - exact h. - - intros Δ t [] h'. + - intros Δ [] h'. all: cbn. + rewrite app_context_nil_l. assumption. + rewrite app_context_nil_l. assumption. @@ -359,7 +357,7 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). All_local_rel P [] Γ -> All_local_env P Γ. Proof. intro X. eapply All_local_env_impl. exact X. - intros Γ0 t [] XX; cbn in XX; rewrite app_context_nil_l in XX; assumption. + intros Γ0 [] XX; cbn in XX; rewrite app_context_nil_l in XX; assumption. Defined. Lemma All_local_app_rel {P Γ Γ'} : @@ -392,22 +390,9 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). (** Well-formedness of local environments embeds a sorting for each variable *) - Definition on_local_decl (P : context -> term -> typ_or_sort -> Type) Γ d := - match d.(decl_body) with - | Some b => P Γ b (Typ d.(decl_type)) - | None => P Γ d.(decl_type) (SortRel d.(decl_name).(binder_relevance)) - end. - - Lemma All_local_env_inv P (d : context_decl) (Γ : context) (X : All_local_env P (Γ ,, d)) : - on_local_decl P Γ d * All_local_env P Γ. - Proof. - inv X; intuition; red; simpl; eauto. - Qed. + Definition on_decl (P : context -> judgment -> Type) Γ d := + P Γ (TripleOpt d.(decl_body) d.(decl_type) (Some d.(decl_name).(binder_relevance))). - Definition on_decl (P : context -> term -> typ_or_sort -> Type) Γ d := - P Γ d.(decl_type) (SortRel d.(decl_name).(binder_relevance)) × - option_default (fun body => P Γ body (Typ d.(decl_type))) d.(decl_body) True. - Lemma nth_error_All_local_env {P Γ n d} : nth_error Γ n = Some d -> All_local_env P Γ -> @@ -417,58 +402,56 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). induction 1 in n, Heq |- *. - destruct n; simpl in Heq; discriminate. - destruct n. - + inv Heq. simpl. split => //. + + inv Heq. simpl => //. + simpl in Heq. simpl. apply IHX => //. - destruct n. - + inv Heq. simpl. split => //. + + inv Heq. simpl => //. + simpl in Heq. simpl. apply IHX => //. Qed. - Definition on_def_type (P : context -> term -> typ_or_sort -> Type) Γ d := - P Γ d.(dtype) (SortRel d.(dname).(binder_relevance)). - - Definition on_def_body (P : context -> term -> typ_or_sort -> Type) types Γ d := - P (Γ ,,, types) d.(dbody) (Typ (lift0 #|types| d.(dtype))). + Lemma All_local_env_inv P (d : context_decl) (Γ : context) (X : All_local_env P (Γ ,, d)) : + on_decl P Γ d * All_local_env P Γ. + Proof. + inv X; intuition; red; simpl; eauto. + Qed. + + Definition on_def (P : context -> judgment -> Type) types Γ d := + P (Γ ,,, types) (TripleRel d.(dbody) (lift0 #|types| d.(dtype)) d.(dname).(binder_relevance)). Definition lift_judgment - (check : context -> term -> term -> Type) - (infer_sort : context -> term -> option relevance -> Type) : - (context -> term -> typ_or_sort -> Type) := - fun Γ t T => + (check : term -> term -> Type) + (triple : option term -> term -> option relevance -> Type) : + judgment -> Type := + fun T => match T with - | Typ T => check Γ t T - | Sort relopt => infer_sort Γ t relopt + | Typ t T => check t T + | TripleOpt t T relopt => triple t T relopt end. - - Lemma lift_judgment_impl {P Ps Q Qs Γ Γ' t t' T} : - lift_judgment P Ps Γ t T -> - (forall T, P Γ t T -> Q Γ' t' T) -> - (forall relopt, Ps Γ t relopt -> Qs Γ' t' relopt) -> - lift_judgment Q Qs Γ' t' T. - Proof. - intros HT HPQ HPsQs. - destruct T; simpl. - * apply HPQ, HT. - * apply HPsQs, HT. - Qed. (* Common uses *) - Definition lift_wf_term wf_term := (lift_judgment (fun Γ t T => wf_term Γ t × wf_term Γ T) (fun Γ t relopt => wf_term Γ t)). + Definition lift_wf_term wf_term := (lift_judgment (fun t T => wf_term t × wf_term T) (fun t T relopt => option_default wf_term t unit × wf_term T)). + Definition lift_on_term on_term := (fun (Γ : context) => lift_judgment (fun t T => on_term Γ t × on_term Γ T) (fun t T relopt => option_default (on_term Γ) t unit × on_term Γ T)). - Definition infer_sort (sorting : global_env_ext -> context -> term -> Universe.t -> Type) := - (fun Σ Γ T relopt => { s : Universe.t & isSortRelOpt s relopt × sorting Σ Γ T s }). - Notation typing_sort typing := (fun Σ Γ T s => typing Σ Γ T (tSort s)). + Notation infer_sort sorting := (fun Σ Γ T relopt => { s : Universe.t & isSortRelOpt s relopt × sorting Σ Γ T s }). + + Definition on_triple (typing : global_env_ext -> context -> term -> term -> Type) sorting := (fun (Σ : global_env_ext) Γ t T relopt => + option_default (fun t => typing Σ Γ t T × isTermRelOpt Σ (marks_of_context Γ) t relopt) t unit × infer_sort sorting Σ Γ T relopt). - Definition lift_typing typing Σ := lift_judgment (typing Σ) (infer_sort (typing_sort typing) Σ). - Definition lift_sorting checking sorting Σ := lift_judgment (checking Σ) (infer_sort sorting Σ). + Definition lift_relation (typing : global_env_ext -> context -> term -> term -> Type) (sorting : global_env_ext -> context -> term -> Universe.t -> Type) Σ Γ := + lift_judgment (typing Σ Γ) (on_triple typing sorting Σ Γ). + Notation typing_sort typing := (fun Σ Γ T s => typing Σ Γ T (tSort s)). + + Definition lift_typing typing Σ := lift_relation typing (typing_sort typing) Σ. + Definition lift_sorting checking sorting Σ := lift_relation checking sorting Σ. + Notation Prop_conj P Q := (fun Σ Γ t T => P Σ Γ t T × Q Σ Γ t T). Definition lift_typing2 P Q := lift_typing (Prop_conj P Q). Lemma infer_sort_impl {P Q} {Σ Σ' : global_env_ext} {Γ Γ' : context} {t t' : term} {relopt} : - forall f, (forall s r, isSortRelOpt s r -> isSortRelOpt (f s) r) -> + forall f, (forall s, isSortRelOpt s relopt -> isSortRelOpt (f s) relopt) -> forall tu: infer_sort P Σ Γ t relopt, let s := tu.π1 in (P Σ Γ t s -> Q Σ' Γ' t' (f s)) -> @@ -478,7 +461,7 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). Qed. Lemma infer_typing_sort_impl {P Q} {Σ Σ' : global_env_ext} {Γ Γ' : context} {t t' : term} {relopt} : - forall f (Hf : forall s r, isSortRelOpt s r -> isSortRelOpt (f s) r), + forall f (Hf : forall s, isSortRelOpt s relopt -> isSortRelOpt (f s) relopt), forall tu: infer_sort (typing_sort P) Σ Γ t relopt, let s := tu.π1 in (P Σ Γ t (tSort s) -> Q Σ' Γ' t' (tSort (f s))) -> @@ -487,80 +470,104 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). apply (infer_sort_impl (P := typing_sort P) (Q := typing_sort Q)). Qed. - Lemma lift_typing_impl {P Q Σ Σ' Γ Γ' t t' T} : - lift_typing P Σ Γ t T -> - (forall T, P Σ Γ t T -> Q Σ' Γ' t' T) -> - lift_typing Q Σ' Γ' t' T. + Lemma on_triple_impl_id {P Ps Q Qs} {Σ Σ' : global_env_ext} {Γ Γ' : context} {t : option term} {T : term} {relopt} : + forall (Htr : forall t, isTermRelOpt Σ (marks_of_context Γ) t relopt -> isTermRelOpt Σ' (marks_of_context Γ') t relopt), + forall tu: on_triple P Ps Σ Γ t T relopt, + (forall t, P Σ Γ t T -> Q Σ' Γ' t T) -> + let s := tu.2.π1 in + (Ps Σ Γ T s -> Qs Σ' Γ' T s) -> + on_triple Q Qs Σ' Γ' t T relopt. + Proof. + intros Htr [Ht HT] HPQ s HPQs. + split. + - destruct t => //=. + destruct Ht; split; auto. + - apply infer_sort_impl with id HT => //. + Qed. + + Lemma lift_relation_impl {P Ps Q Qs Σ Γ T} : + lift_relation P Ps Σ Γ T -> + (forall t T, P Σ Γ t T -> Q Σ Γ t T) -> + (forall t s, Ps Σ Γ t s -> Qs Σ Γ t s) -> + lift_relation Q Qs Σ Γ T. + Proof. + intros HT HPQ HPsQs. + destruct T => /=; auto. + eapply on_triple_impl_id with HT => //; auto. + Qed. + + Lemma lift_typing_impl {P Q Σ Γ T} : + lift_typing P Σ Γ T -> + (forall t T, P Σ Γ t T -> Q Σ Γ t T) -> + lift_typing Q Σ Γ T. Proof. intros HT HPQ. - apply lift_judgment_impl with (1 := HT); tas. - intros relopt Hs; apply infer_typing_sort_impl with id Hs => //. apply HPQ. + eapply lift_relation_impl with (1 := HT) => //; auto. Qed. Section TypeLocalOver. - Context (checking : context -> term -> term -> Type). - Context (sorting : context -> term -> option relevance -> Type). - Context (cproperty : forall (Γ : context), - All_local_env (lift_judgment checking sorting) Γ -> - forall (t T : term), checking Γ t T -> Type). - Context (sproperty : forall (Γ : context), - All_local_env (lift_judgment checking sorting) Γ -> - forall (t : term) (relopt : option relevance), sorting Γ t relopt -> Type). - - Inductive All_local_env_over_gen : - forall (Γ : context), All_local_env (lift_judgment checking sorting) Γ -> Type := + Context (checking : global_env_ext -> context -> term -> term -> Type). + Context (sorting : global_env_ext -> context -> term -> Universe.t -> Type). + Context (cproperty : forall (Σ : global_env_ext) (Γ : context), + All_local_env (lift_relation checking sorting Σ) Γ -> + forall (t T : term), checking Σ Γ t T -> Type). + Context (sproperty : forall (Σ : global_env_ext) (Γ : context), + All_local_env (lift_relation checking sorting Σ) Γ -> + forall (t : term) (s : Universe.t), sorting Σ Γ t s -> Type). + + Inductive All_local_env_over_gen Σ : + forall (Γ : context), All_local_env (lift_relation checking sorting Σ) Γ -> Type := | localenv_over_nil : - All_local_env_over_gen [] localenv_nil + All_local_env_over_gen Σ [] localenv_nil | localenv_over_cons_abs Γ na t - (all : All_local_env (lift_judgment checking sorting) Γ) : - All_local_env_over_gen Γ all -> - forall (tu : lift_judgment checking sorting Γ t (SortRel na.(binder_relevance))) - (Hs: sproperty Γ all _ _ tu), - All_local_env_over_gen (Γ ,, vass na t) + (all : All_local_env (lift_relation checking sorting Σ) Γ) : + All_local_env_over_gen Σ Γ all -> + forall (tu : lift_relation checking sorting Σ Γ (SortRel t na.(binder_relevance))) + (Hs: sproperty Σ Γ all _ _ tu.2.π2.2), + All_local_env_over_gen Σ (Γ ,, vass na t) (localenv_cons_abs all tu) | localenv_over_cons_def Γ na b t - (all : All_local_env (lift_judgment checking sorting) Γ) (tb : checking Γ b t) : - All_local_env_over_gen Γ all -> - forall (Hc: cproperty Γ all _ _ tb), - forall (tu : lift_judgment checking sorting Γ t (SortRel na.(binder_relevance))) - (Hs: sproperty Γ all _ _ tu), - All_local_env_over_gen (Γ ,, vdef na b t) - (localenv_cons_def all tu tb). + (all : All_local_env (lift_relation checking sorting Σ) Γ) : + All_local_env_over_gen Σ Γ all -> + forall (tu : lift_relation checking sorting Σ Γ (TripleRel b t na.(binder_relevance))) + (Hc: cproperty Σ Γ all _ _ tu.1.1) + (Hs: sproperty Σ Γ all _ _ tu.2.π2.2), + All_local_env_over_gen Σ (Γ ,, vdef na b t) + (localenv_cons_def all tu). End TypeLocalOver. Derive Signature for All_local_env_over_gen. - Notation Proploc_conj P Q := (fun Γ t T => P Γ t T × Q Γ t T). - - Lemma All_local_env_over_gen_2 checking sorting cproperty sproperty Γ wfΓ : - let cproperty_full Γ wfΓ t T check := cproperty Γ t T in - let sproperty_full Γ wfΓ t r check := sproperty Γ t r check in - All_local_env_over_gen checking sorting cproperty_full sproperty_full Γ wfΓ -> - All_local_env (lift_judgment (Proploc_conj checking cproperty) (fun Γ t r => { s : sorting Γ t r & sproperty Γ t r s })) Γ. + Lemma All_local_env_over_gen_2 checking sorting cproperty sproperty Σ Γ wfΓ : + let cproperty_full Σ Γ wfΓ t T check := cproperty Σ Γ t T in + let sproperty_full Σ Γ wfΓ t s check := sproperty Σ Γ t s in + All_local_env_over_gen checking sorting cproperty_full sproperty_full Σ Γ wfΓ -> + All_local_env (lift_relation (Prop_conj checking cproperty) (Prop_conj sorting sproperty) Σ) Γ. Proof. intros cfull sfull. induction 1; constructor => //. - all: cbn in tu |- *; eauto. + all: destruct tu as (H & s & ? & ?); cbn in *. + all: repeat split. + all: eauto. + all: now destruct H. Qed. - Definition All_local_env_over typing property Σ := - (All_local_env_over_gen (typing Σ) (infer_sort (typing_sort typing) Σ) (property Σ) (fun Γ H t relopt tu => property _ _ H _ _ tu.π2.2)). - - Definition All_local_env_over_sorting checking sorting cproperty (sproperty : forall Σ Γ _ t s, sorting Σ Γ t s -> Type) Σ := - (All_local_env_over_gen (checking Σ) (infer_sort sorting Σ) (cproperty Σ) (fun Γ H t relopt tu => sproperty _ Γ H t tu.π1 tu.π2.2)). + Definition All_local_env_over typing property := + (All_local_env_over_gen typing (typing_sort typing) property (fun Σ Γ H t s tu => property Σ Γ H t (tSort s) tu)). + + Definition All_local_env_over_sorting := All_local_env_over_gen. Lemma All_local_env_over_2 typing property (Σ : global_env_ext) (Γ : context) (wfΓ : All_local_env (lift_typing typing Σ) Γ) : let property_full Σ Γ wfΓ t T Hty := property Σ Γ t T in All_local_env_over typing property_full Σ Γ wfΓ -> All_local_env (lift_typing2 typing property Σ) Γ. - Proof. + Proof. intros full; unfold full, All_local_env_over. intro H; eapply All_local_env_over_gen_2 in H. apply All_local_env_impl with (1 := H); intros. - apply lift_judgment_impl with (1 := X); intros => //. - destruct X0 as ((s & e & Hs) & Hs'); exists s; now repeat split. + apply lift_relation_impl with (1 := X); intros => //. Qed. Section TypeCtxInst. @@ -569,7 +576,7 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). (* Γ |- s : Δ, where Δ is a telescope (reverse context) *) Inductive ctx_inst (Γ : context) : list term -> context -> Type := | ctx_inst_nil : ctx_inst Γ [] [] - | ctx_inst_ass na t i inst Δ : + | ctx_inst_ass na t i inst Δ : typing Γ i t -> ctx_inst Γ inst (subst_telescope [i] 0 Δ) -> ctx_inst Γ (i :: inst) (vass na t :: Δ) @@ -579,7 +586,7 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). Derive Signature NoConfusion for ctx_inst. End TypeCtxInst. - Lemma ctx_inst_impl P Q Γ inst Δ : + Lemma ctx_inst_impl P Q Γ inst Δ : ctx_inst P Γ inst Δ -> (forall t T, P Γ t T -> Q Γ t T) -> ctx_inst Q Γ inst Δ. @@ -588,23 +595,23 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). Qed. Section All_local_env_size. - Context {P : forall (Γ : context), term -> typ_or_sort -> Type}. - Context (Psize : forall Γ t T, P Γ t T -> size). + Context {P : forall (Γ : context), judgment -> Type}. + Context (Psize : forall Γ j, P Γ j -> size). Fixpoint All_local_env_size_gen base Γ (w : All_local_env P Γ) : size := match w with | localenv_nil => base - | localenv_cons_abs Γ' na t w' p => Psize _ _ _ p + All_local_env_size_gen base _ w' - | localenv_cons_def Γ' na b t w' pt pb => Psize _ _ _ pt + Psize _ _ _ pb + All_local_env_size_gen base _ w' + | localenv_cons_abs Γ' na t w' p => Psize _ _ p + All_local_env_size_gen base _ w' + | localenv_cons_def Γ' na b t w' p => Psize _ _ p + All_local_env_size_gen base _ w' end. - + Lemma All_local_env_size_pos base Γ w : base <= All_local_env_size_gen base Γ w. Proof using Type. induction w. all: simpl ; lia. Qed. End All_local_env_size. - + Notation ctx_shifted P Γ := (fun Δ => P (Γ ,,, Δ)). Notation All_local_rel_size_gen Psize base := (fun Γ Δ (w : All_local_rel _ Γ Δ) => All_local_env_size_gen (ctx_shifted Psize Γ) base Δ w). @@ -629,28 +636,28 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). Qed. Section lift_judgment_size. - Context {checking : context -> term -> term -> Type}. - Context {sorting : context -> term -> option relevance -> Type}. - Context (csize : forall (Γ : context) (t T : term), checking Γ t T -> size). - Context (ssize : forall (Γ : context) (t : term) (relopt : option relevance), sorting Γ t relopt -> size). - - Definition lift_judgment_size Γ t T (w : lift_judgment checking sorting Γ t T) : size := - match T return lift_judgment checking sorting Γ t T -> size with - | Typ T => csize _ _ _ - | Sort relopt => ssize _ _ _ + Context {checking : global_env_ext -> context -> term -> term -> Type}. + Context {sorting : global_env_ext -> context -> term -> Universe.t -> Type}. + Context (csize : forall (Σ : global_env_ext) (Γ : context) (t T : term), checking Σ Γ t T -> size). + Context (ssize : forall (Σ : global_env_ext) (Γ : context) (t : term) (s : Universe.t), sorting Σ Γ t s -> size). + + Definition lift_judgment_size Σ Γ j (w : (fun Γ => lift_judgment (checking Σ Γ) (on_triple checking sorting Σ Γ)) Γ j) : size := + match j return lift_judgment _ _ j -> size with + | Typ t T => csize _ _ _ _ + | TripleOpt None _ _ => fun tu => ssize _ _ _ _ tu.2.π2.2 + | TripleOpt (Some b) ty _ => fun tu => csize _ _ _ _ tu.1.1 + ssize _ _ _ _ tu.2.π2.2 end w. End lift_judgment_size. Implicit Types (Σ : global_env_ext) (Γ : context) (t : term). - Notation infer_sort_size typing_size := (fun Σ Γ t r (tu: infer_sort _ Σ Γ t r) => let '(s; (e, d)) := tu in typing_size Σ Γ t s d). Notation typing_sort_size typing_size := (fun Σ Γ t s (tu: typing_sort _ Σ Γ t s) => typing_size Σ Γ t (tSort s) tu). Section Regular. Context {typing : global_env_ext -> context -> term -> term -> Type}. Context (typing_size : forall Σ Γ t T, typing Σ Γ t T -> size). - - Definition lift_typing_size Σ := lift_judgment_size (typing_size Σ) (infer_sort_size (typing_sort_size typing_size) Σ). + + Definition lift_typing_size Σ := lift_judgment_size typing_size (typing_sort_size typing_size) Σ. Definition All_local_env_size Σ := All_local_env_size_gen (lift_typing_size Σ) 0. Definition All_local_rel_size Σ := All_local_rel_size_gen (lift_typing_size Σ) 0. End Regular. @@ -660,7 +667,7 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). Context (checking_size : forall Σ Γ t T, checking Σ Γ t T -> size). Context (sorting_size : forall Σ Γ t s, sorting Σ Γ t s -> size). - Definition lift_sorting_size Σ := lift_judgment_size (checking_size Σ) (infer_sort_size sorting_size Σ). + Definition lift_sorting_size := lift_judgment_size checking_size sorting_size. Definition All_local_env_sorting_size Σ := All_local_env_size_gen (lift_sorting_size Σ) 1. Definition All_local_rel_sorting_size Σ := All_local_rel_size_gen (lift_sorting_size Σ) 1. End Bidirectional. @@ -683,10 +690,10 @@ Module Conversion (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E) (ET : E (eqna : eq_binder_annot na na') (eqt : P pb t t') : All_decls_alpha_pb (vass na t) (vass na' t') - + | all_decls_alpha_vdef {na na' : binder_annot name} {b t b' t' : term} (eqna : eq_binder_annot na na') - (eqb : P Conv b b') (* Note that definitions must be convertible, otherwise this notion + (eqb : P Conv b b') (* Note that definitions must be convertible, otherwise this notion of cumulativity is useless *) (eqt : P pb t t') : All_decls_alpha_pb (vdef na b t) (vdef na' b' t'). @@ -694,13 +701,13 @@ Module Conversion (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E) (ET : E Derive Signature NoConfusion for All_decls_alpha_pb. Arguments All_decls_alpha_pb pb P : clear implicits. - + Definition cumul_pb_decls pb (Σ : global_env_ext) (Γ Γ' : context) : forall (x y : context_decl), Type := All_decls_alpha_pb pb (cumul_gen Σ Γ). - Definition cumul_pb_context pb (Σ : global_env_ext) := + Definition cumul_pb_context pb (Σ : global_env_ext) := All2_fold (cumul_pb_decls pb Σ). - + Definition cumul_ctx_rel Σ Γ Δ Δ' := All2_fold (fun Δ Δ' => cumul_pb_decls Cumul Σ (Γ ,,, Δ) (Γ ,,, Δ')) Δ Δ'. End Conversion. @@ -729,7 +736,7 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT Context {cf: checker_flags}. Context (Pcmp: global_env_ext -> context -> conv_pb -> term -> term -> Type). - Context (P : global_env_ext -> context -> term -> typ_or_sort -> Type). + Context (P : global_env_ext -> context -> judgment -> Type). Definition on_context Σ ctx := All_local_env (P Σ) ctx. @@ -741,25 +748,25 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT match Δ with | [] => wf_universe Σ u | {| decl_name := na; decl_body := None; decl_type := t |} :: Δ => - type_local_ctx Σ Γ Δ u × isSortRel u na.(binder_relevance) × P Σ (Γ ,,, Δ) t (Typ (tSort u)) + type_local_ctx Σ Γ Δ u × isSortRel u na.(binder_relevance) × P Σ (Γ ,,, Δ) (Typ t (tSort u)) | {| decl_name := na; decl_body := Some b; decl_type := t |} :: Δ => - type_local_ctx Σ Γ Δ u × P Σ (Γ ,,, Δ) t (SortRel na.(binder_relevance)) × P Σ (Γ ,,, Δ) b (Typ t) + type_local_ctx Σ Γ Δ u × P Σ (Γ ,,, Δ) (SortRel t na.(binder_relevance)) × P Σ (Γ ,,, Δ) (Typ b t) end. Fixpoint sorts_local_ctx Σ (Γ Δ : context) (us : list Universe.t) : Type := match Δ, us with | [], [] => unit - | {| decl_name := na; decl_body := None; decl_type := t |} :: Δ, u :: us => - sorts_local_ctx Σ Γ Δ us × isSortRel u na.(binder_relevance) × P Σ (Γ ,,, Δ) t (Typ (tSort u)) - | {| decl_name := na; decl_body := Some b; decl_type := t |} :: Δ, us => - sorts_local_ctx Σ Γ Δ us × P Σ (Γ ,,, Δ) t (SortRel na.(binder_relevance)) × P Σ (Γ ,,, Δ) b (Typ t) + | {| decl_name := na; decl_body := None; decl_type := t |} :: Δ, u :: us => + sorts_local_ctx Σ Γ Δ us × isSortRel u na.(binder_relevance) × P Σ (Γ ,,, Δ) (Typ t (tSort u)) + | {| decl_name := na; decl_body := Some b; decl_type := t |} :: Δ, us => + sorts_local_ctx Σ Γ Δ us × P Σ (Γ ,,, Δ) (SortRel t na.(binder_relevance)) × P Σ (Γ ,,, Δ) (Typ b t) | _, _ => False end. Implicit Types (mdecl : mutual_inductive_body) (idecl : one_inductive_body) (cdecl : constructor_body). - Definition on_type Σ Γ T := P Σ Γ T (Sort None). - Definition on_type_rel Σ Γ T rel := P Σ Γ T (SortRel rel). + Definition on_type Σ Γ T := P Σ Γ (Sort T). + Definition on_type_rel Σ Γ T rel := P Σ Γ (SortRel T rel). Open Scope type_scope. @@ -786,7 +793,7 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT /\ satisfiable_udecl univs udecl /\ valid_on_mono_udecl univs udecl. - (** Positivity checking of the inductive, ensuring that the inductive itself + (** Positivity checking of the inductive, ensuring that the inductive itself can only appear at the right of an arrow in each argument's types. *) (* Definition positive_cstr_arg ninds npars narg (arg : term) : bool := @@ -798,14 +805,14 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT alli (fun i d => noccur_between (npars + narg + i) ninds d.(decl_type)) 0 (List.rev ctx) && let (hd, args) := decompose_app concl in match hd with - | tRel i => - if noccur_between (npars + narg + #|ctx|) ninds (tRel i) then + | tRel i => + if noccur_between (npars + narg + #|ctx|) ninds (tRel i) then (* Call to an unrelated variable *) true else (* Recursive call to the inductive *) (* Coq disallows the inductive to be applied to another inductive in the block *) forallb (noccur_between (npars + narg + #|ctx|) ninds) args - | tInd ind u => + | tInd ind u => if forallb (noccur_between (npars + narg + #|ctx|) ninds) args then (* Unrelated inductive *) true @@ -815,36 +822,36 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT Definition positive_cstr_args ninds npars (args : context) : bool := alli (fun i decl => positive_cstr_arg nind npars i decl.(decl_type)) - (* We smash the context, just as Coq's kernel computes positivity on + (* We smash the context, just as Coq's kernel computes positivity on weak-head normalized types *) (List.rev (smash_context [] args)) *) (** A constructor argument type [t] is positive w.r.t. an inductive block [mdecl] - when it's zeta-normal form is of the shape Π Δ. concl and: + when it's zeta-normal form is of the shape Π Δ. concl and: - [t] does not refer to any inductive in the block. In that case [t] must be a closed type under the context of parameters and previous arguments. - - None of the variable assumptions in Δ refer to any inductive in the block, - but the conclusion [concl] is of the form [mkApps (tRel k) args] for k + - None of the variable assumptions in Δ refer to any inductive in the block, + but the conclusion [concl] is of the form [mkApps (tRel k) args] for k refering to an inductive in the block, and none of the arguments [args] - refer to the inductive. #|args| must be the length of the full inductive application. - + refer to the inductive. #|args| must be the length of the full inductive application. + Let-in assumptions in Δ are systematically unfolded, i.e. we really consider: the zeta-reduction of [t]. *) - - Definition ind_realargs (o : one_inductive_body) := + + Definition ind_realargs (o : one_inductive_body) := match destArity [] o.(ind_type) with | Some (ctx, _) => #|smash_context [] ctx| | _ => 0 end. Inductive positive_cstr_arg mdecl ctx : term -> Type := - | positive_cstr_arg_closed t : + | positive_cstr_arg_closed t : closedn #|ctx| t -> positive_cstr_arg mdecl ctx t - | positive_cstr_arg_concl l k i : + | positive_cstr_arg_concl l k i : (** Mutual inductive references in the conclusion are ok *) #|ctx| <= k -> k < #|ctx| + #|mdecl.(ind_bodies)| -> All (closedn #|ctx|) l -> @@ -854,7 +861,7 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT | positive_cstr_arg_let na b ty t : positive_cstr_arg mdecl ctx (subst [b] 0 t) -> - positive_cstr_arg mdecl ctx (tLetIn na b ty t) + positive_cstr_arg mdecl ctx (tLetIn na b ty t) | positive_cstr_arg_ass na ty t : closedn #|ctx| ty -> @@ -862,23 +869,23 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT positive_cstr_arg mdecl ctx (tProd na ty t). (** A constructor type [t] is positive w.r.t. an inductive block [mdecl] - and inductive [i] when it's zeta normal-form is of the shape Π Δ. concl and: + and inductive [i] when it's zeta normal-form is of the shape Π Δ. concl and: - All of the arguments in Δ are positive. - - The conclusion is of the shape [mkApps (tRel k) indices] + - The conclusion is of the shape [mkApps (tRel k) indices] where [k] refers to the current inductive [i] and [indices] does not mention any of the inductive types in the block. I.e. [indices] are closed terms in [params ,,, args]. *) - + Inductive positive_cstr mdecl i (ctx : context) : term -> Type := | positive_cstr_concl indices : - let headrel : nat := + let headrel : nat := (#|mdecl.(ind_bodies)| - S i + #|ctx|)%nat in All (closedn #|ctx|) indices -> positive_cstr mdecl i ctx (mkApps (tRel headrel) indices) | positive_cstr_let na b ty t : positive_cstr mdecl i ctx (subst [b] 0 t) -> - positive_cstr mdecl i ctx (tLetIn na b ty t) + positive_cstr mdecl i ctx (tLetIn na b ty t) | positive_cstr_ass na ty t : positive_cstr_arg mdecl ctx ty -> @@ -886,7 +893,7 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT positive_cstr mdecl i ctx (tProd na ty t). Definition lift_level n l := - match l with + match l with | Level.lzero | Level.Level _ => l | Level.Var k => Level.Var (n + k) end. @@ -908,7 +915,7 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT Fixpoint variance_cstrs (v : list Variance.t) (u u' : Instance.t) := match v, u, u' with | _, [], [] => ConstraintSet.empty - | v :: vs, u :: us, u' :: us' => + | v :: vs, u :: us, u' :: us' => match v with | Variance.Irrelevant => variance_cstrs vs us us' | Variance.Covariant => ConstraintSet.add (u, ConstraintType.Le 0, u') (variance_cstrs vs us us') @@ -917,7 +924,7 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT | _, _, _ => (* Impossible due to on_variance invariant *) ConstraintSet.empty end. - (** This constructs a duplication of the polymorphic universe context of the inductive, + (** This constructs a duplication of the polymorphic universe context of the inductive, where the two instances are additionally related according to the variance information. *) @@ -934,10 +941,10 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT Some (Polymorphic_ctx auctx', u, u') end. - (** A constructor type respects the given variance [v] if each constructor + (** A constructor type respects the given variance [v] if each constructor argument respects it and each index (in the conclusion) does as well. We formalize this by asking for a cumulativity relation between the contexts - of arguments and conversion of the lists of indices instanciated with [u] and + of arguments and conversion of the lists of indices instanciated with [u] and [u'] where [u `v` u']. *) Definition ind_arities mdecl := arities_context (ind_bodies mdecl). @@ -959,7 +966,7 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT cumul_ctx_rel Pcmp (Σ, univs) (ind_arities mdecl ,,, smash_context [] (ind_params mdecl))@[u] (expand_lets_ctx (ind_params mdecl) (smash_context [] (cstr_args cs)))@[u] (expand_lets_ctx (ind_params mdecl) (smash_context [] (cstr_args cs)))@[u'] * - All2 + All2 (Pcmp (Σ, univs) (ind_arities mdecl ,,, smash_context [] (ind_params mdecl ,,, cstr_args cs))@[u] Conv) (map (subst_instance u ∘ expand_lets (ind_params mdecl ,,, cstr_args cs)) (cstr_indices cs)) (map (subst_instance u' ∘ expand_lets (ind_params mdecl ,,, cstr_args cs)) (cstr_indices cs)) @@ -971,20 +978,20 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT tRel (#|mdecl.(ind_bodies)| - S i + #|mdecl.(ind_params)| + #|cstr_args cdecl|). (* Constructor conclusion shape: the inductives type applied to variables for - the (non-let) parameters + the (non-let) parameters followed by the indices *) Definition cstr_concl mdecl i cdecl := (mkApps (cstr_concl_head mdecl i cdecl) (to_extended_list_k mdecl.(ind_params) #|cstr_args cdecl| ++ cstr_indices cdecl)). - + Record on_constructor Σ mdecl i idecl ind_indices cdecl cunivs := { (* cdecl.1 fresh ?? *) cstr_args_length : context_assumptions (cstr_args cdecl) = cstr_arity cdecl; cstr_eq : cstr_type cdecl = - it_mkProd_or_LetIn mdecl.(ind_params) - (it_mkProd_or_LetIn (cstr_args cdecl) + it_mkProd_or_LetIn mdecl.(ind_params) + (it_mkProd_or_LetIn (cstr_args cdecl) (cstr_concl mdecl i cdecl)); (* The type of the constructor canonically has this shape: parameters, real arguments ending with a reference to the inductive applied to the @@ -995,21 +1002,21 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT on_cargs : sorts_local_ctx Σ (arities_context mdecl.(ind_bodies) ,,, mdecl.(ind_params)) cdecl.(cstr_args) cunivs; - on_cindices : - ctx_inst (fun Γ t T => P Σ Γ t (Typ T)) (arities_context mdecl.(ind_bodies) ,,, mdecl.(ind_params) ,,, cdecl.(cstr_args)) + on_cindices : + ctx_inst (fun Γ t T => P Σ Γ (Typ t T)) (arities_context mdecl.(ind_bodies) ,,, mdecl.(ind_params) ,,, cdecl.(cstr_args)) cdecl.(cstr_indices) (List.rev (lift_context #|cdecl.(cstr_args)| 0 ind_indices)); on_ctype_positive : (* The constructor type is positive *) positive_cstr mdecl i [] (cstr_type cdecl); - on_ctype_variance : (* The constructor type respect the variance annotation + on_ctype_variance : (* The constructor type respect the variance annotation on polymorphic universes, if any. *) - forall v, ind_variance mdecl = Some v -> + forall v, ind_variance mdecl = Some v -> cstr_respects_variance Σ mdecl v cdecl; - on_lets_in_type : if lets_in_constructor_types - then True else is_true (is_assumption_context (cstr_args cdecl)) + on_lets_in_type : if lets_in_constructor_types + then True else is_true (is_assumption_context (cstr_args cdecl)) }. Arguments on_ctype {Σ mdecl i idecl ind_indices cdecl cunivs}. @@ -1021,30 +1028,30 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT Definition on_constructors Σ mdecl i idecl ind_indices := All2 (on_constructor Σ mdecl i idecl ind_indices). - (** Each projection type corresponds to a non-let argument of the - corresponding constructor. It is parameterized over the + (** Each projection type corresponds to a non-let argument of the + corresponding constructor. It is parameterized over the parameters of the inductive type and all the preceding arguments of the constructor. When computing the type of a projection for argument [n] at a given instance of the parameters and a given term [t] in the inductive type, we instantiate the argument context by corresponsping projections [t.π1 ... t.πn-1]. This is essential for subject reduction to hold: each projections type can only refer to the record object through projections. - + Projection types have their parameter and argument contexts smashed to avoid costly computations during type-checking and reduction: we can just substitute - the instances of parameters and the inductive value without considering the + the instances of parameters and the inductive value without considering the presence of let bindings. *) Record on_proj mdecl mind i k (p : projection_body) decl := { on_proj_name : (* All projections are be named after a constructor argument. *) binder_name (decl_name decl) = nNamed p.(proj_name); - on_proj_type : + on_proj_type : (** The stored projection type already has the references to the inductive type substituted along with the previous arguments replaced by projections. *) let u := abstract_instance mdecl.(ind_universes) in let ind := {| inductive_mind := mind; inductive_ind := i |} in p.(proj_type) = subst (inds mind u mdecl.(ind_bodies)) (S (ind_npars mdecl)) - (subst (projs ind mdecl.(ind_npars) k) 0 + (subst (projs ind mdecl.(ind_npars) k) 0 (lift 1 k (decl_type decl))); on_proj_relevance : p.(proj_relevance) = decl.(decl_name).(binder_relevance) }. @@ -1071,7 +1078,7 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT on_projs : Alli (on_projection mdecl mind i cdecl) 0 idecl.(ind_projs) }. Definition check_constructors_smaller φ cunivss ind_sort := - Forall (fun cunivs => + Forall (fun cunivs => Forall (fun argsort => leq_universe φ argsort ind_sort) cunivs) cunivss. (** This ensures that all sorts in kelim are lower @@ -1091,7 +1098,7 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT IntoPropSProp (* Squashed: some arguments are higher than Prop, restrict to Prop *) | _ => (* Squashed: at least 2 constructors *) IntoPropSProp end. - + Definition elim_sort_sprop_ind (ind_ctors_sort : list constructor_univs) := match ind_ctors_sort with | [] => (* Empty inductive strict proposition: *) IntoAny @@ -1117,7 +1124,7 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT × if indices_matter then type_local_ctx Σ params ind_indices ind_sort else True. - + Record on_ind_body Σ mind mdecl i idecl := { (** The type of the inductive must be an arity, sharing the same params as the rest of the block, and maybe having a context of indices. *) @@ -1151,30 +1158,30 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT ind_sorts : 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 : + onIndices : (* The inductive type respect the variance annotation on polymorphic universes, if any. *) - forall v, ind_variance mdecl = Some v -> + forall v, ind_variance mdecl = Some v -> ind_respects_variance Σ mdecl v idecl.(ind_indices) }. Definition on_variance Σ univs (variances : option (list Variance.t)) := match univs return Type with | Monomorphic_ctx => variances = None - | Polymorphic_ctx auctx => + | Polymorphic_ctx auctx => match variances with | None => unit - | Some v => - ∑ univs' i i', + | Some v => + ∑ univs' i i', [× (variance_universes univs v = Some (univs', i, i')), consistent_instance_ext (Σ, univs') univs i, consistent_instance_ext (Σ, univs') univs i' & List.length v = #|UContext.instance (AUContext.repr auctx)|] end end. - + (** We allow empty blocks for simplicity (no well-typed reference to them can be made). *) @@ -1192,8 +1199,7 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT (** *** Typing of constant declarations *) Definition on_constant_decl Σ d := - on_type_rel Σ [] d.(cst_type) d.(cst_relevance) × - option_default (fun trm => P Σ [] trm (Typ d.(cst_type))) d.(cst_body) True. + P Σ [] (TripleRelOpt d.(cst_body) d.(cst_type) d.(cst_relevance)). Definition on_global_decl Σ kn decl := match decl with @@ -1222,10 +1228,10 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT on_global_decls univs (Σ ,, (kn, d)). Derive Signature for on_global_decls. - Definition on_global_univs (c : ContextSet.t) := + Definition on_global_univs (c : ContextSet.t) := let levels := global_levels c in let cstrs := ContextSet.constraints c in - ConstraintSet.For_all (declared_cstr_levels levels) cstrs /\ + ConstraintSet.For_all (declared_cstr_levels levels) cstrs /\ LS.For_all (negb ∘ Level.is_var) levels /\ consistent cstrs. @@ -1277,9 +1283,9 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT Arguments onVariance {_ Pcmp P Σ mind mdecl}. - Lemma type_local_ctx_impl (P Q : global_env_ext -> context -> term -> typ_or_sort -> Type) Σ Γ Δ u : + Lemma type_local_ctx_impl (P Q : global_env_ext -> context -> judgment -> Type) Σ Γ Δ u : type_local_ctx P Σ Γ Δ u -> - (forall Γ t T, P Σ Γ t T -> Q Σ Γ t T) -> + (forall Γ j, P Σ Γ j -> Q Σ Γ j) -> type_local_ctx Q Σ Γ Δ u. Proof. intros HP HPQ. revert HP; induction Δ in Γ, HPQ |- *; simpl; auto. @@ -1287,9 +1293,9 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT intros. intuition auto. intuition auto. Qed. - Lemma sorts_local_ctx_impl (P Q : global_env_ext -> context -> term -> typ_or_sort -> Type) Σ Γ Δ u : + Lemma sorts_local_ctx_impl (P Q : global_env_ext -> context -> judgment -> Type) Σ Γ Δ u : sorts_local_ctx P Σ Γ Δ u -> - (forall Γ t T, P Σ Γ t T -> Q Σ Γ t T) -> + (forall Γ j, P Σ Γ j -> Q Σ Γ j) -> sorts_local_ctx Q Σ Γ Δ u. Proof. intros HP HPQ. revert HP; induction Δ in Γ, HPQ, u |- *; simpl; auto. @@ -1299,17 +1305,15 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT Qed. Lemma on_global_decl_impl {cf : checker_flags} Pcmp P Q Σ kn d : - (forall Γ t T, + (forall Γ j, on_global_env Pcmp P Σ.1 -> - P Σ Γ t T -> Q Σ Γ t T) -> + P Σ Γ j -> Q Σ Γ j) -> on_global_env Pcmp P Σ.1 -> on_global_decl Pcmp P Σ kn d -> on_global_decl Pcmp Q Σ kn d. Proof. intros X X0. destruct d; simpl. - - destruct 1; split. - * eapply X => //. - * destruct cst_body => //. now eapply X. + - apply X => //. - intros [onI onP onNP]. constructor; auto. -- eapply Alli_impl; tea. intros. @@ -1330,7 +1334,7 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT revert on_cindices0. generalize (List.rev (lift_context #|cstr_args x0| 0 (ind_indices x))). generalize (cstr_indices x0). - induction 1; simpl; constructor; auto. + induction 1; simpl; constructor; auto. --- simpl; intros. pose (onProjections X1 H). simpl in *; auto. --- destruct X1. simpl. unfold check_ind_sorts in *. destruct Universe.is_prop; auto. @@ -1346,10 +1350,10 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT Qed. Lemma on_global_env_impl {cf : checker_flags} Pcmp P Q : - (forall Σ Γ t T, + (forall Σ Γ j, on_global_env Pcmp P Σ.1 -> on_global_env Pcmp Q Σ.1 -> - P Σ Γ t T -> Q Σ Γ t T) -> + P Σ Γ j -> Q Σ Γ j) -> forall Σ, on_global_env Pcmp P Σ -> on_global_env Pcmp Q Σ. Proof. intros X [univs Σ] [cu X0]; split => /= //. cbn in *. @@ -1368,7 +1372,7 @@ End GlobalMapsSig. Module Type ConversionParSig (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E) (ET : EnvTypingSig T E TU). Import T E TU ET. - + Parameter Inline cumul_gen : forall {cf : checker_flags}, global_env_ext -> context -> conv_pb -> term -> term -> Type. End ConversionParSig. @@ -1389,12 +1393,12 @@ End Typing. Module DeclarationTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E) (ET : EnvTypingSig T E TU) (CT : ConversionSig T E TU ET) - (CS : ConversionParSig T E TU ET) (Ty : Typing T E TU ET CT CS) + (CS : ConversionParSig T E TU ET) (Ty : Typing T E TU ET CT CS) (L : LookupSig T E) (GM : GlobalMapsSig T E TU ET CT L). Import T E L TU ET CT GM CS Ty. - Notation isTypeRelOpt Σ Γ t relopt := (lift_typing typing Σ Γ t (Sort relopt)). + Notation isTypeRelOpt Σ Γ t relopt := (lift_typing typing Σ Γ (TripleOpt None t relopt)). Definition isType `{checker_flags} (Σ : global_env_ext) (Γ : context) (t : term) := isTypeRelOpt Σ Γ t None. @@ -1406,7 +1410,7 @@ Module DeclarationTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E) Definition isType_of_isTypeRel `{checker_flags} {Σ Γ t rel} (u: isTypeRel Σ Γ t rel) : isType Σ Γ t := match u with - | existT s (_, σ) => existT _ s (I, σ) + | (p, existT s (_, σ)) => (p, existT _ s (I, σ)) end. (** This predicate enforces that there exists typing derivations for every typable term in env. *) @@ -1418,10 +1422,7 @@ Module DeclarationTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E) (** *** Typing of local environments *) Definition type_local_decl `{checker_flags} Σ Γ d := - match d.(decl_body) with - | None => isType Σ Γ d.(decl_type) - | Some body => Σ ;;; Γ |- body : d.(decl_type) - end. + lift_typing typing Σ Γ (Triple d.(decl_type) d.(decl_type)). (** ** Induction principle for typing up-to a global environment *) @@ -1430,13 +1431,13 @@ Module DeclarationTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E) Definition wf_local_rel `{checker_flags} Σ := All_local_rel (lift_typing typing Σ). - (** Functoriality of global environment typing derivations + folding of the well-formed + (** Functoriality of global environment typing derivations + folding of the well-formed environment assumption. *) Lemma on_wf_global_env_impl `{checker_flags} {Σ : global_env} {wfΣ : on_global_env cumul_gen (lift_typing typing) Σ} P Q : - (forall Σ Γ t T, on_global_env cumul_gen (lift_typing typing) Σ.1 -> - on_global_env cumul_gen P Σ.1 -> + (forall Σ Γ j, on_global_env cumul_gen (lift_typing typing) Σ.1 -> + on_global_env cumul_gen P Σ.1 -> on_global_env cumul_gen Q Σ.1 -> - P Σ Γ t T -> Q Σ Γ t T) -> + P Σ Γ j -> Q Σ Γ j) -> on_global_env cumul_gen P Σ -> on_global_env cumul_gen Q Σ. Proof. unfold on_global_env in *. @@ -1446,13 +1447,11 @@ Module DeclarationTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E) induction 1; constructor; auto. { depelim wfΣ. eauto. } depelim wfΣ. specialize (IHX0 cu wfΣ). - assert (X' := fun Γ t T => X ({| universes := univs; declarations := Σ |}, udecl0) Γ t T + assert (X' := fun Γ j => X ({| universes := univs; declarations := Σ |}, udecl0) Γ j (cu, wfΣ) (cu, X0) (cu, IHX0)); clear X. rename X' into X. clear IHX0. destruct d; simpl. - - destruct o0; split. - * eapply X => //. - * destruct cst_body => //. now eapply X. + - apply X => //. - red in o. simpl in *. destruct o0 as [onI onP onNP]. constructor; auto. diff --git a/template-coq/theories/EtaExpand.v b/template-coq/theories/EtaExpand.v index 1426f085c..67b35eec5 100644 --- a/template-coq/theories/EtaExpand.v +++ b/template-coq/theories/EtaExpand.v @@ -1148,11 +1148,10 @@ Proof. { apply andb_and in H2. destruct H2 as [isl _]. solve_all. } solve_all. { now eapply isLambda_lift, isLambda_eta_expand. } - destruct a as (s & e & Hs1 & Hs2). - destruct a0 as (? & ?). + destruct a as (((Hbo1 & Hbo2) & et) & s & e & Hs1 & Hs2). rewrite !firstn_length. rewrite !Nat.min_l; try lia. eapply expanded_lift'. - 5: eapply e0. 2: reflexivity. 2: now len. + 5: eapply Hbo2. 2: reflexivity. 2: now len. 2: now len. { rewrite map_app. f_equal. rewrite map_rev. f_equal. now rewrite !mapi_map, map_mapi. } eapply Forall2_app; solve_all. @@ -1170,9 +1169,9 @@ Proof. len; lia. rewrite repeat_length. len; lia. + cbn - [rev_map seq]. rewrite rev_map_spec. cbn. rewrite Nat.sub_0_r. cbn. destruct List.rev; cbn; congruence. - - cbn. econstructor; eauto. eapply All_Forall, All_map, All_impl. eapply (All_mix X X0). intros ? ((? & ? & ? & ?) & ? & ?). cbn. - specialize (e0 (repeat None #|mfix| ++ Γ'))%list. - rewrite map_app, map_repeat in e0. len. eapply e0. + - cbn. econstructor; eauto. eapply All_Forall, All_map, All_impl. eapply X. intros ? (((Hbo1 & Hbo2) & et) & s & e & Hs1 & Hs2). cbn. + specialize (Hbo2 (repeat None #|mfix| ++ Γ'))%list. + rewrite map_app, map_repeat in Hbo2. len. eapply Hbo2. eapply Forall2_app; eauto. unfold types. assert (#|Typing.fix_context mfix| = #|mfix|). { unfold Typing.fix_context. now len. } revert H4. generalize (Typing.fix_context mfix). clear. @@ -1345,8 +1344,8 @@ Proof. cbn. constructor. cbn. constructor. len. rewrite app_nil_r. - red in t0, t1. - forward (eta_expand_expanded (Σ := Σ) Γ (repeat None #|Γ|) _ _ wfΣ t1). + destruct t0 as ((t0 & _) & t1). + forward (eta_expand_expanded (Σ := Σ) Γ (repeat None #|Γ|) _ _ wfΣ t0). clear. induction Γ; cbn; constructor; auto. now rewrite map_repeat. Qed. @@ -1383,10 +1382,10 @@ Lemma eta_expand_global_decl_expanded {cf : checker_flags} g kn d : Proof. intros wf ond. destruct d; cbn in *. - - destruct ond as (onty & onbo). + - destruct ond as (onbo & onty). destruct c as [na body ty rel]; cbn in *. destruct body. constructor => //; cbn. - apply (eta_expand_expanded (Σ := g) [] [] t na wf onbo). constructor. + apply (eta_expand_expanded (Σ := g) [] [] t na wf onbo.1). constructor. destruct onty as (s & e & Hs). constructor => //. - destruct ond as [onI onP onN onV]. constructor. cbn. @@ -1401,7 +1400,7 @@ Proof. pose proof onc.(on_cargs). eapply eta_expand_context_sorts in X0. now len in X0. len. len. - pose proof onc.(on_ctype). destruct X0 as (s & e & Hs). + pose proof onc.(on_ctype). destruct X0 as (_ & s & e & Hs). epose proof (eta_expand_expanded (Σ := g) _ (repeat None #|ind_bodies m|) _ _ wf Hs). forward H. rewrite -arities_context_length. clear. induction (arities_context _); constructor; auto. diff --git a/template-coq/theories/Typing.v b/template-coq/theories/Typing.v index 598fa92b7..bcd475ce8 100644 --- a/template-coq/theories/Typing.v +++ b/template-coq/theories/Typing.v @@ -831,8 +831,7 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> fix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> wf_local Σ Γ -> - All (on_def_type (lift_typing typing Σ) Γ) mfix -> - All (on_def_body (lift_typing typing Σ) (fix_context mfix) Γ) mfix -> + All (on_def (lift_typing typing Σ) (fix_context mfix) Γ) mfix -> wf_fixpoint Σ mfix -> Σ ;;; Γ |- tFix mfix n : decl.(dtype) @@ -840,8 +839,7 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> cofix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> wf_local Σ Γ -> - All (on_def_type (lift_typing typing Σ) Γ) mfix -> - All (on_def_body (lift_typing typing Σ) (fix_context mfix) Γ) mfix -> + All (on_def (lift_typing typing Σ) (fix_context mfix) Γ) mfix -> wf_cofixpoint Σ mfix -> Σ ;;; Γ |- tCoFix mfix n : decl.(dtype) @@ -954,8 +952,8 @@ Proof. - exact (S (Nat.max d1 (Nat.max d2 (Nat.max (ctx_inst_size _ typing_size c1) (all2i_size _ (fun _ x y p => Nat.max (typing_size _ _ _ _ p.1.2) (typing_size _ _ _ _ p.2)) a0))))). - - exact (S (Nat.max (Nat.max (All_local_env_size typing_size _ _ a) (all_size _ (fun x p => lift_typing_size typing_size Σ _ _ _ p) a0)) (all_size _ (fun x p => typing_size Σ _ _ _ p) a1))). - - exact (S (Nat.max (Nat.max (All_local_env_size typing_size _ _ a) (all_size _ (fun x p => lift_typing_size typing_size Σ _ _ _ p) a0)) (all_size _ (fun x p => typing_size Σ _ _ _ p) a1))). + - exact (S (Nat.max (All_local_env_size typing_size _ _ a) (all_size _ (fun x p => lift_typing_size typing_size Σ _ _ p) a0))). + - exact (S (Nat.max (All_local_env_size typing_size _ _ a) (all_size _ (fun x p => lift_typing_size typing_size Σ _ _ p) a0))). Defined. Lemma typing_size_pos `{checker_flags} {Σ Γ t T} (d : Σ ;;; Γ |- t : T) : typing_size d > 0. @@ -1063,13 +1061,12 @@ Proof. destruct w. - simpl. congruence. - intros [=]. subst d Γ0. - exists w. simpl. destruct l as (x & ? & t0). exists x. exists t0. pose (typing_size_pos t0). + exists w. simpl. destruct l as (? & x & ? & t0). exists x. exists t0. pose (typing_size_pos t0). cbn. split. + lia. + auto with arith. - intros [=]. subst d Γ0. - exists w. simpl. simpl in l. destruct l as (u & ? & h). - simpl in l0. + exists w. simpl. simpl in l. destruct l as ((l0 & ?) & u & ? & h). exists u, l0, h. cbn. pose (typing_size_pos h). pose (typing_size_pos l0). @@ -1217,8 +1214,7 @@ Lemma typing_ind_env `{cf : checker_flags} : fix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> PΓ Σ Γ wfΓ -> - All (on_def_type (lift_typing2 typing P Σ) Γ) mfix -> - All (on_def_body (lift_typing2 typing P Σ) types Γ) mfix -> + All (on_def (lift_typing2 typing P Σ) types Γ) mfix -> wf_fixpoint Σ.1 mfix -> P Σ Γ (tFix mfix n) decl.(dtype)) -> @@ -1227,8 +1223,7 @@ Lemma typing_ind_env `{cf : checker_flags} : cofix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> PΓ Σ Γ wfΓ -> - All (on_def_type (lift_typing2 typing P Σ) Γ) mfix -> - All (on_def_body (lift_typing2 typing P Σ) types Γ) mfix -> + All (on_def (lift_typing2 typing P Σ) types Γ) mfix -> wf_cofixpoint Σ.1 mfix -> P Σ Γ (tCoFix mfix n) decl.(dtype)) -> @@ -1283,12 +1278,10 @@ Proof. apply IH'; auto. - simpl. simpl in *. destruct d. - + destruct Xg; split. 2: destruct c, cst_body0 => //. - 1: rename o1 into Xg. 2: rename o2 into Xg. - all: apply lift_typing_impl with (1 := Xg); intros ? Hs. - all: specialize (IH ((Σ', udecl); wfΣ; _; _; _; Hs)). - all: forward IH; [constructor 1; simpl; subst Σ' Σg; cbn; lia|]. - all: apply IH. + + apply lift_typing_impl with (1 := Xg) => ?? Hs. + specialize (IH ((Σ', udecl); wfΣ; _; _; _; Hs)). + forward IH; [constructor 1; simpl; subst Σ' Σg; cbn; lia|]. + apply IH. + red in Xg. destruct Xg as [onI onP onnp]; constructor; eauto. * unshelve eset (IH' := fun p => IH ((Σ', udecl); wfΣ; p) _). @@ -1298,36 +1291,35 @@ Proof. refine {| ind_arity_eq := Xg.(ind_arity_eq); ind_cunivs := Xg.(ind_cunivs) |}. -- apply onArity in Xg. - apply lift_typing_impl with (1 := Xg); intros ? Hs. + apply lift_typing_impl with (1 := Xg) => ?? Hs. apply (IH (_; _; _; Hs)). -- pose proof Xg.(onConstructors) as Xg'. eapply All2_impl; eauto. intros. destruct X13 as [cass tyeq onctyp oncargs oncind]. unshelve econstructor; eauto. - { apply lift_typing_impl with (1 := onctyp); intros ? Hs. + { apply lift_typing_impl with (1 := onctyp) => ?? Hs. apply (IH (_; _; _; Hs)). } - { apply sorts_local_ctx_impl with (1 := oncargs); intros Γ' t' T' HT'. - apply lift_typing_impl with (1 := HT'); intros ? Hs. + { apply sorts_local_ctx_impl with (1 := oncargs) => Γ' j Hj. + apply lift_typing_impl with (1 := Hj) => ?? Hs. apply (IH (_; _; _; Hs)). } { clear -IH oncind. revert oncind. generalize (List.rev (lift_context #|cstr_args x0| 0 (ind_indices x))). generalize (cstr_indices x0). induction 1; constructor; auto. - do 2 red in t0 |- *. + do 3 red in t0 |- *. apply (IH (_; (_; (_; t0)))). } -- intros Hprojs; pose proof (onProjections Xg Hprojs); auto. -- destruct Xg. simpl. unfold check_ind_sorts in *. - destruct Universe.is_prop; auto. - destruct Universe.is_sprop; auto. + destruct (ind_sort x) => //=. split. apply ind_sorts0. destruct indices_matter; auto. - eapply type_local_ctx_impl. eapply ind_sorts0. intros ??? HT. - apply lift_typing_impl with (1 := HT); intros ? Hs. + eapply type_local_ctx_impl. eapply ind_sorts0. intros ?? HT. + apply lift_typing_impl with (1 := HT) => ?? Hs. apply (IH (_; _; _; Hs)). -- apply Xg.(ind_relevance_compat). -- apply (onIndices Xg). * red in onP |- *. - apply All_local_env_impl with (1 := onP); intros ??? HT. - apply lift_typing_impl with (1 := HT); intros ? Hs. + apply All_local_env_impl with (1 := onP); intros ?? HT. + apply lift_typing_impl with (1 := HT) => ?? Hs. apply (IH ((Σ', udecl); (wfΣ; _; _; _; Hs))). constructor 1. simpl. subst Σ' Σg; cbn; lia. @@ -1353,11 +1345,11 @@ Proof. clearbody foo. assert (All_local_env_size (@typing_size cf) Σ Γ' foo < typing_size H) by lia. clear H1 H0 Hty. revert foo H2. - induction foo; simpl in *; try destruct t3 as (u & e & Hu); cbn in *; constructor. + induction foo; simpl in *; try destruct t3 as (? & u & e & Hu); cbn in *; constructor. - simpl in *. apply IHfoo. lia. - red. apply (X14 _ _ _ Hu). lia. - simpl in *. apply IHfoo. lia. - - red. apply (X14 _ _ _ t4). lia. + - red. apply (X14 _ _ _ o.1). lia. - red. simpl. apply (X14 _ _ _ Hu). lia. } clear IH. @@ -1446,63 +1438,46 @@ Proof. -- clear X X0 Xcast X1 X2 X3 X4 X5 X6 X7 X8 X9 X11 X12. eapply X10; eauto; clear X10. - * assert (forall t T (Hty : Σ;;; Γ |- t : T), - typing_size Hty < S (all_size _ (fun x p => lift_typing_size (@typing_size _) Σ _ _ _ p) a0) -> - on_global_env cumul_gen (lift_typing P) Σ.1 * P Σ Γ t T). - { intros; eauto. eapply (X14 _ _ _ Hty); eauto. simpl in *. lia. } - clear X13 X14 a pΓ. - clear -a0 X. - induction a0; constructor. - destruct p as (s & e & Hs). exists s; repeat split; auto. - apply (X (dtype x) (tSort s) Hs). simpl. lia. - apply IHa0. intros. eapply (X _ _ Hty); eauto. - cbn [all_size]. lia. - * simpl in X14. - assert (forall Γ' : context, - wf_local Σ Γ' -> - forall (t T : term) (Hty : Σ;;; Γ' |- t : T), - typing_size Hty < S (all_size _ (fun x p => typing_size p) a1) -> - on_global_env cumul_gen (lift_typing P) Σ.1 * P Σ Γ' t T). - { intros. eapply (X14 _ _ _ Hty); eauto. lia. } - clear X14 X13. - clear e decl i a0 i0 pΓ. - remember (fix_context mfix) as mfixcontext. clear Heqmfixcontext. - - induction a1; econstructor; eauto. - ++ split; auto. - eapply (X _ (typing_wf_local p) _ _ p). simpl. lia. - ++ eapply IHa1. intros. - eapply (X _ X0 _ _ Hty). simpl; lia. - - -- clear X X0 Xcast X1 X2 X3 X4 X5 X6 X7 X8 X9 X10 X12. - eapply X11; eauto; clear X11. - * assert (forall t T (Hty : Σ;;; Γ |- t : T), - typing_size Hty < S (all_size _ (fun x p => lift_typing_size (@typing_size _) Σ _ _ _ p) a0) -> - on_global_env cumul_gen (lift_typing P) Σ.1 * P Σ Γ t T). - { intros; eauto. eapply (X14 _ _ _ Hty); eauto. simpl in *. lia. } - clear X13 X14 a pΓ. - clear -a0 X. - induction a0; constructor. - destruct p as (s & e & Hs). exists s; repeat split; auto. - apply (X (dtype x) (tSort s) Hs). simpl. lia. - apply IHa0. intros. eapply (X _ _ Hty); eauto. - cbn [all_size]. lia. - * simpl in X14. - assert (forall Γ' : context, - wf_local Σ Γ' -> - forall (t T : term) (Hty : Σ;;; Γ' |- t : T), - typing_size Hty < S (all_size _ (fun x p => typing_size p) a1) -> - on_global_env cumul_gen (lift_typing P) Σ.1 * P Σ Γ' t T). - { intros. eapply (X14 _ _ _ Hty); eauto. lia. } - clear X14 X13. - clear e decl a0 i i0 pΓ. - remember (fix_context mfix) as mfixcontext. clear Heqmfixcontext. - - induction a1; econstructor; eauto. - ++ split; auto. - eapply (X _ (typing_wf_local p) _ _ p). simpl. lia. - ++ eapply IHa1. intros. - eapply (X _ X0 _ _ Hty). simpl; lia. + simpl in X14. + assert (forall Γ' : context, + wf_local Σ Γ' -> + forall (t T : term) (Hty : Σ;;; Γ' |- t : T), + typing_size Hty < S (all_size _ (fun x p => lift_typing_size (@typing_size _) Σ _ _ p) a0) -> + on_global_env cumul_gen (lift_typing P) Σ.1 * P Σ Γ' t T). + { intros. eapply (X14 _ _ _ Hty); eauto. simpl in *. lia. } + clear X14 X13. + clear -a0 X. + remember (fix_context mfix) as mfixcontext. clear Heqmfixcontext. + + induction a0; econstructor; eauto. + ++ destruct p as ((Hbo & Hborel) & s & e & Hs). + repeat split; auto. + eapply (X _ (typing_wf_local Hbo) _ _ Hbo). simpl. lia. + exists s; repeat split; auto. + apply (X _ (typing_wf_local Hs) _ _ Hs). simpl. lia. + ++ eapply IHa0. intros. + eapply (X _ X0 _ _ Hty). cbn [all_size]; lia. + + -- clear X X0 Xcast X1 X2 X3 X4 X5 X6 X7 X8 X9 X10 X12. + eapply X11; eauto; clear X11. + assert (forall Γ' : context, + wf_local Σ Γ' -> + forall (t T : term) (Hty : Σ;;; Γ' |- t : T), + typing_size Hty < S (all_size _ (fun x p => lift_typing_size (@typing_size _) Σ _ _ p) a0) -> + on_global_env cumul_gen (lift_typing P) Σ.1 * P Σ Γ' t T). + { intros. eapply (X14 _ _ _ Hty); eauto. simpl in *. lia. } + clear X14 X13. + clear -a0 X. + remember (fix_context mfix) as mfixcontext. clear Heqmfixcontext. + + induction a0; econstructor; eauto. + ++ destruct p as ((Hbo & Hborel) & s & e & Hs). + repeat split; auto. + eapply (X _ (typing_wf_local Hbo) _ _ Hbo). simpl. lia. + exists s; repeat split; auto. + apply (X _ (typing_wf_local Hs) _ _ Hs). simpl. lia. + ++ eapply IHa0. intros. + eapply (X _ X0 _ _ Hty). cbn [all_size]; lia. Qed. (** * Lemmas about All_local_env *) @@ -1528,10 +1503,9 @@ Proof. now exists ((kn, d) :: Σ''). Qed. -Lemma All_local_env_app_inv - (P : context -> term -> typ_or_sort -> Type) l l' : +Lemma All_local_env_app_inv (P : context -> judgment -> Type) l l' : All_local_env P (l ,,, l') -> - All_local_env P l * All_local_env (fun Γ t T => P (l ,,, Γ) t T) l'. + All_local_env P l * All_local_env (fun Γ j => P (l ,,, Γ) j) l'. Proof. induction l'; simpl; split; auto. - constructor. @@ -1546,13 +1520,12 @@ Proof. + eapply localenv_cons_def. * apply IHl'. apply X0. * apply X1. - * eapply X2. Qed. Lemma All_local_env_lookup {P Γ n} {decl} : All_local_env P Γ -> nth_error Γ n = Some decl -> - on_local_decl P (skipn (S n) Γ) decl. + on_decl P (skipn (S n) Γ) decl. Proof. induction 1 in n, decl |- *. - simpl. destruct n; simpl; congruence. @@ -1564,17 +1537,17 @@ Proof. + apply IHX. Qed. -Lemma All_local_env_app `{checker_flags} (P : context -> term -> typ_or_sort -> Type) l l' : - All_local_env P l * All_local_env (fun Γ t T => P (l ,,, Γ) t T) l' -> +Lemma All_local_env_app `{checker_flags} (P : context -> judgment -> Type) l l' : + All_local_env P l * All_local_env (fun Γ j => P (l ,,, Γ) j) l' -> All_local_env P (l ,,, l'). Proof. induction l'; simpl; auto. intuition. intuition. destruct a. destruct decl_body. inv b. econstructor; eauto. inv b; econstructor; eauto. Qed. -Lemma All_local_env_map `{checker_flags} (P : context -> term -> typ_or_sort -> Type) f l : +Lemma All_local_env_map `{checker_flags} (P : context -> judgment -> Type) f l : (forall u, f (tSort u) = tSort u) -> - All_local_env (fun Γ t T => P (map (map_decl f) Γ) (f t) (typ_or_sort_map f T)) l -> All_local_env P (map (map_decl f) l). + All_local_env (fun Γ j => P (map (map_decl f) Γ) (judgment_map f j)) l -> All_local_env P (map (map_decl f) l). Proof. intros Hf. induction 1; econstructor; eauto. Qed. @@ -1597,7 +1570,7 @@ Defined. Definition lookup_wf_local_decl {Γ P} (wfΓ : All_local_env P Γ) (n : nat) {decl} (eq : nth_error Γ n = Some decl) : ∑ Pskip : All_local_env P (skipn (S n) Γ), - on_local_decl P (skipn (S n) Γ) decl. + on_decl P (skipn (S n) Γ) decl. Proof. induction wfΓ in n, decl, eq |- *; simpl. - elimtype False. destruct n; inversion eq. @@ -1605,19 +1578,19 @@ Proof. + simpl. exists wfΓ. injection eq; intros <-. apply t0. + apply IHwfΓ. auto with arith. - destruct n. - + exists wfΓ. injection eq; intros <-. apply t1. + + exists wfΓ. injection eq; intros <-. apply t0. + apply IHwfΓ. apply eq. Defined. Definition on_wf_local_decl `{checker_flags} {Σ Γ} (P : forall Σ Γ (wfΓ : wf_local Σ Γ) t T, Σ ;;; Γ |- t : T -> Type) (wfΓ : wf_local Σ Γ) {d} - (H : on_local_decl (lift_typing typing Σ) Γ d) := - match d as d' return (on_local_decl (lift_typing typing Σ) Γ d') -> Type with + (H : on_decl (lift_typing typing Σ) Γ d) := + match d as d' return (on_decl (lift_typing typing Σ) Γ d') -> Type with | {| decl_name := na; decl_body := Some b; decl_type := ty |} - => fun H => P Σ Γ wfΓ b ty H + => fun H => P Σ Γ wfΓ b ty H.1.1 | {| decl_name := na; decl_body := None; decl_type := ty |} - => fun H => P Σ Γ wfΓ _ _ (H.π2.2) + => fun H => P Σ Γ wfΓ _ _ (H.2.π2.2) end H. diff --git a/template-coq/theories/TypingWf.v b/template-coq/theories/TypingWf.v index 3696d2dc6..7ed097003 100644 --- a/template-coq/theories/TypingWf.v +++ b/template-coq/theories/TypingWf.v @@ -16,6 +16,9 @@ Existing Class wf. and the global context. *) +Definition wf_decl_pred Σ (Γ : context) T : Type := + lift_wf_term (WfAst.wf Σ) T. + Lemma All_local_env_wf_decl Σ : forall (Γ : context), All (wf_decl Σ) Γ -> All_local_env (wf_decl_pred Σ) Γ. @@ -28,29 +31,28 @@ Proof. * apply IHΓ. inv X; eauto. * red. inv X. split. -- apply X0. - -- constructor. - * red. inv X. eauto. + -- apply X0. + econstructor. * apply IHΓ. inv X; eauto. * red. inv X. split. -- apply X0. - -- constructor. + -- apply X0. Qed. Lemma on_global_decl_impl `{checker_flags} Σ P Q kn d : - (forall Σ Γ t T, on_global_env cumul_gen P Σ.1 -> P Σ Γ t T -> Q Σ Γ t T) -> + (forall Σ Γ j, on_global_env cumul_gen P Σ.1 -> P Σ Γ j -> Q Σ Γ j) -> on_global_env cumul_gen P Σ.1 -> on_global_decl cumul_gen P Σ kn d -> on_global_decl cumul_gen Q Σ kn d. Proof. intros. eapply on_global_decl_impl; tea. apply (X Σ). Qed. Lemma on_global_env_impl `{checker_flags} Σ P Q : - (forall Σ Γ t T, on_global_env cumul_gen P Σ.1 -> P Σ Γ t T -> Q Σ Γ t T) -> + (forall Σ Γ j, on_global_env cumul_gen P Σ.1 -> P Σ Γ j -> Q Σ Γ j) -> on_global_env cumul_gen P Σ -> on_global_env cumul_gen Q Σ. Proof. intros. eapply on_global_env_impl; tea. auto. Qed. Lemma All_local_env_wf_decl_inv Σ (a : context_decl) (Γ : list context_decl) (X : All_local_env (wf_decl_pred Σ) (a :: Γ)) : - on_local_decl (wf_decl_pred Σ) Γ a * All_local_env (wf_decl_pred Σ) Γ. + on_decl (wf_decl_pred Σ) Γ a * All_local_env (wf_decl_pred Σ) Γ. Proof. apply All_local_env_inv, X. Qed. Lemma unfold_fix_wf: @@ -202,8 +204,8 @@ Proof. eapply nth_error_alli in Hidecl; eauto. pose proof (onArity Hidecl). rewrite Hidecl.(ind_arity_eq) in X0. - destruct X0 as [s Hs]; wf. - eapply wf_it_mkProd_or_LetIn in s as [? H]. + destruct X0 as (_ & Hs); wf. + eapply wf_it_mkProd_or_LetIn in Hs as [? H]. eapply wf_it_mkProd_or_LetIn in H as []. solve_all. now eapply wf_decl_extends. Qed. @@ -244,7 +246,6 @@ Lemma All_local_env_wf_decls Σ ctx : All (wf_decl Σ) ctx. Proof. induction 1; constructor; auto. - destruct t0 as [Hw Hs]. now split. Qed. Lemma on_global_inductive_wf_params {cf:checker_flags} {Σ : global_env_ext} {kn mdecl} : @@ -441,7 +442,10 @@ Section WfAst. on_global_env cumul_gen wf_decl_pred Σ. Proof using Type. apply on_global_env_impl => Σ' Γ t []; simpl; unfold wf_decl_pred; - intros; auto. destruct X0 as [s ?]; now split. + intros; auto. + destruct t. + - destruct X as [s ?] => //. + - destruct X as (? & (_ & _ & ? & _)) => //. split => //. destruct t => //. destruct o as ((? & _) & _) => //. Qed. (* Hint Resolve on_global_wf_Forall_decls : wf. *) @@ -461,7 +465,7 @@ Section WfAst. intros oib. apply onParams in oib. red in oib. induction (ind_params mdecl) as [|[? [] ?] ?]; - inv oib; constructor; auto. now destruct X0. + inv oib; constructor; auto. Qed. Lemma declared_inductive_wf_params {ind mdecl idecl} : @@ -587,7 +591,7 @@ Section WfLookup. split => //. - now destruct onArity. - rewrite ind_arity_eq in onArity . - destruct onArity as [ona _]. + destruct onArity as (_ & ona). eapply wf_it_mkProd_or_LetIn in ona as [_ ona]. now eapply wf_it_mkProd_or_LetIn in ona as []. - unfold on_constructors in onConstructors. @@ -600,7 +604,7 @@ Section WfLookup. induction onConstructors; constructor; auto. destruct r. rewrite cstr_eq in on_ctype. - destruct on_ctype as [wf _]. + destruct on_ctype as [_ wf]. eapply wf_it_mkProd_or_LetIn in wf as [_ wf]. eapply wf_it_mkProd_or_LetIn in wf as [_ wf]. rewrite /cstr_concl in wf. @@ -793,12 +797,12 @@ Section WfRed. Proof using Type. intros wΣ h. unfold declared_constant in h. - destruct (lookup_on_global_env wΣ h) as [Σ' [wΣ' [ext (hty & hbo)]]]. + destruct (lookup_on_global_env wΣ h) as [Σ' [wΣ' [ext (hbo & hty)]]]. split. - - rewrite /on_type_rel /wf_decl_pred in hty. destruct hty. eauto using wf_extends. + - rewrite /on_type_rel /wf_decl_pred in hty. eauto using wf_extends. - destruct cst_body => //. rewrite /option_default /wf_decl_pred in hbo. - rewrite /on_option. destruct hbo. eauto using wf_extends. + rewrite /on_option. eauto using wf_extends. Qed. Lemma wf_it_mkProd_or_LetIn_inv (Σ' : global_env_ext) Γ (wfΓ : wf_local Σ' Γ) @@ -867,9 +871,12 @@ Section TypingWf. try solve [split; try constructor; intuition auto with wf]. - eapply All_local_env_wf_decls. - induction X; constructor; auto; red; - unfold wf_decl_pred, typ_or_sort_default; - intuition auto. + apply All_local_env_over_2 in X. + apply All_local_env_impl with (1 := X) => ?[??|t??] H. + + destruct H => //. + + destruct H as (? & _ & _ & _ & HT & _). + split => //. + destruct t => //. destruct o as ((_ & ? & _) & _) => //. - split; wf. apply wf_lift. apply (nth_error_all H X). - split. constructor; auto. wf. @@ -880,7 +887,7 @@ Section TypingWf. induction X1; auto. apply IHX1. apply wf_subst. now destruct p0. destruct p. now inv w. - split. wf. apply wf_subst_instance. wf. - destruct (lookup_on_global_env X H) as [Σ' [wfΣ' [ext (prfty & prfbo)]]]; eauto. + destruct (lookup_on_global_env X H) as [Σ' [wfΣ' [ext (prfbo & prfty)]]]; eauto. rewrite /on_type_rel //= in prfty. destruct prfty as (? & _ & []). wf. @@ -915,17 +922,22 @@ Section TypingWf. clear H. split. + constructor. - solve_all; destruct a, b. - all: intuition. + apply All_impl with (1 := X0) => d H. + destruct H as (((_ & (? & _)) & _) & _ & _ & _ & ? & _). + split => //. now eapply wf_lift_wf. + eapply All_nth_error in X0; eauto. - destruct X0 as [s ?]; intuition. + destruct X0 as (((_ & (? & _)) & _) & _ & _ & _ & ? & _). + now eapply wf_lift_wf. - subst types. split. + constructor. - solve_all; destruct a, b. - all: intuition. - + eapply All_nth_error in X0; eauto. destruct X0 as [s ?]; intuition. + apply All_impl with (1 := X0) => d Hs. + destruct Hs as (((_ & (? & _)) & _) & _ & _ & _ & ? & _). + split => //. now eapply wf_lift_wf. + + eapply All_nth_error in X0; eauto. + destruct X0 as (((_ & (? & _)) & _) & _ & _ & _ & ? & _). + now eapply wf_lift_wf. Qed. Lemma typing_all_wf_decl Σ (wfΣ : wf Σ.1) Γ (wfΓ : wf_local Σ Γ) : @@ -943,8 +955,10 @@ Section TypingWf. do 2 red in wfΣ. eapply on_global_env_impl; eauto; simpl; intros. unfold wf_decl_pred. - destruct T. apply X1. - now destruct X1 as (x & a & wf1 & wf2). + destruct j. apply X1. + destruct X1 as (o & _ & _ & wf2 & _). + split => //. + destruct t => //. destruct o as ((? & _) & _) => //. Qed. Lemma typing_wf Σ (wfΣ : wf Σ.1) Γ t T : diff --git a/template-coq/theories/WfAst.v b/template-coq/theories/WfAst.v index 076119bb4..490af3aa1 100644 --- a/template-coq/theories/WfAst.v +++ b/template-coq/theories/WfAst.v @@ -184,10 +184,7 @@ Proof. Qed. Definition wf_decl Σ d := - option_default (wf Σ) (decl_body d) True * wf Σ (decl_type d). - -Definition wf_decl_pred Σ (Γ : context) t T : Type := - wf Σ t × typ_or_sort_default (wf Σ) T True. + option_default (wf Σ) (decl_body d) unit * wf Σ (decl_type d). Lemma wf_mkApp Σ u a : wf Σ u -> wf Σ a -> wf Σ (mkApp u a). Proof. diff --git a/template-coq/theories/utils.v b/template-coq/theories/utils.v index c1f72bdf1..513be97c3 100644 --- a/template-coq/theories/utils.v +++ b/template-coq/theories/utils.v @@ -34,3 +34,5 @@ Global Open Scope metacoq_scope. Arguments Z.add : simpl nomatch. Arguments Nat.leb : simpl nomatch. Arguments Nat.eqb : simpl nomatch. + +#[global] Hint Constructors unit : core. From 6adeff3c206c9eb8057713401ba2e09bbd5deefe Mon Sep 17 00:00:00 2001 From: Yann Leray Date: Tue, 5 Jul 2022 16:02:32 +0200 Subject: [PATCH 17/17] Prepare for rebase --- pcuic/theories/Bidirectional/BDTyping.v | 157 +++++----- pcuic/theories/Conversion/PCUICClosedConv.v | 28 +- .../Conversion/PCUICUnivSubstitutionConv.v | 61 ++-- .../theories/Conversion/PCUICWeakeningConv.v | 4 +- .../Conversion/PCUICWeakeningEnvConv.v | 2 +- pcuic/theories/PCUICAst.v | 52 +++- pcuic/theories/PCUICGeneration.v | 8 +- pcuic/theories/PCUICRelevance.v | 45 --- pcuic/theories/PCUICRelevanceTerm.v | 23 +- pcuic/theories/PCUICTyping.v | 277 ++++++++---------- pcuic/theories/PCUICWcbvEval.v | 4 +- pcuic/theories/PCUICWeakeningEnv.v | 11 +- pcuic/theories/Syntax/PCUICDepth.v | 10 +- pcuic/theories/Syntax/PCUICInduction.v | 10 +- pcuic/theories/Syntax/PCUICLiftSubst.v | 21 ++ pcuic/theories/Typing/PCUICClosedTyp.v | 76 ++--- pcuic/theories/Typing/PCUICRenameTyp.v | 57 ++-- .../Typing/PCUICUnivSubstitutionTyp.v | 112 +++---- pcuic/theories/Typing/PCUICWeakeningEnvTyp.v | 221 +++++--------- template-coq/theories/Ast.v | 5 +- template-coq/theories/Environment.v | 2 +- template-coq/theories/EnvironmentTyping.v | 186 ++++++++++-- template-coq/theories/EtaExpand.v | 9 +- template-coq/theories/Typing.v | 32 +- template-coq/theories/TypingWf.v | 18 +- 25 files changed, 769 insertions(+), 662 deletions(-) diff --git a/pcuic/theories/Bidirectional/BDTyping.v b/pcuic/theories/Bidirectional/BDTyping.v index e873e5b1b..c9f73e7df 100644 --- a/pcuic/theories/Bidirectional/BDTyping.v +++ b/pcuic/theories/Bidirectional/BDTyping.v @@ -47,6 +47,7 @@ Inductive infering `{checker_flags} (Σ : global_env_ext) (Γ : context) : term | infer_LetIn na b B t s A : isSortRel s na.(binder_relevance) -> + isTermRel Σ (marks_of_context Γ) b na.(binder_relevance) -> Σ ;;; Γ |- B ▹□ s -> Σ ;;; Γ |- b ◃ B -> Σ ;;; Γ ,, vdef na b B |- t ▹ A -> @@ -111,16 +112,14 @@ Inductive infering `{checker_flags} (Σ : global_env_ext) (Γ : context) : term | infer_Fix mfix n decl : fix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> - All (on_def_type (lift_sorting checking infering_sort Σ) Γ) mfix -> - All (on_def_body (lift_sorting checking infering_sort Σ) (fix_context mfix) Γ) mfix -> + All (on_def (lift_sorting checking infering_sort Σ) (fix_context mfix) Γ) mfix -> wf_fixpoint Σ mfix -> Σ ;;; Γ |- tFix mfix n ▹ dtype decl | infer_CoFix mfix n decl : cofix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> - All (on_def_type (lift_sorting checking infering_sort Σ) Γ) mfix -> - All (on_def_body (lift_sorting checking infering_sort Σ) (fix_context mfix) Γ) mfix -> + All (on_def (lift_sorting checking infering_sort Σ) (fix_context mfix) Γ) mfix -> wf_cofixpoint Σ mfix -> Σ ;;; Γ |- tCoFix mfix n ▹ dtype decl @@ -203,10 +202,8 @@ Proof. | |- _ => exact 1 end. - exact (S (Nat.max a0 (Nat.max i (Nat.max i0 (Nat.max (ctx_inst_size _ (checking_size _) c1) (branches_size (checking_size _) (infering_sort_size _) a2)))))). - - exact (S (Nat.max (all_size _ (fun d p => infering_sort_size _ _ _ _ _ p.π2.2) a) - (all_size _ (fun x p => checking_size _ _ _ _ _ p) a0))). - - exact (S (Nat.max (all_size _ (fun d p => infering_sort_size _ _ _ _ _ p.π2.2) a) - (all_size _ (fun x => checking_size _ _ _ _ _) a0))). + - exact (S (all_size _ (fun d p => lift_sorting_size (checking_size _) (infering_sort_size _) _ _ _ p.1 + lift_sorting_size (checking_size _) (infering_sort_size _) _ _ _ p.2) a)). + - exact (S (all_size _ (fun d p => lift_sorting_size (checking_size _) (infering_sort_size _) _ _ _ p.1 + lift_sorting_size (checking_size _) (infering_sort_size _) _ _ _ p.2) a)). Defined. Fixpoint infering_size_pos `{checker_flags} {Σ Γ t T} (d : Σ ;;; Γ |- t ▹ T) @@ -315,8 +312,8 @@ Section BidirectionalInduction. Lemma bidir_ind_env : let Pdecl_check := fun Σ Γ wfΓ t T tyT => Pcheck Γ t T in let Pdecl_sort := fun Σ Γ wfΓ t u tyT => Psort Γ t u in - let Pdecl_check_rel Γ := fun _ Δ _ t T _ => Pcheck (Γ,,,Δ) t T in - let Pdecl_sort_rel Γ := fun _ Δ _ t u _ => Psort (Γ,,,Δ) t u in + let Pdecl_check_rel := fun _ Γ Δ _ t T _ => Pcheck (Γ,,,Δ) t T in + let Pdecl_sort_rel := fun _ Γ Δ _ t u _ => Psort (Γ,,,Δ) t u in let PcheckΣ := fun Σ => Pcheck in let PsortΣ := fun Σ => Psort in @@ -324,14 +321,14 @@ Section BidirectionalInduction. All_local_env_over_sorting checking infering_sort Pdecl_check Pdecl_sort Σ Γ wfΓ -> PΓ Γ) -> (forall (Γ Γ' : context) (wfΓ' : wf_local_bd_rel Σ Γ Γ'), - All_local_env_over_sorting (fun Σ Δ => checking Σ (Γ,,,Δ)) (fun Σ Δ => infering_sort Σ (Γ,,,Δ)) (Pdecl_check_rel Γ) (Pdecl_sort_rel Γ) Σ Γ' wfΓ' -> PΓ_rel Γ Γ') -> + All_local_rel_over_sorting checking infering_sort Pdecl_check_rel Pdecl_sort_rel Σ Γ Γ' wfΓ' -> PΓ_rel Γ Γ') -> (forall (Γ : context) (n : nat) decl, nth_error Γ n = Some decl -> Pinfer Γ (tRel n) (lift0 (S n) (decl_type decl))) -> (forall (Γ : context) (s : Universe.t), - wf_universe Σ s-> + wf_universe Σ s -> Pinfer Γ (tSort s) (tSort (Universe.super s))) -> (forall (Γ : context) (na : aname) (t b : term) (s1 s2 : Universe.t), @@ -342,20 +339,14 @@ Section BidirectionalInduction. Psort (Γ,, vass na t) b s2 -> Pinfer Γ (tProd na t b) (tSort (Universe.sort_of_product s1 s2))) -> - (forall (Γ : context) (na : aname) (t b : term) (s : Universe.t) (bty : term), - isSortRel s na.(binder_relevance) -> - Σ ;;; Γ |- t ▹□ s -> - Psort Γ t s -> + (forall (Γ : context) (na : aname) (t b : term) (bty : term), + on_decl (lift_sorting (Prop_conj checking PcheckΣ) (Prop_conj infering_sort PsortΣ) Σ) Γ (vass na t) -> Σ ;;; Γ,, vass na t |- b ▹ bty -> Pinfer (Γ,, vass na t) b bty -> Pinfer Γ (tLambda na t b) (tProd na t bty)) -> - (forall (Γ : context) (na : aname) (b B t : term) (s : Universe.t) (A : term), - isSortRel s na.(binder_relevance) -> - Σ ;;; Γ |- B ▹□ s -> - Psort Γ B s -> - Σ ;;; Γ |- b ◃ B -> - Pcheck Γ b B -> + (forall (Γ : context) (na : aname) (b B t : term) (A : term), + on_decl (lift_sorting (Prop_conj checking PcheckΣ) (Prop_conj infering_sort PsortΣ) Σ) Γ (vdef na b B) -> Σ ;;; Γ,, vdef na b B |- t ▹ A -> Pinfer (Γ,, vdef na b B) t A -> Pinfer Γ (tLetIn na b B t) (tLetIn na b B A)) -> @@ -433,16 +424,14 @@ Section BidirectionalInduction. (forall (Γ : context) (mfix : mfixpoint term) (n : nat) decl, fix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> - All (on_def_type (lift_sorting (Prop_conj checking PcheckΣ) (Prop_conj infering_sort PsortΣ) Σ) Γ) mfix -> - All (on_def_body (lift_sorting (Prop_conj checking PcheckΣ) (Prop_conj infering_sort PsortΣ) Σ) (fix_context mfix) Γ) mfix -> + All (on_def (lift_sorting (Prop_conj checking PcheckΣ) (Prop_conj infering_sort PsortΣ) Σ) (fix_context mfix) Γ) mfix -> wf_fixpoint Σ mfix -> Pinfer Γ (tFix mfix n) (dtype decl)) -> (forall (Γ : context) (mfix : mfixpoint term) (n : nat) decl, cofix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> - All (on_def_type (lift_sorting (Prop_conj checking PcheckΣ) (Prop_conj infering_sort PsortΣ) Σ) Γ) mfix -> - All (on_def_body (lift_sorting (Prop_conj checking PcheckΣ) (Prop_conj infering_sort PsortΣ) Σ) (fix_context mfix) Γ) mfix -> + All (on_def (lift_sorting (Prop_conj checking PcheckΣ) (Prop_conj infering_sort PsortΣ) Σ) (fix_context mfix) Γ) mfix -> wf_cofixpoint Σ mfix -> Pinfer Γ (tCoFix mfix n) (dtype decl)) -> @@ -498,7 +487,7 @@ Section BidirectionalInduction. - eapply HΓ. dependent induction wfΓ. + constructor. - + destruct t0 as (s & e & Hs). + + destruct t0 as (Hb & s & e & Hs). constructor. * apply IHwfΓ. intros ; apply IH. @@ -508,18 +497,18 @@ Section BidirectionalInduction. cbn. assert (0 < wfl_size wfΓ) by apply All_local_env_size_pos. lia. - + destruct t0 as (s & e & Hs). + + destruct t0 as ((Hb & mk) & s & e & Hs). constructor. * apply IHwfΓ. intros ; apply IH. cbn in H |- *. pose proof (infering_sort_size_pos Hs). lia. * red. applyIH. pose proof (infering_sort_size_pos Hs). cbn in H |- *. lia. - * red. applyIH. pose proof (checking_size_pos t1). cbn in H |- *. lia. + * red. applyIH. pose proof (checking_size_pos Hb). cbn in H |- *. lia. - eapply HΓRel. dependent induction wfΓ'. + constructor. - + destruct t0 as (s & e & Hs). + + destruct t0 as (Hb & s & e & Hs). constructor. * apply IHwfΓ'. intros ; apply IH. @@ -532,7 +521,7 @@ Section BidirectionalInduction. assert (0 < wfl_size_rel wfΓ') by apply All_local_env_size_pos. unfold wfl_size_rel in H. lia. - + destruct t0 as (s & e & Hs). + + destruct t0 as ((Hb & mk) & s & e & Hs). constructor. * apply IHwfΓ'. intros ; apply IH. @@ -540,7 +529,7 @@ Section BidirectionalInduction. fold (wfl_size_rel wfΓ'). pose proof (infering_sort_size_pos Hs). lia. * red. applyIH. pose proof (infering_sort_size_pos Hs). cbn in H |- *. lia. - * red. applyIH. pose proof (checking_size_pos t1). cbn in H |- *. lia. + * red. applyIH. pose proof (checking_size_pos Hb). cbn in H |- *. lia. - destruct c. @@ -557,9 +546,15 @@ Section BidirectionalInduction. all: applyIH. - unshelve eapply HLambda ; auto. + repeat split; eauto. + eexists; repeat split; eauto. + unfold PsortΣ. all: applyIH. - unshelve eapply HLetIn ; auto. + repeat split; eauto. + 2: eexists; repeat split; eauto. + 1,2: unfold PcheckΣ, PsortΣ. all: applyIH. - eapply HApp ; eauto. @@ -617,76 +612,56 @@ Section BidirectionalInduction. - unshelve eapply HFix ; eauto. - all: cbn [typing_sum_size infering_size] in IH. - all: remember (fix_context mfix) as mfixcontext. - all: unfold on_def_body, PcheckΣ in a0 |- *; cbn in a0 |- *. - + remember (all_size _ _ a0) as size. - clear -IH. - dependent induction a. - 1: by constructor. - constructor. - * destruct p as (s & e & Hs). - exists s; - repeat split; [apply e|apply Hs|hnf]. - applyIH. - * apply (IHa size). - intros. - apply IH. - cbn [all_size]. lia. - - + remember (all_size _ (fun d p => _ p.π2.2) a) as size. - clear -IH. - induction a0. - 1: by constructor. - constructor. - * intuition. - applyIH. - * apply IHa0. - intros. - apply IH. - cbn. lia. + cbn [typing_sum_size infering_size] in IH. + remember (fix_context mfix) as mfixcontext. + unfold on_def, PcheckΣ in a |- *; cbn in a |- *. + clear -IH a. + dependent induction a. + 1: by constructor. + constructor. + * destruct p as (((Hb & mk) & _s & _e & _Hs) & tt & s & e & Hs). + repeat split; [apply Hb|hnf|apply mk|..]. + 2:exists _s; + repeat split; [apply _e|apply _Hs|hnf]. + 3:exists s; + repeat split; [apply e|apply Hs|hnf]. + all: applyIH. + * apply IHa. + intros. + apply IH. + cbn [all_size]. lia. - unshelve eapply HCoFix ; eauto. - all: cbn [typing_sum_size infering_size] in IH. - all: remember (fix_context mfix) as mfixcontext. - all: unfold on_def_body, PcheckΣ in a0 |- *; cbn in a0 |- *. - + remember (all_size _ _ a0) as size. - clear -IH. - dependent induction a. - 1: by constructor. - constructor. - * destruct p as (s & e & Hs). - exists s; - repeat split; [apply e|apply Hs|hnf]. - applyIH. - * apply (IHa size). - intros. - apply IH. - cbn [all_size]. lia. - - + remember (all_size _ (fun d p => _ p.π2.2) a) as size. - clear -IH. - induction a0 as [| ? ? ?]. - 1: by constructor. - constructor. - * intuition. - applyIH. - * apply IHa0. - intros. - apply IH. - cbn. lia. + cbn [typing_sum_size infering_size] in IH. + remember (fix_context mfix) as mfixcontext. + unfold on_def, PcheckΣ in a |- *; cbn in a |- *. + clear -IH a. + dependent induction a. + 1: by constructor. + constructor. + * destruct p as (((Hb & mk) & _s & _e & _Hs) & tt & s & e & Hs). + repeat split; [apply Hb|hnf|apply mk|..]. + 2:exists _s; + repeat split; [apply _e|apply _Hs|hnf]. + 3:exists s; + repeat split; [apply e|apply Hs|hnf]. + all: applyIH. + * apply IHa. + intros. + apply IH. + cbn [all_size]. lia. - destruct i. - unshelve (eapply HiSort ; try eassumption) ; try eassumption. + unshelve (eapply HiSort ; tea) ; tea. all:applyIH. - destruct i. - unshelve (eapply HiProd ; try eassumption) ; try eassumption. + unshelve (eapply HiProd ; tea) ; tea. all: applyIH. - destruct i. - unshelve (eapply HiInd ; try eassumption) ; try eassumption. + unshelve (eapply HiInd ; tea) ; tea. all: applyIH. Qed. diff --git a/pcuic/theories/Conversion/PCUICClosedConv.v b/pcuic/theories/Conversion/PCUICClosedConv.v index c63c2b0ee..85fa385a5 100644 --- a/pcuic/theories/Conversion/PCUICClosedConv.v +++ b/pcuic/theories/Conversion/PCUICClosedConv.v @@ -2,7 +2,7 @@ From Coq Require Import Morphisms. From MetaCoq.Template Require Import config utils. From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICCases PCUICInduction - PCUICLiftSubst PCUICUnivSubst PCUICSigmaCalculus PCUICClosed + PCUICLiftSubst PCUICUnivSubst PCUICSigmaCalculus PCUICClosed PCUICRelevanceTerm PCUICOnFreeVars PCUICTyping PCUICReduction PCUICGlobalEnv PCUICWeakeningEnv PCUICEquality. @@ -19,7 +19,8 @@ Proof. induction Δ; simpl; auto. destruct a as [na [b|] ty]; intros wfΓ wfctx; constructor; intuition auto. - exists s; auto. + split; cbn; auto. + exists s; auto. Qed. Lemma sorts_local_ctx_All_local_env {cf} P Σ Γ Δ s : @@ -31,7 +32,7 @@ Proof. destruct a as [na [b|] ty]; intros wfΓ wfctx; constructor; intuition eauto. destruct s => //. destruct wfctx; eauto. - destruct s => //. destruct wfctx. exists t; auto. + destruct s => //. destruct wfctx. split; cbn; auto. exists t; auto. Qed. Lemma type_local_ctx_Pclosed Σ Γ Δ s : @@ -42,7 +43,7 @@ Proof. destruct a as [? [] ?]; intuition auto. - apply Alli_app_inv; auto. constructor. simpl. rewrite List.rev_length. 2:constructor. - unfold closed_decl. unfold Pclosed in b0. simpl. + unfold closed_decl. destruct b as ((b0&_)&_). unfold Pclosed in b0. simpl. rewrite app_context_length in b0. now rewrite Nat.add_comm. - apply Alli_app_inv; auto. constructor. simpl. rewrite List.rev_length. 2:constructor. @@ -59,9 +60,9 @@ Proof. destruct a as [? [] ?]; intuition auto. - apply Alli_app_inv; eauto. constructor. simpl. rewrite List.rev_length. 2:constructor. - unfold closed_decl. unfold Pclosed in b0. simpl. + unfold closed_decl. destruct b as ((b0&_)&_). unfold Pclosed in b0. simpl. rewrite app_context_length in b0. now rewrite Nat.add_comm. - - destruct s as [|u us]; auto. destruct X as [X b]. + - destruct s as [|u us]; auto. destruct X as [X (_&b)]. apply Alli_app_inv; eauto. constructor. simpl. rewrite List.rev_length. 2:constructor. unfold closed_decl. unfold Pclosed in b. simpl. @@ -78,17 +79,23 @@ Proof. - apply Alli_app_inv; auto. constructor. simpl. rewrite List.rev_length. 2:constructor. unfold closed_decl. unfold Pclosed in l. simpl. red in l. - destruct l as [s H]. + destruct l as (_ & [s H]). now rewrite andb_true_r in H. - apply Alli_app_inv; auto. constructor. simpl. rewrite List.rev_length. 2:constructor. + destruct l as ((?&_)&_). now simpl. Qed. Lemma weaken_env_prop_closed {cf} : weaken_env_prop cumulSpec0 (lift_typing typing) (lift_typing (fun (_ : global_env_ext) (Γ : context) (t T : term) => closedn #|Γ| t && closedn #|Γ| T)). -Proof. repeat red. intros. destruct t; red in X0; eauto. Qed. +Proof. + intros Σ **. + destruct T => //. + apply on_triple_impl_id with X2 => // _T. + eauto with extends. +Qed. Lemma closedn_ctx_alpha {k ctx ctx'} : @@ -103,11 +110,10 @@ Qed. Lemma closedn_All_local_env (ctx : list context_decl) : All_local_env - (fun (Γ : context) (b : term) (t : typ_or_sort) => - closedn #|Γ| b && typ_or_sort_default (closedn #|Γ|) t true) ctx -> + (fun (Γ : context) => lift_wf_term (closedn #|Γ|)) ctx -> closedn_ctx 0 ctx. Proof. - induction 1; auto; rewrite closedn_ctx_cons IHX /=; now move/andP: t0 => []. + induction 1; auto; rewrite closedn_ctx_cons IHX /=; apply andb_and; move: t0 => [] //=. Qed. Lemma declared_minductive_closed_inds {cf} {Σ ind mdecl u} {wfΣ : wf Σ} : diff --git a/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v b/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v index 10fccbb3b..475ce7dd9 100644 --- a/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v +++ b/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v @@ -3,7 +3,7 @@ From Coq Require Import ssreflect CRelationClasses. From MetaCoq.Template Require Import utils config Universes uGraph. From MetaCoq.PCUIC Require Import PCUICAst PCUICOnOne PCUICAstUtils PCUICInduction PCUICLiftSubst PCUICEquality PCUICUnivSubst - PCUICCases PCUICRelevance PCUICCumulativity PCUICTyping + PCUICCases PCUICRelevance PCUICRelevanceTerm PCUICCumulativity PCUICTyping PCUICReduction PCUICWeakeningEnv PCUICClosed PCUICPosition PCUICGuardCondition. @@ -851,6 +851,24 @@ Proof. * now rewrite !subst_instance_make'_make in H. Qed. +Lemma isTermRel_subst_instance Σ Γ u t rel : + isTermRel Σ Γ t rel -> isTermRel Σ Γ t@[u] rel. +Proof. + intro h. + induction t using term_forall_list_ind in Γ, h |- *; depelim h. + all: try solve [ try rewrite H; econstructor => //; eauto ]. + - cbn. erewrite map_dname. econstructor. rewrite nth_error_map e => //. + - cbn. erewrite map_dname. econstructor. rewrite nth_error_map e => //. +Qed. + +Lemma isTermRelOpt_subst_instance Σ Γ u t relopt : + isTermRelOpt Σ (marks_of_context Γ) t relopt -> isTermRelOpt Σ (marks_of_context Γ@[u]) t@[u] relopt. +Proof. + rewrite marks_of_context_univ_subst. + destruct relopt => //. + apply isTermRel_subst_instance. +Qed. + Definition precompose_subst_instance_global__1 Σ Re Rle gr napp u i i' := fst (precompose_subst_instance_global Σ Re Rle gr napp u i i'). @@ -1616,8 +1634,8 @@ 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 infer_typing_sort_impl with _ tu; [apply relevance_subst_opt|]. - all: cbn in *; eauto. + - eapply on_sortrel_impl with _ tu => //?; [apply relevance_subst_opt|]; eauto. + - eapply on_triplefull_impl with _ tu => //?; [apply relevance_subst_opt|apply isTermRelOpt_subst_instance|..]; eauto. Qed. #[global] Hint Resolve All_local_env_over_subst_instance : univ_subst. @@ -1691,16 +1709,6 @@ Qed. Definition wf_global_ext {cf : checker_flags} Σ ext := wf_ext_wk (Σ, ext). -Lemma relevance_of_term_subst_instance Σ Γ u t rel : - isTermRel Σ Γ t rel -> isTermRel Σ Γ t@[u] rel. -Proof. - intro h. - induction t using term_forall_list_ind in Γ, h |- *; depelim h. - all: try solve [ try rewrite H; econstructor => //; eauto ]. - - erewrite map_dname. econstructor. rewrite nth_error_map e => //. - - erewrite map_dname. econstructor. rewrite nth_error_map e => //. -Qed. - Require Import Morphisms. Require Import ssreflect. Set SimplIsCbn. @@ -1970,6 +1978,23 @@ Section SubstIdentity. - rewrite product_subst_instance. f_equal; intuition eauto; now noconf b0; noconf b1. + - destruct X as (_ & s & _ & _ & Hs). + intuition auto. + + - destruct X as (_ & s & _ & _ & Hs). + f_equal; + intuition auto. + + - destruct X as (((_ & IHb) & _) & _). + intuition auto. + + - destruct X as (_ & s & _ & _ & IHs). + intuition auto. + + - destruct X as (((_ & IHb) & _) & s & _ & _ & IHs). + f_equal; + intuition auto. + - intuition auto. noconf a; noconf b; noconf b0. rewrite subst_instance_subst /= /subst1. repeat (f_equal; simpl; auto). @@ -2012,10 +2037,12 @@ Section SubstIdentity. rewrite /subst_instance_list. now rewrite map_rev Hpars. * rewrite [subst_instance_constr _ _]subst_instance_two. noconf Hi. now rewrite [subst_instance _ u]H. - - solve_all. now destruct a as (s & ? & ? & ?). - - clear X0. now eapply nth_error_all in X as (s & e & Hs & IHs & _). - - solve_all. now destruct a as (s & ? & ? & ?). - - clear X0. now eapply nth_error_all in X as (s & e & Hs & IHs & _). + - solve_all. + all: destruct X as ((((_ & Hb) & _) & _) & _ & ? & _ & _ & Hs); intuition auto. + - eapply nth_error_all in X as ((((_ & Hb) & _) & _) & _ & ? & _ & _ & Hs); tea; intuition auto. + - solve_all. + all: destruct X as ((((_ & Hb) & _) & _) & _ & ? & _ & _ & Hs); intuition auto. + - eapply nth_error_all in X as ((((_ & Hb) & _) & _) & _ & ? & _ & _ & Hs); tea; intuition auto. Qed. Lemma typed_subst_abstract_instance Σ Γ t T : diff --git a/pcuic/theories/Conversion/PCUICWeakeningConv.v b/pcuic/theories/Conversion/PCUICWeakeningConv.v index d6793f50f..bbb86f155 100644 --- a/pcuic/theories/Conversion/PCUICWeakeningConv.v +++ b/pcuic/theories/Conversion/PCUICWeakeningConv.v @@ -350,7 +350,7 @@ Qed. Lemma isType_on_free_vars {cf} {Σ : global_env_ext} {wfΣ : wf Σ} {Γ T} : isType Σ Γ T -> on_free_vars xpredT T. Proof. - intros (s & e & Hs). + intros (Hb & s & e & Hs). eapply subject_closed in Hs. rewrite closedP_on_free_vars in Hs. eapply on_free_vars_impl; tea => //. @@ -359,7 +359,7 @@ Qed. Lemma isType_on_ctx_free_vars {cf} {Σ : global_env_ext} {wfΣ : wf Σ} {Γ T} : isType Σ Γ T -> on_ctx_free_vars xpredT Γ. Proof. - intros (s & e & Hs). + intros (Hb & s & e & Hs). eapply typing_wf_local in Hs. eapply closed_wf_local in Hs; tea. eapply (closed_ctx_on_free_vars xpredT) in Hs. diff --git a/pcuic/theories/Conversion/PCUICWeakeningEnvConv.v b/pcuic/theories/Conversion/PCUICWeakeningEnvConv.v index 1cd60da16..7178b78b2 100644 --- a/pcuic/theories/Conversion/PCUICWeakeningEnvConv.v +++ b/pcuic/theories/Conversion/PCUICWeakeningEnvConv.v @@ -49,7 +49,7 @@ Qed. Section ExtendsWf. Context {cf : checker_flags}. Context {Pcmp: global_env_ext -> context -> conv_pb -> term -> term -> Type}. - Context {P: global_env_ext -> context -> term -> typ_or_sort -> Type}. + Context {P: global_env_ext -> context -> judgment -> Type}. Let wf := on_global_env Pcmp P. diff --git a/pcuic/theories/PCUICAst.v b/pcuic/theories/PCUICAst.v index 7e999a7c5..172530e17 100644 --- a/pcuic/theories/PCUICAst.v +++ b/pcuic/theories/PCUICAst.v @@ -544,6 +544,8 @@ Module PCUICEnvironment := Environment PCUICTerm. Export PCUICEnvironment. (* Do NOT `Include` this module, as this would sadly duplicate the rewrite database... *) +Module PCUICLookup := EnvironmentTyping.Lookup PCUICTerm PCUICEnvironment. +Include PCUICLookup. (** Decompose an arity into a context and a sort *) @@ -564,10 +566,57 @@ Definition inds ind u (l : list one_inductive_body) := end in aux (List.length l). +Definition mark_context := list relevance. +Definition marks_of_context (Γ: context) := List.map (binder_relevance ∘ decl_name) Γ. + +Inductive isTermRel (Σ : global_env) (Γ : mark_context) : term -> relevance -> Type := + | rel_Rel n rel : + nth_error Γ n = Some rel -> isTermRel Σ Γ (tRel n) rel + + | rel_Lambda na A t rel : + isTermRel Σ (Γ ,, na.(binder_relevance)) t rel -> isTermRel Σ Γ (tLambda na A t) rel + + | rel_LetIn na b B t rel : + isTermRel Σ (Γ ,, na.(binder_relevance)) t rel -> isTermRel Σ Γ (tLetIn na b B t) rel + + | rel_App t u rel : + isTermRel Σ Γ t rel -> isTermRel Σ Γ (tApp t u) rel + + | rel_Const kn u decl : + declared_constant Σ kn decl -> isTermRel Σ Γ (tConst kn u) decl.(cst_relevance) + + | rel_Construct ind i u mdecl idecl cdecl : + declared_constructor Σ (ind, i) mdecl idecl cdecl -> isTermRel Σ Γ (tConstruct ind i u) idecl.(ind_relevance) + + | rel_Case ci p c brs : + isTermRel Σ Γ (tCase ci p c brs) ci.(ci_relevance) + + | rel_Proj p u mdecl idecl cdecl pdecl : + declared_projection Σ p mdecl idecl cdecl pdecl -> isTermRel Σ Γ (tProj p u) pdecl.(proj_relevance) + + | rel_Fix mfix n def : + nth_error mfix n = Some def -> isTermRel Σ Γ (tFix mfix n) def.(dname).(binder_relevance) + + | rel_CoFix mfix n def : + nth_error mfix n = Some def -> isTermRel Σ Γ (tCoFix mfix n) def.(dname).(binder_relevance) + + | rel_Sort s : + isTermRel Σ Γ (tSort s) Relevant + + | rel_Prod na A B : + isTermRel Σ Γ (tProd na A B) Relevant + + | rel_Ind ind u : + isTermRel Σ Γ (tInd ind u) Relevant. + +Derive Signature for isTermRel. + + Module PCUICTermUtils <: TermUtils PCUICTerm PCUICEnvironment. Definition destArity := destArity. Definition inds := inds. +Definition isTermRel := isTermRel. End PCUICTermUtils. @@ -1308,9 +1357,6 @@ Proof. + constructor. Qed. -Module PCUICLookup := EnvironmentTyping.Lookup PCUICTerm PCUICEnvironment. -Include PCUICLookup. - Derive NoConfusion for global_decl. Module PCUICGlobalMaps := EnvironmentTyping.GlobalMaps diff --git a/pcuic/theories/PCUICGeneration.v b/pcuic/theories/PCUICGeneration.v index 4da6b95dd..25bfb2067 100644 --- a/pcuic/theories/PCUICGeneration.v +++ b/pcuic/theories/PCUICGeneration.v @@ -29,11 +29,11 @@ Section Generation. Proof using Type. intros Ht Hsp. revert t Ht. induction Hsp; simpl; auto. - intros t Ht. eapply type_Cumul; eauto. eapply i.π2. + intros t Ht. eapply type_Cumul; eauto. eapply i.2.π2. intros. specialize (IHHsp (tApp t0 hd)). apply IHHsp. - destruct i as (s & e & Hs). + destruct i as (Hb & s & e & Hs). eapply type_App; eauto. eapply type_Cumul; eauto. Qed. @@ -49,12 +49,12 @@ Section Generation. - simpl. cbn. eapply ih. simpl in h. pose proof (typing_wf_local h) as hc. dependent induction hc. - destruct t0 as (s & e & Hs). + destruct t0 as ((Hb & mk) & s & e & Hs). econstructor; try eassumption. - simpl. cbn. eapply ih. pose proof (typing_wf_local h) as hc. cbn in hc. dependent induction hc. - destruct t0 as (s & e & Hs). + destruct t0 as (Hb & s & e & Hs). econstructor; try eassumption. Qed. diff --git a/pcuic/theories/PCUICRelevance.v b/pcuic/theories/PCUICRelevance.v index 4ee26fd9b..8723cc445 100644 --- a/pcuic/theories/PCUICRelevance.v +++ b/pcuic/theories/PCUICRelevance.v @@ -4,51 +4,6 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils. Require Import ssreflect. -Definition mark_context := list relevance. -Definition marks_of_context (Γ: context) := List.map (binder_relevance ∘ decl_name) Γ. - -Inductive isTermRel (Σ : global_env) (Γ : mark_context) : term -> relevance -> Type := - | rel_Rel n rel : - nth_error Γ n = Some rel -> isTermRel Σ Γ (tRel n) rel - - | rel_Lambda na A t rel : - isTermRel Σ (Γ ,, na.(binder_relevance)) t rel -> isTermRel Σ Γ (tLambda na A t) rel - - | rel_LetIn na b B t rel : - isTermRel Σ (Γ ,, na.(binder_relevance)) t rel -> isTermRel Σ Γ (tLetIn na b B t) rel - - | rel_App t u rel : - isTermRel Σ Γ t rel -> isTermRel Σ Γ (tApp t u) rel - - | rel_Const kn u decl : - declared_constant Σ kn decl -> isTermRel Σ Γ (tConst kn u) decl.(cst_relevance) - - | rel_Construct ind i u mdecl idecl cdecl : - declared_constructor Σ (ind, i) mdecl idecl cdecl -> isTermRel Σ Γ (tConstruct ind i u) idecl.(ind_relevance) - - | rel_Case ci p c brs : - isTermRel Σ Γ (tCase ci p c brs) ci.(ci_relevance) - - | rel_Proj p u mdecl idecl cdecl pdecl : - declared_projection Σ p mdecl idecl cdecl pdecl -> isTermRel Σ Γ (tProj p u) pdecl.(proj_relevance) - - | rel_Fix mfix n def : - nth_error mfix n = Some def -> isTermRel Σ Γ (tFix mfix n) def.(dname).(binder_relevance) - - | rel_CoFix mfix n def : - nth_error mfix n = Some def -> isTermRel Σ Γ (tCoFix mfix n) def.(dname).(binder_relevance) - - | rel_Sort s : - isTermRel Σ Γ (tSort s) Relevant - - | rel_Prod na A B : - isTermRel Σ Γ (tProd na A B) Relevant - - | rel_Ind ind u : - isTermRel Σ Γ (tInd ind u) Relevant. - -Derive Signature for isTermRel. - Lemma mark_of_context_app Γ Γ' Δ Δ' : marks_of_context Γ = marks_of_context Γ' -> marks_of_context Δ = marks_of_context Δ' -> marks_of_context (Γ,,, Δ) = marks_of_context (Γ',,, Δ'). Proof. diff --git a/pcuic/theories/PCUICRelevanceTerm.v b/pcuic/theories/PCUICRelevanceTerm.v index a7e5bbb04..afbcf8891 100644 --- a/pcuic/theories/PCUICRelevanceTerm.v +++ b/pcuic/theories/PCUICRelevanceTerm.v @@ -20,14 +20,27 @@ Lemma mark_inst_case_branch_context p br : marks_of_context (inst_case_branch_context p br) = marks_of_context br.(bcontext). Proof. apply mark_inst_case_context. Qed. -Lemma extends_irrelevant {cf : checker_flags} {Pcmp P} Σ Σ' Γ t : +Lemma mark_fix_context f g mfix : + marks_of_context (fix_context (map (map_def f g) mfix)) = marks_of_context (fix_context mfix). +Proof. + rewrite /fix_context /marks_of_context !mapi_map !map_rev !map_mapi //. +Qed. + +Lemma marks_of_context_univ_subst Γ u : marks_of_context Γ@[u] = marks_of_context Γ. +Proof. + rewrite /marks_of_context /subst_instance /subst_instance_context /map_context map_map //=. +Qed. + +Lemma extends_isTermRelOpt {cf : checker_flags} {Pcmp P} Σ Σ' Γ t relopt : on_global_env Pcmp P Σ' -> extends Σ Σ' -> - isTermRel Σ Γ t Irrelevant -> - isTermRel Σ' Γ t Irrelevant. + PCUICEnvTyping.isTermRelOpt Σ Γ t relopt -> + PCUICEnvTyping.isTermRelOpt Σ' Γ t relopt. Proof. + destruct relopt => //=. induction t in Γ |- * using term_forall_list_ind; - intros wfΣ' ext Hirr; depelim Hirr; try econstructor; eauto. - all: solve [ rewrite H; econstructor => //; eauto with extends ]. + intros wfΣ' ext Hirr; depelim Hirr; try econstructor; eauto with extends. Qed. +#[global] Hint Resolve extends_isTermRelOpt : extends. + diff --git a/pcuic/theories/PCUICTyping.v b/pcuic/theories/PCUICTyping.v index 2da8507f8..99c43b77a 100644 --- a/pcuic/theories/PCUICTyping.v +++ b/pcuic/theories/PCUICTyping.v @@ -192,13 +192,16 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> Σ ;;; Γ |- tProd na A B : tSort (Universe.sort_of_product s1 s2) | type_Lambda : forall na A t s1 B, + (* on_decl (lift_typing typing Σ) Γ (vass na A) -> *) isSortRel s1 na.(binder_relevance) -> Σ ;;; Γ |- A : tSort s1 -> Σ ;;; Γ ,, vass na A |- t : B -> Σ ;;; Γ |- tLambda na A t : tProd na A B | type_LetIn : forall na b B t s1 A, + (* on_decl (lift_typing typing Σ) Γ (vdef na b B) -> *) isSortRel s1 na.(binder_relevance) -> + isTermRel Σ (marks_of_context Γ) b na.(binder_relevance) -> Σ ;;; Γ |- B : tSort s1 -> Σ ;;; Γ |- b : B -> Σ ;;; Γ ,, vdef na b B |- t : A -> @@ -252,8 +255,7 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> wf_local Σ Γ -> fix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> - All (on_def_type (lift_typing typing Σ) Γ) mfix -> - All (on_def_body (lift_typing typing Σ) (fix_context mfix) Γ) mfix -> + All (on_def (lift_typing typing Σ) (fix_context mfix) Γ) mfix -> wf_fixpoint Σ mfix -> Σ ;;; Γ |- tFix mfix n : decl.(dtype) @@ -261,8 +263,7 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> wf_local Σ Γ -> cofix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> - All (on_def_type (lift_typing typing Σ) Γ) mfix -> - All (on_def_body (lift_typing typing Σ) (fix_context mfix) Γ) mfix -> + All (on_def (lift_typing typing Σ) (fix_context mfix) Γ) mfix -> wf_cofixpoint Σ mfix -> Σ ;;; Γ |- tCoFix mfix n : decl.(dtype) @@ -299,6 +300,24 @@ Proof. intros h [] [] []; assumption. Qed. +Lemma type_Lambda' {cf} {Σ Γ na A t B} : + on_decl (lift_typing typing Σ) Γ (vass na A) -> + Σ ;;; Γ ,, vass na A |- t : B -> + Σ ;;; Γ |- tLambda na A t : tProd na A B. +Proof. + intros (_ & s1 & e & Hs) Ht. + now econstructor. +Qed. + +Lemma type_LetIn' {cf} {Σ Γ na b B t A} : + on_decl (lift_typing typing Σ) Γ (vdef na b B) -> + Σ ;;; Γ ,, vdef na b B |- t : A -> + Σ ;;; Γ |- tLetIn na b B t : tLetIn na b B A. +Proof. + intros ((Hb & mk) & s1 & e & Hs) Ht. + now econstructor. +Qed. + (** ** Typechecking of global environments *) Definition has_nparams npars ty := @@ -335,7 +354,7 @@ Arguments iswelltyped {cf Σ Γ t A} h. Definition isType_welltyped {cf Σ} {Γ T} : isType Σ Γ T -> welltyped Σ Γ T. Proof. - intros []. now econstructor. + intros [_ []]. now econstructor. Qed. Definition isTypeRel_welltyped {cf Σ} {Γ T rel} @@ -405,10 +424,8 @@ Proof. - exact (S (Nat.max (All_local_env_size typing_size _ _ wf_pctx) (Nat.max (ctx_inst_size _ typing_size ind_inst) (Nat.max d2 (Nat.max d3 (branches_size typing_size brs_ty)))))). - - exact (S (Nat.max (Nat.max (All_local_env_size typing_size _ _ a) - (all_size _ (fun x p => (lift_typing_size typing_size) Σ _ _ _ p) a0)) (all_size _ (fun x p => typing_size Σ _ _ _ p) a1))). - - exact (S (Nat.max (Nat.max (All_local_env_size typing_size _ _ a) - (all_size _ (fun x p => (lift_typing_size typing_size) Σ _ _ _ p) a0)) (all_size _ (fun x p => typing_size Σ _ _ _ p) a1))). + - exact (S (Nat.max (All_local_env_size typing_size _ _ a) (all_size _ (fun x p => (lift_typing_size typing_size Σ _ _ p.1 + lift_typing_size typing_size Σ _ _ p.2)%nat) a0))). + - exact (S (Nat.max (All_local_env_size typing_size _ _ a) (all_size _ (fun x p => (lift_typing_size typing_size Σ _ _ p.1 + lift_typing_size typing_size Σ _ _ p.2)%nat) a0))). Defined. Lemma typing_size_pos `{checker_flags} {Σ Γ t T} (d : Σ ;;; Γ |- t : T) : typing_size d > 0. @@ -525,7 +542,7 @@ Lemma type_Cumul' {cf:checker_flags} {Σ Γ t} T {T'} : Σ ;;; Γ |- T <=s T' -> Σ ;;; Γ |- t : T'. Proof. - intros Ht (s & e & Hs) cum. + intros Ht (Hb & s & e & Hs) cum. eapply type_Cumul; eauto. Qed. @@ -571,18 +588,17 @@ Proof. destruct w. - simpl. congruence. - intros [=]. subst d Γ0. - destruct l as (s & e & Hs). + destruct l as (Hb & s & e & Hs). exists w. simpl. exists s. exists Hs. pose (typing_size_pos Hs). cbn. split. + lia. + auto with arith. - intros [=]. subst d Γ0. - destruct l as (s & e & Hs). + destruct l as ((Hb & mk) & s & e & Hs). exists w. simpl. - simpl in l0. - exists s, l0, Hs. cbn. + exists s, Hb, Hs. cbn. pose (typing_size_pos Hs). - pose (typing_size_pos l0). + pose (typing_size_pos Hb). intuition eauto. all: try lia. Qed. @@ -623,22 +639,14 @@ Lemma typing_ind_env_app_size `{cf : checker_flags} : Σ ;;; Γ,, vass na t |- b : tSort s2 -> P Σ (Γ,, vass na t) b (tSort s2) -> P Σ Γ (tProd na t b) (tSort (Universe.sort_of_product s1 s2))) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (t b : term) - (s1 : Universe.t) (bty : term), - isSortRel s1 na.(binder_relevance) -> + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (t b : term) (bty : term), PΓ Σ Γ -> - Σ ;;; Γ |- t : tSort s1 -> - P Σ Γ t (tSort s1) -> + on_decl (lift_typing2 typing P Σ) Γ (vass na t) -> Σ ;;; Γ,, vass na t |- b : bty -> P Σ (Γ,, vass na t) b bty -> P Σ Γ (tLambda na t b) (tProd na t bty)) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (b b_ty b' : term) - (s1 : Universe.t) (b'_ty : term), - isSortRel s1 na.(binder_relevance) -> + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (b b_ty b' : term) (b'_ty : term), PΓ Σ Γ -> - Σ ;;; Γ |- b_ty : tSort s1 -> - P Σ Γ b_ty (tSort s1) -> - Σ ;;; Γ |- b : b_ty -> - P Σ Γ b b_ty -> + on_decl (lift_typing2 typing P Σ) Γ (vdef na b b_ty) -> Σ ;;; Γ,, vdef na b b_ty |- b' : b'_ty -> P Σ (Γ,, vdef na b b_ty) b' b'_ty -> P Σ Γ (tLetIn na b b_ty b') (tLetIn na b b_ty b'_ty)) -> @@ -718,8 +726,7 @@ Lemma typing_ind_env_app_size `{cf : checker_flags} : fix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> PΓ Σ (Γ ,,, types) -> - All (on_def_type (lift_typing2 typing P Σ) Γ) mfix -> - All (on_def_body (lift_typing2 typing P Σ) types Γ) mfix -> + All (on_def (lift_typing2 typing P Σ) types Γ) mfix -> wf_fixpoint Σ mfix -> P Σ Γ (tFix mfix n) decl.(dtype)) -> @@ -728,8 +735,7 @@ Lemma typing_ind_env_app_size `{cf : checker_flags} : cofix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> PΓ Σ (Γ ,,, types) -> - All (on_def_type (lift_typing2 typing P Σ) Γ) mfix -> - All (on_def_body (lift_typing2 typing P Σ) types Γ) mfix -> + All (on_def (lift_typing2 typing P Σ) types Γ) mfix -> wf_cofixpoint Σ mfix -> P Σ Γ (tCoFix mfix n) decl.(dtype)) -> @@ -787,12 +793,10 @@ Proof. * simpl. simpl in *. destruct d. - + destruct Xg; split. 2: destruct c, cst_body0 => //. - 1: rename o1 into Xg. 2: rename o2 into Xg. - all: apply lift_typing_impl with (1 := Xg); intros ? Hs. - all: specialize (IH ((Σ', udecl); wfΣ; _; _; _; Hs)). - all: forward IH; [constructor 1; simpl; subst Σ' Σg; cbn; lia|]. - all: apply IH. + + apply lift_typing_impl with (1 := Xg) => // ?? Hs. + specialize (IH ((Σ', udecl); wfΣ; _; _; _; Hs)). + forward IH; [constructor 1; simpl; subst Σ' Σg; cbn; lia|]. + apply IH. + red in Xg. destruct Xg as [onI onP onnp]; constructor; eauto. { unshelve eset (IH' := fun p => IH ((Σ', udecl); wfΣ; p) _). @@ -802,36 +806,36 @@ Proof. refine {| ind_arity_eq := Xg.(ind_arity_eq); ind_cunivs := Xg.(ind_cunivs) |}. - apply onArity in Xg. - apply lift_typing_impl with (1 := Xg); intros ? Hs. + apply lift_typing_impl with (1 := Xg) => // ?? Hs. apply (IH (_; _; _; Hs)). - pose proof Xg.(onConstructors) as Xg'. eapply All2_impl; eauto. intros. destruct X as [cass tyeq onctyp oncargs oncind]. unshelve econstructor; eauto. - { apply lift_typing_impl with (1 := onctyp); intros ? Hs. + { apply lift_typing_impl with (1 := onctyp) => // ?? Hs. apply (IH (_; _; _; Hs)). } - { apply sorts_local_ctx_impl with (1 := oncargs); intros Γ' t' T' HT'. - apply lift_typing_impl with (1 := HT'); intros ? Hs. + { apply sorts_local_ctx_impl with (1 := oncargs) => // Γ' j Hj. + apply lift_typing_impl with (1 := Hj) => // ?? Hs. apply (IH (_; _; _; Hs)). } { clear -IH oncind. revert oncind. generalize (List.rev (lift_context #|cstr_args x0| 0 (ind_indices x))). generalize (cstr_indices x0). induction 1; constructor; auto. - do 2 red in t0 |- *. + do 3 red in t0 |- *. apply (IH (_; (_; (_; t0)))). } - intros Hprojs; pose proof (onProjections Xg Hprojs); auto. - destruct Xg. simpl. unfold check_ind_sorts in *. destruct Universe.is_prop; auto. destruct Universe.is_sprop; auto. split. apply ind_sorts. destruct indices_matter; auto. - eapply type_local_ctx_impl. eapply ind_sorts. intros ??? HT. - apply lift_typing_impl with (1 := HT); intros ? Hs. + eapply type_local_ctx_impl. eapply ind_sorts. auto. intros ?? HT. + apply lift_typing_impl with (1 := HT) => // ?? Hs. apply (IH (_; _; _; Hs)). - apply Xg.(ind_relevance_compat). - apply Xg.(onIndices). } { red in onP |- *. - apply All_local_env_impl with (1 := onP); intros ??? HT. - apply lift_typing_impl with (1 := HT); intros ? Hs. + apply All_local_env_impl with (1 := onP) => ?? HT. + apply lift_typing_impl with (1 := HT) => // ?? Hs. apply (IH ((Σ', udecl); (wfΣ; _; _; _; Hs))). constructor 1. simpl. subst Σ' Σg; cbn; lia. } @@ -855,11 +859,11 @@ Proof. { intros. eapply (XΓ _ _ _ Hwf); auto. clear -Pdecl wfΣ X14 H0. revert Hwf H0. - induction Hwf; cbn in *; try destruct t2 as (u & e & Hu); simpl in *; constructor. + induction Hwf; cbn in *; try destruct t2 as (Hb & u & e & Hu); simpl in *; constructor. - simpl in *. apply IHHwf. lia. - red. apply (X14 _ _ _ Hu). lia. - simpl in *. apply IHHwf. lia. - - red. apply (X14 _ _ _ t3). lia. + - red. apply (X14 _ _ _ Hb.1). lia. - red. simpl. apply (X14 _ _ _ Hu). lia. } assert (Htywf : forall Γ' t T (Hty : Σ ;;; Γ' |- t : T), @@ -881,10 +885,18 @@ Proof. unshelve eapply X14; simpl; auto with arith]. -- match reverse goal with - H : _ |- _ => eapply H - end; eauto; - unshelve eapply X14; simpl; eauto with arith wf. - + H : _ |- _ => eapply H + end; eauto. + repeat split; eauto. eexists; repeat split; eauto. + all: unshelve eapply X14; simpl; [ auto | lia ]. + + -- match reverse goal with + H : _ |- _ => eapply H + end; eauto. + repeat split; eauto. + 2: eexists; repeat split; eauto. + all: unshelve eapply X14; simpl; [ auto | lia ]. + -- match reverse goal with H : _ |- _ => eapply H end; eauto. all:try unshelve eapply X14; simpl; auto; try lia. @@ -951,77 +963,56 @@ Proof. -- clear X X0 X1 X2 X3 X4 X5 X6 X7 X8 X9 X11 X12. eapply X10; eauto; clear X10. simpl in *. - * assert(forall Γ0 (t T : term) (Hty : Σ;;; Γ0 |- t : T), - typing_size Hty < S (all_size _ (fun x p => typing_size p) a1) -> - PΓ Σ Γ0). - { intros. eapply (Htywf _ _ _ Hty); eauto. lia. } + * assert(forall Γ (t T : term) (Hty : Σ;;; Γ |- t : T), + typing_size Hty < S (all_size _ (fun x p => (lift_typing_size (@typing_size _) Σ _ _ p.1 + lift_typing_size (@typing_size _) Σ _ _ p.2)%nat) a0) -> + PΓ Σ Γ). + { intros. eapply (Htywf _ _ _ Hty); eauto. cbn in H. lia. } destruct mfix. now rewrite nth_error_nil in e. - depelim a1. - eapply (X _ _ _ o). simpl. lia. - * assert(forall (t T : term) (Hty : Σ;;; Γ |- t : T), - typing_size Hty < S (all_size _ (fun x '(s; (e, d)) => typing_size d) a0) -> + depelim a0. + eapply (X _ _ _ o.1.1.1). simpl. lia. + * assert(forall Γ (t T : term) (Hty : Σ;;; Γ |- t : T), + typing_size Hty < S (all_size _ (fun x p => (lift_typing_size (@typing_size _) Σ _ _ p.1 + lift_typing_size (@typing_size _) Σ _ _ p.2)%nat) a0) -> Forall_decls_typing P Σ.1 * P Σ Γ t T). - intros; eauto. eapply (X14 _ _ _ Hty); eauto. simpl in *. unfold infer_sort. cbn. lia. + intros; eauto. eapply (X14 _ _ _ Hty); eauto. simpl in *. lia. clear Hwf Htywf X14 a pΓ Hdecls. clear -a0 X. + remember (fix_context mfix) as mfixcontext. clear Heqmfixcontext. induction a0; constructor. - destruct p as (s & e & Hs). exists s; repeat split; auto. - apply (X (dtype x) (tSort s) Hs). simpl. lia. - apply IHa0. intros. eapply (X _ _ Hty); eauto. + destruct p as (((Hb & mk) & _s & _e & _Hs) & tt & s & e & Hs). + repeat split; auto. + 1: apply (X _ _ _ Hb); simpl; lia. + exists _s; repeat split; auto. + apply (X _ _ (tSort _s) _Hs). simpl. lia. + exists s; repeat split; auto. + apply (X _ _ (tSort s) Hs). simpl. lia. + apply IHa0. intros. eapply (X _ _ _ Hty); eauto. cbn [all_size]. lia. - * simpl in X14. - assert(forall Γ0 : context, - wf_local Σ Γ0 -> - forall (t T : term) (Hty : Σ;;; Γ0 |- t : T), - typing_size Hty < S (all_size _ (fun x p => typing_size p) a1) -> - Forall_decls_typing P Σ.1 * P Σ Γ0 t T). - { intros. eapply (X14 _ _ _ Hty); eauto. lia. } - clear X14 Hwf Htywf. - clear e decl f a0 Hdecls i. - remember (fix_context mfix) as mfixcontext. clear Heqmfixcontext. - - induction a1; econstructor; eauto. - ++ split; auto. - eapply (X _ (typing_wf_local p) _ _ p). simpl. lia. - ++ eapply IHa1. intros. - eapply (X _ X0 _ _ Hty). simpl; lia. - -- clear X X0 X1 X2 X3 X4 X5 X6 X7 X8 X9 X10 X12. eapply X11; eauto; clear X11. simpl in *. - * assert(forall Γ0 (t T : term) (Hty : Σ;;; Γ0 |- t : T), - typing_size Hty < S (all_size _ (fun x p => typing_size p) a1) -> - PΓ Σ Γ0). - { intros. eapply (Htywf _ _ _ Hty); eauto. lia. } + * assert(forall Γ (t T : term) (Hty : Σ;;; Γ |- t : T), + typing_size Hty < S (all_size _ (fun x p => (lift_typing_size (@typing_size _) Σ _ _ p.1 + lift_typing_size (@typing_size _) Σ _ _ p.2)%nat) a0) -> + PΓ Σ Γ). + { intros. eapply (Htywf _ _ _ Hty); eauto. cbn in H. lia. } destruct mfix. now rewrite nth_error_nil in e. - depelim a1. - eapply (X _ _ _ o). simpl. lia. - * assert(forall (t T : term) (Hty : Σ;;; Γ |- t : T), - typing_size Hty < S (all_size _ (fun x '(s; (e, d)) => typing_size d) a0) -> - Forall_decls_typing P Σ.1 * P Σ Γ t T). - intros; eauto. eapply (X14 _ _ _ Hty); eauto. simpl in *. unfold infer_sort. cbn. lia. + depelim a0. + eapply (X _ _ _ o.1.1.1). simpl. lia. + * assert(forall Γ (t T : term) (Hty : Σ;;; Γ |- t : T), + typing_size Hty < S (all_size _ (fun x p => (lift_typing_size (@typing_size _) Σ _ _ p.1 + lift_typing_size (@typing_size _) Σ _ _ p.2)%nat) a0) -> + Forall_decls_typing P Σ.1 * P Σ Γ t T). + intros; eauto. eapply (X14 _ _ _ Hty); eauto. simpl in *. lia. clear Hwf Htywf X14 a pΓ Hdecls. clear -a0 X. + remember (fix_context mfix) as mfixcontext. clear Heqmfixcontext. induction a0; constructor. - destruct p as (s & e & Hs). exists s; repeat split; auto. - apply (X (dtype x) (tSort s) Hs). simpl. lia. - apply IHa0. intros. eapply (X _ _ Hty); eauto. + destruct p as (((Hb & mk) & _s & _e & _Hs) & tt & s & e & Hs). + repeat split; auto. + 1: apply (X _ _ _ Hb). simpl. lia. + exists _s; repeat split; auto. + 1: apply (X _ _ _ _Hs). simpl. lia. + exists s; repeat split; auto. + 1: apply (X _ _ _ Hs). simpl. lia. + apply IHa0. intros. eapply (X _ _ _ Hty); eauto. cbn [all_size]. lia. - * simpl in X14. - assert(forall Γ0 : context, - wf_local Σ Γ0 -> - forall (t T : term) (Hty : Σ;;; Γ0 |- t : T), - typing_size Hty < S (all_size _ (fun x p => typing_size p) a1) -> - Forall_decls_typing P Σ.1 * P Σ Γ0 t T). - { intros. eapply (X14 _ _ _ Hty); eauto. lia. } - clear X14 Hwf Htywf. - clear e decl c a0 Hdecls i. - remember (fix_context mfix) as mfixcontext. clear Heqmfixcontext. - - induction a1; econstructor; eauto. - ++ split; auto. - eapply (X _ (typing_wf_local p) _ _ p). simpl. lia. - ++ eapply IHa1. intros. - eapply (X _ X0 _ _ Hty). simpl; lia. Qed. Lemma typing_ind_env `{cf : checker_flags} : @@ -1050,22 +1041,14 @@ Lemma typing_ind_env `{cf : checker_flags} : Σ ;;; Γ,, vass na t |- b : tSort s2 -> P Σ (Γ,, vass na t) b (tSort s2) -> P Σ Γ (tProd na t b) (tSort (Universe.sort_of_product s1 s2))) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (t b : term) - (s1 : Universe.t) (bty : term), - isSortRel s1 na.(binder_relevance) -> + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (t b : term) (bty : term), PΓ Σ Γ -> - Σ ;;; Γ |- t : tSort s1 -> - P Σ Γ t (tSort s1) -> + on_decl (lift_typing2 typing P Σ) Γ (vass na t) -> Σ ;;; Γ,, vass na t |- b : bty -> P Σ (Γ,, vass na t) b bty -> P Σ Γ (tLambda na t b) (tProd na t bty)) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (b b_ty b' : term) - (s1 : Universe.t) (b'_ty : term), - isSortRel s1 na.(binder_relevance) -> + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (na : aname) (b b_ty b' : term) (b'_ty : term), PΓ Σ Γ -> - Σ ;;; Γ |- b_ty : tSort s1 -> - P Σ Γ b_ty (tSort s1) -> - Σ ;;; Γ |- b : b_ty -> - P Σ Γ b b_ty -> + on_decl (lift_typing2 typing P Σ) Γ (vdef na b b_ty) -> Σ ;;; Γ,, vdef na b b_ty |- b' : b'_ty -> P Σ (Γ,, vdef na b b_ty) b' b'_ty -> P Σ Γ (tLetIn na b b_ty b') (tLetIn na b b_ty b'_ty)) -> @@ -1141,8 +1124,7 @@ Lemma typing_ind_env `{cf : checker_flags} : fix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> PΓ Σ (Γ ,,, types) -> - All (on_def_type (lift_typing2 typing P Σ) Γ) mfix -> - All (on_def_body (lift_typing2 typing P Σ) types Γ) mfix -> + All (on_def (lift_typing2 typing P Σ) types Γ) mfix -> wf_fixpoint Σ mfix -> P Σ Γ (tFix mfix n) decl.(dtype)) -> @@ -1151,8 +1133,7 @@ Lemma typing_ind_env `{cf : checker_flags} : cofix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> PΓ Σ (Γ ,,, types) -> - All (on_def_type (lift_typing2 typing P Σ) Γ) mfix -> - All (on_def_body (lift_typing2 typing P Σ) types Γ) mfix -> + All (on_def (lift_typing2 typing P Σ) types Γ) mfix -> wf_cofixpoint Σ mfix -> P Σ Γ (tCoFix mfix n) decl.(dtype)) -> @@ -1194,14 +1175,14 @@ Section All_local_env. Context {cf: checker_flags}. - Lemma All_local_env_app (P : context -> term -> typ_or_sort -> Type) l l' : - All_local_env P l * All_local_env (fun Γ t T => P (l ,,, Γ) t T) l' -> + Lemma All_local_env_app (P : context -> judgment -> Type) l l' : + All_local_env P l * All_local_env (fun Γ T => P (l ,,, Γ) T) l' -> All_local_env P (l ,,, l'). Proof using Type. intros [H H']. apply All_local_rel_app => //. Qed. - Lemma All_local_env_app_inv (P : context -> term -> typ_or_sort -> Type) l l' : + Lemma All_local_env_app_inv (P : context -> judgment -> Type) l l' : All_local_env P (l ,,, l') -> - All_local_env P l * All_local_env (fun Γ t T => P (l ,,, Γ) t T) l'. + All_local_env P l * All_local_env (fun Γ T => P (l ,,, Γ) T) l'. Proof using Type. apply All_local_app_rel. Qed. Definition wf_local_rel_app_inv {Σ Γ1 Γ2 Γ3} : @@ -1218,7 +1199,7 @@ Section All_local_env. Lemma All_local_env_lookup {P Γ n} {decl} : All_local_env P Γ -> nth_error Γ n = Some decl -> - on_local_decl P (skipn (S n) Γ) decl. + on_decl P (skipn (S n) Γ) decl. Proof using Type. induction 1 in n, decl |- *. simpl. destruct n; simpl; congruence. destruct n. red. simpl. intros [= <-]. simpl. apply t0. @@ -1262,9 +1243,9 @@ Section All_local_env. apply wf_local_app; auto. Qed. - Lemma All_local_env_map (P : context -> term -> typ_or_sort -> Type) f l : + Lemma All_local_env_map (P : context -> judgment -> Type) f l : (forall u, f (tSort u) = tSort u) -> - All_local_env (fun Γ t T => P (map (map_decl f) Γ) (f t) (typ_or_sort_map f T)) l + All_local_env (fun Γ T => P (map (map_decl f) Γ) (judgment_map f T)) l -> All_local_env P (map (map_decl f) l). Proof using Type. intros Hf. induction 1; econstructor; eauto. @@ -1300,16 +1281,10 @@ Section All_local_env. now eapply wf_local_app_l in wf. Qed. - Definition on_local_decl_glob (P : term -> typ_or_sort -> Type) d := - match d.(decl_body) with - | Some b => P b (Typ d.(decl_type)) × P d.(decl_type) (SortRel d.(decl_name).(binder_relevance)) - | None => P d.(decl_type) (SortRel d.(decl_name).(binder_relevance)) - end. - Definition lookup_wf_local_decl {Γ P} (wfΓ : All_local_env P Γ) (n : nat) {decl} (eq : nth_error Γ n = Some decl) : ∑ Pskip : All_local_env P (skipn (S n) Γ), - on_local_decl_glob (P (skipn (S n) Γ)) decl. + on_decl P (skipn (S n) Γ) decl. Proof. induction wfΓ in n, decl, eq |- *; simpl. - elimtype False. destruct n; depelim eq. @@ -1318,17 +1293,17 @@ Section All_local_env. + apply IHwfΓ. auto with arith. - destruct n. + exists wfΓ. injection eq; intros <-. - simpl. split; auto. + simpl. auto. + apply IHwfΓ. apply eq. Defined. Definition on_wf_local_decl {Σ Γ} (P : forall Σ Γ (wfΓ : wf_local Σ Γ) t T, Σ ;;; Γ |- t : T -> Type) - (wfΓ : wf_local Σ Γ) {d} (H : on_local_decl_glob (lift_typing typing Σ Γ) d) := - match d as d' return (on_local_decl_glob (lift_typing typing Σ Γ) d') -> Type with + (wfΓ : wf_local Σ Γ) {d} (H : on_decl (lift_typing typing Σ) Γ d) := + match d return (on_decl (lift_typing typing Σ) Γ d) -> Type with | {| decl_name := na; decl_body := Some b; decl_type := ty |} => - fun H => (P Σ Γ wfΓ b ty H.1 * P Σ Γ wfΓ _ _ H.2.π2.2)%type - | {| decl_name := na; decl_body := None; decl_type := ty |} => fun H => P Σ Γ wfΓ _ _ H.π2.2 + fun H => P Σ Γ wfΓ b ty H.1.1 × P Σ Γ wfΓ _ _ H.2.π2.2 + | {| decl_name := na; decl_body := None; decl_type := ty |} => fun H => P Σ Γ wfΓ _ _ H.2.π2.2 end H. Lemma nth_error_All_local_env_over {P Σ Γ n decl} (eq : nth_error Γ n = Some decl) {wfΓ : wf_local Σ Γ} : @@ -1347,7 +1322,7 @@ Section All_local_env. Lemma All_local_env_prod_inv : forall P Q Γ, - All_local_env (fun Δ A t => P Δ A t × Q Δ A t) Γ -> + All_local_env (fun Δ j => P Δ j × Q Δ j) Γ -> All_local_env P Γ × All_local_env Q Γ. Proof using Type. intros P Q Γ h. @@ -1355,7 +1330,7 @@ Section All_local_env. - split ; constructor. - destruct IHh, t0. split ; constructor ; auto. - - destruct IHh, t0, t1. + - destruct IHh, t0. split ; constructor ; auto. Qed. @@ -1369,12 +1344,12 @@ Section All_local_env. - split ; constructor. - destruct IHh as [IHP IHQ]. split; constructor; tas. - all: eapply lift_typing_impl; tea. - all: now intros T [HP HQ]. + all: eapply lift_typing_impl; eauto. + all: now intros ?? [HP HQ]. - destruct IHh as [IHP IHQ]. split; constructor; tas. - all: eapply lift_typing_impl; unfold lift_typing2 in t0, t1; tea. - all: now intros T [HP HQ]. + all: eapply lift_typing_impl; unfold lift_typing2 in t0; eauto. + all: now intros ?? [HP HQ]. Qed. End All_local_env. diff --git a/pcuic/theories/PCUICWcbvEval.v b/pcuic/theories/PCUICWcbvEval.v index d67bdca09..2227919d6 100644 --- a/pcuic/theories/PCUICWcbvEval.v +++ b/pcuic/theories/PCUICWcbvEval.v @@ -625,8 +625,8 @@ Section Wcbv. Proof using Type. move=> wfΣ Hc Hb. rewrite PCUICClosed.closedn_subst_instance. - apply declared_decl_closed in Hc as (Hty & Hbo) => //. - rewrite Hb in Hbo. now move/andP: Hbo. + apply declared_decl_closed in Hc as (Hbo & Hty) => //. + now rewrite Hb in Hbo. Qed. Lemma closed_iota ci ind p c u args brs br : diff --git a/pcuic/theories/PCUICWeakeningEnv.v b/pcuic/theories/PCUICWeakeningEnv.v index 4d85bb618..7bce85063 100644 --- a/pcuic/theories/PCUICWeakeningEnv.v +++ b/pcuic/theories/PCUICWeakeningEnv.v @@ -260,7 +260,7 @@ Qed. Section ExtendsWf. Context {cf : checker_flags}. Context {Pcmp: global_env_ext -> context -> conv_pb -> term -> term -> Type}. - Context {P: global_env_ext -> context -> term -> typ_or_sort -> Type}. + Context {P: global_env_ext -> context -> judgment -> Type}. Let wf := on_global_env Pcmp P. @@ -477,12 +477,12 @@ Definition weaken_env_prop_full forall Γ t T, P Σ Γ t T -> P (Σ', Σ.2) Γ t T. Definition weaken_env_prop - (P : global_env_ext -> context -> term -> typ_or_sort -> Type) := - forall Σ Σ' φ, wf Σ -> wf Σ' -> extends Σ Σ' -> forall Γ t T, P (Σ, φ) Γ t T -> P (Σ', φ) Γ t T. + (P : global_env_ext -> context -> judgment -> Type) := + forall Σ Σ' φ, wf Σ -> wf Σ' -> extends Σ Σ' -> forall Γ T, P (Σ, φ) Γ T -> P (Σ', φ) Γ T. Definition weaken_env_decls_prop - (P : global_env_ext -> context -> term -> typ_or_sort -> Type) := - forall Σ Σ' φ, wf Σ' -> extends_decls Σ Σ' -> forall Γ t T, P (Σ, φ) Γ t T -> P (Σ', φ) Γ t T. + (P : global_env_ext -> context -> judgment -> Type) := + forall Σ Σ' φ, wf Σ' -> extends_decls Σ Σ' -> forall Γ T, P (Σ, φ) Γ T -> P (Σ', φ) Γ T. Lemma extends_decls_wf Σ Σ' : wf Σ' -> extends_decls Σ Σ' -> wf Σ. @@ -503,6 +503,7 @@ Arguments weaken_env_prop_full {cf} (Pcmp P)%function_scope _%function_scope. Arguments weaken_env_decls_prop {cf} (Pcmp P)%function_scope _%function_scope. Arguments weaken_env_prop {cf} (Pcmp P)%function_scope _%function_scope. +#[global] Hint Resolve extends_decls_extends : extends. #[global] Hint Resolve extends_lookup : extends. #[global] Hint Resolve weakening_env_declared_constant : extends. #[global] Hint Resolve weakening_env_declared_minductive : extends. diff --git a/pcuic/theories/Syntax/PCUICDepth.v b/pcuic/theories/Syntax/PCUICDepth.v index 81d87aa4a..48cd349ee 100644 --- a/pcuic/theories/Syntax/PCUICDepth.v +++ b/pcuic/theories/Syntax/PCUICDepth.v @@ -255,7 +255,7 @@ Proof. Defined.*) Definition onctx_rel (P : context -> term -> Type) (Γ Δ : context) := - All_local_env (lift_wf_term (fun Δ => P (Γ ,,, Δ))) Δ. + All_local_env (fun Δ => lift_wf_term (P (Γ ,,, Δ))) Δ. Definition CasePredProp (P : context -> term -> Type) Γ (p : predicate term) := All (P Γ) p.(pparams) × onctx_rel P Γ (inst_case_context p.(pparams) p.(puinst) p.(pcontext)) × @@ -330,10 +330,10 @@ Lemma term_forall_ctx_list_ind : P Γ (tCase ci p t brs)) -> (forall Γ (s : projection) (t : term), P Γ t -> P Γ (tProj s t)) -> (forall Γ (m : mfixpoint term) (n : nat), - All_local_env (lift_wf_term (ctx_shifted P Γ)) (fix_context m) -> + All_local_env (fun Δ => lift_wf_term (ctx_shifted P Γ Δ)) (fix_context m) -> tFixProp (P Γ) (P (Γ ,,, fix_context m)) m -> P Γ (tFix m n)) -> (forall Γ (m : mfixpoint term) (n : nat), - All_local_env (lift_wf_term (ctx_shifted P Γ)) (fix_context m) -> + All_local_env (fun Δ => lift_wf_term (ctx_shifted P Γ Δ)) (fix_context m) -> tFixProp (P Γ) (P (Γ ,,, fix_context m)) m -> P Γ (tCoFix m n)) -> (* (forall Γ p, P Γ (tPrim p)) -> *) forall Γ (t : term), P Γ t. @@ -368,12 +368,12 @@ Proof. - case: a => [na [b|] ty] /=; rewrite {1}/decl_depth_gen /context_depth_gen /= => Hlt; constructor; auto. + eapply IHΔ => //. unfold context_depth. lia. - + simpl. apply aux => //; red. lia. + simpl. split. * apply aux => //; red. lia. * apply aux => //; red; lia. + apply IHΔ => //; unfold context_depth; lia. - + simpl. apply aux => //; red. lia. } + + simpl. split; auto. + apply aux => //; red. lia. } assert (forall m, list_depth_gen (fun x : def term => depth (dtype x)) m < S (mfixpoint_depth_gen depth m)). { clear. unfold mfixpoint_depth_gen, def_depth_gen. induction m. simpl. auto. simpl. lia. } assert (forall m, list_depth_gen (fun x : def term => depth (dbody x)) m < S (mfixpoint_depth_gen depth m)). diff --git a/pcuic/theories/Syntax/PCUICInduction.v b/pcuic/theories/Syntax/PCUICInduction.v index b65fcb5ae..e93066703 100644 --- a/pcuic/theories/Syntax/PCUICInduction.v +++ b/pcuic/theories/Syntax/PCUICInduction.v @@ -444,7 +444,7 @@ Proof. Qed. Definition onctx_rel (P : context -> term -> Type) (Γ Δ : context) := - All_local_env (lift_wf_term (ctx_shifted P Γ)) Δ. + All_local_env (fun Δ => lift_wf_term (ctx_shifted P Γ Δ)) Δ. Definition CasePredProp (P : context -> term -> Type) Γ (p : predicate term) := All (P Γ) p.(pparams) × onctx_rel P Γ (pcontext p) × @@ -475,10 +475,10 @@ Lemma term_forall_ctx_list_ind : P Γ (tCase ci p t brs)) -> (forall Γ (s : projection) (t : term), P Γ t -> P Γ (tProj s t)) -> (forall Γ (m : mfixpoint term) (n : nat), - All_local_env (lift_wf_term (fun Γ' t => P (Γ ,,, Γ') t)) (fix_context m) -> + All_local_env (fun Γ' => lift_wf_term (fun t => P (Γ ,,, Γ') t)) (fix_context m) -> tFixProp (P Γ) (P (Γ ,,, fix_context m)) m -> P Γ (tFix m n)) -> (forall Γ (m : mfixpoint term) (n : nat), - All_local_env (lift_wf_term (fun Γ' t => P (Γ ,,, Γ') t)) (fix_context m) -> + All_local_env (fun Γ' => lift_wf_term (fun t => P (Γ ,,, Γ') t)) (fix_context m) -> tFixProp (P Γ) (P (Γ ,,, fix_context m)) m -> P Γ (tCoFix m n)) -> (* (forall Γ p, P Γ (tPrim p)) -> *) forall Γ (t : term), P Γ t. @@ -512,12 +512,12 @@ Proof. - case: a => [na [b|] ty] /=; rewrite {1}/decl_size /context_size /= => Hlt; constructor; auto. + eapply IHΔ => //. unfold context_size. lia. - + simpl. apply aux => //. red. lia. + simpl. split. * apply aux => //. red. lia. * apply aux => //; red; lia. + apply IHΔ => //; unfold context_size; lia. - + simpl. apply aux => //. red. lia. } + + simpl. split; auto. + apply aux => //. red. lia. } assert (forall m, list_size (fun x : def term => size (dtype x)) m < S (mfixpoint_size size m)). { clear. unfold mfixpoint_size, def_size. induction m. simpl. auto. simpl. lia. } assert (forall m, list_size (fun x : def term => size (dbody x)) m < S (mfixpoint_size size m)). diff --git a/pcuic/theories/Syntax/PCUICLiftSubst.v b/pcuic/theories/Syntax/PCUICLiftSubst.v index 544f7a65f..12c1aa65a 100644 --- a/pcuic/theories/Syntax/PCUICLiftSubst.v +++ b/pcuic/theories/Syntax/PCUICLiftSubst.v @@ -108,6 +108,27 @@ Ltac nth_leb_simpl := | _ => lia || congruence || solve [repeat (f_equal; try lia)] end. +Lemma lift_inj : forall M M' i j, lift i j M = lift i j M' -> M = M'. +Proof. + intros M. + induction M using term_forall_list_ind; destruct M' => //; simpl; intros i j e; inversion e; clear e; f_equal; eauto. + - move: H0. + case: (Nat.leb_spec j n); intro; case: (Nat.leb_spec j n0); lia. + - apply All2_eq. apply All2_eq_eq, All2_map_inv in H1. solve_all. + - destruct p, p0; simpl in *. f_equal; auto. + + apply All2_eq. apply All2_eq_eq, All2_map_inv in H1. solve_all. + + destruct X as (_ & _ & ?); eauto. + - apply All2_eq. apply All2_eq_eq, All2_map_inv in H6. solve_all. + destruct x, y; simpl in *; inversion b; clear b; f_equal; eauto. + unfold id in H6. rewrite H6 in H7. eauto. + - apply All2_eq. apply All2_eq_eq, All2_map_inv in H0. apply All2_length in H0 as Hlen. solve_all. + rewrite Hlen in b. + destruct x, y; simpl in *; inversion b; clear b; f_equal; eauto. + - apply All2_eq. apply All2_eq_eq, All2_map_inv in H0. apply All2_length in H0 as Hlen. solve_all. + rewrite Hlen in b. + destruct x, y; simpl in *; inversion b; clear b; f_equal; eauto. +Qed. + Lemma lift0_id : forall M k, lift 0 k M = M. Proof. intros M. diff --git a/pcuic/theories/Typing/PCUICClosedTyp.v b/pcuic/theories/Typing/PCUICClosedTyp.v index f0ce359d7..b5ecb55a2 100644 --- a/pcuic/theories/Typing/PCUICClosedTyp.v +++ b/pcuic/theories/Typing/PCUICClosedTyp.v @@ -75,14 +75,19 @@ Qed. Lemma declared_decl_closed_ind {cf : checker_flags} {Σ : global_env} {wfΣ : wf Σ} {cst decl} : lookup_env Σ cst = Some decl -> Forall_decls_typing (fun (_ : global_env_ext) (Γ : context) (t T : term) => closedn #|Γ| t && closedn #|Γ| T) Σ -> - on_global_decl cumulSpec0 (fun Σ Γ b t => closedn #|Γ| b && typ_or_sort_default (closedn #|Γ|) t true) + on_global_decl cumulSpec0 (fun Σ Γ => lift_wf_term (closedn #|Γ|)) (Σ, universes_decl_of_decl decl) cst decl. Proof. intros. eapply weaken_lookup_on_global_env; eauto. red; eauto. eapply @on_global_env_impl with (Σ := (empty_ext Σ)); cycle 1. tea. - red; intros. destruct T; intuition auto with wf. - destruct X2 as [s0 Hs0]. simpl. rtoProp; intuition. + intros. + destruct j. + - cbn in X2. now rtoProp. + - destruct X2 as (Hb & s & e & Ht); rtoProp. + split => //. + destruct t; cbn in Hb |- *; auto. + destruct Hb; now rtoProp. Qed. Lemma declared_minductive_closed_ind {cf:checker_flags} {Σ : global_env} {wfΣ : wf Σ}{mdecl mind} : @@ -114,18 +119,18 @@ Proof. apply andb_and; split. apply andb_and. split. - rewrite andb_and. split. * apply onArity in oib. hnf in oib. - now move/andP: oib => [] /= ->. + now move: oib => [] /= _ ->. * pose proof (onArity oib). rewrite oib.(ind_arity_eq) in X. - red in X. simpl in X. + destruct X as (_&X). simpl in X. rewrite !closedn_it_mkProd_or_LetIn /= !andb_true_r in X. now move/andP: X. - pose proof (onConstructors oib). red in X. eapply All_forallb. eapply All2_All_left; eauto. intros cdecl cs X0; - move/andP: (on_ctype X0) => []. + move: (on_ctype X0) => []. simpl. unfold closed_constructor_body. - intros Hty _. + intros _ Hty. rewrite arities_context_length in Hty. rewrite Hty. rewrite X0.(cstr_eq) closedn_it_mkProd_or_LetIn in Hty. @@ -183,6 +188,11 @@ Proof. rewrite closedn_ctx_cons in H0. now move/andP: H0 => [H0 _]. simpl in H2; lia. + - destruct X0 as (_ & _ & _ & _ & (Hs & _)%andb_and). + intuition auto. + - destruct X0 as (((_ & (Hb & _)%andb_and) & _) & _ & _ & _ & (Hs & _)%andb_and). + intuition auto. + - intuition auto. generalize (closedn_subst [u] #|Γ| 0 B). rewrite Nat.add_0_r. move=> Hs. apply: Hs => /=. simpl. rewrite H1 => //. @@ -190,14 +200,14 @@ Proof. - rewrite closedn_subst_instance. eapply lookup_on_global_env in X0; eauto. - destruct X0 as [Σ' [hext [onu HΣ'] (IHty & IHbo)]]. + destruct X0 as [Σ' [hext [onu HΣ'] (IHbo & IHty)]]. destruct IHty as (_ & _ & (IH & _)%andb_and). eauto using closed_upwards with arith. - rewrite closedn_subst_instance. eapply declared_inductive_inv in X0; eauto. apply onArity in X0. repeat red in X0. - destruct X0 as (_ & _ & (X0 & _)%andb_and). + destruct X0 as (_ & _ & _ & (X0 & _)%andb_and). intuition eauto using closed_upwards with arith. - destruct isdecl as [Hidecl Hcdecl]. @@ -210,7 +220,7 @@ Proof. pose proof X0.(onConstructors) as XX. eapply All2_nth_error_Some in Hcdecl; eauto. destruct Hcdecl as [? [? ?]]. cbn in *. - destruct o as [? ? (_ & _ & (Hs & _)%andb_and) _]. + destruct o as [? ? (_ & _ & _ & (Hs & _)%andb_and) _]. rewrite arities_context_length in Hs. eauto using closed_upwards with arith. @@ -262,26 +272,24 @@ Proof. eapply closed_upwards; eauto. lia. - clear H. - split. solve_all. destruct b. - destruct x; simpl in *. + split. solve_all. destruct X0 as ((([_ Hb] & _) & _) & _ & s & _ & [_ Hs]). + simpl in *. unfold test_def. simpl. rtoProp. - split. - rewrite -> app_context_length in H, H3. rewrite -> Nat.add_comm in *. - eapply closedn_lift_inv in H3; eauto. lia. + split => //. + rewrite -> app_context_length in H4. subst types. - now rewrite app_context_length fix_context_length in H. - eapply nth_error_all in X0; eauto. simpl in X0. intuition auto. rtoProp. - now destruct X0 as (_ & _ & _ & (cl & _)%andb_and). + now rewrite fix_context_length in H4. + eapply nth_error_all in X0 as (_ & _ & _ & _ & [_ [Ht _]%andb_and]); eauto. - - split. solve_all. destruct b. - destruct x; simpl in *. + - clear H. + split. solve_all. destruct X0 as ((([_ Hb] & _) & _) & _ & s & _ & [_ Hs]). + simpl in *. unfold test_def. simpl. rtoProp. - split. - now destruct a as (_ & _ & _ & (cl & _)%andb_and). - rewrite -> app_context_length in H3, H4. rewrite -> Nat.add_comm in *. - subst types. now rewrite fix_context_length in H3. - eapply nth_error_all in X0; eauto. - now destruct X0 as (_ & _ & _ & (cl & _)%andb_and). + split => //. + rewrite -> app_context_length in H4. + subst types. + now rewrite fix_context_length in H4. + eapply nth_error_all in X0 as (_ & _ & _ & _ & [_ [Ht _]%andb_and]); eauto. Qed. Lemma declared_minductive_closed {cf:checker_flags} {Σ : global_env} {wfΣ : wf Σ} {mdecl mind} : @@ -373,7 +381,7 @@ Qed. #[global] Hint Extern 10 => progress unfold PCUICTypingDef.typing in * : fvs. Lemma isTypeRelOpt_closed {cf:checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ.1} {Γ T rel} : isTypeRelOpt Σ Γ T rel -> closedn #|Γ| T. -Proof. intros (s & e & Hs); fvs. Qed. +Proof. intros (_ & s & e & Hs); fvs. Qed. Lemma isType_closed {cf:checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ.1} {Γ T} : isType Σ Γ T -> closedn #|Γ| T. Proof. apply isTypeRelOpt_closed. Qed. @@ -470,7 +478,7 @@ Qed. Lemma declared_decl_closed `{checker_flags} {Σ : global_env} {cst decl} : wf Σ -> lookup_env Σ cst = Some decl -> - on_global_decl cumulSpec0 (fun Σ Γ b t => closedn #|Γ| b && typ_or_sort_default (closedn #|Γ|) t true) + on_global_decl cumulSpec0 (fun Σ Γ => lift_wf_term (closedn #|Γ|)) (Σ, universes_decl_of_decl decl) cst decl. Proof. intros. @@ -487,7 +495,7 @@ Proof. intros h. unfold declared_constant in h. eapply lookup_on_global_env in h. 2: eauto. - destruct h as [Σ' [ext wfΣ' (declty' & declbo')]]. + destruct h as [Σ' [ext wfΣ' (declbo' & declty')]]. destruct declty' as (? & _ & Hs). now eapply subject_closed in Hs. Unshelve. all:tea. Qed. @@ -503,8 +511,8 @@ Proof. intros Σ cst decl body hΣ h e. unfold declared_constant in h. eapply lookup_on_global_env in h. 2: eauto. - destruct h as [Σ' [ext wfΣ' (declty' & declbo')]]. - rewrite e in declbo'. + destruct h as [Σ' [ext wfΣ' (declbo' & declty')]]. + rewrite e in declbo'; destruct declbo' as (declbo' & _). now eapply subject_closed in declbo'. Unshelve. all:tea. Qed. @@ -524,7 +532,7 @@ Proof. destruct h1 as [Σ' [ext wfΣ' decl']]. red in decl'. destruct decl' as [h ? ? ?]. eapply Alli_nth_error in h. 2: eassumption. - simpl in h. destruct h as [? (? & ? & h) ? ? ?]. + simpl in h. destruct h as [? (_ & ? & _ & h) ? ? ?]. eapply typecheck_closed in h as [? e]. 2: auto. now move: e => [_ /andP []]. Qed. @@ -973,7 +981,7 @@ Proof. fix auxl' 1; destruct l; [constructor|]; simpl; rewrite ?Nat.sub_0_r. move/andP => [] tl /andP [tdef tty]. constructor. + rewrite Nat.sub_0_r. simpl. split; [apply auxt|]; tas. - destruct (decl_body c); simpl in *; auto. exact tt. + destruct (decl_body c); simpl in *; auto. + eapply Alli_shift, Alli_impl; eauto. simpl. intros n' x. now replace (Nat.pred #|l| - n' + #|pparams p|) with (#|l| - S n' + #|pparams p|). @@ -990,7 +998,7 @@ Proof. fix auxl'' 1; destruct l; [constructor|]; simpl; rewrite ?Nat.sub_0_r. move/andP => [] tl /andP [tdef tty]. constructor. + rewrite Nat.sub_0_r. simpl. split; [apply auxt|]; tas. - destruct (decl_body c); simpl in *; auto. exact tt. + destruct (decl_body c); simpl in *; auto. + eapply Alli_shift, Alli_impl; eauto. simpl. intros n' x. now replace (Nat.pred #|l| - n' + #|pparams p|) with (#|l| - S n' + #|pparams p|). diff --git a/pcuic/theories/Typing/PCUICRenameTyp.v b/pcuic/theories/Typing/PCUICRenameTyp.v index b159f0886..56d1f1c0a 100644 --- a/pcuic/theories/Typing/PCUICRenameTyp.v +++ b/pcuic/theories/Typing/PCUICRenameTyp.v @@ -606,7 +606,7 @@ Lemma rename_predicate_preturn f p : preturn (rename_predicate f p). Proof. reflexivity. Qed. -Lemma wf_local_app_renaming P Σ Γ Δ : +(* Lemma wf_local_app_renaming P Σ {wfΣ : wf Σ.1} Γ Δ : All_local_env (lift_typing (fun (Σ : global_env_ext) (Γ' : context) (t T : term) => forall P (Δ : PCUICEnvironment.context) (f : nat -> nat), renaming (shiftnP #|Γ ,,, Γ'| P) Σ Δ (Γ ,,, Γ') f -> Σ ;;; Δ |- rename f t : rename f T) Σ) @@ -619,14 +619,29 @@ Proof. induction X. - apply a. - rewrite rename_context_snoc /=. constructor; auto. - apply infer_typing_sort_impl with id t0 => //; intros Hs. + eapply on_sortrel_impl with _ t0 => // Hs. eapply (Hs P (Δ' ,,, rename_context f Γ0) (shiftn #|Γ0| f)). split => //. eapply urenaming_ext. { now rewrite app_length -shiftnP_add. } { reflexivity. } now eapply urenaming_context. - rewrite rename_context_snoc /=. constructor; auto. - * apply infer_typing_sort_impl with id t0 => //; intros Hs. + destruct t0 as ((Hb & mk) & s & e & Ht). + eassert (renaming (shiftnP #|Γ,,, Γ0| P) Σ (Δ' ,,, rename_context f Γ0) (Γ,,, Γ0) _). { + split => //. + eapply urenaming_ext. + { now rewrite app_length -shiftnP_add. } + { reflexivity. } now eapply urenaming_context. + } + specialize (Hb P (Δ' ,,, rename_context f Γ0) (shiftn #|Γ0| f) X0). + specialize (Ht P (Δ' ,,, rename_context f Γ0) (shiftn #|Γ0| f) X0). + repeat split; tea. + 2: eexists; repeat split; tea. + eapply rename_relevance_of_term; tea. + 2: eapply closedn_on_free_vars, subject_closed; tea. + 1: eapply urenaming_ext. + eapply on_triplefull_impl with _ t0 => //. + 1: intros t1; eapply rename_relevance_of_term. apply (Hs P (Δ' ,,, rename_context f Γ0) (shiftn #|Γ0| f)). split => //. eapply urenaming_ext. @@ -636,7 +651,7 @@ Proof. eapply urenaming_ext. { now rewrite app_length -shiftnP_add. } { reflexivity. } now eapply urenaming_context. -Qed. +Qed. *) Lemma rename_decompose_prod_assum f Γ t : decompose_prod_assum (rename_context f Γ) (rename (shiftn #|Γ| f) t) @@ -806,8 +821,8 @@ Proof. - intros Σ wfΣ Γ wfΓ HΓ. split; auto. induction HΓ; constructor; tas. - all: apply infer_typing_sort_impl with id tu => //; intros Hty. - all: eauto. + + eapply on_sortrel_impl_id with tu => //. + + eapply on_triplefull_impl_id with tu => //. - intros Σ wfΣ Γ wfΓ n decl isdecl ihΓ P Δ f hf. simpl in *. @@ -830,31 +845,33 @@ Proof. eapply renaming_vass. 2: eauto. constructor. * destruct hf as [hΔ hf]. auto. - * simpl. exists s1. split; [apply e|]; eapply ihA; eauto. - - intros Σ wfΣ Γ wfΓ na A t s1 B e X hA ihA ht iht P Δ f hf. + * split; cbn; auto. exists s1. split; [apply e|]; eapply ihA; eauto. + - intros Σ wfΣ Γ wfΓ na A t B X hA ht iht P Δ f hf. simpl. - (* /andP [_ havB]. *) - simpl. econstructor. - + apply e. - + eapply ihA; eauto. + apply type_Lambda'. + + eapply on_sortrel_impl with _ hA => //. + now intros []. + eapply iht; eauto; simpl. eapply renaming_extP. { now rewrite -(shiftnP_add 1). } eapply renaming_vass. 2: eauto. constructor. * destruct hf as [hΔ hf]. auto. - * simpl. exists s1. split; [apply e|]; eapply ihA; eauto. - - intros Σ wfΣ Γ wfΓ na b B t s1 A e X hB ihB hb ihb ht iht P Δ f hf. - simpl. econstructor. - + apply e. + * destruct hA as (_ & s1 & e & _ & Hs1). split; cbn; auto. exists s1. split; [apply e|]; eapply Hs1; eauto. + - intros Σ wfΣ Γ wfΓ na b B t A X hB ht iht P Δ f hf. + simpl. + apply type_LetIn'. + + eapply on_triplefull_impl with _ hB => //. + 1: intros; eapply rename_relevance_of_term => //; tea. + eapply ihB; tea. + eapply ihb; tea. + eapply iht; tea. eapply renaming_extP. { now rewrite -(shiftnP_add 1). } eapply renaming_vdef. 2: eauto. - constructor. - * destruct hf. assumption. - * simpl. exists s1. split; [apply e|]; eapply ihB; tea. - * simpl. eapply ihb; tea. + constructor; [apply hf|]. + repeat split. + * eapply ihb; tea. + * eapply rename_relevance_of_term; tea. + * exists s1. split; [apply e|]; eapply ihB; tea. - intros Σ wfΣ Γ wfΓ t na A B s u X hty ihty ht iht hu ihu P Δ f hf. simpl. eapply meta_conv. + eapply type_App. diff --git a/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v b/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v index 38ece128d..5758fadc3 100644 --- a/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v +++ b/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v @@ -222,12 +222,12 @@ Proof using Type. induction 1. + constructor. + simpl. constructor; auto. - eapply infer_typing_sort_impl with _ tu; [apply relevance_subst_opt|]; intros Hty. + eapply on_sortrel_impl with _ tu => //?; [apply relevance_subst_opt|]. eapply Hs; auto. + simpl. constructor; auto. - ++ eapply infer_typing_sort_impl with _ tu; [apply relevance_subst_opt|]; intros Hty. - eapply Hs; auto. + eapply on_triplefull_impl with _ tu => //?; [apply relevance_subst_opt|apply isTermRelOpt_subst_instance|..]. ++ apply Hc; auto. + ++ apply Hs; auto. - intros n decl eq X u univs wfΣ' H. rewrite subst_instance_lift. rewrite map_decl_type. econstructor; aa. @@ -244,17 +244,18 @@ Proof using Type. + apply relevance_subst; apply Xe. + eapply X1; eauto. + eapply X3; eauto. - - intros n t0 b s1 bty Xe X X0 X1 X2 X3 u univs wfΣ' H. - econstructor. - + apply relevance_subst; apply Xe. - + eapply X1; aa. - + eapply X3; aa. - - intros n b b_ty b' s1 b'_ty Xe X X0 X1 X2 X3 X4 X5 u univs wfΣ' H. - econstructor. - + eapply relevance_subst. apply Xe. - + eauto. - + eauto. - + eapply X5; aa. + - intros n t0 b bty X X0 X1 X2 u univs wfΣ' H. + apply type_Lambda'. + + eapply on_sortrel_impl with _ X0 => //; + eauto using relevance_subst_opt, isTermRelOpt_subst_instance; + now intros []. + + eapply X2; aa. + - intros n b b_ty b' b'_ty X X0 X1 X2 u univs wfΣ' H. + eapply type_LetIn'. + + eapply on_triplefull_impl with _ X0 => //; + eauto using relevance_subst_opt, isTermRelOpt_subst_instance; + now intros []. + + eapply X2; aa. - intros t0 na A B s u X X0 X1 X2 X3 X4 X5 u0 univs wfΣ' H. rewrite subst_instance_subst. cbn. econstructor. + eapply X1; eauto. @@ -322,26 +323,26 @@ Proof using Type. eapply X2 in H0; tas. rewrite subst_instance_mkApps in H0. eassumption. - - intros mfix n decl H H0 H1 X X0 wffix u univs wfΣ'. + - intros mfix n decl H H0 H1 X wffix u univs wfΣ'. rewrite (map_dtype _ (subst_instance u)). econstructor. + specialize (H1 u univs wfΣ' H2). rewrite subst_instance_app in H1. now eapply wf_local_app_inv in H1 as []. + now eapply fix_guard_subst_instance. + rewrite nth_error_map H0. reflexivity. - + apply All_map, (All_impl X); simpl; intuition auto. - eapply infer_typing_sort_impl with _ X1; [apply relevance_subst_opt|]. - intros [_ Hs]; now apply Hs. - + eapply All_map, All_impl; tea. - intros x [X1 X3]. cbn. - specialize (X3 u univs wfΣ' H2). - rewrite (map_dbody (subst_instance u)) in X3. - rewrite subst_instance_lift in X3. - rewrite fix_context_length ?map_length in X0, X1, X3. - rewrite (map_dtype _ (subst_instance u) x) in X3. - rewrite subst_instance_app in X3. - rewrite <- (fix_context_subst_instance u mfix). - now len. + + apply All_map, (All_impl X); simpl. intros d X0. + split; [apply fst in X0|apply snd in X0]. + * rewrite -subst_instance_lift. + rewrite !fix_context_length ?map_length in X0 |- *. + rewrite <- (fix_context_subst_instance u mfix). + unfold app_context. + rewrite <- subst_instance_app. + eapply on_triplefull_impl with (tu := X0) => //; + eauto using relevance_subst_opt, isTermRelOpt_subst_instance; + now intros []. + * eapply on_sortrel_impl with _ X0 => //; + eauto using relevance_subst_opt. + now intros []. + red; rewrite <- wffix. unfold wf_fixpoint, wf_fixpoint_gen. f_equal. @@ -350,26 +351,26 @@ Proof using Type. rewrite map_map_compose. now rewrite subst_instance_check_one_fix. - - intros mfix n decl H H0 H1 X X0 wffix u univs wfΣ'. + - intros mfix n decl H H0 H1 X wffix u univs wfΣ'. rewrite (map_dtype _ (subst_instance u)). econstructor. + specialize (H1 u univs wfΣ' H2). rewrite subst_instance_app in H1. now eapply wf_local_app_inv in H1 as []. + now eapply cofix_guard_subst_instance. + rewrite nth_error_map H0. reflexivity. - + apply All_map, (All_impl X); simpl; intuition auto. - eapply infer_typing_sort_impl with _ X1; [apply relevance_subst_opt|]. - intros [_ Hs]; now apply Hs. - + eapply All_map, All_impl; tea. - intros x [X1 X3]; cbn. - specialize (X3 u univs wfΣ' H2). - rewrite (map_dbody (subst_instance u)) in X3. - rewrite subst_instance_lift in X3. - rewrite fix_context_length ?map_length in X0, X1, X3. - rewrite (map_dtype _ (subst_instance u) x) in X3. - rewrite subst_instance_app in X3. - rewrite <- (fix_context_subst_instance u mfix). - now len. + + apply All_map, (All_impl X); simpl. intros d X0. + split; [apply fst in X0|apply snd in X0]. + * rewrite -subst_instance_lift. + rewrite !fix_context_length ?map_length in X0 |- *. + rewrite <- (fix_context_subst_instance u mfix). + unfold app_context. + rewrite <- subst_instance_app. + eapply on_triplefull_impl with (tu := X0) => //; + eauto using relevance_subst_opt, isTermRelOpt_subst_instance; + now intros []. + * eapply on_sortrel_impl with _ X0 => //; + eauto using relevance_subst_opt. + now intros []. + red; rewrite <- wffix. unfold wf_cofixpoint, wf_cofixpoint_gen. rewrite map_map_compose. @@ -481,8 +482,8 @@ Lemma isType_subst_instance_decl Σ Γ T c decl u : isType Σ (subst_instance u Γ) (subst_instance u T). Proof using Type. intros wfΣ look isty cu. - eapply infer_typing_sort_impl with _ isty; [apply relevance_subst_opt|]. - intros Hs; now eapply (typing_subst_instance_decl _ _ _ (tSort _)). + eapply on_sortrel_impl with _ isty => //?. + now eapply (typing_subst_instance_decl _ _ _ (tSort _)). Qed. Lemma isTypeRel_subst_instance_decl {Σ Γ T r c decl u} : @@ -493,9 +494,8 @@ Lemma isTypeRel_subst_instance_decl {Σ Γ T r c decl u} : isTypeRel Σ (subst_instance u Γ) (subst_instance u T) r. Proof using Type. intros wfΣ look isty cu. - eapply infer_typing_sort_impl with _ isty; [apply relevance_subst_opt|]. - intros Hs. - eapply (typing_subst_instance_decl _ _ _ (tSort _)) in Hs; tea. + eapply on_sortrel_impl with _ isty => //?; [apply relevance_subst_opt|]. + now eapply (typing_subst_instance_decl _ _ _ (tSort _)). Qed. Lemma isArity_subst_instance u T : @@ -513,9 +513,9 @@ Lemma wf_local_subst_instance Σ Γ ext u : Proof using Type. destruct Σ as [Σ φ]. intros X X0 X1. simpl in *. induction X1; cbn; constructor; auto. - 1,2: eapply infer_typing_sort_impl with _ t0; [apply relevance_subst_opt|]; intros Hs. - 3: rename t1 into Hs. - all: eapply typing_subst_instance'' in Hs; eauto; apply X. + 1: eapply on_sortrel_impl with _ t0 => //; eauto using relevance_subst_opt. + 2: eapply on_triplefull_impl with _ t0 => //; eauto using relevance_subst_opt, isTermRelOpt_subst_instance. + all: intros Hs; eapply typing_subst_instance'' in Hs; eauto; apply X. Qed. Lemma wf_local_subst_instance_decl Σ Γ c decl u : @@ -527,9 +527,9 @@ Lemma wf_local_subst_instance_decl Σ Γ c decl u : Proof using Type. destruct Σ as [Σ φ]. intros X X0 X1 X2. induction X1; cbn; constructor; auto. - 1,2: eapply infer_typing_sort_impl with _ t0; [apply relevance_subst_opt|]; intros Hs. - 3: rename t1 into Hs. - all: eapply typing_subst_instance_decl in Hs; eauto; apply X. + 1: eapply on_sortrel_impl with _ t0 => //; eauto using relevance_subst_opt. + 2: eapply on_triplefull_impl with _ t0 => //; eauto using relevance_subst_opt, isTermRelOpt_subst_instance. + all: intros Hs; eapply typing_subst_instance_decl in Hs; eauto; apply X. Qed. Lemma subst_instance_ind_sort_id Σ mdecl ind idecl : @@ -542,7 +542,7 @@ Qed. pose proof (on_declared_inductive decli) as [onmind oib]. pose proof (onArity oib) as ona. rewrite (oib.(ind_arity_eq)) in ona. - red in ona. destruct ona as (s & e & t). + red in ona. destruct ona as (Hb & s & e & t). eapply typed_subst_abstract_instance in t. 2:split; simpl; auto. - rewrite !subst_instance_it_mkProd_or_LetIn in t. @@ -564,7 +564,7 @@ Qed. pose proof (on_declared_inductive decli) as [_ oib]. pose proof (onArity oib) as ona. rewrite (oib.(ind_arity_eq)) in ona |- *. - red in ona. destruct ona as (s & e & t). + red in ona. destruct ona as (Hb & s & e & t). eapply typed_subst_abstract_instance in t; eauto. destruct decli as [declm _]. eapply declared_inductive_wf_global_ext in declm; auto. @@ -576,7 +576,7 @@ Qed. isType Σ Γ T -> subst_instance u T = T. Proof using Type. intros wf_ext u isT. - destruct isT as (s & e & t). + destruct isT as (Hb & s & e & t). eapply typed_subst_abstract_instance in t; auto. Qed. diff --git a/pcuic/theories/Typing/PCUICWeakeningEnvTyp.v b/pcuic/theories/Typing/PCUICWeakeningEnvTyp.v index 8b60e43fe..efe232cf5 100644 --- a/pcuic/theories/Typing/PCUICWeakeningEnvTyp.v +++ b/pcuic/theories/Typing/PCUICWeakeningEnvTyp.v @@ -2,7 +2,7 @@ From MetaCoq.Template Require Import config utils. From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICEquality PCUICContextSubst PCUICUnivSubst PCUICCases - PCUICReduction PCUICCumulativity PCUICTyping + PCUICReduction PCUICCumulativity PCUICTyping PCUICRelevanceTerm PCUICGuardCondition PCUICGlobalEnv PCUICWeakeningEnv PCUICWeakeningEnvConv. From Equations Require Import Equations. @@ -35,8 +35,8 @@ Lemma extends_wf_local `{cf : checker_flags} Σ Γ (wfΓ : wf_local Σ Γ) : Proof. intros X0 Σ' H0. induction X0 in H0 |- *; try econstructor; simpl in *; intuition auto. - - apply infer_typing_sort_impl with id tu => //. intro; eauto. - - apply infer_typing_sort_impl with id tu => //. intro; eauto. + - apply on_sortrel_impl_id with tu => //; intro; eauto with extends. + - apply on_triplefull_impl_id with tu => //; intro; eauto with extends. Qed. #[global] Hint Resolve extends_wf_local : extends. @@ -85,11 +85,16 @@ Proof. rename_all_hyps; try solve [econstructor; eauto 2 with extends]. - induction X; constructor; eauto 2 with extends. - + apply infer_typing_sort_impl with id tu => //; intro; auto. - + apply infer_typing_sort_impl with id tu => //; intro; auto. - + eapply Hc; eauto. + + apply on_sortrel_impl_id with tu => //; intro; auto. + + apply on_triplefull_impl_id with tu => //; intro; eauto with extends. - econstructor; eauto 2 with extends. now apply extends_wf_universe. + - eapply type_Lambda'; eauto. + apply lift_typing_impl with (1 := X0); eauto with extends. + intros ?? []; eauto. + - eapply type_LetIn'; eauto. + apply lift_typing_impl with (1 := X0); eauto with extends. + intros ?? []; eauto. - econstructor; eauto 2 with extends. all: econstructor; eauto 2 with extends. * revert X5. clear -Σ' wfΣ' extΣ. induction 1; constructor; try destruct t0; eauto with extends. @@ -99,17 +104,15 @@ Proof. now apply wf_local_app_inv in forall_Σ'. + eapply fix_guard_extends; eauto. + eapply (All_impl X0); intros d X. - apply (lift_typing_impl X); now intros ? []. - + eapply (All_impl X1); intros d X. - apply (lift_typing_impl X); now intros ? []. + split; [apply fst in X|apply snd in X]. + all: apply (lift_typing_impl X); [eauto with extends | now intros ?? []]. - econstructor; eauto with extends. + specialize (forall_Σ' _ wfΣ' extΣ). now apply wf_local_app_inv in forall_Σ'. + eapply cofix_guard_extends; eauto. + eapply (All_impl X0); intros d X. - apply (lift_typing_impl X); now intros ? []. - + eapply (All_impl X1); intros d X. - apply (lift_typing_impl X); now intros ? []. + split; [apply fst in X|apply snd in X]. + all: apply (lift_typing_impl X); [eauto with extends | now intros ?? []]. - econstructor. 1: eauto. + eapply forall_Σ'1; eauto. + destruct Σ as [Σ φ]. eapply weakening_env_cumulSpec in cumulA; eauto. @@ -123,69 +126,37 @@ Lemma weakening_on_global_decl `{checker_flags} P Σ Σ' φ kn decl : Proof. unfold weaken_env_prop. intros HPΣ wfΣ wfΣ' Hext Hdecl. - destruct decl. - 1:{ - destruct Hdecl; split. - - eapply (HPΣ Σ Σ'); eauto. - - destruct cst_body => //. - eapply (HPΣ Σ Σ'); eauto. - } - simpl in *. - destruct Hdecl as [onI onP onnP]; constructor; eauto. - - eapply Alli_impl; eauto. intros. - destruct X. unshelve econstructor; eauto. - + unfold on_type_rel in *; intuition eauto. - + unfold on_constructors in *. eapply All2_impl; eauto. - intros. - destruct X as [? ? ? ?]. unshelve econstructor; eauto. - * unfold on_type_rel in *; eauto. - * clear on_cindices cstr_eq cstr_args_length. - revert on_cargs. - induction (cstr_args x0) in y |- *; destruct y; simpl in *; eauto. - ** destruct a as [na [b|] ty]; simpl in *; intuition eauto. - ** destruct a as [na [b|] ty]; simpl in *; intuition eauto. - * clear on_ctype on_cargs. - revert on_cindices. - generalize (List.rev (lift_context #|cstr_args x0| 0 (ind_indices x))). - generalize (cstr_indices x0). - induction 1; constructor; eauto. - * simpl. - intros v indv. specialize (on_ctype_variance v indv). - simpl in *. move: on_ctype_variance. - unfold cstr_respects_variance. destruct variance_universes as [[[univs u] u']|]; auto. - intros [args idxs]. split. - ** eapply (All2_fold_impl args); intros. - inversion X; constructor; auto. - ++ eapply weakening_env_cumulSpec; eauto. - ++ eapply weakening_env_convSpec; eauto. - ++ eapply weakening_env_cumulSpec; eauto. - ** eapply (All2_impl idxs); intros. - eapply weakening_env_convSpec; eauto. - + unfold check_ind_sorts in *. - destruct Universe.is_prop; auto. - destruct Universe.is_sprop; auto. - split; [apply fst in ind_sorts|apply snd in ind_sorts]. - * eapply Forall_impl; tea; cbn. - intros. eapply Forall_impl; tea; simpl; intros. - eapply leq_universe_subset; tea. - apply weakening_env_global_ext_constraints; tea. - * destruct indices_matter; [|trivial]. clear -ind_sorts HPΣ wfΣ wfΣ' Hext. - induction ind_indices; simpl in *; auto. - -- eapply (extends_wf_universe (Σ:=(Σ,φ)) Σ'); auto. - -- destruct a as [na [b|] ty]; simpl in *; intuition eauto. - + intros v onv. - move: (onIndices v onv). unfold ind_respects_variance. - destruct variance_universes as [[[univs u] u']|] => //. - intros idx; eapply (All2_fold_impl idx); simpl. - intros par par' t t' d. - inv d; constructor; auto. + eapply on_global_decl_impl. + 8: eassumption. + 7: eassumption. + all: simpl; eauto. + - intros m v c. + unfold cstr_respects_variance. destruct variance_universes as [[[univs u] u']|]; auto. + intros [args idxs]. split. + ** eapply (All2_fold_impl args); intros. + inversion X; constructor; auto. ++ eapply weakening_env_cumulSpec; eauto. ++ eapply weakening_env_convSpec; eauto. ++ eapply weakening_env_cumulSpec; eauto. - - red in onP |- *. eapply All_local_env_impl; eauto. - - move: onVariance. - rewrite /on_variance. destruct ind_universes => //. - destruct ind_variance => //. + ** eapply (All2_impl idxs); intros. + eapply weakening_env_convSpec; eauto. + - intros m v i. + unfold ind_respects_variance. + destruct variance_universes as [[[univs u] u']|] => //. + intros idx; eapply (All2_fold_impl idx); simpl. + intros par par' t t' d. + inv d; constructor; auto. + ++ eapply weakening_env_cumulSpec; eauto. + ++ eapply weakening_env_convSpec; eauto. + ++ eapply weakening_env_cumulSpec; eauto. + - intros u. eapply (extends_wf_universe (Σ:=(Σ,φ)) Σ'); auto. + - intros l s Hc. + eapply Forall_impl; tea; cbn. + intros. eapply Forall_impl; tea; simpl; intros. + eapply leq_universe_subset; tea. + apply weakening_env_global_ext_constraints; tea. + - rewrite /on_variance. + move=> [|?] // [?|] //. intros [univs' [i [i' []]]]. exists univs', i, i'. split => //. all:eapply weakening_env_consistent_instance; tea. Qed. @@ -198,72 +169,40 @@ Lemma weakening_on_global_decl_ext `{checker_flags} P Σ Σ' φ kn decl : Proof. unfold weaken_env_prop. intros HPΣ wfΣ' Hext Hdecl. - pose proof (wfΣ := extends_decls_wf _ _ wfΣ' Hext). - destruct decl. - 1:{ - destruct Hdecl; split. - - eapply (HPΣ Σ Σ'); eauto. - - destruct cst_body => //. - eapply (HPΣ Σ Σ'); eauto. - } - simpl in *. - destruct Hdecl as [onI onP onnP]; constructor; eauto. - - eapply Alli_impl; eauto. intros. - destruct X. unshelve econstructor; eauto. - + unfold on_type_rel in *; intuition eauto. - + unfold on_constructors in *. eapply All2_impl; eauto. - intros. - destruct X as [? ? ? ?]. unshelve econstructor; eauto. - * unfold on_type_rel in *; eauto. - * clear on_cindices cstr_eq cstr_args_length. - revert on_cargs. - induction (cstr_args x0) in y |- *; destruct y; simpl in *; eauto. - ** destruct a as [na [b|] ty]; simpl in *; intuition eauto. - ** destruct a as [na [b|] ty]; simpl in *; intuition eauto. - * clear on_ctype on_cargs. - revert on_cindices. - generalize (List.rev (lift_context #|cstr_args x0| 0 (ind_indices x))). - generalize (cstr_indices x0). - induction 1; constructor; eauto. - * simpl. - intros v indv. specialize (on_ctype_variance v indv). - simpl in *. move: on_ctype_variance. - unfold cstr_respects_variance. destruct variance_universes as [[[univs u] u']|]; auto. - intros [args idxs]. split. - ** eapply (All2_fold_impl args); intros. - inversion X; constructor; auto. - ++ eapply weakening_env_cumulSpec; eauto. tc. - ++ eapply weakening_env_convSpec; eauto. tc. - ++ eapply weakening_env_cumulSpec; eauto. tc. - ** eapply (All2_impl idxs); intros. - eapply weakening_env_convSpec; eauto. tc. - + unfold check_ind_sorts in *. - destruct Universe.is_prop; auto. - destruct Universe.is_sprop; auto. - split; [apply fst in ind_sorts|apply snd in ind_sorts]. - * eapply Forall_impl; tea; cbn. - intros. eapply Forall_impl; tea; simpl; intros. - eapply leq_universe_subset; tea. - apply weakening_env_global_ext_constraints; tea. tc. - * destruct indices_matter; [|trivial]. clear -ind_sorts HPΣ wfΣ wfΣ' Hext. - induction ind_indices; simpl in *; auto. - -- eapply (extends_wf_universe (Σ:=(Σ,φ)) Σ'); auto. tc. - -- destruct a as [na [b|] ty]; simpl in *; intuition eauto. - + intros v onv. - move: (onIndices v onv). unfold ind_respects_variance. - destruct variance_universes as [[[univs u] u']|] => //. - intros idx; eapply (All2_fold_impl idx); simpl. - intros par par' t t' d. - inv d; constructor; auto. - ++ eapply weakening_env_cumulSpec; eauto; tc. - ++ eapply weakening_env_convSpec; eauto; tc. - ++ eapply weakening_env_cumulSpec; eauto; tc. - - red in onP |- *. eapply All_local_env_impl; eauto. - - move: onVariance. - rewrite /on_variance. destruct ind_universes => //. - destruct ind_variance => //. + assert (Hext' : extends Σ Σ') by apply extends_decls_extends, Hext. + eapply on_global_decl_impl. + 8: tea. + 7: eapply (extends_decls_wf _ _ wfΣ' Hext). + all: eauto. + - intros m v c. + unfold cstr_respects_variance. destruct variance_universes as [[[univs u] u']|]; auto. + intros [args idxs]. split. + ** eapply (All2_fold_impl args); intros. + inversion X; constructor; auto. + ++ eapply weakening_env_cumulSpec; eauto. + ++ eapply weakening_env_convSpec; eauto. + ++ eapply weakening_env_cumulSpec; eauto. + ** eapply (All2_impl idxs); intros. + eapply weakening_env_convSpec; eauto. + - intros m v i. + unfold ind_respects_variance. + destruct variance_universes as [[[univs u] u']|] => //. + intros idx; eapply (All2_fold_impl idx); simpl. + intros par par' t t' d. + inv d; constructor; auto. + ++ eapply weakening_env_cumulSpec; eauto. + ++ eapply weakening_env_convSpec; eauto. + ++ eapply weakening_env_cumulSpec; eauto. + - intros u. eapply (extends_wf_universe (Σ:=(Σ,φ)) Σ'); auto. + - intros l s Hc. + eapply Forall_impl; tea; cbn. + intros. eapply Forall_impl; tea; simpl; intros. + eapply leq_universe_subset; tea. + apply weakening_env_global_ext_constraints; tea. + - rewrite /on_variance. + move=> [|?] // [?|] //. intros [univs' [i [i' []]]]. exists univs', i, i'. split => //. - all:eapply weakening_env_consistent_instance; tea; tc. + all:eapply weakening_env_consistent_instance; tea. Qed. Lemma weakening_env_decls_lookup_on_global_env `{checker_flags} P Σ Σ' c decl : @@ -468,8 +407,8 @@ Qed. Lemma weaken_env_prop_typing `{checker_flags} : weaken_env_prop cumulSpec0 (lift_typing typing) (lift_typing typing). Proof. - red. intros * wfΣ wfΣ' Hext Γ t T HT. - apply lift_typing_impl with (1 := HT); intros ? Hty. + red. intros * wfΣ wfΣ' Hext Γ T HT. + apply lift_typing_impl with (1 := HT) => ?? Hty; eauto with extends. eapply (weakening_env (_, _)). 2: eauto. all: auto. @@ -477,8 +416,8 @@ Qed. Lemma weaken_env_decls_prop_typing `{checker_flags} : weaken_env_decls_prop cumulSpec0 (lift_typing typing) (lift_typing typing). Proof. - red. intros * wfΣ' Hext Γ t T HT. - apply lift_typing_impl with (1 := HT); intros ? Hty. + red. intros * wfΣ' Hext Γ T HT. + apply lift_typing_impl with (1 := HT) => // ?? Hty; eauto with extends. eapply (weakening_env (_, _)). 2-4: eauto. * cbn; now eapply extends_decls_wf. diff --git a/template-coq/theories/Ast.v b/template-coq/theories/Ast.v index 4776fcb8d..39a211cf9 100644 --- a/template-coq/theories/Ast.v +++ b/template-coq/theories/Ast.v @@ -675,14 +675,13 @@ Inductive isTermRel (Σ : global_env) (Γ : mark_context) : term -> relevance -> | rel_Ind ind u : isTermRel Σ Γ (tInd ind u) Relevant. Derive Signature for isTermRel. -Definition isTermRelOpt Σ Γ t relopt := option_default (isTermRel Σ Γ t) relopt unit. Module TemplateTermUtils <: TermUtils TemplateTerm Env. Definition destArity := destArity. Definition inds := inds. -Definition isTermRelOpt := isTermRelOpt. +Definition isTermRel := isTermRel. End TemplateTermUtils. @@ -694,7 +693,7 @@ Ltac unf_term := unfold TemplateTerm.term in *; unfold TemplateTerm.tRel in *; unfold TemplateTerm.closedn in *; unfold TemplateTerm.noccur_between in *; unfold TemplateTerm.subst_instance_constr in *; unfold TemplateTermUtils.destArity; unfold TemplateTermUtils.inds; - unfold TemplateTermUtils.isTermRelOpt. + unfold TemplateTermUtils.isTermRel. Definition tDummy := tVar "". diff --git a/template-coq/theories/Environment.v b/template-coq/theories/Environment.v index f38e5c4cd..b4dcfc81e 100644 --- a/template-coq/theories/Environment.v +++ b/template-coq/theories/Environment.v @@ -727,6 +727,6 @@ Module Type TermUtils (T: Term) (E: EnvironmentSig T). Parameter Inline destArity : context -> term -> option (context × Universe.t). Parameter Inline inds : kername -> Instance.t -> list one_inductive_body -> list term. - Parameter Inline isTermRelOpt : global_env -> mark_context -> term -> option relevance -> Type. + Parameter Inline isTermRel : global_env -> mark_context -> term -> relevance -> Type. End TermUtils. diff --git a/template-coq/theories/EnvironmentTyping.v b/template-coq/theories/EnvironmentTyping.v index ee9e2d0d6..93b1316b5 100644 --- a/template-coq/theories/EnvironmentTyping.v +++ b/template-coq/theories/EnvironmentTyping.v @@ -416,7 +416,7 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). Qed. Definition on_def (P : context -> judgment -> Type) types Γ d := - P (Γ ,,, types) (TripleRel d.(dbody) (lift0 #|types| d.(dtype)) d.(dname).(binder_relevance)). + P (Γ ,,, types) (TripleRel d.(dbody) (lift0 #|types| d.(dtype)) d.(dname).(binder_relevance)) × P Γ (SortRel d.(dtype) d.(dname).(binder_relevance)). Definition lift_judgment (check : term -> term -> Type) @@ -430,6 +430,8 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). (* Common uses *) + Definition isTermRelOpt Σ Γ t relopt := option_default (isTermRel Σ Γ t) relopt unit. + Definition lift_wf_term wf_term := (lift_judgment (fun t T => wf_term t × wf_term T) (fun t T relopt => option_default wf_term t unit × wf_term T)). Definition lift_on_term on_term := (fun (Γ : context) => lift_judgment (fun t T => on_term Γ t × on_term Γ T) (fun t T relopt => option_default (on_term Γ) t unit × on_term Γ T)). @@ -470,38 +472,106 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). apply (infer_sort_impl (P := typing_sort P) (Q := typing_sort Q)). Qed. - Lemma on_triple_impl_id {P Ps Q Qs} {Σ Σ' : global_env_ext} {Γ Γ' : context} {t : option term} {T : term} {relopt} : - forall (Htr : forall t, isTermRelOpt Σ (marks_of_context Γ) t relopt -> isTermRelOpt Σ' (marks_of_context Γ') t relopt), + Lemma on_triple_impl {P Ps Q Qs} {Σ Σ' : global_env_ext} {Γ Γ' : context} {t : option term} {T : term} {relopt} : + forall f (Hf : forall s, isSortRelOpt s relopt -> isSortRelOpt (f s) relopt) + (Htr : forall t, isTermRelOpt Σ (marks_of_context Γ) t relopt -> isTermRelOpt Σ' (marks_of_context Γ') t relopt), forall tu: on_triple P Ps Σ Γ t T relopt, (forall t, P Σ Γ t T -> Q Σ' Γ' t T) -> let s := tu.2.π1 in - (Ps Σ Γ T s -> Qs Σ' Γ' T s) -> + (Ps Σ Γ T s -> Qs Σ' Γ' T (f s)) -> on_triple Q Qs Σ' Γ' t T relopt. Proof. - intros Htr [Ht HT] HPQ s HPQs. + intros f Hf Htr [Ht HT] HPQ s HPQs. split. - destruct t => //=. destruct Ht; split; auto. - - apply infer_sort_impl with id HT => //. + - apply infer_sort_impl with f HT => //. + Qed. + + Lemma on_sortrel_impl {P Ps Q Qs} {Σ Σ' : global_env_ext} {Γ Γ' : context} {T : term} {relopt} : + forall f fs + (Hf : forall s, tSort (fs s) = f (tSort s)) + (Hfs : forall s, isSortRelOpt s relopt -> isSortRelOpt (fs s) relopt), + forall tu: on_triple P Ps Σ Γ None T relopt, + let s := tu.2.π1 in + (Ps Σ Γ T s -> Qs Σ' Γ' (f T) (fs s)) -> + on_triple Q Qs Σ' Γ' None (f T) relopt. + Proof. + intros f fs Hf Hfs [Ht HT] s HPQs. + split; cbn; auto. + apply infer_sort_impl with fs HT => //. + Qed. + + Lemma on_sortrel_impl_id {P Ps Q Qs} {Σ Σ' : global_env_ext} {Γ Γ' : context} {T : term} {relopt} : + forall tu: on_triple P Ps Σ Γ None T relopt, + let s := tu.2.π1 in + (Ps Σ Γ T s -> Qs Σ' Γ' T s) -> + on_triple Q Qs Σ' Γ' None T relopt. + Proof. + intros [Ht HT] s HPQs. + split; cbn; auto. + apply infer_sort_impl with id HT => //. Qed. - Lemma lift_relation_impl {P Ps Q Qs Σ Γ T} : + Lemma on_triplefull_impl {P Ps Q Qs} {Σ Σ' : global_env_ext} {Γ Γ' : context} {t T : term} {relopt} : + forall f fs + (Hf : forall s, tSort (fs s) = f (tSort s)) + (Hfs : forall s, isSortRelOpt s relopt -> isSortRelOpt (fs s) relopt) + (Htr : forall t, isTermRelOpt Σ (marks_of_context Γ) t relopt -> isTermRelOpt Σ' (marks_of_context Γ') (f t) relopt), + forall tu: on_triple P Ps Σ Γ (Some t) T relopt, + (P Σ Γ t T -> Q Σ' Γ' (f t) (f T)) -> + let s := tu.2.π1 in + (Ps Σ Γ T s -> Qs Σ' Γ' (f T) (fs s)) -> + on_triple Q Qs Σ' Γ' (Some (f t)) (f T) relopt. + Proof. + intros f fs Hf Hfs Htr [(Ht & mk) HT] s HPQs. + split. + 1: split; auto. + apply infer_sort_impl with fs HT => //. + Qed. + + Lemma on_triplefull_impl_id {P Ps Q Qs} {Σ Σ' : global_env_ext} {Γ Γ' : context} {t T : term} {relopt} : + forall (Htr : forall t, isTermRelOpt Σ (marks_of_context Γ) t relopt -> isTermRelOpt Σ' (marks_of_context Γ') t relopt), + forall tu: on_triple P Ps Σ Γ (Some t) T relopt, + (P Σ Γ t T -> Q Σ' Γ' t T) -> + let s := tu.2.π1 in + (Ps Σ Γ T s -> Qs Σ' Γ' T s) -> + on_triple Q Qs Σ' Γ' (Some t) T relopt. + Proof. + intros Htr [(Ht & mk) HT] s HPQs. + split. + 1: split; auto. + apply infer_sort_impl with id HT => //. + Qed. + + Lemma on_triple_impl_id {P Ps Q Qs} {Σ Σ' : global_env_ext} {Γ Γ' : context} {t : option term} {T : term} {relopt} : + forall (Htr : forall t, isTermRelOpt Σ (marks_of_context Γ) t relopt -> isTermRelOpt Σ' (marks_of_context Γ') t relopt), + forall tu: on_triple P Ps Σ Γ t T relopt, + (forall t, P Σ Γ t T -> Q Σ' Γ' t T) -> + let s := tu.2.π1 in + (Ps Σ Γ T s -> Qs Σ' Γ' T s) -> + on_triple Q Qs Σ' Γ' t T relopt. + Proof. intros. apply on_triple_impl with id tu => //. Qed. + + Lemma lift_relation_impl {P Ps Q Qs} {Σ Σ' : global_env_ext} {Γ Γ' T} : lift_relation P Ps Σ Γ T -> - (forall t T, P Σ Γ t T -> Q Σ Γ t T) -> - (forall t s, Ps Σ Γ t s -> Qs Σ Γ t s) -> - lift_relation Q Qs Σ Γ T. + forall (Htr: forall t relopt, isTermRelOpt Σ (marks_of_context Γ) t relopt -> isTermRelOpt Σ' (marks_of_context Γ') t relopt), + (forall t T, P Σ Γ t T -> Q Σ' Γ' t T) -> + (forall t s, Ps Σ Γ t s -> Qs Σ' Γ' t s) -> + lift_relation Q Qs Σ' Γ' T. Proof. - intros HT HPQ HPsQs. + intros HT Htr HPQ HPsQs. destruct T => /=; auto. eapply on_triple_impl_id with HT => //; auto. Qed. - Lemma lift_typing_impl {P Q Σ Γ T} : + Lemma lift_typing_impl {P Q} {Σ Σ' : global_env_ext} {Γ Γ' T} : lift_typing P Σ Γ T -> - (forall t T, P Σ Γ t T -> Q Σ Γ t T) -> - lift_typing Q Σ Γ T. + (forall t relopt, isTermRelOpt Σ (marks_of_context Γ) t relopt -> isTermRelOpt Σ' (marks_of_context Γ') t relopt) -> + (forall t T, P Σ Γ t T -> Q Σ' Γ' t T) -> + lift_typing Q Σ' Γ' T. Proof. - intros HT HPQ. + intros HT Htr HPQ. eapply lift_relation_impl with (1 := HT) => //; auto. Qed. @@ -537,8 +607,42 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). All_local_env_over_gen Σ (Γ ,, vdef na b t) (localenv_cons_def all tu). + End TypeLocalOver. + Section TypeLocalOver. + Context (checking : global_env_ext -> context -> term -> term -> Type). + Context (sorting : global_env_ext -> context -> term -> Universe.t -> Type). + Context (cproperty_rel : forall (Σ : global_env_ext) (Γ' Γ : context), + All_local_rel (lift_relation checking sorting Σ) Γ' Γ -> + forall (t T : term), checking Σ (Γ' ,,, Γ) t T -> Type). + Context (sproperty_rel : forall (Σ : global_env_ext) (Γ' Γ : context), + All_local_rel (lift_relation checking sorting Σ) Γ' Γ -> + forall (t : term) (s : Universe.t), sorting Σ (Γ' ,,, Γ) t s -> Type). + + Inductive All_local_env_over_rel_gen Σ Γ' : + forall (Γ : context), All_local_rel (lift_relation checking sorting Σ) Γ' Γ -> Type := + | localenv_over_rel_nil : + All_local_env_over_rel_gen Σ Γ' [] localenv_nil + + | localenv_over_rel_cons_abs Γ na t + (all : All_local_rel (lift_relation checking sorting Σ) Γ' Γ) : + All_local_env_over_rel_gen Σ Γ' Γ all -> + forall (tu : lift_relation checking sorting Σ (Γ' ,,, Γ) (SortRel t na.(binder_relevance))) + (Hs: sproperty_rel Σ Γ' Γ all _ _ tu.2.π2.2), + All_local_env_over_rel_gen Σ Γ' (Γ ,, vass na t) + (localenv_cons_abs all tu) + + | localenv_over_rel_cons_def Γ na b t + (all : All_local_rel (lift_relation checking sorting Σ) Γ' Γ) : + All_local_env_over_rel_gen Σ Γ' Γ all -> + forall (tu : lift_relation checking sorting Σ (Γ' ,,, Γ) (TripleRel b t na.(binder_relevance))) + (Hc: cproperty_rel Σ Γ' Γ all _ _ tu.1.1) + (Hs: sproperty_rel Σ Γ' Γ all _ _ tu.2.π2.2), + All_local_env_over_rel_gen Σ Γ' (Γ ,, vdef na b t) + (localenv_cons_def all tu). + End TypeLocalOver. Derive Signature for All_local_env_over_gen. + Derive Signature for All_local_env_over_rel_gen. Lemma All_local_env_over_gen_2 checking sorting cproperty sproperty Σ Γ wfΓ : let cproperty_full Σ Γ wfΓ t T check := cproperty Σ Γ t T in @@ -558,6 +662,7 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T) (TU : TermUtils T E). (All_local_env_over_gen typing (typing_sort typing) property (fun Σ Γ H t s tu => property Σ Γ H t (tSort s) tu)). Definition All_local_env_over_sorting := All_local_env_over_gen. + Definition All_local_rel_over_sorting := All_local_env_over_rel_gen. Lemma All_local_env_over_2 typing property (Σ : global_env_ext) (Γ : context) (wfΓ : All_local_env (lift_typing typing Σ) Γ) : let property_full Σ Γ wfΓ t T Hty := property Σ Γ t T in @@ -750,7 +855,7 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT | {| decl_name := na; decl_body := None; decl_type := t |} :: Δ => type_local_ctx Σ Γ Δ u × isSortRel u na.(binder_relevance) × P Σ (Γ ,,, Δ) (Typ t (tSort u)) | {| decl_name := na; decl_body := Some b; decl_type := t |} :: Δ => - type_local_ctx Σ Γ Δ u × P Σ (Γ ,,, Δ) (SortRel t na.(binder_relevance)) × P Σ (Γ ,,, Δ) (Typ b t) + type_local_ctx Σ Γ Δ u × P Σ (Γ ,,, Δ) (TripleRel b t na.(binder_relevance)) end. Fixpoint sorts_local_ctx Σ (Γ Δ : context) (us : list Universe.t) : Type := @@ -759,7 +864,7 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT | {| decl_name := na; decl_body := None; decl_type := t |} :: Δ, u :: us => sorts_local_ctx Σ Γ Δ us × isSortRel u na.(binder_relevance) × P Σ (Γ ,,, Δ) (Typ t (tSort u)) | {| decl_name := na; decl_body := Some b; decl_type := t |} :: Δ, us => - sorts_local_ctx Σ Γ Δ us × P Σ (Γ ,,, Δ) (SortRel t na.(binder_relevance)) × P Σ (Γ ,,, Δ) (Typ b t) + sorts_local_ctx Σ Γ Δ us × P Σ (Γ ,,, Δ) (TripleRel b t na.(binder_relevance)) | _, _ => False end. @@ -1283,20 +1388,22 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT Arguments onVariance {_ Pcmp P Σ mind mdecl}. - Lemma type_local_ctx_impl (P Q : global_env_ext -> context -> judgment -> Type) Σ Γ Δ u : + Lemma type_local_ctx_impl (P Q : global_env_ext -> context -> judgment -> Type) Σ Σ' Γ Δ u : type_local_ctx P Σ Γ Δ u -> - (forall Γ j, P Σ Γ j -> Q Σ Γ j) -> - type_local_ctx Q Σ Γ Δ u. + (forall u, wf_universe Σ u -> wf_universe Σ' u) -> + (forall Γ j, P Σ Γ j -> Q Σ' Γ j) -> + type_local_ctx Q Σ' Γ Δ u. Proof. intros HP HPQ. revert HP; induction Δ in Γ, HPQ |- *; simpl; auto. destruct a as [na [b|] ty]; simpl; auto. intros. intuition auto. intuition auto. Qed. - Lemma sorts_local_ctx_impl (P Q : global_env_ext -> context -> judgment -> Type) Σ Γ Δ u : + Lemma sorts_local_ctx_impl (P Q : global_env_ext -> context -> judgment -> Type) Σ Σ' Γ Δ u : sorts_local_ctx P Σ Γ Δ u -> - (forall Γ j, P Σ Γ j -> Q Σ Γ j) -> - sorts_local_ctx Q Σ Γ Δ u. + (forall u, wf_universe Σ u -> wf_universe Σ' u) -> + (forall Γ j, P Σ Γ j -> Q Σ' Γ j) -> + sorts_local_ctx Q Σ' Γ Δ u. Proof. intros HP HPQ. revert HP; induction Δ in Γ, HPQ, u |- *; simpl; auto. destruct a as [na [b|] ty]; simpl; auto. @@ -1304,14 +1411,20 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT destruct u; auto. intuition eauto. Qed. - Lemma on_global_decl_impl {cf : checker_flags} Pcmp P Q Σ kn d : + Lemma on_global_decl_impl {cf : checker_flags} Pcmp P P' Q Σ Σ' kn d : (forall Γ j, on_global_env Pcmp P Σ.1 -> - P Σ Γ j -> Q Σ Γ j) -> + P' Σ Γ j -> Q Σ' Γ j) -> + (forall m v c, cstr_respects_variance Pcmp Σ.1 m v c -> cstr_respects_variance Pcmp Σ'.1 m v c) -> + (forall m v i, ind_respects_variance Pcmp Σ.1 m v i -> ind_respects_variance Pcmp Σ'.1 m v i) -> + (forall u, wf_universe Σ u -> wf_universe Σ' u) -> + (forall l s, check_constructors_smaller (global_ext_constraints Σ) l s -> + check_constructors_smaller (global_ext_constraints Σ') l s) -> + (forall u l, on_variance Σ.1 u l -> on_variance Σ'.1 u l) -> on_global_env Pcmp P Σ.1 -> - on_global_decl Pcmp P Σ kn d -> on_global_decl Pcmp Q Σ kn d. + on_global_decl Pcmp P' Σ kn d -> on_global_decl Pcmp Q Σ' kn d. Proof. - intros X X0. + intros X Xr Xr' Xu Xc Xv X0. destruct d; simpl. - apply X => //. - intros [onI onP onNP]. @@ -1340,15 +1453,26 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT destruct Universe.is_prop; auto. destruct Universe.is_sprop; auto. split. - * apply ind_sorts0. + * apply Xc, ind_sorts0. * destruct indices_matter; auto. - eapply type_local_ctx_impl. eapply ind_sorts0. eauto. + eapply type_local_ctx_impl. eapply ind_sorts0. all: eauto. --- apply X1.(ind_relevance_compat). - --- apply X1.(onIndices). + --- intros v e. apply Xr', X1.(onIndices), e. -- red in onP. red. eapply All_local_env_impl; eauto. Qed. + Lemma on_global_decl_impl' {cf : checker_flags} Pcmp P Q Σ kn d : + (forall Γ j, + on_global_env Pcmp P Σ.1 -> + P Σ Γ j -> Q Σ Γ j) -> + on_global_env Pcmp P Σ.1 -> + on_global_decl Pcmp P Σ kn d -> on_global_decl Pcmp Q Σ kn d. + Proof. + intro. + apply on_global_decl_impl => //. + Qed. + Lemma on_global_env_impl {cf : checker_flags} Pcmp P Q : (forall Σ Γ j, on_global_env Pcmp P Σ.1 -> @@ -1358,7 +1482,7 @@ Module GlobalMaps (T: Term) (E: EnvironmentSig T) (TU : TermUtils T E) (ET: EnvT Proof. intros X [univs Σ] [cu X0]; split => /= //. cbn in *. induction X0; constructor; auto. - eapply on_global_decl_impl; tea. + eapply on_global_decl_impl'; tea. - intros; apply X => //. - split => //. Qed. diff --git a/template-coq/theories/EtaExpand.v b/template-coq/theories/EtaExpand.v index 67b35eec5..ad9dc4818 100644 --- a/template-coq/theories/EtaExpand.v +++ b/template-coq/theories/EtaExpand.v @@ -1148,7 +1148,7 @@ Proof. { apply andb_and in H2. destruct H2 as [isl _]. solve_all. } solve_all. { now eapply isLambda_lift, isLambda_eta_expand. } - destruct a as (((Hbo1 & Hbo2) & et) & s & e & Hs1 & Hs2). + destruct a as ((((Hbo1 & Hbo2) & mk) & _) & tt & s & e & Hs1 & Hs2). rewrite !firstn_length. rewrite !Nat.min_l; try lia. eapply expanded_lift'. 5: eapply Hbo2. 2: reflexivity. 2: now len. @@ -1169,7 +1169,7 @@ Proof. len; lia. rewrite repeat_length. len; lia. + cbn - [rev_map seq]. rewrite rev_map_spec. cbn. rewrite Nat.sub_0_r. cbn. destruct List.rev; cbn; congruence. - - cbn. econstructor; eauto. eapply All_Forall, All_map, All_impl. eapply X. intros ? (((Hbo1 & Hbo2) & et) & s & e & Hs1 & Hs2). cbn. + - cbn. econstructor; eauto. eapply All_Forall, All_map, All_impl. eapply X. intros ? ((((Hbo1 & Hbo2) & mk) & _) & tt & s & e & Hs1 & Hs2). cbn. specialize (Hbo2 (repeat None #|mfix| ++ Γ'))%list. rewrite map_app, map_repeat in Hbo2. len. eapply Hbo2. eapply Forall2_app; eauto. unfold types. @@ -1358,11 +1358,10 @@ Proof. eapply All_fold_fold_context_k_defs. cbn. len. induction ctx' in hs, cunivs |- *; cbn; auto. constructor; eauto. - cbn in hs. destruct a as [na [b|] ty]; try destruct hs as [hs ?]. + cbn in hs. destruct a as [na [b|] ty]; try destruct hs as [hs ((Hb&mk)& u&e&Ht)]. specialize (IHctx' cunivs hs). constructor; auto. constructor. len. rewrite repeat_app. - destruct p as [[s Hs] ?]. - epose proof (eta_expand_expanded (Σ := Σ) _ (repeat None (#|ctx'| + #|ctx|)) _ _ wfΣ t). + epose proof (eta_expand_expanded (Σ := Σ) _ (repeat None (#|ctx'| + #|ctx|)) _ _ wfΣ Hb). forward H. clear. rewrite -app_context_length. induction (_ ,,, _); cbn; constructor; auto. diff --git a/template-coq/theories/Typing.v b/template-coq/theories/Typing.v index bcd475ce8..a59a15d65 100644 --- a/template-coq/theories/Typing.v +++ b/template-coq/theories/Typing.v @@ -952,8 +952,8 @@ Proof. - exact (S (Nat.max d1 (Nat.max d2 (Nat.max (ctx_inst_size _ typing_size c1) (all2i_size _ (fun _ x y p => Nat.max (typing_size _ _ _ _ p.1.2) (typing_size _ _ _ _ p.2)) a0))))). - - exact (S (Nat.max (All_local_env_size typing_size _ _ a) (all_size _ (fun x p => lift_typing_size typing_size Σ _ _ p) a0))). - - exact (S (Nat.max (All_local_env_size typing_size _ _ a) (all_size _ (fun x p => lift_typing_size typing_size Σ _ _ p) a0))). + - exact (S (Nat.max (All_local_env_size typing_size _ _ a) (all_size _ (fun x p => (lift_typing_size typing_size Σ _ _ p.1 + lift_typing_size typing_size Σ _ _ p.2)%nat) a0))). + - exact (S (Nat.max (All_local_env_size typing_size _ _ a) (all_size _ (fun x p => (lift_typing_size typing_size Σ _ _ p.1 + lift_typing_size typing_size Σ _ _ p.2)%nat) a0))). Defined. Lemma typing_size_pos `{checker_flags} {Σ Γ t T} (d : Σ ;;; Γ |- t : T) : typing_size d > 0. @@ -1278,7 +1278,7 @@ Proof. apply IH'; auto. - simpl. simpl in *. destruct d. - + apply lift_typing_impl with (1 := Xg) => ?? Hs. + + apply lift_typing_impl with (1 := Xg) => // ?? Hs. specialize (IH ((Σ', udecl); wfΣ; _; _; _; Hs)). forward IH; [constructor 1; simpl; subst Σ' Σg; cbn; lia|]. apply IH. @@ -1291,16 +1291,16 @@ Proof. refine {| ind_arity_eq := Xg.(ind_arity_eq); ind_cunivs := Xg.(ind_cunivs) |}. -- apply onArity in Xg. - apply lift_typing_impl with (1 := Xg) => ?? Hs. + apply lift_typing_impl with (1 := Xg) => // ?? Hs. apply (IH (_; _; _; Hs)). -- pose proof Xg.(onConstructors) as Xg'. eapply All2_impl; eauto. intros. destruct X13 as [cass tyeq onctyp oncargs oncind]. unshelve econstructor; eauto. - { apply lift_typing_impl with (1 := onctyp) => ?? Hs. + { apply lift_typing_impl with (1 := onctyp) => // ?? Hs. apply (IH (_; _; _; Hs)). } - { apply sorts_local_ctx_impl with (1 := oncargs) => Γ' j Hj. - apply lift_typing_impl with (1 := Hj) => ?? Hs. + { apply sorts_local_ctx_impl with (1 := oncargs) => // Γ' j Hj. + apply lift_typing_impl with (1 := Hj) => // ?? Hs. apply (IH (_; _; _; Hs)). } { clear -IH oncind. revert oncind. @@ -1312,14 +1312,14 @@ Proof. -- destruct Xg. simpl. unfold check_ind_sorts in *. destruct (ind_sort x) => //=. split. apply ind_sorts0. destruct indices_matter; auto. - eapply type_local_ctx_impl. eapply ind_sorts0. intros ?? HT. - apply lift_typing_impl with (1 := HT) => ?? Hs. + eapply type_local_ctx_impl. eapply ind_sorts0. auto. intros ?? HT. + apply lift_typing_impl with (1 := HT) => // ?? Hs. apply (IH (_; _; _; Hs)). -- apply Xg.(ind_relevance_compat). -- apply (onIndices Xg). * red in onP |- *. apply All_local_env_impl with (1 := onP); intros ?? HT. - apply lift_typing_impl with (1 := HT) => ?? Hs. + apply lift_typing_impl with (1 := HT) => // ?? Hs. apply (IH ((Σ', udecl); (wfΣ; _; _; _; Hs))). constructor 1. simpl. subst Σ' Σg; cbn; lia. @@ -1442,7 +1442,7 @@ Proof. assert (forall Γ' : context, wf_local Σ Γ' -> forall (t T : term) (Hty : Σ;;; Γ' |- t : T), - typing_size Hty < S (all_size _ (fun x p => lift_typing_size (@typing_size _) Σ _ _ p) a0) -> + typing_size Hty < S (all_size _ (fun x p => (lift_typing_size (@typing_size _) Σ _ _ p.1 + lift_typing_size (@typing_size _) Σ _ _ p.2)%nat) a0) -> on_global_env cumul_gen (lift_typing P) Σ.1 * P Σ Γ' t T). { intros. eapply (X14 _ _ _ Hty); eauto. simpl in *. lia. } clear X14 X13. @@ -1450,9 +1450,11 @@ Proof. remember (fix_context mfix) as mfixcontext. clear Heqmfixcontext. induction a0; econstructor; eauto. - ++ destruct p as ((Hbo & Hborel) & s & e & Hs). + ++ destruct p as (((Hbo & mk) & _s & _e & _Hs) & tt & s & e & Hs). repeat split; auto. eapply (X _ (typing_wf_local Hbo) _ _ Hbo). simpl. lia. + exists _s; repeat split; auto. + apply (X _ (typing_wf_local _Hs) _ _ _Hs). simpl. lia. exists s; repeat split; auto. apply (X _ (typing_wf_local Hs) _ _ Hs). simpl. lia. ++ eapply IHa0. intros. @@ -1463,7 +1465,7 @@ Proof. assert (forall Γ' : context, wf_local Σ Γ' -> forall (t T : term) (Hty : Σ;;; Γ' |- t : T), - typing_size Hty < S (all_size _ (fun x p => lift_typing_size (@typing_size _) Σ _ _ p) a0) -> + typing_size Hty < S (all_size _ (fun x p => (lift_typing_size (@typing_size _) Σ _ _ p.1 + lift_typing_size (@typing_size _) Σ _ _ p.2)%nat) a0) -> on_global_env cumul_gen (lift_typing P) Σ.1 * P Σ Γ' t T). { intros. eapply (X14 _ _ _ Hty); eauto. simpl in *. lia. } clear X14 X13. @@ -1471,9 +1473,11 @@ Proof. remember (fix_context mfix) as mfixcontext. clear Heqmfixcontext. induction a0; econstructor; eauto. - ++ destruct p as ((Hbo & Hborel) & s & e & Hs). + ++ destruct p as (((Hbo & mk) & _s & _e & _Hs) & tt & s & e & Hs). repeat split; auto. eapply (X _ (typing_wf_local Hbo) _ _ Hbo). simpl. lia. + exists _s; repeat split; auto. + apply (X _ (typing_wf_local _Hs) _ _ _Hs). simpl. lia. exists s; repeat split; auto. apply (X _ (typing_wf_local Hs) _ _ Hs). simpl. lia. ++ eapply IHa0. intros. diff --git a/template-coq/theories/TypingWf.v b/template-coq/theories/TypingWf.v index 7ed097003..298125d3a 100644 --- a/template-coq/theories/TypingWf.v +++ b/template-coq/theories/TypingWf.v @@ -43,7 +43,7 @@ Lemma on_global_decl_impl `{checker_flags} Σ P Q kn d : (forall Σ Γ j, on_global_env cumul_gen P Σ.1 -> P Σ Γ j -> Q Σ Γ j) -> on_global_env cumul_gen P Σ.1 -> on_global_decl cumul_gen P Σ kn d -> on_global_decl cumul_gen Q Σ kn d. -Proof. intros. eapply on_global_decl_impl; tea. apply (X Σ). Qed. +Proof. intros. eapply on_global_decl_impl; tea. apply (X Σ). all: auto. Qed. Lemma on_global_env_impl `{checker_flags} Σ P Q : (forall Σ Γ j, on_global_env cumul_gen P Σ.1 -> P Σ Γ j -> Q Σ Γ j) -> @@ -217,7 +217,7 @@ Lemma sorts_local_ctx_All_wf_decl {cf:checker_flags} {Σ} {mdecl} {u: list Unive Proof. induction args as [|[na [b|] ty] args] in u |- * ; constructor. - - constructor; now destruct X as (?&(?&?)&(?&?)). + - constructor; now destruct X as (?&?&?). - eapply IHargs; now apply X. - destruct u => //; constructor; cbnr; now destruct X as (?&?&?&?). - destruct u => //; eapply IHargs; now apply X. @@ -923,21 +923,19 @@ Section TypingWf. split. + constructor. apply All_impl with (1 := X0) => d H. - destruct H as (((_ & (? & _)) & _) & _ & _ & _ & ? & _). - split => //. now eapply wf_lift_wf. + destruct H as ((((_ & ? & _) & _) & _) & (_ & _ & _ & _ & (? & _))). + split => //. + eapply All_nth_error in X0; eauto. - destruct X0 as (((_ & (? & _)) & _) & _ & _ & _ & ? & _). - now eapply wf_lift_wf. + destruct X0 as ((((_ & ? & _) & _) & _) & (_ & _ & _ & _ & (? & _))) => //. - subst types. split. + constructor. apply All_impl with (1 := X0) => d Hs. - destruct Hs as (((_ & (? & _)) & _) & _ & _ & _ & ? & _). - split => //. now eapply wf_lift_wf. + destruct Hs as ((((_ & ? & _) & _) & _) & (_ & _ & _ & _ & (? & _))). + split => //. + eapply All_nth_error in X0; eauto. - destruct X0 as (((_ & (? & _)) & _) & _ & _ & _ & ? & _). - now eapply wf_lift_wf. + destruct X0 as ((((_ & ? & _) & _) & _) & (_ & _ & _ & _ & (? & _))) => //. Qed. Lemma typing_all_wf_decl Σ (wfΣ : wf Σ.1) Γ (wfΓ : wf_local Σ Γ) :