From 45666652abfbc44e3e23289ddfa4d107a2b302af Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Mon, 30 Oct 2023 12:51:51 +0100 Subject: [PATCH 01/17] Support primitive array terms --- common/theories/Environment.v | 47 +++-- common/theories/Primitive.v | 4 +- erasure/theories/EEtaExpandedFix.v | 7 +- pcuic/theories/Bidirectional/BDTyping.v | 18 +- .../Conversion/PCUICUnivSubstitutionConv.v | 46 +++++ .../Conversion/PCUICWeakeningConfigConv.v | 6 + .../Conversion/PCUICWeakeningEnvConv.v | 6 + pcuic/theories/PCUICAst.v | 189 +++++++++++++++++- pcuic/theories/PCUICCSubst.v | 1 + pcuic/theories/PCUICCumulativitySpec.v | 39 +++- pcuic/theories/PCUICEquality.v | 53 ++++- pcuic/theories/PCUICReduction.v | 66 +++++- pcuic/theories/PCUICSigmaCalculus.v | 2 + pcuic/theories/PCUICTyping.v | 71 ++++++- pcuic/theories/Syntax/PCUICClosed.v | 1 + pcuic/theories/Syntax/PCUICInduction.v | 38 +++- pcuic/theories/Syntax/PCUICLiftSubst.v | 4 +- pcuic/theories/Syntax/PCUICOnFreeVars.v | 10 +- pcuic/theories/Syntax/PCUICReflect.v | 14 +- pcuic/theories/Typing/PCUICClosedTyp.v | 24 ++- .../Typing/PCUICUnivSubstitutionTyp.v | 32 ++- .../theories/Typing/PCUICWeakeningConfigTyp.v | 6 +- pcuic/theories/Typing/PCUICWeakeningEnvTyp.v | 14 +- pcuic/theories/utils/PCUICAstUtils.v | 35 +++- pcuic/theories/utils/PCUICPretty.v | 4 +- pcuic/theories/utils/PCUICPrimitive.v | 80 ++++++-- pcuic/theories/utils/PCUICSize.v | 8 + template-coq/src/ast_denoter.ml | 5 +- template-coq/src/ast_quoter.ml | 12 +- template-coq/src/constr_denoter.ml | 36 ++-- template-coq/src/constr_quoter.ml | 12 +- template-coq/src/constr_reification.ml | 10 +- template-coq/src/denoter.ml | 21 +- template-coq/src/quoter.ml | 89 +++++---- template-coq/src/reification.ml | 1 + template-coq/src/tm_util.ml | 61 +++--- template-coq/theories/Ast.v | 21 +- template-coq/theories/AstUtils.v | 4 + template-coq/theories/Checker.v | 2 +- template-coq/theories/Constants.v | 1 + template-coq/theories/EtaExpand.v | 19 +- template-coq/theories/Induction.v | 2 + template-coq/theories/Pretty.v | 2 + template-coq/theories/ReflectAst.v | 10 + template-coq/theories/TermEquality.v | 10 +- template-coq/theories/Typing.v | 46 ++++- template-coq/theories/TypingWf.v | 4 + template-coq/theories/WfAst.v | 8 +- 48 files changed, 995 insertions(+), 206 deletions(-) diff --git a/common/theories/Environment.v b/common/theories/Environment.v index aaa015c1b..caa3517ba 100644 --- a/common/theories/Environment.v +++ b/common/theories/Environment.v @@ -40,51 +40,58 @@ Module Retroknowledge. Record t := mk_retroknowledge { retro_int63 : option kername; retro_float64 : option kername; + retro_array : option kername; }. - Definition empty := {| retro_int63 := None; retro_float64 := None |}. + Definition empty := {| retro_int63 := None; retro_float64 := None; retro_array := None |}. Definition extends (x y : t) := option_extends x.(retro_int63) y.(retro_int63) /\ - option_extends x.(retro_float64) y.(retro_float64). + option_extends x.(retro_float64) y.(retro_float64) /\ + option_extends x.(retro_array) y.(retro_array). Existing Class extends. Definition extendsb (x y : t) := option_extendsb x.(retro_int63) y.(retro_int63) && - option_extendsb x.(retro_float64) y.(retro_float64). + option_extendsb x.(retro_float64) y.(retro_float64) && + option_extendsb x.(retro_array) y.(retro_array). Lemma extendsT x y : reflect (extends x y) (extendsb x y). Proof. - rewrite /extends/extendsb; do 2 case: option_extendsT; cbn; constructor; intuition. + rewrite /extends/extendsb; do 3 case: option_extendsT; cbn; constructor; intuition. Qed. Lemma extends_spec x y : extendsb x y <-> extends x y. Proof. rewrite /extends/extendsb -!option_extends_spec /is_true !Bool.andb_true_iff //=. + intuition auto. Qed. #[global] Instance extends_refl x : extends x x. Proof. - split; apply option_extends_refl. + repeat split; apply option_extends_refl. Qed. #[global] Instance extends_trans : RelationClasses.Transitive Retroknowledge.extends. Proof. - intros x y z [] []; split; cbn; now etransitivity; tea. + intros x y z [? []] [? []]; repeat split; cbn; now etransitivity; tea. Qed. #[export,program] Instance reflect_t : ReflectEq t := { eqb x y := (x.(retro_int63) == y.(retro_int63)) && - (x.(retro_float64) == y.(retro_float64)) + (x.(retro_float64) == y.(retro_float64)) && + (x.(retro_array) == y.(retro_array)) }. Next Obligation. - do 2 case: eqb_spec; destruct x, y; cbn; intros; subst; constructor; congruence. + do 3 case: eqb_spec; destruct x, y; cbn; intros; subst; constructor; congruence. Qed. (** This operation is asymmetric; it perfers the first argument when the arguments are incompatible, but otherwise takes the join *) Definition merge (r1 r2 : t) : t := {| retro_int63 := match r1.(retro_int63) with Some v => Some v | None => r2.(retro_int63) end - ; retro_float64 := match r1.(retro_float64) with Some v => Some v | None => r2.(retro_float64) end |}. + ; retro_float64 := match r1.(retro_float64) with Some v => Some v | None => r2.(retro_float64) end + ; retro_array := match r1.(retro_array) with Some v => Some v | None => r2.(retro_array) end + |}. Lemma extends_l_merge r1 r2 : extends r1 (merge r1 r2). @@ -102,7 +109,8 @@ Module Retroknowledge. Definition compatible (x y : t) : bool := match x.(retro_int63), y.(retro_int63) with Some x, Some y => x == y | _, _ => true end - && match x.(retro_float64), y.(retro_float64) with Some x, Some y => x == y | _, _ => true end. + && match x.(retro_float64), y.(retro_float64) with Some x, Some y => x == y | _, _ => true end + && match x.(retro_array), y.(retro_array) with Some x, Some y => x == y | _, _ => true end. Lemma extends_r_merge r1 r2 : compatible r1 r2 -> extends r2 (merge r1 r2). @@ -836,12 +844,25 @@ Module Environment (T : Term). match p with | primInt => Σ.(retroknowledge).(Retroknowledge.retro_int63) | primFloat => Σ.(retroknowledge).(Retroknowledge.retro_float64) + | primArray => Σ.(retroknowledge).(Retroknowledge.retro_array) end. - Definition primitive_invariants (cdecl : constant_body) := - ∑ s, [/\ cdecl.(cst_type) = tSort s, cdecl.(cst_body) = None & - cdecl.(cst_universes) = Monomorphic_ctx]. + Definition tImpl (dom codom : term) : term := + tProd {| binder_name := nAnon; binder_relevance := Relevant |} + dom (lift 1 0 codom). + Definition array_uctx := ([nAnon], ConstraintSet.empty). + + Definition primitive_invariants (p : prim_tag) (cdecl : constant_body) := + match p with + | primInt | primFloat => + ∑ s, [/\ cdecl.(cst_type) = tSort s, cdecl.(cst_body) = None & + cdecl.(cst_universes) = Monomorphic_ctx] + | primArray => + let s := Universe.make (Level.lvar 0) in + [/\ cdecl.(cst_type) = tImpl (tSort s) (tSort s), cdecl.(cst_body) = None & + cdecl.(cst_universes) = Polymorphic_ctx array_uctx] + end. (** A context of global declarations + global universe constraints, i.e. a global environment *) diff --git a/common/theories/Primitive.v b/common/theories/Primitive.v index d8c6e95c2..c3de3770a 100644 --- a/common/theories/Primitive.v +++ b/common/theories/Primitive.v @@ -6,8 +6,8 @@ Local Open Scope bs. Variant prim_tag := | primInt - | primFloat. - (* | primArray. *) + | primFloat + | primArray. Derive NoConfusion EqDec for prim_tag. Definition string_of_prim_int (i:Uint63.int) : string := diff --git a/erasure/theories/EEtaExpandedFix.v b/erasure/theories/EEtaExpandedFix.v index 0d68fb076..951bdf971 100644 --- a/erasure/theories/EEtaExpandedFix.v +++ b/erasure/theories/EEtaExpandedFix.v @@ -525,7 +525,7 @@ Section isEtaExp. #|Γ| = k -> nth_error mfix idx = Some d -> closed (EAst.tFix mfix idx) -> - forallb (fun x => isLambda x.(dbody) && isEtaExp (rev_map (S ∘ rarg) mfix) x.(dbody)) mfix -> + forallb (fun x => isEtaExp (rev_map (S ∘ rarg) mfix) x.(dbody)) mfix -> isEtaExp (Γ ++ [1 + d.(EAst.rarg)] ++ Δ) b -> isEtaExp (Γ ++ Δ) (ECSubst.csubst (EAst.tFix mfix idx) k b). Proof using Type*. intros Hk Hnth Hcl. @@ -631,13 +631,14 @@ Section isEtaExp. Qed. Lemma isEtaExp_fixsubstl Δ mfix t : - forallb (fun x => isLambda x.(dbody) && + forallb (fun x => + (* isLambda x.(dbody) && *) isEtaExp (rev_map (S ∘ rarg) mfix) x.(dbody)) mfix -> isEtaExp ((rev_map (S ∘ rarg) mfix) ++ Δ) t -> isEtaExp Δ (substl (fix_subst mfix) t). Proof using Type*. intros Hall Heta. - assert (Hcl : closed (EAst.tFix mfix 0) ). { cbn. solve_all. rtoProp; intuition auto. eapply isEtaExp_closed in H0. revert H0. now len. } + assert (Hcl : closed (EAst.tFix mfix 0) ). { cbn. solve_all. rtoProp; intuition auto. eapply isEtaExp_closed in H. revert H. now len. } revert Hcl Hall Heta. intros Hcl Hall Heta. cbn in Hcl. solve_all. diff --git a/pcuic/theories/Bidirectional/BDTyping.v b/pcuic/theories/Bidirectional/BDTyping.v index 5b4cd20b2..c36912543 100644 --- a/pcuic/theories/Bidirectional/BDTyping.v +++ b/pcuic/theories/Bidirectional/BDTyping.v @@ -124,8 +124,9 @@ Inductive infering `{checker_flags} (Σ : global_env_ext) (Γ : context) : term | infer_Prim p prim_ty cdecl : primitive_constant Σ (prim_val_tag p) = Some prim_ty -> declared_constant Σ prim_ty cdecl -> - primitive_invariants cdecl -> - Σ ;;; Γ |- tPrim p ▹ tConst prim_ty [] + primitive_invariants (prim_val_tag p) cdecl -> + primitive_typing_hyps checking Σ Γ p -> + Σ ;;; Γ |- tPrim p ▹ prim_type p prim_ty with infering_sort `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> Universe.t -> Type := | infer_sort_Sort t T u: @@ -196,6 +197,7 @@ Proof. | H : checking _ _ _ _ |- _ => apply checking_size in H | H : wf_local_bd _ _ |- _ => apply (All_local_env_sorting_size _ _ (checking_size _) (infering_sort_size _) _ _) in H | H : wf_local_bd_rel _ _ _ |- _ => apply (All_local_rel_sorting_size (checking_size _) (infering_sort_size _) _ _) in H + | H : primitive_typing_hyps _ _ _ _ |- _ => apply (primitive_typing_hyps_size _ (checking_size _)) in H end ; match goal with | H : All2i _ _ _ _ |- _ => idtac @@ -448,8 +450,10 @@ Section BidirectionalInduction. (forall (Γ : context) p prim_ty cdecl, primitive_constant Σ (prim_val_tag p) = Some prim_ty -> declared_constant Σ prim_ty cdecl -> - primitive_invariants cdecl -> - Pinfer Γ (tPrim p) (tConst prim_ty [])) -> + primitive_invariants (prim_val_tag p) cdecl -> + primitive_typing_hyps checking Σ Γ p -> + primitive_typing_hyps (fun _ => Pcheck) Σ Γ p -> + Pinfer Γ (tPrim p) (prim_type p prim_ty)) -> (forall (Γ : context) (t T : term) (u : Universe.t), Σ ;;; Γ |- t ▹ T -> @@ -681,6 +685,12 @@ Section BidirectionalInduction. cbn. lia. - unshelve eapply HPrim; eauto. + simpl in IH. + destruct p1; constructor; eauto. + applyIH. applyIH. simpl in IH. + clear -IH. induction hvalue; constructor; eauto. + eapply (IH (check_cons _ _ _ p)). simpl; lia. + eapply IHhvalue. intros; simpl. eapply IH. simpl. lia. - destruct i. unshelve (eapply HiSort ; try eassumption) ; try eassumption. diff --git a/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v b/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v index 3ef9f0181..77e366f26 100644 --- a/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v +++ b/pcuic/theories/Conversion/PCUICUnivSubstitutionConv.v @@ -190,6 +190,11 @@ Proof. * solve_all. reflexivity. * eapply X => //. - solve_all. reflexivity. + - destruct p as [? []]; cbn in X; constructor; cbn; intuition eauto; cbn. + * rewrite -!subst_instance_univ_make. now eapply hRe. + * now eapply a1. + * now eapply a0. + * solve_all. eapply All_All2; tea. intuition eauto. now apply X. Qed. #[global] @@ -355,6 +360,8 @@ Proof. rewrite map_def_map_def; solve_all. - rewrite map_map. apply All_map_eq. solve_all. rewrite map_def_map_def; solve_all. + - rewrite !mapu_prim_compose_rew. solve_all. + intro. eapply subst_instance_level_two. Qed. Lemma subst_instance_two_context u1 u2 (Γ : context) : @@ -880,6 +887,11 @@ Proof. repeat split; simpl; eauto; solve_all. * eapply precompose_subst_instance. eapply R_universe_instance_impl; eauto. + - destruct p as [? []]; depelim X1; cbn in X; try constructor; intuition eauto. + * cbn. rewrite -!subst_instance_univ_make. now eapply he. + * now eapply a2. + * now eapply a0. + * cbn. solve_all. Qed. Lemma leq_term_subst_instance {cf : checker_flags} Σ : SubstUnivPreserved (leq_term Σ). @@ -1442,6 +1454,7 @@ Proof. red. split; cbn; eauto. rewrite subst_instance_app in r0. now rewrite <- (fix_context_subst_instance u mfix0). + - cbn. eapply array_red_val. eapply OnOne2_map. cbn. solve_all. Qed. Lemma subst_instance_ws_cumul_pb {cf : checker_flags} (Σ : global_env_ext) Γ u A B univs : @@ -1819,6 +1832,30 @@ Section SubstIdentity. simpl in lin, onu. lsets. Qed. + + Lemma consistent_instance_ext_subst_abs_level Σ decl u : + wf_ext_wk Σ -> + consistent_instance_ext Σ decl [u] -> + subst_instance_level (abstract_instance Σ.2) u = u. + Proof using Type. + intros [wfΣ onu] cu. + destruct decl. + - simpl in cu. destruct u; simpl in *; try discriminate; auto. + - destruct cu as [decl' [sizeu vc]]. + clear sizeu vc. + destruct u; simpl; auto. cbn -[global_ext_levels] in decl'. + rewrite andb_true_r in decl'. + eapply LevelSet.mem_spec in decl'. + eapply in_var_global_ext in decl'; auto. + destruct (udecl_prop_in_var_poly onu decl') as [[univs csts] eq]. + rewrite eq in decl' |- *. simpl in *. + rewrite mapi_unfold in decl' |- *. + eapply LevelSet_In_fold_right_add in decl'. + eapply In_unfold_inj in decl'; try congruence. + eapply (nth_error_unfold Level.lvar) in decl'. + now rewrite (nth_error_nth _ _ _ decl'). + Qed. + Lemma consistent_instance_ext_subst_abs Σ decl u : wf_ext_wk Σ -> consistent_instance_ext Σ decl u -> @@ -2004,6 +2041,15 @@ Section SubstIdentity. - clear X0. eapply nth_error_all in X as [s [Hs [IHs _]]]; eauto. - solve_all. destruct a as [s [? ?]]. solve_all. - clear X0. eapply nth_error_all in X as [s [Hs [IHs _]]]; eauto. + - destruct p as [? []]; cbn => //. do 2 f_equal. + depelim X1. specialize (hty X2); specialize (hdef X2). + unfold mapu_array_model; destruct a; cbn -[Universe.make] in * => //=; f_equal; intuition eauto. + * destruct array_level => //. + rewrite subst_instance_univ_make in b. now injection b. + * solve_all. + - depelim X1; cbn => //=. destruct X. simp prim_type. cbn. f_equal; intuition eauto. + do 2 f_equal. cbn -[Universe.make] in b. rewrite subst_instance_univ_make in b. + now injection b. Qed. Lemma typed_subst_abstract_instance Σ Γ t T : diff --git a/pcuic/theories/Conversion/PCUICWeakeningConfigConv.v b/pcuic/theories/Conversion/PCUICWeakeningConfigConv.v index 4b0a0668a..7829564da 100644 --- a/pcuic/theories/Conversion/PCUICWeakeningConfigConv.v +++ b/pcuic/theories/Conversion/PCUICWeakeningConfigConv.v @@ -97,6 +97,8 @@ Proof. eapply All2_impl'; tea. eapply All_impl; eauto. cbn. intros x [? ?] y [[[? ?] ?] ?]. repeat split; eauto. + - intros h; depelim h; constructor; cbn in X; intuition eauto. + depelim e; cbn in X; constructor; intuition eauto. solve_all. Qed. Lemma weakening_config_cumul_gen {cf1 cf2} pb Σ Γ M N : @@ -142,6 +144,10 @@ Proof. * solve_all. - eapply cumul_Fix. solve_all. - eapply cumul_CoFix; solve_all. + - eapply cumul_Prim; solve_all. + destruct X; constructor; eauto. + * now eapply (@cmp_universe_config_impl cf1 cf2). + * solve_all. - eapply cumul_Ind; eauto. 2:solve_all. eapply @R_global_instance_weaken_subrel; [ .. | eassumption ]. all: repeat change (@eq_universe ?cf) with (@compare_universe cf Conv). diff --git a/pcuic/theories/Conversion/PCUICWeakeningEnvConv.v b/pcuic/theories/Conversion/PCUICWeakeningEnvConv.v index e6a8b8809..1f2ea9c14 100644 --- a/pcuic/theories/Conversion/PCUICWeakeningEnvConv.v +++ b/pcuic/theories/Conversion/PCUICWeakeningEnvConv.v @@ -125,6 +125,8 @@ Proof using P Pcmp cf. eapply All2_impl'; tea. eapply All_impl; eauto. cbn. intros x [? ?] y [[[? ?] ?] ?]. repeat split; eauto. + - inversion 1; subst; constructor. + depelim X1; constructor; cbn in X; intuition eauto. solve_all. Qed. Lemma weakening_env_red1 Σ Σ' Γ M N : @@ -198,6 +200,10 @@ Proof. * solve_all. - eapply cumul_Fix; solve_all. - eapply cumul_CoFix; solve_all. + - eapply cumul_Prim; solve_all. + destruct X; constructor; eauto. + * unfold compare_universe. eapply subrel_extends_eq; tea. + * solve_all. - eapply cumul_Ind; eauto. 2:solve_all. eapply @R_global_instance_weaken_env. 1,2,6:eauto. all: tc. - eapply cumul_Construct; eauto. 2:solve_all. diff --git a/pcuic/theories/PCUICAst.v b/pcuic/theories/PCUICAst.v index d1d985dce..8f03ffa5b 100644 --- a/pcuic/theories/PCUICAst.v +++ b/pcuic/theories/PCUICAst.v @@ -1,7 +1,7 @@ (* Distributed under the terms of the MIT license. *) -From Coq Require Import ssreflect Morphisms. +From Coq Require Import ssreflect ssrbool Morphisms. From MetaCoq.Utils Require Import utils. -From MetaCoq.Common Require Export Universes BasicAst Environment Reflect. +From MetaCoq.Common Require Export Primitive Universes BasicAst Environment Reflect. From MetaCoq.Common Require EnvironmentTyping. From MetaCoq.PCUIC Require Export PCUICPrimitive. From Equations Require Import Equations. @@ -238,6 +238,38 @@ Definition isLambda t := Lemma isLambda_inv t : isLambda t -> exists na ty bod, t = tLambda na ty bod. Proof. destruct t => //; eauto. Qed. +Set Equations Transparent. + +Definition mapu_array_model {term term'} (fl : Level.t -> Level.t) (f : term -> term') + (ar : array_model term) : array_model term' := + {| array_level := fl ar.(array_level); + array_value := map f ar.(array_value); + array_default := f ar.(array_default); + array_type := f ar.(array_type) |}. + +Equations mapu_prim {term term'} (f : Level.t -> Level.t) (g : term -> term') + (p : PCUICPrimitive.prim_val term) : PCUICPrimitive.prim_val term' := +| _, _, (primInt; primIntModel i) => (primInt; primIntModel i) +| _, _, (primFloat; primFloatModel fl) => (primFloat; primFloatModel fl) +| f, g, (primArray; primArrayModel ar) => + (primArray; primArrayModel (mapu_array_model f g ar)). + +Notation map_array_model := (mapu_array_model id). +Notation map_prim := (mapu_prim id). + +Equations test_prim (p : term -> bool) (p : prim_val) : bool := +| p, (primInt; _) => true +| p, (primFloat; _) => true +| p, (primArray; primArrayModel ar) => + List.forallb p ar.(array_value) && p ar.(array_default) && p ar.(array_type). + +Equations test_primu (p : Level.t -> bool) (t : term -> bool) (p : prim_val) : bool := +| _, _, (primInt; _) => true +| _, _, (primFloat; _) => true +| p, pt, (primArray; primArrayModel ar) => + p ar.(array_level) && forallb pt ar.(array_value) && + pt ar.(array_default) && pt ar.(array_type). + (** Basic operations on the AST: lifting, substitution and tests for variable occurrences. *) Fixpoint lift n k t : term := @@ -261,6 +293,7 @@ Fixpoint lift n k t : term := let k' := List.length mfix + k in let mfix' := List.map (map_def (lift n k) (lift n k')) mfix in tCoFix mfix' idx + | tPrim p => tPrim (map_prim (lift n k) p) | x => x end. @@ -296,6 +329,7 @@ Fixpoint subst s k u := let k' := List.length mfix + k in let mfix' := List.map (map_def (subst s k) (subst s k')) mfix in tCoFix mfix' idx + | tPrim p => tPrim (map_prim (subst s k) p) | x => x end. @@ -323,6 +357,7 @@ Fixpoint closedn k (t : term) : bool := | tCoFix mfix idx => let k' := List.length mfix + k in List.forallb (test_def (closedn k) (closedn k')) mfix + | tPrim p => test_prim (closedn k) p | _ => true end. @@ -360,6 +395,7 @@ Fixpoint nlict (t : term) : bool := List.forallb (test_def nlict nlict) mfix | tCoFix mfix idx => List.forallb (test_def nlict nlict) mfix + | tPrim p => test_prim nlict p | _ => true end. @@ -386,6 +422,7 @@ Fixpoint noccur_between k n (t : term) : bool := | tCoFix mfix idx => let k' := List.length mfix + k in List.forallb (test_def (noccur_between k n) (noccur_between k' n)) mfix + | tPrim p => test_prim (noccur_between k n) p | _ => true end. @@ -420,7 +457,7 @@ Instance subst_instance_constr : UnivSubst term := | tCoFix mfix idx => let mfix' := List.map (map_def (subst_instance_constr u) (subst_instance_constr u)) mfix in tCoFix mfix' idx - | tPrim _ => c + | tPrim p => tPrim (mapu_prim (subst_instance_level u) (subst_instance_constr u) p) end. (** Tests that the term is closed over [k] universe variables *) @@ -444,6 +481,7 @@ Fixpoint closedu (k : nat) (t : term) : bool := forallb (test_def (closedu k) (closedu k)) mfix | tCoFix mfix idx => forallb (test_def (closedu k) (closedu k)) mfix + | tPrim p => test_primu (closedu_level k) (closedu k) p | _ => true end. @@ -924,6 +962,151 @@ Proof. eapply map_ext => b. apply map_branch_map_branch. Qed. +Lemma mapu_prim_compose {term term' term''} + f (g : term' -> term'') f' (g' : term -> term') : mapu_prim f g ∘ mapu_prim f' g' =1 mapu_prim (f ∘ f') (g ∘ g'). +Proof. + intros [? []]; cbn => //. do 3 f_equal. + unfold mapu_array_model; destruct a => //=. now rewrite map_map_compose. +Qed. + +Lemma mapu_prim_compose_rew {term term' term''} + f (g : term' -> term'') f' (g' : term -> term') : + forall x, mapu_prim f g (mapu_prim f' g' x) = mapu_prim (f ∘ f') (g ∘ g') x. +Proof. intros x. now rewrite (mapu_prim_compose _ _ _ _ x). Qed. + +#[global] Hint Rewrite @mapu_prim_compose_rew : map. + +Definition tPrimProp {term} (P : term -> Type) (p : PCUICPrimitive.prim_val term) : Type := + match p.π2 return Type with + | primIntModel f => unit + | primFloatModel f => unit + | primArrayModel a => P a.(array_type) × P a.(array_default) × All P a.(array_value) + end. + +Lemma mapu_array_model_proper {term term'} (l l' : Level.t -> Level.t) (f g : term -> term') a : + l =1 l' -> f =1 g -> + mapu_array_model l f a = mapu_array_model l' g a. +Proof. + destruct a; cbn ; rewrite /mapu_array_model /=. intros; f_equal; eauto. now eapply map_ext. +Qed. + +Lemma mapu_array_model_proper_cond {term term'} (P : term -> Type) (l l' : Level.t -> Level.t) (f g : term -> term') a : + l =1 l' -> (forall x, P x -> f x = g x) -> + P a.(array_type) × P a.(array_default) × All P a.(array_value) -> + mapu_array_model l f a = mapu_array_model l' g a. +Proof. + destruct a; cbn ; rewrite /mapu_array_model /=. intros hl hf [? []]; f_equal; eauto. + induction a; cbn => //. rewrite IHa. rewrite hf //. +Qed. + +Lemma primProp_map_eq {term term'} P p l l' (f g : term -> term') : + tPrimProp P p -> + l =1 l' -> + (forall x, P x -> f x = g x) -> + mapu_prim l f p = mapu_prim l' g p. +Proof. + destruct p as [? []]; cbn => //. + intros [? []] hl hp. do 2 f_equal. + eapply mapu_array_model_proper_cond; tea. intuition auto. +Qed. + +Lemma primProp_map_id {term} P p (f : term -> term) : + tPrimProp P p -> + (forall x, P x -> f x = x) -> + map_prim f p = p. +Proof. + intros hp hf. + rewrite (primProp_map_eq P p id id f id) //. + destruct p as [? []]; cbn => //. + destruct a => //=. rewrite /mapu_array_model /=. rewrite map_id //. +Qed. + +Lemma test_prim_primProp {P p} : test_prim P p -> tPrimProp P p. +Proof. + destruct p as [? []]; cbn => //. + move/andP => [] /andP[]. intuition auto. + now eapply forallb_All in a0. +Qed. + +Lemma primProp_test_prim {P : term -> bool} {p} : tPrimProp P p -> test_prim P p. +Proof. + destruct p as [? []]; cbn => //. intros. rtoProp. + intuition auto. + now eapply All_forallb. +Qed. + +Lemma tPrimProp_impl {term} {P Q : term -> Type} {p} : tPrimProp P p -> (forall x, P x -> Q x) -> tPrimProp Q p. +Proof. + intros hp hpq; destruct p as [? []]; cbn in * => //. intuition auto. + eapply All_impl; tea. +Qed. + +Lemma tPrimProp_prod {term} {P Q : term -> Type} {p} : tPrimProp P p -> tPrimProp Q p -> tPrimProp (fun x => P x × Q x) p. +Proof. + destruct p as [? []]; cbn => //. intuition auto. + now eapply All_prod. +Qed. + +Lemma primProp_map {term} (P : term -> Type) f (p : PCUICPrimitive.prim_val term) : + tPrimProp (fun x => P (f x)) p -> + tPrimProp P (map_prim f p). +Proof. + destruct p as [? []]; cbn => //. intuition eauto. now eapply All_map. +Qed. + +Lemma primProp_map_inv {term} (P : term -> Type) f (p : PCUICPrimitive.prim_val term) : + tPrimProp P (map_prim f p) -> + tPrimProp (fun x => P (f x)) p. +Proof. + destruct p as [? []]; cbn => //. intuition eauto; now eapply All_map_inv. +Qed. + +Lemma primProp_mapu_id {P : term -> Type} {pu put} p f g : + tPrimProp P p -> test_primu pu put p -> + (forall u, pu u -> f u = u) -> + (forall t, P t -> put t -> g t = t) -> + mapu_prim f g p = p. +Proof. + intros hp ht hf hg. + destruct p as [? []]; cbn => //. f_equal. f_equal. + destruct a; destruct hp; cbn in *. unfold mapu_array_model; cbn. + rtoProp. destruct p0. f_equal; intuition eauto. + eapply forallb_All in H2. eapply All_prod in a; tea. + eapply All_map_id, All_impl; tea. intuition eauto. apply hg; intuition auto. +Qed. + +Lemma test_primu_test_primu_tPrimProp {P : term -> Type} {pu put} {pu' : Level.t -> bool} {put' : term -> bool} p f g : + tPrimProp P p -> test_primu pu put p -> + (forall u, pu u -> pu' (f u)) -> + (forall t, P t -> put t -> put' (g t)) -> + test_primu pu' put' (mapu_prim f g p). +Proof. + intros hp ht hf hg. + destruct p as [? []]; cbn => //. + destruct a; destruct hp; cbn in *. + rtoProp. destruct p0. intuition eauto. + eapply forallb_All in H2. eapply All_prod in a; tea. + eapply All_forallb, All_map, All_impl; tea. intuition eauto. apply hg; intuition auto. +Qed. + +Lemma test_prim_mapu {put l f p} : + test_prim put (mapu_prim l f p) = test_prim (fun x => put (f x)) p. +Proof. + destruct p as [? []]; cbn; eauto. + now rewrite forallb_map. +Qed. +#[global] Hint Rewrite @test_prim_mapu : map. + +Lemma test_prim_eq_spec {P : term -> Type} {put} {put' : term -> bool} p : + tPrimProp P p -> + (forall t, P t -> put t = put' t) -> + test_prim put p = test_prim put' p. +Proof. + destruct p as [? []]; cbn => //. intuition eauto. + f_equal; eauto. f_equal; eauto. + eapply All_forallb_eq_forallb; tea. +Qed. + Definition tCaseBrsProp {A} (P : A -> Type) (l : list (branch A)) := All (fun x => onctx P (bcontext x) × P (bbody x)) l. diff --git a/pcuic/theories/PCUICCSubst.v b/pcuic/theories/PCUICCSubst.v index 9a9f66faf..7aa1b949f 100644 --- a/pcuic/theories/PCUICCSubst.v +++ b/pcuic/theories/PCUICCSubst.v @@ -40,6 +40,7 @@ Fixpoint csubst t k u := let k' := List.length mfix + k in let mfix' := List.map (map_def (csubst t k) (csubst t k')) mfix in tCoFix mfix' idx + | tPrim p => tPrim (map_prim (csubst t k) p) | x => x end. diff --git a/pcuic/theories/PCUICCumulativitySpec.v b/pcuic/theories/PCUICCumulativitySpec.v index e7c81aaef..12bf4aff5 100644 --- a/pcuic/theories/PCUICCumulativitySpec.v +++ b/pcuic/theories/PCUICCumulativitySpec.v @@ -34,6 +34,21 @@ Proof. all: repeat first [ assumption | toAll ]. Qed. +Inductive eq_prim_dep (eq_term : term -> term -> Type) (Re : Universe.t -> Universe.t -> Prop) (eq_term_dep : forall x y, eq_term x y -> Type) (Re' : forall a b, Re a b -> Type) : forall x y : prim_val, eq_prim eq_term Re x y -> Type := + | eqPrimInt_dep i : eq_prim_dep eq_term Re eq_term_dep Re' (primInt; i) (primInt; i) (eqPrimInt eq_term Re i) + | eqPrimFloat_dep f : eq_prim_dep eq_term Re eq_term_dep Re' (primFloat; f) (primFloat; f) (eqPrimFloat _ _ f) + | eqPrimArray_dep a a' : + forall (hre : Re (Universe.make a.(array_level)) (Universe.make a'.(array_level))) + (eqdef : eq_term a.(array_default) a'.(array_default)) + (eqty : eq_term a.(array_type) a'.(array_type)) + (eqt : All2 eq_term a.(array_value) a'.(array_value)), + Re' _ _ hre -> + eq_term_dep _ _ eqdef -> + eq_term_dep _ _ eqty -> + All2_dep eq_term_dep eqt -> + eq_prim_dep eq_term Re eq_term_dep Re' + (primArray; primArrayModel a) (primArray; primArrayModel a') (eqPrimArray _ _ a a' hre eqdef eqty eqt). + Reserved Notation " Σ ;;; Γ ⊢ t ≤s[ pb ] u" (at level 50, Γ, t, u at next level, format "Σ ;;; Γ ⊢ t ≤s[ pb ] u"). @@ -147,6 +162,10 @@ Inductive cumulSpec0 {cf : checker_flags} (Σ : global_env_ext) Γ (pb : conv_pb ) mfix mfix' -> Σ ;;; Γ ⊢ tCoFix mfix idx ≤s[pb] tCoFix mfix' idx +| cumul_Prim p p' : + eq_prim (fun x y => Σ ;;; Γ ⊢ x ≤s[Conv] y) (compare_universe Conv Σ) p p' -> + Σ ;;; Γ ⊢ tPrim p ≤s[pb] tPrim p' + (** Reductions *) (** Beta red *) @@ -425,6 +444,10 @@ Lemma cumulSpec0_rect : P cf Σ Γ pb (tCoFix mfix idx) (tCoFix mfix' idx) (cumul_CoFix _ _ _ _ _ _ H)) -> + (forall Γ pb p p' (e : eq_prim (cumulSpec0 Σ Γ Conv) (eq_universe Σ) p p'), + eq_prim_dep (cumulSpec0 Σ Γ Conv) (eq_universe Σ) (P cf Σ Γ Conv) (fun _ _ _ => True) p p' e -> + P cf Σ Γ pb (tPrim p) (tPrim p') (cumul_Prim _ _ _ _ _ e)) -> + (* cumulativity rules *) (forall (Γ : context) (pb : conv_pb) (i : inductive) (u u' : list Level.t) @@ -463,14 +486,14 @@ Proof. - eapply X8; eauto. - eapply X9; eauto. - eapply X10; eauto. - - eapply X20; eauto. clear -a aux. + - eapply X21; eauto. clear -a aux. revert args args' a. fix aux' 3; destruct a; constructor; auto. - - eapply X21; eauto. clear -a aux. + - eapply X22; eauto. clear -a aux. revert args args' a. fix aux' 3; destruct a; constructor; auto. - - eapply X22; eauto. - eapply X23; eauto. + - eapply X24; eauto. - eapply X11; eauto. revert args args' a. fix aux' 3; destruct a; constructor; auto. @@ -503,6 +526,11 @@ Proof. fix aux' 3; destruct a; constructor. + intuition auto. + auto. + - eapply X20; eauto. + clear -e aux. + induction e; constructor; auto. + clear -a0 aux. revert a0. + induction a0; constructor; auto. - eapply X; eauto. - eapply X0; eauto. - eapply X1; eauto. @@ -695,6 +723,10 @@ Lemma convSpec0_ind_all : P cf Σ Γ Conv (tCoFix mfix idx) (tCoFix mfix' idx) (cumul_CoFix _ _ _ _ _ _ H)) -> + (forall Γ p p' (e : eq_prim (cumulSpec0 Σ Γ Conv) (eq_universe Σ) p p'), + eq_prim_dep (cumulSpec0 Σ Γ Conv) (eq_universe Σ) (P cf Σ Γ Conv) (fun _ _ _ => True) p p' e -> + P cf Σ Γ Conv (tPrim p) (tPrim p') (cumul_Prim _ _ _ _ _ e)) -> + (* cumulativity rules *) (forall (Γ : context) (i : inductive) (u u' : list Level.t) @@ -729,6 +761,7 @@ Proof. remember Conv as pb eqn:Hpb in Ht |- *. induction Ht; eauto; subst. all: exactly_once (idtac; multimatch goal with H : _ |- _ => eapply H end; eauto). + 6:{ destruct X25; constructor; auto. eapply All2_dep_impl; tea; intuition auto. } all: cbv [cumul_predicate_dep] in *. all: repeat destruct ?; subst. all: destruct_head'_prod. diff --git a/pcuic/theories/PCUICEquality.v b/pcuic/theories/PCUICEquality.v index d4a3058fa..5396bc44e 100644 --- a/pcuic/theories/PCUICEquality.v +++ b/pcuic/theories/PCUICEquality.v @@ -275,6 +275,17 @@ Definition eq_predicate (eq_term : term -> term -> Type) Re p p' := ((eq_context_gen eq eq p.(pcontext) p'.(pcontext)) * eq_term p.(preturn) p'.(preturn))). +Inductive eq_prim (eq_term : term -> term -> Type) Re : prim_val -> prim_val -> Type := + | eqPrimInt i : eq_prim eq_term Re (primInt; i) (primInt; i) + | eqPrimFloat f : eq_prim eq_term Re (primFloat; f) (primFloat; f) + | eqPrimArray a a' : + Re (Universe.make a.(array_level)) (Universe.make a'.(array_level)) -> + eq_term a.(array_default) a'.(array_default) -> + eq_term a.(array_type) a'.(array_type) -> + All2 eq_term a.(array_value) a'.(array_value) -> + eq_prim eq_term Re (primArray; primArrayModel a) (primArray; primArrayModel a'). +Derive Signature for eq_prim. + (** ** Syntactic ws_cumul_pb up-to universes We don't look at printing annotations *) @@ -285,7 +296,6 @@ Definition eq_predicate (eq_term : term -> term -> Type) Re p p' := Reserved Notation " Σ ⊢ t <==[ Rle , napp ] u" (at level 50, t, u at next level, format "Σ ⊢ t <==[ Rle , napp ] u"). - Inductive eq_term_upto_univ_napp Σ (Re Rle : Universe.t -> Universe.t -> Prop) (napp : nat) : term -> term -> Type := | eq_Rel : forall n, Σ ⊢ tRel n <==[ Rle , napp ] tRel n @@ -368,7 +378,9 @@ Inductive eq_term_upto_univ_napp Σ (Re Rle : Universe.t -> Universe.t -> Prop) ) 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) +| eq_Prim i i' : + eq_prim (eq_term_upto_univ_napp Σ Re Re 0) Re i i' -> + eq_term_upto_univ_napp Σ Re Rle napp (tPrim i) (tPrim i') where " Σ ⊢ t <==[ Rle , napp ] u " := (eq_term_upto_univ_napp Σ _ Rle napp t u) : type_scope. Notation eq_term_upto_univ Σ Re Rle := (eq_term_upto_univ_napp Σ Re Rle 0). @@ -530,6 +542,8 @@ Proof. eapply onctx_eq_ctx in a; eauto. - eapply All_All2; eauto; simpl; intuition eauto. - eapply All_All2; eauto; simpl; intuition eauto. + - destruct p as [? []]; constructor; cbn in X; intuition eauto. + eapply All_All2; eauto. Qed. #[global] @@ -608,6 +622,10 @@ Proof. + constructor. + destruct r as [[h1 h2] [[[h3 h4] h5] h6]]. eapply h1 in h3 ; auto. constructor; auto. + - econstructor. + depelim e; cbn in X; constructor; intuition eauto. + eapply All2_All_mix_left in a0 as h; eauto. cbn in h. + eapply All2_sym; solve_all. Qed. #[global] @@ -739,6 +757,16 @@ Proof. + dependent destruction a0. constructor ; eauto. intuition eauto. transitivity (rarg y); auto. + - dependent destruction e2; constructor. + depelim e; intuition eauto. depelim e1; constructor; cbn in X; intuition eauto. + eapply All2_All_mix_left in b0 as h; eauto. + clear b0 a0. clear -he hle a2 h. revert h a2. + generalize (array_value a) (array_value a') (array_value a'0). clear -he hle. + intros l l0 l1. + induction 1 in l1 |- *; intros. + + assumption. + + dependent destruction a2. constructor ; eauto. + destruct r as [h1 h2]. eauto. Qed. #[global] @@ -823,7 +851,7 @@ Proof. induction u in v, n, h, h' |- * using term_forall_list_ind. all: simpl ; inversion h ; subst; inversion h' ; subst ; try constructor ; auto. - all: eapply RelationClasses.antisymmetry; eauto. + all: try eapply RelationClasses.antisymmetry; eauto. Qed. #[global] @@ -953,6 +981,8 @@ Proof. eapply All2_impl'; tea. eapply All_impl; eauto. cbn. intros x [? ?] y [[[? ?] ?] ?]. repeat split; eauto. + - intros h; depelim h. depelim e; constructor; cbn in X; constructor; intuition eauto. + solve_all. Qed. #[global] @@ -984,6 +1014,8 @@ Proof. eapply All2_impl'; tea. eapply All_impl; eauto. cbn. intros x [? ?] y [[[? ?] ?] ?]. repeat split; eauto. + - intros h; depelim h. constructor. depelim e; cbn in X; constructor; intuition eauto. + solve_all. Qed. #[global] @@ -1039,6 +1071,11 @@ Proof. - cbn. constructor. pose proof (All2_length a). solve_all. rewrite H. eauto. + - cbn. constructor. depelim e; cbn in X; intuition eauto. + * depelim i; cbn in X; constructor; eauto. + * depelim f; cbn in X; constructor; eauto. + * constructor; cbn; eauto. + solve_all. Qed. Lemma lift_compare_term `{checker_flags} pb Σ ϕ n k t t' : @@ -1117,6 +1154,11 @@ Proof. - cbn. constructor ; try sih ; eauto. pose proof (All2_length a). solve_all. now rewrite H. + - cbn; constructor. depelim e. + * cbn. depelim i; cbn in X; constructor; intuition eauto. + * cbn. depelim f; cbn in X; constructor; intuition eauto. + * cbn. depelim a; cbn in X |- *; constructor; cbn; intuition eauto. + solve_all. Qed. Lemma eq_term_upto_univ_subst Σ Re Rle : @@ -1670,6 +1712,11 @@ Proof. now eapply eq_term_upto_univ_sym. now eapply eq_term_upto_univ_sym. now symmetry. + - depelim e; constructor; eauto. + now eapply eq_term_upto_univ_sym. + now eapply eq_term_upto_univ_sym. + eapply All2_sym; solve_all. + now eapply eq_term_upto_univ_sym. Qed. Lemma eq_univ_make : diff --git a/pcuic/theories/PCUICReduction.v b/pcuic/theories/PCUICReduction.v index 347a08f2f..92f19547c 100644 --- a/pcuic/theories/PCUICReduction.v +++ b/pcuic/theories/PCUICReduction.v @@ -1,8 +1,8 @@ (* Distributed under the terms of the MIT license. *) From MetaCoq.Utils Require Import utils. -From MetaCoq.Common Require Import config. +From MetaCoq.Common Require Import config Primitive. From MetaCoq.PCUIC Require Import PCUICAst PCUICOnOne PCUICAstUtils - PCUICLiftSubst PCUICUnivSubst PCUICInduction + PCUICLiftSubst PCUICUnivSubst PCUICInduction PCUICPrimitive PCUICCases PCUICClosed PCUICTactics. Require Import ssreflect. @@ -16,6 +16,25 @@ Reserved Notation " Σ ;;; Γ |- t ⇝ u " (at level 50, Γ, t, u at next level) Local Open Scope type_scope. +Definition set_array_default (ar : array_model term) (v : term) := + {| array_level := ar.(array_level); + array_default := v; + array_type := ar.(array_type); + array_value := ar.(array_value) |}. + +Definition set_array_type (ar : array_model term) (v : term) := + {| array_level := ar.(array_level); + array_default := ar.(array_default); + array_type := v; + array_value := ar.(array_value) |}. + +Definition set_array_value (ar : array_model term) (v : list term) := + {| array_level := ar.(array_level); + array_default := ar.(array_default); + array_type := ar.(array_type); + array_value := v |}. + + Inductive red1 (Σ : global_env) (Γ : context) : term -> term -> Type := (** Reductions *) (** Beta *) @@ -116,6 +135,22 @@ Inductive red1 (Σ : global_env) (Γ : context) : term -> term -> Type := | cofix_red_body mfix0 mfix1 idx : OnOne2 (on_Trel_eq (fun t u => Σ ;;; Γ ,,, fix_context mfix0 |- t ⇝ u) dbody (fun x => (dname x, dtype x, rarg x))) mfix0 mfix1 -> Σ ;;; Γ |- tCoFix mfix0 idx ⇝ tCoFix mfix1 idx + +| array_red_val arr value : + OnOne2 (fun t u => Σ ;;; Γ |- t ⇝ u) arr.(array_value) value -> + Σ ;;; Γ |- tPrim (primArray; primArrayModel arr) ⇝ + tPrim (primArray; primArrayModel (set_array_value arr value)) + +| array_red_def arr def : + Σ ;;; Γ |- arr.(array_default) ⇝ def -> + Σ ;;; Γ |- tPrim (primArray; primArrayModel arr) ⇝ + tPrim (primArray; primArrayModel (set_array_default arr def)) + +| array_red_type arr ty : + Σ ;;; Γ |- arr.(array_type) ⇝ ty -> + Σ ;;; Γ |- tPrim (primArray; primArrayModel arr) ⇝ + tPrim (primArray; primArrayModel (set_array_type arr ty)) + where " Σ ;;; Γ |- t ⇝ u " := (red1 Σ Γ t u). Definition red1_ctx Σ := (OnOne2_local_env (on_one_decl (fun Δ t t' => red1 Σ Δ t t'))). @@ -240,9 +275,28 @@ Lemma red1_ind_all : (fun x => (dname x, dtype x, rarg x))) mfix0 mfix1 -> P Γ (tCoFix mfix0 idx) (tCoFix mfix1 idx)) -> + + (forall (Γ : context) (arr : array_model term) + (value : list term), + OnOne2 (Trel_conj (red1 Σ Γ) (P Γ)) (array_value arr) value -> + P Γ (tPrim (primArray; primArrayModel arr)) + (tPrim (primArray; primArrayModel (set_array_value arr value)))) -> + + (forall (Γ : context) (arr : array_model term) + (def : term), Σ;;; Γ |- array_default arr ⇝ def -> + P Γ (array_default arr) def -> + P Γ (tPrim (primArray; primArrayModel arr)) + (tPrim (primArray; primArrayModel (set_array_default arr def)))) -> + + (forall (Γ : context) (arr : array_model term) + (ty : term), Σ;;; Γ |- array_type arr ⇝ ty -> + P Γ (array_type arr) ty -> + P Γ (tPrim (primArray; primArrayModel arr)) + (tPrim (primArray; primArrayModel (set_array_type arr ty)))) -> + forall (Γ : context) (t t0 : term), red1 Σ Γ t t0 -> P Γ t t0. Proof. - intros. rename X27 into Xlast. revert Γ t t0 Xlast. + intros. rename X30 into Xlast. revert Γ t t0 Xlast. fix aux 4. intros Γ t T. move aux at top. destruct 1; match goal with @@ -296,6 +350,10 @@ Proof. revert o. generalize (fix_context mfix0). intros c new. revert mfix0 mfix1 new; fix auxl 3; intros l l' Hl; destruct Hl; constructor; try split; auto; intuition. + + - revert value o. generalize (array_value arr). + fix auxl 3; intros l l' Hl; destruct Hl; + constructor; try split; auto; intuition. Defined. #[global] @@ -639,7 +697,7 @@ Section ReductionCongruence. revert Γ y r. eapply (fill_context_elim x P P' P''); subst P P' P''; cbv beta. all: intros **; simp fill_context; cbn in *; auto; try solve [constructor; eauto]. - Qed. + Qed. Theorem red_contextual_closure_equiv Γ t u : red Σ Γ t u <~> contextual_closure (red Σ) Γ t u. Proof using Type. diff --git a/pcuic/theories/PCUICSigmaCalculus.v b/pcuic/theories/PCUICSigmaCalculus.v index c161da5e7..f5ebf69e0 100644 --- a/pcuic/theories/PCUICSigmaCalculus.v +++ b/pcuic/theories/PCUICSigmaCalculus.v @@ -116,6 +116,7 @@ Fixpoint rename (f : renamingT) t : term := | tCoFix mfix idx => let mfix' := List.map (map_def (rename f) (rename (shiftn (List.length mfix) f))) mfix in tCoFix mfix' idx + | tPrim p => tPrim (map_prim (rename f) p) | x => x end. @@ -583,6 +584,7 @@ Fixpoint inst s u := | tCoFix mfix idx => let mfix' := map (map_def (inst s) (inst (up (List.length mfix) s))) mfix in tCoFix mfix' idx + | tPrim p => tPrim (map_prim (inst s) p) | x => x end. diff --git a/pcuic/theories/PCUICTyping.v b/pcuic/theories/PCUICTyping.v index 3e836690b..a5844cdb5 100644 --- a/pcuic/theories/PCUICTyping.v +++ b/pcuic/theories/PCUICTyping.v @@ -173,6 +173,25 @@ Variant case_branch_typing `{checker_flags} wf_local_fun typing Σ Γ (ci:case_i (typing Σ (Γ ,,, brctxty.1) brctxty.2 (tSort ps))))) 0 idecl.(ind_ctors) brs). +Variant primitive_typing_hyps `{checker_flags} + (typing : forall (Σ : global_env_ext) (Γ : context), term -> term -> Type) + Σ Γ : prim_val term -> Type := +| prim_int_hyps i : primitive_typing_hyps typing Σ Γ (primInt; i) +| prim_float_hyps f : primitive_typing_hyps typing Σ Γ (primFloat; f) +| prim_array_hyps a + (wfl : wf_universe Σ (Universe.make a.(array_level))) + (hty : typing Σ Γ a.(array_type) (tSort (Universe.make a.(array_level)))) + (hdef : typing Σ Γ a.(array_default) a.(array_type)) + (hvalue : All (fun x => typing Σ Γ x a.(array_type)) a.(array_value)) : + primitive_typing_hyps typing Σ Γ (primArray; primArrayModel a). +Derive Signature for primitive_typing_hyps. + +Equations prim_type (p : prim_val term) (cst : kername) : term := +prim_type (primInt; _) cst := tConst cst []; +prim_type (primFloat; _) cst := tConst cst []; +prim_type (primArray; primArrayModel a) cst := tApp (tConst cst [a.(array_level)]) a.(array_type). +Transparent prim_type. + Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> term -> Type := | type_Rel : forall n decl, wf_local Σ Γ -> @@ -266,8 +285,9 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> wf_local Σ Γ -> primitive_constant Σ (prim_val_tag p) = Some prim_ty -> declared_constant Σ prim_ty cdecl -> - primitive_invariants cdecl -> - Σ ;;; Γ |- tPrim p : tConst prim_ty [] + primitive_invariants (prim_val_tag p) cdecl -> + primitive_typing_hyps typing Σ Γ p -> + Σ ;;; Γ |- tPrim p : prim_type p prim_ty | type_Cumul : forall t A B s, Σ ;;; Γ |- t : A -> @@ -375,6 +395,19 @@ Section CtxInstSize. end. End CtxInstSize. +Section PrimitiveSize. + Context {cf} (typing : global_env_ext -> context -> term -> term -> Type) + (typing_size : forall {Σ Γ t T}, typing Σ Γ t T -> size). + + Definition primitive_typing_hyps_size Σ Γ p (h : primitive_typing_hyps typing Σ Γ p) : size. + destruct h. + - exact 0. + - exact 0. + - exact ((typing_size _ _ _ _ hty) + (typing_size _ _ _ _ hdef) + + all_size _ (fun x p => typing_size _ _ _ _ p) hvalue). + Defined. +End PrimitiveSize. + Definition typing_size `{checker_flags} {Σ Γ t T} (d : Σ ;;; Γ |- t : T) : size. Proof. revert Σ Γ t T d. @@ -406,7 +439,7 @@ Proof. (all_size _ (fun x p => (infer_sort_size (typing_sort_size typing_size)) Σ _ _ p) a0)) (all_size _ (fun x p => typing_size Σ _ _ _ p) a1))). - exact (S (Nat.max (Nat.max (All_local_env_size typing_size _ _ a) (all_size _ (fun x p => (infer_sort_size (typing_sort_size typing_size)) Σ _ _ p) a0)) (all_size _ (fun x p => typing_size Σ _ _ _ p) a1))). - - exact (S (All_local_env_size typing_size _ _ a)). + - exact (S (Nat.max (All_local_env_size typing_size _ _ a) (primitive_typing_hyps_size typing typing_size _ _ _ p1))). Defined. Lemma typing_size_pos `{checker_flags} {Σ Γ t T} (d : Σ ;;; Γ |- t : T) : typing_size d > 0. @@ -445,8 +478,8 @@ Definition wf_ext `{checker_flags} := on_global_env_ext cumulSpec0 (lift_typing Existing Class wf_ext. #[global] Hint Mode wf_ext + + : typeclass_intances. - Lemma wf_ext_wf {cf:checker_flags} Σ : wf_ext Σ -> wf Σ. + Proof. intro H; apply H. Qed. #[global] Existing Instance wf_ext_wf. @@ -585,6 +618,16 @@ Proof. all: try lia. Qed. +Lemma All_map_size {A} {P Q : A -> Type} (sizep : forall x, P x -> size) l + (a : All P l) : + (forall x (p : P x), sizep _ p < S (all_size P sizep a) -> Q x) -> + All Q l. +Proof. + induction a; constructor; auto. + - apply (X x p); simpl; lia. + - apply IHa. intros. eapply (X _ p0); simpl; lia. +Qed. + (** *** An induction principle ensuring the Σ declarations enjoy the same properties. Also theads the well-formedness of the local context and the induction principle for it, and gives the right induction hypothesis on typing judgments in application spines, @@ -731,8 +774,10 @@ Lemma typing_ind_env_app_size `{cf : checker_flags} : PΓ Σ Γ -> primitive_constant Σ (prim_val_tag p) = Some prim_ty -> declared_constant Σ prim_ty cdecl -> - primitive_invariants cdecl -> - P Σ Γ (tPrim p) (tConst prim_ty [])) -> + primitive_invariants (prim_val_tag p) cdecl -> + primitive_typing_hyps typing Σ Γ p -> + primitive_typing_hyps P Σ Γ p -> + P Σ Γ (tPrim p) (prim_type p prim_ty)) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (t A B : term) s, PΓ Σ Γ -> @@ -1051,6 +1096,14 @@ Proof. eapply (X _ (typing_wf_local p) _ _ p). simpl. lia. ++ eapply IHa1. intros. eapply (X _ X0 _ _ Hty). simpl; lia. + + -- eapply X12; tea. + clear -p2 X14. + simpl in X14. destruct p2; constructor; auto. + * eapply (X14 _ _ _ hty). simpl. lia. + * eapply (X14 _ _ _ hdef). simpl. lia. + * simpl in X14. eapply (All_map_size (fun t p => @typing_size cf Σ Γ t (array_type a0) p) _ hvalue); tea. + intros x p hs. apply (X14 _ _ _ p). simpl. lia. Qed. Lemma typing_ind_env `{cf : checker_flags} : @@ -1185,8 +1238,10 @@ Lemma typing_ind_env `{cf : checker_flags} : PΓ Σ Γ -> primitive_constant Σ (prim_val_tag p) = Some prim_ty -> declared_constant Σ prim_ty cdecl -> - primitive_invariants cdecl -> - P Σ Γ (tPrim p) (tConst prim_ty [])) -> + primitive_invariants (prim_val_tag p) cdecl -> + primitive_typing_hyps typing Σ Γ p -> + primitive_typing_hyps P Σ Γ p -> + P Σ Γ (tPrim p) (prim_type p prim_ty)) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (t A B : term) s, PΓ Σ Γ -> diff --git a/pcuic/theories/Syntax/PCUICClosed.v b/pcuic/theories/Syntax/PCUICClosed.v index cf7f0cf2b..8cddd17e5 100644 --- a/pcuic/theories/Syntax/PCUICClosed.v +++ b/pcuic/theories/Syntax/PCUICClosed.v @@ -121,6 +121,7 @@ Proof. autorewrite with map; simpl closed in *; solve_all; unfold test_def, test_snd, test_predicate_k, test_branch_k in *; + try eapply primProp_map; try solve [simpl lift; simpl closed; f_equal; auto; repeat (rtoProp; simpl in *; solve_all)]; try easy. - elim (Nat.leb_spec k' n0); intros. simpl. diff --git a/pcuic/theories/Syntax/PCUICInduction.v b/pcuic/theories/Syntax/PCUICInduction.v index 8e511aff8..c2bb01af3 100644 --- a/pcuic/theories/Syntax/PCUICInduction.v +++ b/pcuic/theories/Syntax/PCUICInduction.v @@ -41,7 +41,7 @@ Lemma term_forall_list_ind : (forall (s : projection) (t : term), P t -> P (tProj s t)) -> (forall (m : mfixpoint term) (n : nat), tFixProp P P m -> P (tFix m n)) -> (forall (m : mfixpoint term) (n : nat), tFixProp P P m -> P (tCoFix m n)) -> - (forall p, P (tPrim p)) -> + (forall p, tPrimProp P p -> P (tPrim p)) -> forall t : term, P t. Proof. intros until t. revert t. @@ -85,6 +85,12 @@ Proof. fix auxm 1. destruct mfix; constructor; [|apply auxm]. split; apply auxt. + + * destruct prim. destruct p; cbn; intuition auto. + destruct a. cbn. + revert array_value. + fix auxm 1; destruct array_value; constructor; [|apply auxm]. + apply auxt. Defined. Lemma size_decompose_app_rec t L : @@ -259,7 +265,7 @@ Lemma term_forall_mkApps_ind : (forall (s : projection) (t : term), P t -> P (tProj s t)) -> (forall (m : mfixpoint term) (n : nat), tFixProp P P m -> P (tFix m n)) -> (forall (m : mfixpoint term) (n : nat), tFixProp P P m -> P (tCoFix m n)) -> - (forall i, P (tPrim i)) -> + (forall p, tPrimProp P p -> P (tPrim p)) -> forall t : term, P t. Proof. intros until t. @@ -343,6 +349,15 @@ Proof. apply auxt. hnf; cbn. unfold def_size. lia. apply auxt'. intros. apply auxt. hnf in *; cbn in *. unfold mfixpoint_size, def_size in *. lia. + + - eapply Pprim. + destruct prim. destruct p; cbn; intuition auto. + * eapply auxt. red. cbn. lia. + * eapply auxt. red. cbn. lia. + * destruct a; cbn in *. unfold MR in auxt; cbn in auxt. clear -auxt array_value. + revert array_value auxt. + fix auxt' 1; destruct array_value; constructor => //. + eapply auxt. cbn. lia. eapply auxt'. intros ? ?; eapply auxt. cbn; lia. Defined. Lemma liftP_ctx (P : term -> Type) : @@ -431,6 +446,8 @@ Proof. apply list_size_map_hom. intros. simpl. destruct x. simpl. unfold def_size. simpl. f_equal; symmetry; apply size_lift. + - destruct prim. destruct p; cbn => //. unfold prim_size. simp map_prim => /=. + rewrite !size_lift. rewrite list_size_map_hom; auto. Qed. Definition on_local_decl (P : context -> term -> Type) @@ -485,7 +502,7 @@ Lemma term_forall_ctx_list_ind : (forall Γ (m : mfixpoint term) (n : nat), All_local_env (on_local_decl (fun Γ' t => P (Γ ,,, Γ') t)) (fix_context m) -> tFixProp (P Γ) (P (Γ ,,, fix_context m)) m -> P Γ (tCoFix m n)) -> - (forall Γ p, P Γ (tPrim p)) -> + (forall Γ p, tPrimProp (P Γ) p -> P Γ (tPrim p)) -> forall Γ (t : term), P Γ t. Proof. intros ????????????????? Γ t. @@ -563,6 +580,13 @@ Proof. - eapply X13; try (apply aux; red; simpl; lia). apply auxΓ => //. simpl. specialize (H mfix). lia. red. apply All_pair. split; apply auxl; simpl; auto. + + - eapply X14. + destruct prim. destruct p; cbn; intuition auto; destruct a; cbn in *. + * eapply aux. cbn. lia. + * eapply aux. cbn. lia. + * eapply (auxl _ _ array_value id). unfold id. + change (fun x => size x) with size. lia. Defined. (** This induction principle gives a general induction hypothesis for applications, @@ -592,7 +616,7 @@ Lemma term_ind_size_app : tFixProp P P m -> P (tFix m n)) -> (forall (m : mfixpoint term) (n : nat), tFixProp (P) P m -> P (tCoFix m n)) -> - (forall p, P (tPrim p)) -> + (forall p, tPrimProp P p -> P (tPrim p)) -> forall (t : term), P t. Proof. intros. @@ -644,4 +668,10 @@ Proof. * eapply X13; try (apply aux; red; simpl; lia). red. apply All_pair. split; apply auxl; simpl; auto. + + * eapply X14. + destruct hh, p; cbn; intuition auto. + 1-2:(eapply aux; cbn; lia). + eapply (auxl _ (array_value a) id). + change (fun x => size (id x)) with size. cbn; lia. Defined. diff --git a/pcuic/theories/Syntax/PCUICLiftSubst.v b/pcuic/theories/Syntax/PCUICLiftSubst.v index 1c0ad92a4..222d76ab5 100644 --- a/pcuic/theories/Syntax/PCUICLiftSubst.v +++ b/pcuic/theories/Syntax/PCUICLiftSubst.v @@ -113,7 +113,7 @@ Proof. intros M. elim M using term_forall_list_ind; simpl in |- *; intros; try easy ; try (try rewrite H; try rewrite H0 ; try rewrite H1 ; easy); - try (f_equal; auto; solve_all). + try solve [(f_equal; auto; solve_all)]. now elim (leb k n). Qed. @@ -134,7 +134,7 @@ Proof. intros; simpl; autorewrite with map; try (rewrite -> H, ?H0, ?H1; auto); try (f_equal; auto; solve_all). - elim (leb_spec k n); intros. + - elim (leb_spec k n); intros. + elim (leb_spec i (n0 + n)); intros; lia. + elim (leb_spec i n); intros; lia. Qed. diff --git a/pcuic/theories/Syntax/PCUICOnFreeVars.v b/pcuic/theories/Syntax/PCUICOnFreeVars.v index d04c01f40..95a4f2d4b 100644 --- a/pcuic/theories/Syntax/PCUICOnFreeVars.v +++ b/pcuic/theories/Syntax/PCUICOnFreeVars.v @@ -88,7 +88,7 @@ Fixpoint on_free_vars (p : nat -> bool) (t : term) : bool := | tFix mfix idx | tCoFix mfix idx => List.forallb (test_def (on_free_vars p) (on_free_vars (shiftnP #|mfix| p))) mfix | tVar _ | tSort _ | tConst _ _ | tInd _ _ | tConstruct _ _ _ => true - | tPrim _ => true + | tPrim pr => test_prim (on_free_vars p) pr end. Lemma on_free_vars_ext (p q : nat -> bool) t : @@ -115,6 +115,7 @@ Proof. eapply b; rewrite H //. - simpl; intuition auto. f_equal; eauto 2. eapply b; rewrite H //. + - solve_all. Qed. #[global] @@ -197,6 +198,7 @@ Proof. - unfold test_def; solve_all. rewrite shiftnP_closedP shiftnP_xpredT. now len in b. + - solve_all. Qed. Lemma closedn_on_free_vars {P n t} : closedn n t -> on_free_vars (shiftnP n P) t. @@ -228,6 +230,7 @@ Proof. - repeat (solve_all; f_equal). - unfold test_def. solve_all. - unfold test_def; solve_all. + - solve_all. Qed. Definition on_free_vars_decl P d := @@ -373,6 +376,7 @@ Proof. len; rewrite shiftnP_strengthenP. f_equal; eauto. - unfold test_def in *. simpl; intros ? []. len; rewrite shiftnP_strengthenP. f_equal; eauto. + - solve_all. Qed. Definition on_free_vars_terms p s := @@ -403,7 +407,8 @@ Proof. revert t p k. induction t using PCUICInduction.term_forall_list_ind; simpl => //; intros; simpl. - all:try (rtoProp; rewrite ?shiftnP_substP; now rewrite ?IHt1 ?IHt2 ?IHt3). + all:try (rtoProp; rewrite ?shiftnP_substP; now rewrite ?IHt1 ?IHt2 ?IHt3); + try solve [solve_all]. - intros. destruct (Nat.leb_spec k n). * destruct nth_error eqn:eq. + unfold on_free_vars_terms in *. toAll. @@ -422,7 +427,6 @@ Proof. now rewrite H0. * simpl. rewrite /substP /strengthenP /=. rewrite H0. now nat_compare_specs. - - solve_all. - rtoProp. destruct X. solve_all. * len. rewrite shiftnP_substP. solve_all. * len. rewrite shiftnP_substP; solve_all. diff --git a/pcuic/theories/Syntax/PCUICReflect.v b/pcuic/theories/Syntax/PCUICReflect.v index a28cedb55..92ae085e0 100644 --- a/pcuic/theories/Syntax/PCUICReflect.v +++ b/pcuic/theories/Syntax/PCUICReflect.v @@ -115,7 +115,7 @@ Fixpoint eqb_term (u v : term) : bool := eqb x.(rarg) y.(rarg) && eqb x.(dname) y.(dname)) mfix mfix' - | tPrim p, tPrim p' => eqb p p' + | tPrim p, tPrim p' => @eqb_prim_val _ eqb_term p p' | _, _ => false end. @@ -194,6 +194,8 @@ Proof. destruct (eqb_spec na na'); t'; destruct (r ty'); t'; destruct (o b'); t'. Qed. +Transparent eqb_prim_model eqb_prim_val. + #[program,global] Instance eqb_term_reflect : ReflectEq term := {| eqb := eqb_term |}. Next Obligation. @@ -243,6 +245,16 @@ Proof. destruct (r0 dbody0); t'. destruct (eqb_spec rarg rarg0); t'. destruct (eqb_spec dname dname0); t'. } + - destruct p, p; cbn in *; destruct prim, p; cbn; nodec. + case: eqb_spec; intros; subst; nodec. constructor; auto. + case: eqb_spec; intros; subst; nodec. constructor; auto. + destruct X as [eqty [eqd eqv]]. unfold eqb_array. + destruct (eqty (array_type a0)); t'. + destruct (eqd (array_default a0)); t'. 2:noconf H1; auto. + 2:{ noconf H0; auto. } + case: eqb_spec; intros; subst; nodec; cbn. + case: (reflect_prop_list eqv) => //. constructor; auto. f_equal. f_equal. + destruct a, a0; cbn in *; congruence. nodec. Qed. #[global] diff --git a/pcuic/theories/Typing/PCUICClosedTyp.v b/pcuic/theories/Typing/PCUICClosedTyp.v index ca5ee2893..e158ac0d7 100644 --- a/pcuic/theories/Typing/PCUICClosedTyp.v +++ b/pcuic/theories/Typing/PCUICClosedTyp.v @@ -295,6 +295,11 @@ Proof. destruct X0 as [s [Hs cl]]. now rewrite andb_true_r in cl. Unshelve. all:eauto. + + - destruct p as [[] pv]; cbn in X0 |- *; simp prim_type => //. + depelim pv. simp prim_type. cbn. depelim X2. + move/andP: hdef => [] -> ->; rewrite !andb_true_r. split => //. + solve_all. Qed. Lemma declared_minductive_closed {cf:checker_flags} {Σ : global_env} {wfΣ : wf Σ} {mdecl mind} : @@ -806,6 +811,8 @@ Proof. rewrite -(fix_context_length mfix0). rewrite on_ctx_free_vars_extend // hctx. now apply on_free_vars_fix_context. + - cbn. toAll. solve_all. + eapply OnOne2_impl_All_r; eauto; solve_all. Qed. Lemma red_on_free_vars {cf} {P : nat -> bool} {Σ Γ u v} {wfΣ : wf Σ} : @@ -839,7 +846,7 @@ Lemma term_closedn_list_ind : (forall k (s : projection) (t : term), P k t -> P k (tProj s t)) -> (forall k (m : mfixpoint term) (n : nat), tFixProp (P k) (P (#|fix_context m| + k)) m -> P k (tFix m n)) -> (forall k (m : mfixpoint term) (n : nat), tFixProp (P k) (P (#|fix_context m| + k)) m -> P k (tCoFix m n)) -> - (forall k p, P k (tPrim p)) -> + (forall k p, tPrimProp (P k) p -> P k (tPrim p)) -> forall k (t : term), closedn k t -> P k t. Proof. intros until t. revert k t. @@ -928,6 +935,12 @@ Proof. simpl in clt. move/andP: clt => [clt cll]. simpl in clt. move/andP: clt. intuition auto. move/andP: clt => [cd cmfix]. apply auxm; auto. + + - destruct prim as [? []]; cbn in clt |- *; rtoProp; intuition eauto. + move: H. clear -auxt. generalize (array_value a). + fix auxl 1. destruct l; intro. + * constructor. + * cbn in H; move/andP: H => []; eauto. Defined. Lemma term_noccur_between_list_ind : @@ -951,7 +964,7 @@ Lemma term_noccur_between_list_ind : (forall k n (s : projection) (t : term), P k n t -> P k n (tProj s t)) -> (forall k n (m : mfixpoint term) (i : nat), tFixProp (P k n) (P (#|fix_context m| + k) n) m -> P k n (tFix m i)) -> (forall k n (m : mfixpoint term) (i : nat), tFixProp (P k n) (P (#|fix_context m| + k) n) m -> P k n (tCoFix m i)) -> - (forall k n p, P k n (tPrim p)) -> + (forall k n p, tPrimProp (P k n) p -> P k n (tPrim p)) -> forall k n (t : term), noccur_between k n t -> P k n t. Proof. intros until t. revert k n t. @@ -1034,4 +1047,11 @@ Proof. simpl in clt. move/andP: clt => [clt cll]. simpl in clt. move/andP: clt. intuition auto. move/andP: clt => [cd cmfix]. apply auxm; auto. + + + - destruct prim as [? []]; cbn in clt |- *; rtoProp; intuition eauto. + move: H. clear -auxt. generalize (array_value a). + fix auxl 1. destruct l; intro. + * constructor. + * cbn in H; move/andP: H => []; eauto. Defined. diff --git a/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v b/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v index 9cfe38061..f61e294b4 100644 --- a/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v +++ b/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v @@ -119,6 +119,13 @@ Proof. - eapply cumul_CoFix. apply All2_map. repeat toAll. eapply All2_impl. 1: tea. cbn; intros; intuition. rewrite -> subst_instance_app, fix_context_subst_instance in *; eauto. + - eapply cumul_Prim. depelim e0. + * depelim i; depelim X. cbn in H. cbn. constructor. + * depelim f; depelim X. cbn in H. constructor. + * depelim X. cbn in H. noconf H; cbn in H. constructor; cbn -[Universe.make]; eauto. + + rewrite -!subst_instance_univ_make. + eapply eq_universe_subst_instance; tea. + + solve_all. - repeat rewrite subst_instance_mkApps. eapply cumul_Ind. * apply precompose_subst_instance_global. rewrite map_length. eapply R_global_instance_impl_same_napp; try eapply H; eauto. @@ -190,6 +197,18 @@ Proof using Type. now eapply cumul_decls_subst_instance. Qed. +Lemma subst_instance_prim_type p prim_ty u : (prim_type p prim_ty)@[u] = prim_type (mapu_prim (subst_instance_level u) (subst_instance u) p) prim_ty. +Proof. + destruct p as [? []]; simp prim_type => //=. +Qed. + +Lemma subst_instance_prim_val_tag (p : PCUICPrimitive.prim_val term) u : + prim_val_tag (mapu_prim (subst_instance_level u) (subst_instance u) p) = + prim_val_tag p. +Proof. + destruct p as [? []]; simp prim_type => //=. +Qed. + Hint Resolve subst_instance_cstrs_two satisfies_equal_sets satisfies_subsets : univ_subst. Hint Resolve monomorphic_global_constraint monomorphic_global_constraint_ext : univ_subst. @@ -374,7 +393,18 @@ Proof using Type. rewrite map_map_compose. now rewrite subst_instance_check_one_cofix. - - econstructor; eauto. + - intros. + rewrite subst_instance_prim_type. + econstructor. + + eauto. + + now rewrite subst_instance_prim_val_tag. + + exact H0. + + now rewrite subst_instance_prim_val_tag. + + destruct p as [? []]; depelim X2; constructor; eauto. + * rewrite -subst_instance_univ_make. eapply wf_universe_subst_instance => //. + * cbn -[Universe.make] in hty. + specialize (hty u univs). rewrite subst_instance_univ_make in hty. now eapply hty. + * cbn. solve_all. - intros t0 A B X X0 X1 X2 X3 X4 cum u univs wfΣ' H. econstructor. diff --git a/pcuic/theories/Typing/PCUICWeakeningConfigTyp.v b/pcuic/theories/Typing/PCUICWeakeningConfigTyp.v index 714fc955e..68d149d01 100644 --- a/pcuic/theories/Typing/PCUICWeakeningConfigTyp.v +++ b/pcuic/theories/Typing/PCUICWeakeningConfigTyp.v @@ -20,11 +20,11 @@ Local Ltac constructor_per_goal _ := unshelve constructor_per_goal' (); shelve_unifiable. Lemma weakening_config_wf_local_sized {cf1 cf2 : checker_flags} Σ Γ - (Hwf : match cf1 with _ => wf_local Σ Γ end) + (Hwf : All_local_env (lift_typing (@typing cf1) Σ) Γ) (IH : forall Γ0 t0 T0 (H0 : @typing cf1 Σ Γ0 t0 T0), typing_size H0 < S (All_local_env_size (fun _ _ _ _ H => typing_size H) Σ Γ Hwf) -> @typing cf2 Σ Γ0 t0 T0) - : match cf2 with _ => wf_local Σ Γ end. + : wf_local Σ Γ. Proof. simpl in *. induction Hwf; [ constructor 1 | constructor 2 | constructor 3 ]. @@ -64,6 +64,7 @@ Proof. intros (Σ & Γ & t & T & H). simpl. intros IH. specialize (fun Σ Γ t T H => IH (Σ; Γ; t; T; H)). simpl in IH. destruct H; constructor_per_goal (); try eassumption. + 5:destruct p1; constructor; eauto; solve_all. all: try (unshelve eapply IH; [ eassumption .. | ]; try solve [ constructor; simpl; lia ]). all: try (eapply (@weakening_config_cumulSpec cf1 cf2); eassumption). @@ -112,6 +113,7 @@ Proof. end. all: repeat match goal with H : Forall2 _ (_ :: _) (_ :: _) |- _ => depelim H end. all: try assumption. + 2:{ eapply (All_map_size (fun x H => typing_size H) _ hvalue). intros. eapply (IH _ _ _ p). lia. } all: [ > unshelve eapply (@weakening_config_wf_local_sized cf1 cf2); [ eassumption | ] .. ]. all: intros; unshelve eapply IH; [ eassumption | ]. all: simpl in *; try lia. diff --git a/pcuic/theories/Typing/PCUICWeakeningEnvTyp.v b/pcuic/theories/Typing/PCUICWeakeningEnvTyp.v index 51f6dabb2..4cccc11ea 100644 --- a/pcuic/theories/Typing/PCUICWeakeningEnvTyp.v +++ b/pcuic/theories/Typing/PCUICWeakeningEnvTyp.v @@ -86,11 +86,11 @@ Lemma extends_primitive_constant Σ Σ' p t : Proof. intros [_ _ ext]. unfold primitive_constant. - case: ext. - destruct p; case => //. - - move=> _. case => //. - - move=> _; case => //. - - case => //. + case: ext => o [] o' o''. + destruct p => //. + - intros hr. now rewrite hr in o; depelim o. + - intros hr. now rewrite hr in o'; depelim o'. + - intros hr. now rewrite hr in o''; depelim o''. Qed. Local Hint Resolve extends_primitive_constant : extends. @@ -129,6 +129,10 @@ Proof. apply (lift_typing_impl X); now intros ? []. + eapply (All_impl X1); intros d X. apply (lift_typing_impl X); now intros ? []. + - econstructor; eauto with extends. + destruct X2; constructor; eauto with extends. + * now eapply extends_wf_universe. + * solve_all. - econstructor. 1: eauto. + eapply forall_Σ'1; eauto. + destruct Σ as [Σ φ]. eapply weakening_env_cumulSpec in cumulA; eauto. diff --git a/pcuic/theories/utils/PCUICAstUtils.v b/pcuic/theories/utils/PCUICAstUtils.v index c83af52f8..177805b18 100644 --- a/pcuic/theories/utils/PCUICAstUtils.v +++ b/pcuic/theories/utils/PCUICAstUtils.v @@ -50,11 +50,36 @@ Ltac change_Sk := | |- context [#|?l| + (?x + ?y)] => progress replace (#|l| + (x + y)) with ((#|l| + x) + y) by now rewrite Nat.add_assoc end. +#[global] Hint Rewrite mapu_prim_equation_1 mapu_prim_equation_2 mapu_prim_equation_3 : map. + +Ltac to_prop := + match goal with + | H : is_true (test_prim ?P ?p) |- _ => apply test_prim_primProp in H + | |- is_true (test_prim _ _) => apply primProp_test_prim + | H : tPrimProp ?P (map_prim ?f ?p) |- _ => apply primProp_map_inv in H + | |- tPrimProp _ (map_prim _ _) => apply primProp_map + end. + +Ltac merge_prop := + match goal with + | H : tPrimProp ?P ?p, H' : tPrimProp ?Q ?p |- _ => + pose proof (tPrimProp_prod H H'); clear H H' + end. + +Ltac tPrimProp_impl := + match goal with + | H : tPrimProp ?P ?p |- tPrimProp ?Q ?p => apply (tPrimProp_impl H); clear H + end. + +#[global] Hint Extern 4 (tPrimProp _ _) => tPrimProp_impl : all. + Ltac solve_all_one := try lazymatch goal with | H: tCasePredProp _ _ _ |- _ => destruct H as [? [? ?]] + (* | H: tPrimProp _ ?p |- _ => destruct p as [? []]; cbn in H *) end; unfold tCaseBrsProp, tFixProp in *; + repeat to_prop; repeat merge_prop; autorewrite with map; rtoProp; try ( @@ -75,12 +100,20 @@ Ltac solve_all_one := (eapply onctx_test; [eassumption|eassumption|]) || (eapply test_context_k_eqP_id_spec; [eassumption|eassumption|]) || (eapply test_context_k_eqP_eq_spec; [eassumption|]) || - (eapply map_context_eq_spec; [eassumption|])); + (eapply map_context_eq_spec; [eassumption|]) || + (eapply test_prim_eq_spec; [eassumption|]) || + (eapply primProp_map_eq; [eassumption|idtac..]; cbn) || + (eapply primProp_map_id; [eassumption|]) || + (eapply primProp_mapu_id; [eassumption|eassumption| | ]) || + (eapply test_primu_test_primu_tPrimProp; [eassumption|eassumption| | ])); repeat toAll; try All_map; try close_Forall; + try tPrimProp_impl; change_Sk; auto with all; intuition eauto 4 with all. Ltac solve_all := repeat (progress solve_all_one). + +#[global] Hint Extern 4 (_ =1 _) => intro : all. #[global] Hint Extern 10 => rewrite !map_branch_map_branch : all. #[global] Hint Extern 10 => rewrite !map_predicate_map_predicate : all. diff --git a/pcuic/theories/utils/PCUICPretty.v b/pcuic/theories/utils/PCUICPretty.v index f8b44467b..8a4ae2a42 100644 --- a/pcuic/theories/utils/PCUICPretty.v +++ b/pcuic/theories/utils/PCUICPretty.v @@ -117,11 +117,11 @@ Module PrintTermTree. Import bytestring.Tree. Infix "^" := append. - Definition print_prim {term} (soft : term -> Tree.t) (p : prim_val) : Tree.t := + Definition print_prim (soft : term -> Tree.t) (p : prim_val) : Tree.t := match p.π2 return Tree.t with | primIntModel f => "(int: " ^ Primitive.string_of_prim_int f ^ ")" | primFloatModel f => "(float: " ^ Primitive.string_of_float f ^ ")" - (* | primArrayModel a => "(array:" ^ ")" *) + | primArrayModel a => "(array:" ^ string_of_list soft a.(array_value) ^ ")" end. Section Aux. diff --git a/pcuic/theories/utils/PCUICPrimitive.v b/pcuic/theories/utils/PCUICPrimitive.v index c44da46a3..b3502ce76 100644 --- a/pcuic/theories/utils/PCUICPrimitive.v +++ b/pcuic/theories/utils/PCUICPrimitive.v @@ -6,27 +6,30 @@ From Equations Require Import Equations. From Coq Require Import ssreflect. From Coq Require Import Uint63 SpecFloat. -(** We don't enforce the type of the array here*) -Record array_model (term : Type) := - { array_instance : Instance.t; +Record array_model {term : Type} := + { array_level : Level.t; array_type : term; array_default : term; array_value : list term }. Derive NoConfusion for array_model. + +Arguments array_model : clear implicits. #[global] Instance array_model_eqdec {term} (e : EqDec term) : EqDec (array_model term). Proof. eqdec_proof. Qed. +(* We use this inductive definition instead of the more direct case analysis + [prim_model_of] so that [prim_model] can be used in the inductive definition + of terms, otherwise it results in a non-strictly positive definition. + *) Inductive prim_model (term : Type) : prim_tag -> Type := | primIntModel (i : PrimInt63.int) : prim_model term primInt -| primFloatModel (f : PrimFloat.float) : prim_model term primFloat. +| primFloatModel (f : PrimFloat.float) : prim_model term primFloat +| primArrayModel (a : array_model term) : prim_model term primArray. -(* | primIntModel (i : Int63.t) : prim_model term primInt *) -(* | primFloatModel (f : float64_model) : prim_model term primFloat. *) -(* | primArrayModel (a : array_model term) : prim_model term primArray. *) Arguments primIntModel {term}. Arguments primFloatModel {term}. -(* Arguments primArrayModel {term}. *) +Arguments primArrayModel {term}. Derive Signature NoConfusion for prim_model. @@ -34,7 +37,7 @@ Definition prim_model_of (term : Type) (p : prim_tag) : Type := match p with | primInt => PrimInt63.int | primFloat => PrimFloat.float - (* | primArray => array_model term *) + | primArray => array_model term end. Definition prim_val term := ∑ t : prim_tag, prim_model term t. @@ -45,7 +48,7 @@ Definition prim_model_val {term} (p : prim_val term) : prim_model_of term (prim_ match prim_val_model p in prim_model _ t return prim_model_of term t with | primIntModel i => i | primFloatModel f => f - (* | primArrayModel a => a *) + | primArrayModel a => a end. Lemma exist_irrel_eq {A} (P : A -> bool) (x y : sig P) : proj1_sig x = proj1_sig y -> x = y. @@ -73,40 +76,73 @@ Qed. #[global] Instance reflect_eq_spec_float : ReflectEq SpecFloat.spec_float := EqDec_ReflectEq _. -Equations eqb_prim_model {term} {t : prim_tag} (x y : prim_model term t) : bool := +Import ReflectEq. + +Definition eqb_array {term} {eqt : term -> term -> bool} (x y : array_model term) : bool := + eqb x.(array_level) y.(array_level) && + forallb2 eqt x.(array_value) y.(array_value) && + eqt x.(array_default) y.(array_default) && + eqt x.(array_type) y.(array_type). + +#[program,global] +Instance reflect_eq_array {term} {req : ReflectEq term}: ReflectEq (array_model term) := + { eqb := eqb_array (eqt := eqb) }. +Next Obligation. + intros term req [] []; cbn. unfold eqb_array. cbn. + change (forallb2 eqb) with (eqb (A := list term)). + case: eqb_spec => //=. intros ->. + case: eqb_spec => //=. intros ->. + case: eqb_spec => //=. intros ->. + case: eqb_spec => //=. intros ->. constructor; auto. + all:constructor; congruence. +Qed. + +Equations eqb_prim_model {term} {req : term -> term -> bool} {t : prim_tag} (x y : prim_model term t) : bool := | primIntModel x, primIntModel y := ReflectEq.eqb x y - | primFloatModel x, primFloatModel y := ReflectEq.eqb x y. + | primFloatModel x, primFloatModel y := ReflectEq.eqb x y + | primArrayModel x, primArrayModel y := eqb_array (eqt:=req) x y. #[global, program] -Instance prim_model_reflecteq {term} {p : prim_tag} : ReflectEq (prim_model term p) := - {| ReflectEq.eqb := eqb_prim_model |}. +Instance prim_model_reflecteq {term} {req : ReflectEq term} {p : prim_tag} : ReflectEq (prim_model term p) := + {| ReflectEq.eqb := eqb_prim_model (req := eqb) |}. Next Obligation. intros. depelim x; depelim y; simp eqb_prim_model. case: ReflectEq.eqb_spec; constructor; subst; auto. congruence. case: ReflectEq.eqb_spec; constructor; subst; auto. congruence. + change (eqb_array a a0) with (eqb a a0). + case: ReflectEq.eqb_spec; constructor; subst; auto. congruence. Qed. #[global] -Instance prim_model_eqdec {term} : forall p : prim_tag, EqDec (prim_model term p) := _. +Instance prim_model_eqdec {term} {req : ReflectEq term} : forall p : prim_tag, EqDec (prim_model term p) := _. -Equations eqb_prim_val {term} (x y : prim_val term) : bool := - | (primInt; i), (primInt; i') := ReflectEq.eqb i i' - | (primFloat; f), (primFloat; f') := ReflectEq.eqb f f' +Equations eqb_prim_val {term} {req : term -> term -> bool} (x y : prim_val term) : bool := + | (primInt; i), (primInt; i') := eqb_prim_model (req := req) i i' + | (primFloat; f), (primFloat; f') := eqb_prim_model (req := req) f f' + | (primArray; a), (primArray; a') := eqb_prim_model (req := req) a a' | x, y := false. #[global, program] -Instance prim_val_reflect_eq {term} : ReflectEq (prim_val term) := - {| ReflectEq.eqb := eqb_prim_val |}. +Instance prim_val_reflect_eq {term} {req : ReflectEq term} : ReflectEq (prim_val term) := + {| ReflectEq.eqb := eqb_prim_val (req := eqb) |}. Next Obligation. intros. funelim (eqb_prim_val x y); simp eqb_prim_val. + change (eqb_prim_model i i') with (eqb i i'). case: ReflectEq.eqb_spec; constructor; subst; auto. intros H; noconf H. cbn in n. auto. constructor. intros H; noconf H; auto. constructor. intros H; noconf H; auto. + constructor. intros H; noconf H; auto. + change (eqb_prim_model f f') with (eqb f f'). + case: ReflectEq.eqb_spec; constructor; subst; auto. intros H; noconf H; auto. + constructor. intros H; noconf H; auto. + constructor. intros H; noconf H; auto. + constructor. intros H; noconf H; auto. + change (eqb_prim_model a a') with (eqb a a'). case: ReflectEq.eqb_spec; constructor; subst; auto. intros H; noconf H; auto. Qed. #[global] -Instance prim_tag_model_eqdec term : EqDec (prim_val term) := _. +Instance prim_tag_model_eqdec term {req : ReflectEq term} : EqDec (prim_val term) := _. (** Printing *) @@ -114,5 +150,5 @@ Definition string_of_prim {term} (soft : term -> string) (p : prim_val term) : s match p.π2 return string with | primIntModel f => "(int: " ^ string_of_prim_int f ^ ")" | primFloatModel f => "(float: " ^ string_of_float f ^ ")" - (* | primArrayModel a => "(array:" ^ ")" *) + | primArrayModel a => "(array:" ^ string_of_list soft a.(array_value) ^ ")" end. diff --git a/pcuic/theories/utils/PCUICSize.v b/pcuic/theories/utils/PCUICSize.v index 9ffc21743..3822c0b0d 100644 --- a/pcuic/theories/utils/PCUICSize.v +++ b/pcuic/theories/utils/PCUICSize.v @@ -23,6 +23,13 @@ Definition predicate_size (size : term -> nat) (p : PCUICAst.predicate term) := context_size size p.(pcontext) + size p.(preturn). +Definition prim_size (size : term -> nat) (p : prim_val) : nat := + match p.π2 return nat with + | primIntModel f => 0 + | primFloatModel f => 0 + | primArrayModel a => size a.(array_type) + size a.(array_default) + list_size size a.(array_value) + end. + Fixpoint size t : nat := match t with | tRel i => 1 @@ -36,6 +43,7 @@ Fixpoint size t : nat := | tProj p c => S (size c) | tFix mfix idx => S (mfixpoint_size size mfix) | tCoFix mfix idx => S (mfixpoint_size size mfix) + | tPrim p => S (prim_size size p) | _ => 1 end. diff --git a/template-coq/src/ast_denoter.ml b/template-coq/src/ast_denoter.ml index 25a4ac612..a102c550a 100644 --- a/template-coq/src/ast_denoter.ml +++ b/template-coq/src/ast_denoter.ml @@ -28,6 +28,7 @@ struct type quoted_constraint_type = Universes0.ConstraintType.t type quoted_univ_constraint = Universes0.UnivConstraint.t type quoted_univ_constraints = Universes0.ConstraintSet.t + type quoted_univ_level = Universes0.Level.t type quoted_univ_instance = Universes0.Instance.t type quoted_univ_context = Universes0.UContext.t type quoted_univ_contextset = Universes0.ContextSet.t @@ -105,7 +106,7 @@ struct aci_relevance = x.ci_relevance } let inspect_term (tt: t):(t, quoted_int, quoted_ident, quoted_aname, quoted_sort, quoted_cast_kind, - quoted_kernel_name, quoted_inductive, quoted_relevance, quoted_univ_instance, quoted_proj, + quoted_kernel_name, quoted_inductive, quoted_relevance, quoted_univ_level, quoted_univ_instance, quoted_proj, quoted_int63, quoted_float64) structure_of_term = match tt with | Coq_tRel n -> ACoq_tRel n @@ -127,6 +128,7 @@ struct | Coq_tCoFix (a,b) -> ACoq_tCoFix (List.map unquote_def a,b) | Coq_tInt i -> ACoq_tInt i | Coq_tFloat f -> ACoq_tFloat f + | Coq_tArray (u, arr, def, ty) -> ACoq_tArray (u, Array.of_list arr, def, ty) let unquote_string = Caml_bytestring.caml_string_of_bytestring @@ -223,6 +225,7 @@ struct let u = List.fold_left Univ.Universe.sup (List.hd l) (List.tl l) in evd, Sorts.sort_of_univ u + let unquote_universe_level evm l = evm, unquote_level l let unquote_universe_instance(evm: Evd.evar_map) (l: quoted_univ_instance): Evd.evar_map * Univ.Instance.t = (evm, Univ.Instance.of_array (Array.of_list (List0.map unquote_level l))) diff --git a/template-coq/src/ast_quoter.ml b/template-coq/src/ast_quoter.ml index 6c209d0d5..14462a188 100644 --- a/template-coq/src/ast_quoter.ml +++ b/template-coq/src/ast_quoter.ml @@ -29,6 +29,7 @@ struct type quoted_constraint_type = Universes0.ConstraintType.t type quoted_univ_constraint = Universes0.UnivConstraint.t type quoted_univ_constraints = Universes0.ConstraintSet.t + type quoted_univ_level = Universes0.Level.t type quoted_univ_instance = Universes0.Instance.t type quoted_univ_context = Universes0.UContext.t type quoted_univ_contextset = Universes0.ContextSet.t @@ -156,6 +157,7 @@ struct try ((quote_nonprop_level l, quote_constraint_type ct), quote_nonprop_level l') with e -> assert false + let quote_univ_level = quote_nonprop_level let quote_univ_instance (i : Univ.Instance.t) : quoted_univ_instance = let arr = Univ.Instance.to_array i in (* we assume that valid instances do not contain [Prop] or [SProp] *) @@ -235,6 +237,7 @@ struct let mkLetIn na b t t' = Coq_tLetIn (na,b,t,t') let mkInt i = Coq_tInt i let mkFloat f = Coq_tFloat f + let mkArray u arr ~default ~ty = Coq_tArray (u, Array.to_list arr, default, ty) let rec seq f t = if f < t then @@ -313,11 +316,13 @@ struct type pre_quoted_retroknowledge = { retro_int63 : quoted_kernel_name option; - retro_float64 : quoted_kernel_name option } + retro_float64 : quoted_kernel_name option; + retro_array : quoted_kernel_name option } let quote_retroknowledge r = { Environment.Retroknowledge.retro_int63 = r.retro_int63; - Environment.Retroknowledge.retro_float64 = r.retro_float64 } + Environment.Retroknowledge.retro_float64 = r.retro_float64; + Environment.Retroknowledge.retro_array = r.retro_array } let mk_global_env universes declarations retroknowledge = { universes; declarations; retroknowledge } let mk_program decls tm = (decls, tm) @@ -366,7 +371,8 @@ struct | None -> None *) let inspectTerm (t : term) : (term, quoted_int, quoted_ident, quoted_name, quoted_sort, quoted_cast_kind, - quoted_kernel_name, quoted_inductive, quoted_relevance, quoted_univ_instance, quoted_proj, + quoted_kernel_name, quoted_inductive, quoted_relevance, quoted_univ_level, + quoted_univ_instance, quoted_proj, quoted_int63, quoted_float64) structure_of_term = match t with | Coq_tRel n -> ACoq_tRel n diff --git a/template-coq/src/constr_denoter.ml b/template-coq/src/constr_denoter.ml index 20268ed66..675be7ec1 100644 --- a/template-coq/src/constr_denoter.ml +++ b/template-coq/src/constr_denoter.ml @@ -70,12 +70,12 @@ struct else not_supported_verb trm "from_bool" let unquote_int63 trm = - match Constr.kind trm with + match Constr.kind trm with | Constr.Int i -> i | _ -> not_supported_verb trm "unquote_int63" - + let unquote_float64 trm = - match Constr.kind trm with + match Constr.kind trm with | Constr.Float f -> f | _ -> not_supported_verb trm "unquote_float64" @@ -127,12 +127,12 @@ struct else not_supported_verb trm "unquote_name" - let unquote_evar env evm id args = + let unquote_evar env evm id args = if constr_equall id tfresh_evar_id then let evm, (tyev, s) = Evarutil.new_type_evar env evm Evd.univ_flexible_alg in let evm, ev = Evarutil.new_evar env evm tyev in evm, EConstr.Unsafe.to_constr ev - else + else let id = unquote_nat id in let ev = Evar.unsafe_of_int id in evm, Constr.mkEvar (ev, SList.of_full_list args) @@ -185,7 +185,7 @@ struct | s :: [] -> debug (fun () -> str "Unquoting level " ++ Printer.pr_constr_env (Global.env ()) evm trm); let s = (unquote_string s) in s = "__metacoq_fresh_level__" - | _ -> bad_term_verb trm "unquote_level" + | _ -> bad_term_verb trm "unquote_level" else false let unquote_level evm trm (* of type level *) : Evd.evar_map * Univ.Level.t = @@ -239,7 +239,7 @@ struct let (h,args) = app_full x [] in if constr_equall h tBuild_Universe then (match args with - | x :: _ :: [] -> + | x :: _ :: [] -> (let (h,args) = app_full x [] in if constr_equall h tMktLevelExprSet then match args with @@ -261,6 +261,8 @@ struct else bad_term_verb trm "unquote_universe 4" + let unquote_universe_level evm trm = unquote_level evm trm + let unquote_universe_instance evm trm (* of type universe_instance *) = let l = unquote_list trm in let evm, l = map_evm unquote_level evm l in @@ -342,7 +344,7 @@ struct else not_supported_verb trm "unquote_global_reference" - let unquote_branch trm = + let unquote_branch trm = let (h, args) = app_full trm [] in if constr_equall h tmk_branch then match args with @@ -351,20 +353,20 @@ struct | _ -> bad_term_verb trm "unquote_branch" else not_supported_verb trm "unquote_branch" - let unquote_predicate trm = + let unquote_predicate trm = let (h, args) = app_full trm [] in if constr_equall h tmk_predicate then match args with - | _ty :: auinst :: apars :: apcontext :: apreturn :: [] -> + | _ty :: auinst :: apars :: apcontext :: apreturn :: [] -> let apars = unquote_list apars in let apcontext = unquote_list apcontext in { auinst; apars; apcontext; apreturn } | _ -> bad_term_verb trm "unquote_predicate" else not_supported_verb trm "unquote_predicate" - + let inspect_term (t:Constr.t) - : (Constr.t, quoted_int, quoted_ident, quoted_name, quoted_sort, quoted_cast_kind, quoted_kernel_name, - quoted_inductive, quoted_relevance, quoted_univ_instance, quoted_proj, + : (Constr.t, quoted_int, quoted_ident, quoted_name, quoted_sort, quoted_cast_kind, quoted_kernel_name, + quoted_inductive, quoted_relevance, quoted_univ_level, quoted_univ_instance, quoted_proj, quoted_int63, quoted_float64) structure_of_term = (* debug (fun () -> Pp.(str "denote_term" ++ spc () ++ print_term t)) ; *) let (h,args) = app_full t [] in @@ -459,14 +461,18 @@ struct match args with proj::t::_ -> ACoq_tProj (proj, t) | _ -> CErrors.user_err (print_term t ++ Pp.str ("has bad structure")) - (* else if constr_equall h tInt then + else if constr_equall h tInt then match args with t::_ -> ACoq_tInt t | _ -> CErrors.user_err (print_term t ++ Pp.str ("has bad structure")) else if constr_equall h tFloat then match args with t::_ -> ACoq_tFloat t - | _ -> CErrors.user_err (print_term t ++ Pp.str ("has bad structure")) *) + | _ -> CErrors.user_err (print_term t ++ Pp.str ("has bad structure")) + else if constr_equall h tArray then + match args with + u::v::def::ty::_ -> ACoq_tArray (u, Array.of_list (unquote_list v), def, ty) + | _ -> CErrors.user_err (print_term t ++ Pp.str ("has bad structure")) else CErrors.user_err (str"inspect_term: cannot recognize " ++ print_term t ++ str" (maybe you forgot to reduce it?)") diff --git a/template-coq/src/constr_quoter.ml b/template-coq/src/constr_quoter.ml index 8d2240daa..b007b6655 100644 --- a/template-coq/src/constr_quoter.ml +++ b/template-coq/src/constr_quoter.ml @@ -96,6 +96,9 @@ struct let mkInt i = i let mkFloat f = f + let mkArray u arr ~default ~ty = + constr_mkApp (tArray, [| u; to_coq_listl tTerm (Array.to_list arr); default; ty |]) + let quote_option ty = function | Some tm -> constr_mkApp (cSome, [|ty; tm|]) | None -> constr_mkApp (cNone, [|ty|]) @@ -195,7 +198,7 @@ struct if Level.is_set l then Lazy.force lzero else match Level.var_index l with | Some x -> constr_mkApp (tLevelVar, [| quote_int x |]) - | None -> constr_mkApp (tLevel, [| string_of_level l|]) + | None -> constr_mkApp (tLevel, [| string_of_level l |]) let quote_level l = Tm_util.debug (fun () -> str"quote_level " ++ Level.pr l); @@ -226,6 +229,7 @@ struct let ct = quote_constraint_type ct in constr_mkApp (tmake_univ_constraint, [| l1; ct; l2 |]) + let quote_univ_level u = quote_nonprop_level u (* todo : can be deduced from quote_level, hence shoud be in the Reify module *) let quote_univ_instance u = let arr = Univ.Instance.to_array u in @@ -410,12 +414,14 @@ struct type pre_quoted_retroknowledge = { retro_int63 : quoted_kernel_name option; - retro_float64 : quoted_kernel_name option } + retro_float64 : quoted_kernel_name option; + retro_array : quoted_kernel_name option } let quote_retroknowledge r = let rint63 = to_coq_option (Lazy.force tkername) (fun x -> x) r.retro_int63 in let rfloat64 = to_coq_option (Lazy.force tkername) (fun x -> x) r.retro_float64 in - constr_mkApp (tmk_retroknowledge, [| rint63; rfloat64 |]) + let rarray = to_coq_option (Lazy.force tkername) (fun x -> x) r.retro_array in + constr_mkApp (tmk_retroknowledge, [| rint63; rfloat64; rarray |]) let mk_global_env univs decls retro = constr_mkApp (tBuild_global_env, [| univs; decls; retro |]) diff --git a/template-coq/src/constr_reification.ml b/template-coq/src/constr_reification.ml index ead7b5de5..dc323027f 100644 --- a/template-coq/src/constr_reification.ml +++ b/template-coq/src/constr_reification.ml @@ -23,6 +23,7 @@ struct type quoted_constraint_type = Constr.t (* of type Universes.constraint_type *) type quoted_univ_constraint = Constr.t (* of type Universes.univ_constraint *) type quoted_univ_constraints = Constr.t (* of type Universes.constraints *) + type quoted_univ_level = Constr.t (* of type Universes.Level.t *) type quoted_univ_instance = Constr.t (* of type Universes.universe_instance *) type quoted_univ_context = Constr.t (* of type Universes.UContext.t *) type quoted_univ_contextset = Constr.t (* of type Universes.ContextSet.t *) @@ -149,18 +150,19 @@ struct let tmk_branch = ast "mk_branch" let tmkdecl = ast "mkdecl" let (tTerm,tRel,tVar,tEvar,tSort,tCast,tProd, - tLambda,tLetIn,tApp,tCase,tFix,tConstructor,tConst,tInd,tCoFix,tProj,tInt,tFloat) = + tLambda,tLetIn,tApp,tCase,tFix,tConstructor,tConst,tInd,tCoFix,tProj,tInt,tFloat,tArray) = (ast "term", ast "tRel", ast "tVar", ast "tEvar", ast "tSort", ast "tCast", ast "tProd", ast "tLambda", ast "tLetIn", ast "tApp", ast "tCase", ast "tFix", - ast "tConstruct", ast "tConst", ast "tInd", ast "tCoFix", ast "tProj", ast "tInt", ast "tFloat") + ast "tConstruct", ast "tConst", ast "tInd", ast "tCoFix", ast "tProj", ast "tInt", ast "tFloat", + ast "tArray") let tkername = ast "kername" let tmodpath = ast "modpath" let tMPfile = ast "MPfile" let tMPbound = ast "MPbound" let tMPdot = ast "MPdot" let tfresh_evar_id = ast "fresh_evar_id" - + let tproplevel = ast "level.prop_level_type" let tlevelSProp = ast "level.lsprop" let tlevelProp = ast "level.lprop" @@ -211,7 +213,7 @@ struct let (cFinite,cCoFinite,cBiFinite) = (ast "Finite", ast "CoFinite", ast "BiFinite") let tconstructor_body = ast "constructor_body" let tBuild_constructor_body = ast "Build_constructor_body" - let tprojection_body = ast "projection_body" + let tprojection_body = ast "projection_body" let tBuild_projection_body = ast "Build_projection_body" let tone_inductive_body = ast "one_inductive_body" let tBuild_one_inductive_body = ast "Build_one_inductive_body" diff --git a/template-coq/src/denoter.ml b/template-coq/src/denoter.ml index b6bb08ccc..880033ac2 100644 --- a/template-coq/src/denoter.ml +++ b/template-coq/src/denoter.ml @@ -23,10 +23,11 @@ sig (*val unquote_univ_instance : quoted_univ_instance -> Univ.Instance.t *) val unquote_proj : quoted_proj -> (quoted_inductive * quoted_int * quoted_int) val unquote_universe : Evd.evar_map -> quoted_sort -> Evd.evar_map * Sorts.t + val unquote_universe_level : Evd.evar_map -> quoted_univ_level -> Evd.evar_map * Univ.Level.t val unquote_universe_instance: Evd.evar_map -> quoted_univ_instance -> Evd.evar_map * Univ.Instance.t (* val representsIndConstuctor : quoted_inductive -> Term.constr -> bool *) - val inspect_term : t -> (t, quoted_int, quoted_ident, quoted_aname, quoted_sort, quoted_cast_kind, - quoted_kernel_name, quoted_inductive, quoted_relevance, quoted_univ_instance, quoted_proj, + val inspect_term : t -> (t, quoted_int, quoted_ident, quoted_aname, quoted_sort, quoted_cast_kind, + quoted_kernel_name, quoted_inductive, quoted_relevance, quoted_univ_level, quoted_univ_instance, quoted_proj, quoted_int63, quoted_float64) structure_of_term end @@ -68,7 +69,7 @@ struct match D.inspect_term trm with | ACoq_tRel x -> evm, Constr.mkRel (D.unquote_int x + 1) | ACoq_tVar x -> evm, Constr.mkVar (D.unquote_ident x) - | ACoq_tEvar (n, l) -> + | ACoq_tEvar (n, l) -> let evm, l' = map_evm (aux env) evm l in D.unquote_evar env evm n l' | ACoq_tSort x -> let evm, u = D.unquote_universe evm x in evm, Constr.mkSort u @@ -84,7 +85,7 @@ struct let evm, t = aux env evm t in let evm, b = aux (Environ.push_rel (LocalAssum (n, t)) env) evm b in evm, Constr.mkLambda (n, t, b) - | ACoq_tLetIn (n,e,t,b) -> + | ACoq_tLetIn (n,e,t,b) -> let n = D.unquote_aname n in let evm, e = aux env evm e in let evm, t = aux env evm t in @@ -107,20 +108,20 @@ struct evm, Constr.mkIndU (i, u) | ACoq_tCase (ci, p, c, brs) -> let ind = D.unquote_inductive ci.aci_ind in - let relevance = D.unquote_relevance ci.aci_relevance in + let relevance = D.unquote_relevance ci.aci_relevance in let ci = Inductiveops.make_case_info (Global.env ()) ind relevance Constr.RegularStyle in let evm, puinst = D.unquote_universe_instance evm p.auinst in let evm, pars = map_evm (aux env) evm p.apars in let pars = Array.of_list pars in let napctx = CArray.map_of_list D.unquote_aname (List.rev p.apcontext) in - let pctx = CaseCompat.case_predicate_context env ci puinst pars napctx in + let pctx = CaseCompat.case_predicate_context env ci puinst pars napctx in let evm, pret = aux (Environ.push_rel_context pctx env) evm p.apreturn in let evm, c = aux env evm c in let brs = List.map (fun { abcontext = bctx; abbody = bbody } -> let nabctx = CArray.map_of_list D.unquote_aname (List.rev bctx) in (nabctx, bbody)) brs in let brs = CaseCompat.case_branches_contexts env ci puinst pars (Array.of_list brs) in - let denote_br evm (nas, bctx, bbody) = + let denote_br evm (nas, bctx, bbody) = let evm, bbody = aux (Environ.push_rel_context bctx env) evm bbody in evm, (nas, bbody) in @@ -160,6 +161,12 @@ struct evm, Constr.mkProj (p', t') | ACoq_tInt x -> evm, Constr.mkInt (D.unquote_int63 x) | ACoq_tFloat x -> evm, Constr.mkFloat (D.unquote_float64 x) + | ACoq_tArray (u, arr, def, ty) -> + let evm, u = D.unquote_universe_level evm u in + let evm, arr = Array.fold_left_map (fun evm a -> aux env evm a) evm arr in + let evm, def = aux env evm def in + let evm, ty = aux env evm ty in + evm, Constr.mkArray (Univ.Instance.of_array [|u|], arr, def, ty) in aux env evm trm diff --git a/template-coq/src/quoter.ml b/template-coq/src/quoter.ml index 1fb874a7b..9dbd3018a 100644 --- a/template-coq/src/quoter.ml +++ b/template-coq/src/quoter.ml @@ -15,13 +15,13 @@ let cast_prop = ref (false) (* whether Set Template Cast Propositions is on, as needed for erasure in Certicoq *) let is_cast_prop () = !cast_prop -let warn_primitive_turned_into_axiom = +let warn_primitive_turned_into_axiom = CWarnings.create ~name:"primitive-turned-into-axiom" ~category:"metacoq" Pp.(fun prim -> str "Quoting primitive " ++ str prim ++ str " into an axiom.") let warn_ignoring_private_polymorphic_universes = CWarnings.create ~name:"private-polymorphic-universes-ignored" ~category:"metacoq" Pp.(fun () -> str "Ignoring private polymorphic universes.") - + let toDecl (old: Name.t Context.binder_annot * ((Constr.constr) option) * Constr.constr) : Constr.rel_declaration = let (name,value,typ) = old in match value with @@ -69,6 +69,7 @@ sig val mkCoFix : quoted_int * (quoted_aname array * t array * t array) -> t val mkInt : quoted_int63 -> t val mkFloat : quoted_float64 -> t + val mkArray : quoted_univ_level -> t array -> default:t -> ty:t -> t val mkBindAnn : quoted_name -> quoted_relevance -> quoted_aname val mkName : quoted_ident -> quoted_name @@ -93,6 +94,7 @@ sig val quote_constraint_type : Univ.constraint_type -> quoted_constraint_type val quote_univ_constraint : Univ.univ_constraint -> quoted_univ_constraint + val quote_univ_level : Univ.Level.t -> quoted_univ_level val quote_univ_instance : Univ.Instance.t -> quoted_univ_instance val quote_univ_constraints : Univ.Constraints.t -> quoted_univ_constraints val quote_univ_context : Univ.UContext.t -> quoted_univ_context @@ -122,17 +124,17 @@ sig val quote_context : quoted_context_decl list -> quoted_context val mk_one_inductive_body : - quoted_ident * + quoted_ident * quoted_context (* ind indices context *) * quoted_sort (* ind sort *) * t (* ind type *) * - quoted_sort_family * + quoted_sort_family * (quoted_ident * quoted_context (* arguments context *) * t list (* indices in the conclusion *) * - t (* constr type *) * + t (* constr type *) * quoted_int (* arity (w/o lets) *)) list * (quoted_ident * quoted_relevance * t (* projection type *)) list * - quoted_relevance -> + quoted_relevance -> quoted_one_inductive_body val mk_mutual_inductive_body : @@ -152,10 +154,11 @@ sig val empty_global_declarations : unit -> quoted_global_declarations val add_global_decl : quoted_kernel_name -> quoted_global_decl -> quoted_global_declarations -> quoted_global_declarations - - type pre_quoted_retroknowledge = + + type pre_quoted_retroknowledge = { retro_int63 : quoted_kernel_name option; - retro_float64 : quoted_kernel_name option } + retro_float64 : quoted_kernel_name option; + retro_array : quoted_kernel_name option } val quote_retroknowledge : pre_quoted_retroknowledge -> quoted_retroknowledge @@ -187,12 +190,12 @@ struct let quote_ugraph ?kept (g : UGraph.t) = debug Pp.(fun () -> str"Quoting ugraph"); - let levels, cstrs, eqs = + let levels, cstrs, eqs = match kept with | None -> let cstrs, eqs = UGraph.constraints_of_universes g in UGraph.domain g, cstrs, eqs - | Some l -> + | Some l -> debug Pp.(fun () -> str"Quoting graph restricted to: " ++ Univ.Level.Set.pr Univ.Level.pr l); (* Feedback.msg_debug Pp.(str"Graph is: " ++ UGraph.pr_universes Univ.Level.pr (UGraph.repr g)); *) let dom = UGraph.domain g in @@ -201,7 +204,7 @@ struct let cstrs = time Pp.(str"Computing graph restriction") (UGraph.constraints_for ~kept) g in l, cstrs, [] in - let levels, cstrs = + let levels, cstrs = List.fold_right (fun eqs acc -> match Univ.Level.Set.elements eqs with | [] -> acc @@ -213,7 +216,7 @@ struct in let levels = Univ.Level.Set.add Univ.Level.set levels in debug Pp.(fun () -> str"Universe context: " ++ Univ.pr_universe_context_set Univ.Level.pr (levels, cstrs)); - time (Pp.str"Quoting universe context") + time (Pp.str"Quoting universe context") (fun uctx -> Q.quote_univ_contextset uctx) (levels, cstrs) let quote_inductive' (ind, i) : Q.quoted_inductive = @@ -227,7 +230,7 @@ struct let b', acc = quote_term acc env sigma b in let t', acc = quote_term acc env sigma t in (Q.quote_context_decl (Q.quote_aname na) (Some b') t', acc) - + let quote_rel_context quote_term acc env sigma ctx = let decls, env, acc = List.fold_right (fun decl (ds, env, acc) -> @@ -236,14 +239,14 @@ struct ctx ([],env,acc) in Q.quote_context decls, acc - let quote_binder b = + let quote_binder b = Q.quote_aname b let quote_name_annots nas = Array.map quote_binder nas let quote_terms quote_term acc env sigma ts = - let acc, ts = + let acc, ts = CArray.fold_left_map (fun acc t -> let (x, acc) = quote_term acc env sigma t in acc, x) acc ts in ts, acc @@ -307,13 +310,13 @@ struct Q.quote_int (snd ci.Constr.ci_ind)) in let npar = Q.quote_int ci.Constr.ci_npar in let q_relevance = Q.quote_relevance ci.Constr.ci_relevance in - let acc, q_pars = CArray.fold_left_map (fun acc par -> let (qt, acc) = quote_term acc env sigma par in acc, qt) acc pars in + let acc, q_pars = CArray.fold_left_map (fun acc par -> let (qt, acc) = quote_term acc env sigma par in acc, qt) acc pars in let qu = Q.quote_univ_instance u in - let pctx = CaseCompat.case_predicate_context (snd env) ci u pars predctx in + let pctx = CaseCompat.case_predicate_context (snd env) ci u pars predctx in let qpctx = quote_name_annots predctx in let (qpred,acc) = quote_term acc (push_rel_context pctx env) sigma pred in let (qdiscr,acc) = quote_term acc env sigma discr in - let cbrs = CaseCompat.case_branches_contexts (snd env) ci u pars brs in + let cbrs = CaseCompat.case_branches_contexts (snd env) ci u pars brs in let (branches,acc) = CArray.fold_left2 (fun (bodies,acc) (brnas, brctx, bbody) narg -> let (qbody,acc) = quote_term acc (push_rel_context brctx env) sigma bbody in @@ -335,7 +338,12 @@ struct | Constr.Int i -> (Q.mkInt (Q.quote_int63 i), acc) | Constr.Float f -> (Q.mkFloat (Q.quote_float64 f), acc) | Constr.Meta _ -> failwith "Meta not supported by TemplateCoq" - | Constr.Array _ -> failwith "Primitive arrays not supported by TemplateCoq" + | Constr.Array (u, ar, def, ty) -> + let u = match Univ.Instance.to_array u with [| u |] -> u | _ -> assert false in + let def', acc = quote_term acc env sigma def in + let ty', acc = quote_term acc env sigma ty in + let acc, arr' = Array.fold_left_map (fun acc t -> let t', acc = quote_term acc env sigma t in acc, t') acc ar in + Q.mkArray (Q.quote_univ_level u) arr' ~default:def' ~ty:ty', acc in aux acc env trm and quote_recdecl (acc : 'a) env sigma b (ns,ts,ds) = @@ -378,7 +386,7 @@ struct let ctx = oib.mind_arity_ctxt in CList.chop (List.length ctx - List.length mib.mind_params_ctxt) ctx in - let indices, acc = quote_rel_context quote_term acc (push_rel_context pars env) sigma indices in + let indices, acc = quote_rel_context quote_term acc (push_rel_context pars env) sigma indices in let indty, acc = quote_term acc env sigma indty in let indsort = Q.quote_sort (inductive_sort oib) in let (reified_ctors,acc) = @@ -387,11 +395,11 @@ struct let ty = Inductive.abstract_constructor_type_relatively_to_inductive_types_context ntyps t ty in let ctx, concl = Term.decompose_prod_assum ty in let argctx, parsctx = - CList.chop (List.length ctx - List.length mib.mind_params_ctxt) ctx + CList.chop (List.length ctx - List.length mib.mind_params_ctxt) ctx in let envcstr = push_rel_context parsctx envind in let qargctx, acc = quote_rel_context quote_term acc envcstr sigma argctx in - let qindices, acc = + let qindices, acc = let hd, args = Constr.decompose_appvect concl in let pars, args = CArray.chop mib.mind_nparams args in let envconcl = push_rel_context argctx envcstr in @@ -432,7 +440,7 @@ struct let mind = Q.mk_mutual_inductive_body finite nparams paramsctx bodies uctx var in ref_name, Q.mk_inductive_decl mind, acc in ((fun acc env -> quote_term acc (false, env)), - (fun acc env t mib -> + (fun acc env t mib -> quote_minductive_type acc (false, env) t mib)) let quote_term env sigma trm = @@ -447,17 +455,17 @@ struct Ind of Names.inductive * mutual_inductive_body | Const of KerName.t - let universes_of_ctx ctx = + let universes_of_ctx ctx = Context.Rel.fold_inside (fun univs d -> Univ.Level.Set.union univs (Context.Rel.Declaration.fold_constr (fun c univs -> Univ.Level.Set.union univs (Vars.universes_of_constr c)) d univs)) ~init:Univ.Level.Set.empty ctx - - let universes_of_mib mib = + + let universes_of_mib mib = let pars = mib.mind_params_ctxt in let univs = universes_of_ctx pars in Array.fold_left (fun univs oib -> let univsty = universes_of_ctx oib.mind_arity_ctxt in - Array.fold_left (fun univs cty -> + Array.fold_left (fun univs cty -> let univscty = Vars.universes_of_constr cty in Univ.Level.Set.union univs univscty) univsty oib.mind_user_lc) @@ -495,7 +503,7 @@ struct let cd = Environ.lookup_constant c env in let body = match cd.const_body with | Undef _ -> None - | Primitive t -> + | Primitive t -> warn_primitive_turned_into_axiom (CPrimitives.to_string t); None | Def cs -> Some cs @@ -504,9 +512,9 @@ struct let c, univs = Global.force_proof Library.indirect_accessor lc in match univs with | Opaqueproof.PrivateMonomorphic () -> Some c - | Opaqueproof.PrivatePolymorphic csts -> - let () = - if not (Univ.ContextSet.is_empty csts) then + | Opaqueproof.PrivatePolymorphic csts -> + let () = + if not (Univ.ContextSet.is_empty csts) then warn_ignoring_private_polymorphic_universes () in Some c else None @@ -514,7 +522,7 @@ struct let tm, acc = match body with | None -> None, acc - | Some tm -> + | Some tm -> let univs = Vars.universes_of_constr tm in let () = universes := Univ.Level.Set.union univs !universes in try let (tm, acc) = quote_term acc (Global.env ()) (Evd.from_env @@ Global.env ()) tm in @@ -535,7 +543,7 @@ struct in let rel = Q.quote_relevance cd.const_relevance in let cst_bdy = Q.mk_constant_body ty tm uctx rel in - let decl = Q.mk_constant_decl cst_bdy in + let decl = Q.mk_constant_decl cst_bdy in constants := (Q.quote_kn kn, decl) :: !constants end in @@ -552,22 +560,23 @@ struct in let (tm, _) = quote_rem () env sigma trm in let decls = List.fold_right (fun (kn, d) acc -> Q.add_global_decl kn d acc) !constants (Q.empty_global_declarations ()) in - let univs = - if with_universes then + let univs = + if with_universes then let univs = Univ.Level.Set.union (Vars.universes_of_constr trm) !universes in let univs = Univ.Level.Set.filter (fun l -> Option.is_empty (Univ.Level.var_index l)) univs in quote_ugraph ~kept:univs (Environ.universes env) - else + else (debug Pp.(fun () -> str"Skipping universes: "); - time (Pp.str"Quoting empty universe context") + time (Pp.str"Quoting empty universe context") (fun uctx -> Q.quote_univ_contextset uctx) Univ.ContextSet.empty) in let retro = let retro = env.Environ.retroknowledge in let quote_retro = Option.map (fun c -> Q.quote_kn (Names.Constant.canonical c)) in - let pre = + let pre = { Q.retro_int63 = quote_retro retro.Retroknowledge.retro_int63 ; - Q.retro_float64 = quote_retro retro.Retroknowledge.retro_float64 } + Q.retro_float64 = quote_retro retro.Retroknowledge.retro_float64 ; + Q.retro_array = quote_retro retro.Retroknowledge.retro_array } in Q.quote_retroknowledge pre in let env = Q.mk_global_env univs decls retro in diff --git a/template-coq/src/reification.ml b/template-coq/src/reification.ml index 524fc9779..5eae5f27e 100644 --- a/template-coq/src/reification.ml +++ b/template-coq/src/reification.ml @@ -23,6 +23,7 @@ sig type quoted_constraint_type type quoted_univ_constraint type quoted_univ_constraints + type quoted_univ_level type quoted_univ_instance type quoted_univ_context type quoted_univ_contextset diff --git a/template-coq/src/tm_util.ml b/template-coq/src/tm_util.ml index 0a798f01e..50ec20253 100644 --- a/template-coq/src/tm_util.ml +++ b/template-coq/src/tm_util.ml @@ -1,5 +1,5 @@ open Pp - + let contrib_name = "template-coq" let gen_constant_in_modules s = @@ -9,15 +9,15 @@ let gen_constant_in_modules s = ) (* lazy (Universes.constr_of_global (Coqlib.gen_reference_in_modules locstr dirs s)) *) -(* This allows to load template_plugin and the extractable plugin at the same time +(* This allows to load template_plugin and the extractable plugin at the same time while have option settings apply to both *) let timing_opt = let open Goptions in let key = ["MetaCoq"; "Timing"] in let tables = get_tables () in - try + try let _ = OptionMap.find key tables in - fun () -> + fun () -> let tables = get_tables () in let opt = OptionMap.find key tables in match opt.opt_value with @@ -27,21 +27,21 @@ let gen_constant_in_modules s = declare_bool_option_and_ref ~depr:false ~key ~value:false let time prefix f x = - if timing_opt () then + if timing_opt () then let start = Unix.gettimeofday () in let res = f x in let stop = Unix.gettimeofday () in let () = Feedback.msg_info Pp.(prefix ++ str " executed in: " ++ Pp.real (stop -. start) ++ str "s") in res else f x - + let debug_opt = let open Goptions in let key = ["MetaCoq"; "Debug"] in let tables = get_tables () in - try + try let _ = OptionMap.find key tables in - fun () -> + fun () -> let tables = get_tables () in let opt = OptionMap.find key tables in match opt.opt_value with @@ -65,7 +65,7 @@ let rec filter_map f l = match f x with | Some x' -> x' :: filter_map f xs | None -> filter_map f xs - + let rec app_full trm acc = match Constr.kind trm with Constr.App (f, xs) -> app_full f (Array.to_list xs @ acc) @@ -140,7 +140,7 @@ module CaseCompat = let paramdecl = Vars.subst_instance_context u mib.mind_params_ctxt in let paramsubst = Vars.subst_of_rel_context_instance paramdecl params in case_predicate_context_gen mip ci u paramsubst nas - + let case_branches_contexts_gen mib ci u params brs = (* Γ ⊢ c : I@{u} params args *) (* Γ, indices, self : I@{u} params indices ⊢ p : Type *) @@ -156,7 +156,7 @@ module CaseCompat = (nas, ctx, br) in Array.map2_i build_one_branch brs mip.mind_nf_lc - in + in ebr let case_branches_contexts env ci u pars brs = @@ -171,7 +171,7 @@ module RetypeMindEntry = open Names open Univ - let retype env evm c = + let retype env evm c = Typing.type_of env evm (EConstr.of_constr c) let retype_decl env evm decl = @@ -184,15 +184,15 @@ module RetypeMindEntry = let evm, ty = Typing.solve_evars env evm (EConstr.of_constr ty) in let evm = Typing.check env evm b ty in evm, Context.Rel.Declaration.LocalDef (na, b, ty) - - let retype_context env evm ctx = + + let retype_context env evm ctx = let env, evm, ctx = Context.Rel.fold_outside (fun decl (env, evm, ctx) -> let evm, d = retype_decl env evm decl in - (EConstr.push_rel d env, evm, d :: ctx)) + (EConstr.push_rel d env, evm, d :: ctx)) ctx ~init:(env, evm, []) in evm, ctx - let sup_sort s1 s2 = + let sup_sort s1 s2 = let open Sorts in match s1, s2 with | (_, SProp) -> assert false (* template SProp not allowed *) @@ -207,8 +207,8 @@ module RetypeMindEntry = let evm, pars' = retype_context env evm pars in let envpars = Environ.push_rel_context pars env in let evm, arities = - List.fold_left - (fun (evm, ctx) oib -> + List.fold_left + (fun (evm, ctx) oib -> let ty = oib.mind_entry_arity in let evm, s = retype envpars evm ty in let ty = Term.it_mkProd_or_LetIn ty pars in @@ -218,9 +218,9 @@ module RetypeMindEntry = let env = Environ.push_rel_context arities env in let env = Environ.push_rel_context pars env in let evm = - List.fold_left + List.fold_left (fun evm oib -> - let evm, cstrsort = + let evm, cstrsort = List.fold_left (fun (evm, sort) cstr -> let evm, cstrty = retype env evm cstr in let _, cstrsort = Reduction.dest_arity env (EConstr.to_constr evm cstrty) in @@ -240,7 +240,7 @@ module RetypeMindEntry = let nf_evar c = EConstr.Unsafe.to_constr (Evarutil.nf_evar evm (EConstr.of_constr c)) in let inds = List.map - (fun oib -> + (fun oib -> let arity = nf_evar oib.mind_entry_arity in let cstrs = List.map nf_evar oib.mind_entry_lc in { oib with mind_entry_arity = arity; mind_entry_lc = cstrs }) @@ -248,8 +248,8 @@ module RetypeMindEntry = in { mind with mind_entry_params = pars; mind_entry_inds = inds } - let infer_mentry_univs env evm mind = - let evm = + let infer_mentry_univs env evm mind = + let evm = match mind.mind_entry_universes with | Entries.Monomorphic_ind_entry -> evm | Entries.Template_ind_entry uctx -> evm @@ -260,7 +260,7 @@ module RetypeMindEntry = let evm, mind = infer_mentry_univs env evm mind in let evm = Evd.minimize_universes evm in let mind = nf_mentry_univs evm mind in - let ctx, mind = + let ctx, mind = match mind.mind_entry_universes with | Entries.Monomorphic_ind_entry -> Evd.universe_context_set evm, { mind with mind_entry_universes = Entries.Monomorphic_ind_entry } @@ -270,14 +270,14 @@ module RetypeMindEntry = let uctx' = Evd.to_universe_context evm in Univ.ContextSet.empty, { mind with mind_entry_universes = Entries.Polymorphic_ind_entry uctx' } in ctx, mind -end +end type ('term, 'name, 'nat) adef = { adname : 'name; adtype : 'term; adbody : 'term; rarg : 'nat } type ('term, 'name, 'nat) amfixpoint = ('term, 'name, 'nat) adef list -type ('term, 'name, 'universe_instance) apredicate = - { auinst : 'universe_instance; +type ('term, 'name, 'universe_instance) apredicate = + { auinst : 'universe_instance; apars : 'term list; apcontext : 'name list; apreturn : 'term } @@ -290,8 +290,8 @@ type ('nat, 'inductive, 'relevance) acase_info = { aci_ind : 'inductive; aci_npar : 'nat; aci_relevance : 'relevance } - -type ('term, 'nat, 'ident, 'name, 'quoted_sort, 'cast_kind, 'kername, 'inductive, 'relevance, 'universe_instance, 'projection, 'int63, 'float64) structure_of_term = + +type ('term, 'nat, 'ident, 'name, 'quoted_sort, 'cast_kind, 'kername, 'inductive, 'relevance, 'universe_level, 'universe_instance, 'projection, 'int63, 'float64) structure_of_term = | ACoq_tRel of 'nat | ACoq_tVar of 'ident | ACoq_tEvar of 'nat * 'term list @@ -304,7 +304,7 @@ type ('term, 'nat, 'ident, 'name, 'quoted_sort, 'cast_kind, 'kername, 'inductive | ACoq_tConst of 'kername * 'universe_instance | ACoq_tInd of 'inductive * 'universe_instance | ACoq_tConstruct of 'inductive * 'nat * 'universe_instance - | ACoq_tCase of ('nat, 'inductive, 'relevance) acase_info * + | ACoq_tCase of ('nat, 'inductive, 'relevance) acase_info * ('term, 'name, 'universe_instance) apredicate * 'term * ('term, 'name) abranch list | ACoq_tProj of 'projection * 'term @@ -312,4 +312,5 @@ type ('term, 'nat, 'ident, 'name, 'quoted_sort, 'cast_kind, 'kername, 'inductive | ACoq_tCoFix of ('term, 'name, 'nat) amfixpoint * 'nat | ACoq_tInt of 'int63 | ACoq_tFloat of 'float64 + | ACoq_tArray of 'universe_level * 'term array * 'term * 'term diff --git a/template-coq/theories/Ast.v b/template-coq/theories/Ast.v index 874f82223..dbc367c9d 100644 --- a/template-coq/theories/Ast.v +++ b/template-coq/theories/Ast.v @@ -1,10 +1,10 @@ (* Distributed under the terms of the MIT license. *) +From Equations Require Import Equations. From MetaCoq.Utils Require Import utils. From MetaCoq.Common Require Export Environment EnvironmentTyping Universes BasicAst. (* For primitive integers and floats *) -From Coq Require Uint63 Floats.PrimFloat Floats.SpecFloat. +From Coq Require Uint63 Floats.PrimFloat Floats.SpecFloat PArray. From Coq Require Import ssreflect Morphisms. -From Equations Require Import Equations. (** * AST of Coq kernel terms and kernel data structures @@ -415,7 +415,8 @@ Inductive term : Type := | tFix (mfix : mfixpoint term) (idx : nat) | tCoFix (mfix : mfixpoint term) (idx : nat) | tInt (i : PrimInt63.int) -| tFloat (f : PrimFloat.float). +| tFloat (f : PrimFloat.float) +| tArray (u : Level.t) (arr : list term) (default : term) (type : term). (** This can be used to represent holes, that, when unquoted, turn into fresh existential variables. The fresh evar will depend on the whole context at this point in the term, despite the empty instance. @@ -457,6 +458,8 @@ Fixpoint lift n k t : term := let k' := List.length mfix + k in let mfix' := List.map (map_def (lift n k) (lift n k')) mfix in tCoFix mfix' idx + | tArray u arr def ty => + tArray u (List.map (lift n k) arr) (lift n k def) (lift n k ty) | x => x end. @@ -494,6 +497,8 @@ Fixpoint subst s k u := let k' := List.length mfix + k in let mfix' := List.map (map_def (subst s k) (subst s k')) mfix in tCoFix mfix' idx + | tArray u arr def ty => + tArray u (List.map (subst s k) arr) (subst s k def) (subst s k ty) | x => x end. @@ -523,6 +528,8 @@ Fixpoint closedn k (t : term) : bool := | tCoFix mfix idx => let k' := List.length mfix + k in List.forallb (test_def (closedn k) (closedn k')) mfix + | tArray u arr def ty => + List.forallb (closedn k) arr && closedn k def && closedn k ty | _ => true end. @@ -548,6 +555,9 @@ Fixpoint noccur_between k n (t : term) : bool := | tCoFix mfix idx => let k' := List.length mfix + k in List.forallb (test_def (noccur_between k n) (noccur_between k' n)) mfix + | tArray u arr def ty => + List.forallb (noccur_between k n) arr && + noccur_between k n def && noccur_between k n ty | _ => true end. @@ -556,6 +566,8 @@ Fixpoint noccur_between k n (t : term) : bool := match c with | tRel _ | tVar _ => c | tInt _ | tFloat _ => c + | tArray u' arr def ty => tArray (subst_instance_level u u') (List.map (subst_instance_constr u) arr) + (subst_instance_constr u def) (subst_instance_constr u ty) | tEvar ev args => tEvar ev (List.map (subst_instance_constr u) args) | tSort s => tSort (subst_instance_univ u s) | tConst c u' => tConst c (subst_instance_instance u u') @@ -602,6 +614,8 @@ Fixpoint closedu (k : nat) (t : term) : bool := forallb (test_def (closedu k) (closedu k)) mfix | tCoFix mfix idx => forallb (test_def (closedu k) (closedu k)) mfix + | tArray u arr def ty => + closedu_level k u && forallb (closedu k) arr && closedu k def && closedu k ty | _ => true end. @@ -821,4 +835,3 @@ Definition case_branch_type_gen ind mdecl params puinst ptm i cdecl (br : branch Definition case_branch_type ind mdecl p ptm i cdecl br : context * term := case_branch_type_gen ind mdecl p.(pparams) p.(puinst) ptm i cdecl br. - diff --git a/template-coq/theories/AstUtils.v b/template-coq/theories/AstUtils.v index ea6d72841..07e3c0d16 100644 --- a/template-coq/theories/AstUtils.v +++ b/template-coq/theories/AstUtils.v @@ -68,6 +68,8 @@ Module string_of_term_tree. | tCoFix l n => "CoFix(" ^ (string_of_list (string_of_def string_of_term) l) ^ "," ^ string_of_nat n ^ ")" | tInt i => "Int(" ^ string_of_prim_int i ^ ")" | tFloat f => "Float(" ^ string_of_float f ^ ")" + | tArray u arr def ty => "Array(" ^ string_of_level u ^ "," ^ + string_of_list string_of_term arr ^ "," ^ string_of_term def ^ "," ^ string_of_term ty ^ ")" end. End string_of_term_tree. @@ -252,6 +254,8 @@ Fixpoint strip_casts t := | tCoFix mfix idx => let mfix' := List.map (map_def strip_casts strip_casts) mfix in tCoFix mfix' idx + | tArray u arr def ty => + tArray u (List.map strip_casts arr) (strip_casts def) (strip_casts ty) | tRel _ | tVar _ | tSort _ | tConst _ _ | tInd _ _ | tConstruct _ _ _ => t | tInt _ | tFloat _ => t end. diff --git a/template-coq/theories/Checker.v b/template-coq/theories/Checker.v index 20ed11f70..8a8574408 100644 --- a/template-coq/theories/Checker.v +++ b/template-coq/theories/Checker.v @@ -800,7 +800,7 @@ Section Typecheck. | None => raise (IllFormedFix mfix n) end - | tInt _ | tFloat _ => raise (NotSupported "primitive types") + | tInt _ | tFloat _ | tArray _ _ _ _ => raise (NotSupported "primitive types") end. Definition check (Γ : context) (t : term) (ty : term) : typing_result unit := diff --git a/template-coq/theories/Constants.v b/template-coq/theories/Constants.v index d57ad3e72..1c7834eca 100644 --- a/template-coq/theories/Constants.v +++ b/template-coq/theories/Constants.v @@ -178,6 +178,7 @@ Register MetaCoq.Template.Ast.tFix as metacoq.ast.tFix. Register MetaCoq.Template.Ast.tCoFix as metacoq.ast.tCoFix. Register MetaCoq.Template.Ast.tInt as metacoq.ast.tInt. Register MetaCoq.Template.Ast.tFloat as metacoq.ast.tFloat. +Register MetaCoq.Template.Ast.tArray as metacoq.ast.tArray. (* Local and global declarations *) Register MetaCoq.Template.Ast.parameter_entry as metacoq.ast.parameter_entry. diff --git a/template-coq/theories/EtaExpand.v b/template-coq/theories/EtaExpand.v index 00d8b41cc..96b1acc7f 100644 --- a/template-coq/theories/EtaExpand.v +++ b/template-coq/theories/EtaExpand.v @@ -137,6 +137,7 @@ Section Eta. ++ string_of_kername ind.(inductive_mind)) end | tCast t1 k t2 => tCast (eta_expand Γ t1) k (eta_expand Γ t2) + | tArray u arr def ty => tArray u (List.map (eta_expand Γ) arr) (eta_expand Γ def) (eta_expand Γ ty) | tInt _ | tFloat _ => t end. @@ -292,7 +293,12 @@ Inductive expanded (Γ : list nat): term -> Prop := Forall (expanded Γ) args -> expanded Γ (tApp (tConstruct ind c u) args) | expanded_tInt i : expanded Γ (tInt i) -| expanded_tFloat f : expanded Γ (tFloat f). +| expanded_tFloat f : expanded Γ (tFloat f) +| expanded_tArray u arr def ty : + Forall (expanded Γ) arr -> + expanded Γ def -> + expanded Γ ty -> + expanded Γ (tArray u arr def ty). End expanded. @@ -357,9 +363,14 @@ forall (Σ : global_env) (P : list nat -> term -> Prop), P Γ(tApp (tConstruct ind c u) args)) -> (forall Γ i, P Γ (tInt i)) -> (forall Γ f, P Γ (tFloat f)) -> +(forall Γ u arr def ty, + Forall (P Γ) arr -> + P Γ def -> + P Γ ty -> + P Γ (tArray u arr def ty)) -> forall Γ, forall t : term, expanded Σ Γ t -> P Γ t. Proof. - intros Σ P HRel HRel_app HVar HEvar HSort HCast HProd HLamdba HLetIn HApp HConst HInd HConstruct HCase HProj HFix HCoFix HConstruct_app Hint Hfloat. + intros Σ P HRel HRel_app HVar HEvar HSort HCast HProd HLamdba HLetIn HApp HConst HInd HConstruct HCase HProj HFix HCoFix HConstruct_app Hint Hfloat Harr. fix f 3. intros Γ t Hexp. destruct Hexp; eauto. all: match goal with [H : Forall _ _ |- _] => let all := fresh "all" in rename H into all end. @@ -379,6 +390,7 @@ Proof. generalize mfix at 1 3. intros mfix0 H. induction H; econstructor; cbn in *; eauto; split. - eapply HConstruct_app; eauto. clear - all f. induction all; econstructor; eauto. + - eapply Harr; eauto. clear -all f. induction all; constructor; auto. Qed. Local Hint Constructors expanded : core. @@ -625,6 +637,7 @@ Proof. - destruct t; invs H4. eapply expanded_tConstruct_app; eauto. revert H0. now len. solve_all. + - constructor; eauto. solve_all. Qed. Lemma expanded_lift {Σ : global_env} Γ' Γ'' Γ t : @@ -670,6 +683,7 @@ Proof. shelve. autorewrite with len list in H |- *. eapply H. - eapply expanded_tConstruct_app; eauto. now len. solve_all. + - constructor; eauto. eapply Forall_map. solve_all. Qed. Lemma expanded_lift' {Σ : global_env} Γ' Γ'' Γ t Γassum Γgoal n m : @@ -1220,6 +1234,7 @@ Proof. assert (#|Typing.fix_context mfix| = #|mfix|). { unfold Typing.fix_context. now len. } revert H4. generalize (Typing.fix_context mfix). clear. induction #|mfix|; intros []; cbn; intros; try congruence; econstructor; eauto. + - cbn. econstructor; eauto. solve_all. eapply b; tea. solve_all. - eapply typing_wf_local; eauto. Qed. diff --git a/template-coq/theories/Induction.v b/template-coq/theories/Induction.v index 99d9c779f..d484cef7c 100644 --- a/template-coq/theories/Induction.v +++ b/template-coq/theories/Induction.v @@ -34,6 +34,7 @@ Lemma term_forall_list_ind : (forall (m : mfixpoint term) (n : nat), tFixProp P P m -> P (tCoFix m n)) -> (forall i, P (tInt i)) -> (forall f, P (tFloat f)) -> + (forall u arr def ty, Forall P arr -> P def -> P ty -> P (tArray u arr def ty)) -> forall t : term, P t. Proof. intros until t. revert t. @@ -77,6 +78,7 @@ Lemma term_forall_list_rect : (forall (m : mfixpoint term) (n : nat), tFixType P P m -> P (tCoFix m n)) -> (forall i, P (tInt i)) -> (forall f, P (tFloat f)) -> + (forall u arr def ty, All P arr -> P def -> P ty -> P (tArray u arr def ty)) -> forall t : term, P t. Proof. intros until t. revert t. diff --git a/template-coq/theories/Pretty.v b/template-coq/theories/Pretty.v index 4201cf830..694ce8383 100644 --- a/template-coq/theories/Pretty.v +++ b/template-coq/theories/Pretty.v @@ -252,6 +252,8 @@ Module PrintTermTree. " in " ^ List.nth_default (string_of_nat n) (map (string_of_name ∘ binder_name ∘ dname) l) n) | tInt i => "Int(" ^ string_of_prim_int i ^ ")" | tFloat f => "Float(" ^ string_of_float f ^ ")" + | tArray u arr def ty => "Array(" ^ string_of_level u ^ "," ^ + string_of_list string_of_term arr ^ "," ^ string_of_term def ^ "," ^ string_of_term ty ^ ")" end. Definition pr_context_decl Γ (c : context_decl) : ident * t := diff --git a/template-coq/theories/ReflectAst.v b/template-coq/theories/ReflectAst.v index 3123f91e7..712f4ebd0 100644 --- a/template-coq/theories/ReflectAst.v +++ b/template-coq/theories/ReflectAst.v @@ -158,6 +158,16 @@ Proof. subst. left. reflexivity. - destruct (eq_dec f f0) ; nodec. subst. left. reflexivity. + - destruct (IHx1 t1); subst; nodec. + destruct (IHx2 t2); subst; nodec. + destruct (eq_dec u u0); nodec; subst. + revert arr0. + induction X. + + intros []; auto. nodec. + + intros []; auto. nodec. + destruct (IHX l0). noconf e. + destruct (p t); nodec. subst. left; auto. + nodec. Defined. #[global] Instance reflect_term : ReflectEq term := diff --git a/template-coq/theories/TermEquality.v b/template-coq/theories/TermEquality.v index 24f76ea05..a11f62bb4 100644 --- a/template-coq/theories/TermEquality.v +++ b/template-coq/theories/TermEquality.v @@ -239,7 +239,12 @@ Inductive eq_term_upto_univ_napp Σ (Re Rle : Universe.t -> Universe.t -> Prop) eq_term_upto_univ_napp Σ Re Rle napp (tCast t1 c t2) (tCast t1' c' t2') | eq_Int i : eq_term_upto_univ_napp Σ Re Rle napp (tInt i) (tInt i) -| eq_Float f : eq_term_upto_univ_napp Σ Re Rle napp (tFloat f) (tFloat f). +| eq_Float f : eq_term_upto_univ_napp Σ Re Rle napp (tFloat f) (tFloat f) +| eq_Array u u' arr arr' def def' ty ty' : + All2 (eq_term_upto_univ_napp Σ Re Rle 0) arr arr' -> + eq_term_upto_univ_napp Σ Re Rle 0 def def' -> + eq_term_upto_univ_napp Σ Re Rle 0 ty ty' -> + eq_term_upto_univ_napp Σ Re Rle napp (tArray u arr def ty) (tArray u' arr' def' ty'). Notation eq_term_upto_univ Σ Re Rle := (eq_term_upto_univ_napp Σ Re Rle 0). @@ -313,6 +318,7 @@ Proof. intros x [? ?]. repeat split ; auto. - eapply All_All2. 1: eassumption. intros x [? ?]. repeat split ; auto. + - eapply All_All2; tea. cbn. eauto. Qed. Lemma eq_term_refl `{checker_flags} Σ φ t : eq_term Σ φ t t. @@ -477,6 +483,8 @@ Proof. eapply All2_impl'; tea. eapply All_impl; eauto. cbn. intros x [? ?] y [[[? ?] ?] ?]. repeat split; eauto. + - inversion 1; subst; constructor; eauto. + eapply All2_impl'; tea. eapply All_impl; eauto. Qed. diff --git a/template-coq/theories/Typing.v b/template-coq/theories/Typing.v index bbb9bb543..cc451c39b 100644 --- a/template-coq/theories/Typing.v +++ b/template-coq/theories/Typing.v @@ -845,16 +845,27 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> wf_local Σ Γ -> primitive_constant Σ primInt = Some prim_ty -> declared_constant Σ prim_ty cdecl -> - primitive_invariants cdecl -> + primitive_invariants primInt cdecl -> Σ ;;; Γ |- tInt p : tConst prim_ty [] | type_Float p prim_ty cdecl : wf_local Σ Γ -> primitive_constant Σ primFloat = Some prim_ty -> declared_constant Σ prim_ty cdecl -> - primitive_invariants cdecl -> + primitive_invariants primFloat cdecl -> Σ ;;; Γ |- tFloat p : tConst prim_ty [] +| type_Array prim_ty cdecl u arr def ty : + wf_local Σ Γ -> + primitive_constant Σ primArray = Some prim_ty -> + declared_constant Σ prim_ty cdecl -> + primitive_invariants primArray cdecl -> + let s := Universe.make u in + Σ ;;; Γ |- ty : tSort s -> + Σ ;;; Γ |- def : ty -> + All (fun t => Σ ;;; Γ |- t : ty) arr -> + Σ ;;; Γ |- tArray u arr def ty : tApp (tConst prim_ty [u]) [ty] + | type_Conv t A B s : Σ ;;; Γ |- t : A -> Σ ;;; Γ |- B : tSort s -> @@ -969,6 +980,7 @@ Proof. - exact (S (Nat.max (Nat.max (All_local_env_size typing_size _ _ a) (all_size _ (fun x p => infer_sort_size (typing_sort_size typing_size) Σ _ _ p) a0)) (all_size _ (fun x p => typing_size Σ _ _ _ p) a1))). - exact (S (All_local_env_size typing_size _ _ a)). - exact (S (All_local_env_size typing_size _ _ a)). + - exact (S (Nat.max d2 (Nat.max d3 (all_size _ (fun t => typing_size Σ Γ t ty) a0)))). Defined. Lemma typing_size_pos `{checker_flags} {Σ Γ t T} (d : Σ ;;; Γ |- t : T) : typing_size d > 0. @@ -1245,16 +1257,29 @@ Lemma typing_ind_env `{cf : checker_flags} : PΓ Σ Γ wfΓ -> primitive_constant Σ primInt = Some prim_ty -> declared_constant Σ prim_ty cdecl -> - primitive_invariants cdecl -> + primitive_invariants primInt cdecl -> P Σ Γ (tInt p) (tConst prim_ty [])) -> (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) p prim_ty cdecl, PΓ Σ Γ wfΓ -> primitive_constant Σ primFloat = Some prim_ty -> declared_constant Σ prim_ty cdecl -> - primitive_invariants cdecl -> + primitive_invariants primFloat cdecl -> P Σ Γ (tFloat p) (tConst prim_ty [])) -> + (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) u arr def ty prim_ty cdecl, + PΓ Σ Γ wfΓ -> + primitive_constant Σ primArray = Some prim_ty -> + declared_constant Σ prim_ty cdecl -> + primitive_invariants primArray cdecl -> + let s := Universe.make u in + Σ ;;; Γ |- ty : tSort s -> + P Σ Γ ty (tSort s) -> + Σ ;;; Γ |- def : ty -> + P Σ Γ def ty -> + All (fun t => Σ ;;; Γ |- t : ty × P Σ Γ t ty) arr -> + P Σ Γ (tArray u arr def ty) (tApp (tConst prim_ty [u]) [ty])) -> + (forall Σ (wfΣ : wf Σ) (Γ : context) (wfΓ : wf_local Σ Γ) (t A B : term) s, PΓ Σ Γ wfΓ -> Σ ;;; Γ |- t : A -> @@ -1268,7 +1293,7 @@ Lemma typing_ind_env `{cf : checker_flags} : Proof. intros P Pdecl PΓ; unfold env_prop. intros XΓ. - intros X X0 Xcast X1 X2 X3 X4 X5 X6 X7 X8 X9 X10 X11 Xint Xfloat X12 Σ wfΣ Γ wfΓ t T H. + intros X X0 Xcast X1 X2 X3 X4 X5 X6 X7 X8 X9 X10 X11 Xint Xfloat Xarr X12 Σ wfΣ Γ wfΓ t T H. (* NOTE (Danil): while porting to 8.9, I had to split original "pose" into 2 pieces, otherwise it takes forever to execure the "pose", for some reason *) pose (@Fix_F ({ Σ : global_env_ext & { wfΣ : wf Σ & { Γ & { t & { T & Σ ;;; Γ |- t : T }}}}})) as p0. @@ -1549,6 +1574,17 @@ Proof. eapply (X _ (typing_wf_local p) _ _ p). simpl. lia. ++ eapply IHa1. intros. eapply (X _ X0 _ _ Hty). simpl; lia. + -- eapply Xarr; tea. + * eapply (X14 _ _ _ H). simpl. subst s; lia. + * eapply (X14 _ _ _ H0). simpl. subst s; lia. + * clear -a0 X14. + assert (forall (Γ0 : context) (t T : term) (Hty : Σ;;; Γ0 |- t : T), + typing_size Hty < S (all_size _ (fun t p => typing_size p) a0) -> + on_global_env cumul_gen (lift_typing P) Σ.1 × P Σ Γ0 t T). + { intros ??? Hty ?; eapply (X14 _ _ _ Hty). simpl. lia. } + clear X14. clear -X a0. induction a0; constructor; eauto. + split => //. eapply (X _ _ _ p). cbn. lia. + eapply IHa0. intros. eapply (X _ _ _ Hty). simpl. lia. Qed. (** * Lemmas about All_local_env *) diff --git a/template-coq/theories/TypingWf.v b/template-coq/theories/TypingWf.v index 125210814..a4f97e740 100644 --- a/template-coq/theories/TypingWf.v +++ b/template-coq/theories/TypingWf.v @@ -1031,6 +1031,10 @@ Section TypingWf. solve_all; destruct a, b. all: intuition. + eapply All_nth_error in X0; eauto. destruct X0 as [s ?]; intuition. + + - split => //. + + constructor; intuition auto. solve_all. + + constructor => //. constructor => //. constructor; intuition auto. Qed. Lemma typing_all_wf_decl Σ (wfΣ : wf Σ.1) Γ (wfΓ : wf_local Σ Γ) : diff --git a/template-coq/theories/WfAst.v b/template-coq/theories/WfAst.v index a723d762e..a258e0466 100644 --- a/template-coq/theories/WfAst.v +++ b/template-coq/theories/WfAst.v @@ -57,7 +57,8 @@ Inductive wf {Σ} : term -> Type := wf (tFix mfix k) | wf_tCoFix mfix k : All (fun def => wf def.(dtype) × wf def.(dbody)) mfix -> wf (tCoFix mfix k) | wf_tInt i : wf (tInt i) -| wf_tFloat f : wf (tFloat f). +| wf_tFloat f : wf (tFloat f) +| wf_tArray u arr def ty : All wf arr -> wf def -> wf ty -> wf (tArray u arr def ty). Arguments wf : clear implicits. Derive Signature for wf. @@ -67,6 +68,7 @@ Definition wf_Inv Σ (t : term) : Type := match t with | tRel _ | tVar _ | tSort _ => unit | tInt _ | tFloat _ => unit + | tArray u arr def ty => All (wf Σ) arr * wf Σ def * wf Σ ty | tEvar n l => All (wf Σ) l | tCast t k t' => wf Σ t * wf Σ t' | tProd na t b => wf Σ t * wf Σ b @@ -150,9 +152,10 @@ Lemma term_wf_forall_list_ind Σ : (forall (m : mfixpoint term) (n : nat), tFixProp P P m -> P (tCoFix m n)) -> (forall i, P (tInt i)) -> (forall f, P (tFloat f)) -> + (forall u arr def ty, All P arr -> P def -> P ty -> P (tArray u arr def ty)) -> forall t : term, wf Σ t -> P t. Proof. - intros P H2 H3 H4 H5 H6 H7 H8 H9 H10 H11 H12 H13 H14 H15 H16 H17 H18 H19. + intros P H2 H3 H4 H5 H6 H7 H8 H9 H10 H11 H12 H13 H14 H15 H16 H17 H18 H19 Harr. intros until t. revert t. apply (term_forall_list_rect (fun t => wf Σ t -> P t)); intros; try solve [match goal with @@ -180,6 +183,7 @@ Proof. apply H16. red. red in X. solve_all. - inv X0; auto. apply H17. red. red in X. solve_all. + - inv X2; auto. eapply Harr; eauto using lift_to_wf_list. Qed. Definition wf_decl Σ d := From ed544ea33a9010c9f5a0ed36f2de733dd08d287d Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Mon, 6 Nov 2023 14:32:00 +0100 Subject: [PATCH 02/17] WIP up to confluence for arrays --- pcuic/theories/Conversion/PCUICInstConv.v | 51 ++++++++--- pcuic/theories/Conversion/PCUICNamelessConv.v | 21 ++++- .../theories/Conversion/PCUICOnFreeVarsConv.v | 3 +- pcuic/theories/Conversion/PCUICRenameConv.v | 6 ++ .../Conversion/PCUICWeakeningConfigConv.v | 2 +- pcuic/theories/PCUICAst.v | 6 ++ pcuic/theories/PCUICConfluence.v | 84 +++++++++++++++++- pcuic/theories/PCUICContextConversion.v | 13 +++ pcuic/theories/PCUICContextReduction.v | 5 ++ pcuic/theories/PCUICConversion.v | 37 ++++++++ pcuic/theories/PCUICCumulativitySpec.v | 27 +++--- pcuic/theories/PCUICEquality.v | 47 +++++----- pcuic/theories/PCUICParallelReduction.v | 47 +++++++--- .../PCUICParallelReductionConfluence.v | 36 +++++++- pcuic/theories/PCUICReduction.v | 87 +++++++++++++++++++ pcuic/theories/PCUICSubstitution.v | 6 ++ pcuic/theories/PCUICTyping.v | 4 +- pcuic/theories/PCUICWfUniverses.v | 62 ++++++++++++- pcuic/theories/Syntax/PCUICDepth.v | 31 ++++++- pcuic/theories/Syntax/PCUICNamelessDef.v | 4 +- pcuic/theories/Syntax/PCUICOnFreeVars.v | 16 +++- pcuic/theories/Typing/PCUICInstTyp.v | 14 ++- pcuic/theories/Typing/PCUICRenameTyp.v | 16 +++- .../Typing/PCUICUnivSubstitutionTyp.v | 23 +++-- 24 files changed, 558 insertions(+), 90 deletions(-) diff --git a/pcuic/theories/Conversion/PCUICInstConv.v b/pcuic/theories/Conversion/PCUICInstConv.v index 6c72b81df..840dc29e7 100644 --- a/pcuic/theories/Conversion/PCUICInstConv.v +++ b/pcuic/theories/Conversion/PCUICInstConv.v @@ -1230,7 +1230,7 @@ Proof. intros Hs t Ht. revert P s s' Hs Ht. elim t using term_forall_list_ind; cbn in |- *; intros; try easy. 8-9: rewrite /test_def in Ht. - 1-5,7-9: + 1-5,7-9,10: try rewrite H; try rewrite H0 ; try rewrite H1 ; try easy ; solve [f_equal; solve_all; eauto using up_ext_cond]. @@ -1258,15 +1258,15 @@ Proof. intros H i. rewrite /up. destruct (Nat.leb_spec0 n i) as [Hi|Hi]. - + rewrite on_free_vars_rename. - erewrite on_free_vars_ext. - 2: now eapply addnP_shiftnP. - rewrite /shiftnP => /orP []. - 1: move => /Nat.ltb_spec0 ? ; lia. - intros. - now eapply H. - + rewrite /= /shiftnP. - now move: Hi => /Nat.nle_gt /Nat.ltb_spec0 ->. + + rewrite on_free_vars_rename. + erewrite on_free_vars_ext. + 2: now eapply addnP_shiftnP. + rewrite /shiftnP => /orP []. + 1: move => /Nat.ltb_spec0 ? ; lia. + intros. + now eapply H. + + rewrite /= /shiftnP. + now move: Hi => /Nat.nle_gt /Nat.ltb_spec0 ->. Qed. Lemma on_free_vars_inst (P Q : nat -> bool) s t : @@ -1698,6 +1698,27 @@ Proof. rewrite on_free_vars_ctx_app in H1. solve_all. cbn in *. rewrite shiftnP0 in H2. tea. Defined. +Lemma map_array_model_set_value f arr value : + map_array_model f (set_array_value arr value) = + set_array_value (map_array_model f arr) (List.map f value). +Proof. + destruct arr; cbn => //. +Qed. + +Lemma map_array_model_set_default f arr value : + map_array_model f (set_array_default arr value) = + set_array_default (map_array_model f arr) (f value). +Proof. + destruct arr; cbn => //. +Qed. + +Lemma map_array_model_set_type f arr value : + map_array_model f (set_array_type arr value) = + set_array_type (map_array_model f arr) (f value). +Proof. + destruct arr; cbn => //. +Qed. + Lemma red1_inst {Σ : global_env_ext} {wfΣ : wf Σ} {Γ Δ u v σ} : usubst Γ σ Δ -> is_open_term Γ u -> @@ -1842,6 +1863,14 @@ Proof. + rewrite inst_fix_context_up. eapply usubst_app_up; eauto. + now len. * rewrite shiftnP_add in clb. rewrite <- fix_context_length in clb. rewrite <- app_length in clb. tea. + - cbn. toAll. eapply OnOne2_All_mix_left in X; tea. + rewrite map_array_model_set_value. + eapply red_primArray_one_value. cbn. + eapply OnOne2_map; solve_all. red; simpl; eauto. + - cbn. rewrite map_array_model_set_default. + eapply red_primArray_default; cbn; eauto. + - cbn. rewrite map_array_model_set_type. + eapply red_primArray_type; cbn; eauto. Defined. Lemma eq_term_upto_univ_inst Σ : @@ -1884,6 +1913,8 @@ Proof using Type. apply All2_length in a as e. rewrite <- e. generalize #|m|. intro k. eapply All2_map. simpl. solve_all. + - simpl. constructor. + eapply onPrim_map_prop; tea. cbn; intuition eauto. Qed. Lemma inst_conv {Σ : global_env_ext} {wfΣ : wf Σ} {Γ Δ σ A B} : diff --git a/pcuic/theories/Conversion/PCUICNamelessConv.v b/pcuic/theories/Conversion/PCUICNamelessConv.v index f3e2313c5..a69528123 100644 --- a/pcuic/theories/Conversion/PCUICNamelessConv.v +++ b/pcuic/theories/Conversion/PCUICNamelessConv.v @@ -22,7 +22,6 @@ Local Set Keyed Unification. Set Default Goal Selector "!". - Ltac destruct_one_andb := lazymatch goal with | h : is_true (_ && _) |- _ => @@ -176,6 +175,12 @@ Proof. -- eapply Hbod ; eassumption. -- assumption. * eapply IHm ; assumption. + - f_equal. + destruct o; auto. + f_equal. f_equal. cbn in X, hu, hv. rtoProp. + destruct a, a'; cbn in *. eapply Universe.make_inj in e. f_equal; intuition eauto. + solve_all. induction a0 => //. f_equal; eauto. + eapply r; intuition eauto. Qed. Lemma banon_list l : forallb (banon ∘ anonymize) l. @@ -226,6 +231,7 @@ Proof. repeat (eapply andb_true_intro ; split). all: try assumption. eapply IHm. assumption. + - cbn. solve_all. Qed. Lemma nl_lookup_env : @@ -348,6 +354,13 @@ Proof. intuition auto. * destruct x, y; simpl in *. apply aux; auto. * destruct x, y; simpl in *. apply aux; auto. + - constructor. + destruct i as [? []], i' as [? []]; cbn in o; depelim o; cbn in *; constructor; eauto. + + now eapply aux. + + now eapply aux. + + cbn. induction a2. + ++ constructor. + ++ cbn; constructor; [now eapply aux|]. eapply IHa2. Qed. Lemma eq_context_nl Σ Re Rle ctx ctx' : @@ -609,6 +622,7 @@ Proof. destruct x. simpl in *. unfold map_def, map_def_anon. cbn. rewrite h1 h2. reflexivity. + - simpl. f_equal. solve_all. Qed. Lemma context_position_nlctx : @@ -689,6 +703,7 @@ Proof. * unfold map_def_anon, map_def. simpl. f_equal. all: eapply p. * assumption. + - f_equal. solve_all. Qed. Lemma nlctx_fix_context_alt : @@ -781,6 +796,7 @@ Proof. * unfold map_def_anon, map_def. simpl. f_equal. all: eapply p. * assumption. + - f_equal; solve_all. Qed. Lemma nl_eq_decl {cf:checker_flags} : @@ -1371,6 +1387,8 @@ Proof. + rewrite nlctx_app_context nl_fix_context in r0. assumption. + cbn. congruence. Unshelve. all:eauto. + - cbn. constructor. eapply OnOne2_map, OnOne2_impl; tea. + unfold on_Trel; intuition auto. Qed. Lemma nl_conv {cf:checker_flags} : @@ -1538,6 +1556,7 @@ Proof. - f_equal. induction X; cbnr. f_equal; tas. destruct p, x; unfold map_def_anon; simpl in *. rewrite anonymize_two; congruence. + - f_equal; solve_all. Qed. diff --git a/pcuic/theories/Conversion/PCUICOnFreeVarsConv.v b/pcuic/theories/Conversion/PCUICOnFreeVarsConv.v index eac87fb33..6a0e1dad7 100644 --- a/pcuic/theories/Conversion/PCUICOnFreeVarsConv.v +++ b/pcuic/theories/Conversion/PCUICOnFreeVarsConv.v @@ -52,6 +52,7 @@ Proof. f_equal. 1: auto. by rewrite map_length ebod shiftnP_shiftn. + - solve_all. Qed. Lemma shiftn_ext_cond (P : nat -> bool) f f' n : @@ -75,7 +76,7 @@ Proof. revert P f f' H Ht. elim t using term_forall_list_ind; cbn in |- *; intros; try easy. - 1-6,8: + 1-6,8,11: try rewrite H; try rewrite H0 ; try rewrite H1 ; try easy; solve [f_equal; solve_all; eauto using shiftn_ext_cond]. diff --git a/pcuic/theories/Conversion/PCUICRenameConv.v b/pcuic/theories/Conversion/PCUICRenameConv.v index 6dd91a0a9..0035f456c 100644 --- a/pcuic/theories/Conversion/PCUICRenameConv.v +++ b/pcuic/theories/Conversion/PCUICRenameConv.v @@ -106,6 +106,7 @@ Proof. induction X. 1: reflexivity. destruct p, x. unfold map_def in *. simpl in *. f_equal. all: easy. + - simpl. f_equal. solve_all. Qed. Lemma rename_context_snoc0 : @@ -261,6 +262,8 @@ Proof using Type. simpl. constructor. * unfold map_def. intuition eauto. * eauto. + - simpl. constructor. + eapply onPrim_map_prop; tea. cbn; intuition eauto. Qed. Lemma urenaming_impl : @@ -1052,6 +1055,9 @@ Proof using cf. * rewrite rename_fix_context. rewrite <- fix_context_length. now eapply urenaming_context. * now len. + - constructor. eapply OnOne2_map. cbn in hav. rtoProp. + unfold on_Trel; cbn. eapply forallb_All in H. eapply OnOne2_All_mix_left in X; tea. + solve_all. Qed. Lemma red_on_free_vars {P : nat -> bool} {Σ:global_env_ext} {Γ u v} {wfΣ : wf Σ} : diff --git a/pcuic/theories/Conversion/PCUICWeakeningConfigConv.v b/pcuic/theories/Conversion/PCUICWeakeningConfigConv.v index 7829564da..a3a9c9e6b 100644 --- a/pcuic/theories/Conversion/PCUICWeakeningConfigConv.v +++ b/pcuic/theories/Conversion/PCUICWeakeningConfigConv.v @@ -98,7 +98,7 @@ Proof. eapply All_impl; eauto. cbn. intros x [? ?] y [[[? ?] ?] ?]. repeat split; eauto. - intros h; depelim h; constructor; cbn in X; intuition eauto. - depelim e; cbn in X; constructor; intuition eauto. solve_all. + depelim o; cbn in X; constructor; intuition eauto. solve_all. Qed. Lemma weakening_config_cumul_gen {cf1 cf2} pb Σ Γ M N : diff --git a/pcuic/theories/PCUICAst.v b/pcuic/theories/PCUICAst.v index 8f03ffa5b..a9c3036cf 100644 --- a/pcuic/theories/PCUICAst.v +++ b/pcuic/theories/PCUICAst.v @@ -976,6 +976,12 @@ Proof. intros x. now rewrite (mapu_prim_compose _ _ _ _ x). Qed. #[global] Hint Rewrite @mapu_prim_compose_rew : map. +Lemma prim_val_tag_map {term term'} (p : PCUICPrimitive.prim_val term) fu (ft : term -> term') : + prim_val_tag (mapu_prim fu ft p) = prim_val_tag p. +Proof. + destruct p as [? []] => //. +Qed. + Definition tPrimProp {term} (P : term -> Type) (p : PCUICPrimitive.prim_val term) : Type := match p.π2 return Type with | primIntModel f => unit diff --git a/pcuic/theories/PCUICConfluence.v b/pcuic/theories/PCUICConfluence.v index 5427cdec7..1060284ed 100644 --- a/pcuic/theories/PCUICConfluence.v +++ b/pcuic/theories/PCUICConfluence.v @@ -605,6 +605,40 @@ Proof. eexists. split. + eapply cofix_red_body. eassumption. + constructor; assumption. + - assert (h : ∑ ll, + OnOne2 (red1 Σ Δ) (array_value arr) ll * + All2 (eq_term_upto_univ Σ' Re Re) value ll + ). + { induction X. + - destruct p as [p1 p2]. + eapply p2 in e as hh. destruct hh as [? [? ?]]. + eexists. split. + + constructor. eassumption. + + constructor. + * assumption. + * eapply All2_same. + intros. + eapply eq_term_upto_univ_refl ; eauto. + - destruct IHX as [ll [? ?]]. + eexists. split. + + eapply OnOne2_tl. eassumption. + + constructor ; eauto. + eapply eq_term_upto_univ_refl ; eauto. + } + destruct h as [? [? ?]]. + eexists. split. + + eapply array_red_val; tea. + + do 2 constructor; cbn; eauto; reflexivity. + - specialize (IHh _ e) as [v' []]. + eexists; split. + + eapply array_red_def; eauto. + + do 2 constructor; cbn; eauto. reflexivity. + eapply All2_refl; reflexivity. + - specialize (IHh _ e) as [v' []]. + eexists; split. + + eapply array_red_type; eauto. + + do 2 constructor; cbn; eauto. reflexivity. + eapply All2_refl; reflexivity. Qed. Lemma eq_context_gen_context_assumptions {eq leq Γ Δ} : @@ -1478,9 +1512,42 @@ Proof. eexists. split. + eapply cofix_red_body. eassumption. + constructor. all: eauto. + - depelim e. depelim o. + assert (h : ∑ args, + OnOne2 (red1 Σ Γ) (array_value a') args * + All2 (eq_term_upto_univ Σ' Re Re) value args + ). + { revert X a0. clear r e e0. + generalize (array_value arr), (array_value a'). + intros l l' X a. induction X in l', a |- *. + - destruct p as [p1 p2]. + dependent destruction a. + eapply p2 in e as hh ; eauto. + destruct hh as [? [? ?]]. + eexists. split. + + constructor. eassumption. + + constructor. all:eauto. + + tc. + - dependent destruction a. + destruct (IHX _ a) as [? [? ?]]. + eexists. split. + + eapply OnOne2_tl. eassumption. + + constructor. all: eauto. + } + destruct h as [? [? ?]]. + eexists. split. + + eapply array_red_val. eassumption. + + do 2 constructor. all: eauto. + - depelim e. depelim o. eapply IHh in e as [v' []]; tea; tc. + eexists; split. + + now eapply array_red_def. + + do 2 constructor; eauto. + - depelim e. depelim o. eapply IHh in e0 as [v' []]; tea; tc. + eexists; split. + + now eapply array_red_type. + + do 2 constructor; eauto. Qed. - Lemma Forall2_flip {A} (R : A -> A -> Prop) (x y : list A) : Forall2 (flip R) y x -> Forall2 R x y. Proof. @@ -2065,6 +2132,15 @@ Section RedPred. eapply pred1_refl. apply pred1_refl_gen => //. now rewrite -H; pcuic. + + - constructor; pcuic. constructor; cbn; pcuic. solve_all. + eapply OnOne2_All_mix_left in X; tea. clear a. + eapply OnOne2_All2; pcuic; simpl; + unfold on_Trel; simpl; intros; intuition auto; noconf b0; repeat inv_on_free_vars_xpredT; eauto with fvs. + + + - constructor; pcuic. constructor; cbn; pcuic. + - constructor; pcuic. constructor; cbn; pcuic. Qed. End RedPred. @@ -2234,6 +2310,11 @@ Section PredRed. eapply (All2_impl X3); unfold on_Trel; intuition auto; repeat inv_on_free_vars_xpredT; eauto with fvs. - eapply red_prod; eauto with fvs. - eapply red_evar; eauto with fvs. solve_all. + - depelim X1; try solve [repeat constructor]; eauto. + depelim X2; cbn in H0; rtoProp. + eapply red_primArray_congr; eauto. + + now eapply Universe.make_inj in e. + + solve_all. Qed. Lemma pred1_red_r_gen P Γ Γ' Δ Δ' : forall M N, @@ -3801,6 +3882,7 @@ Proof. - rewrite -(All2_length a). solve_all. apply/andP; split; eauto. len in b2. eapply b2. eauto. + - depelim o; cbn in *; solve_all. Qed. Arguments red1_ctx _ _ _ : clear implicits. diff --git a/pcuic/theories/PCUICContextConversion.v b/pcuic/theories/PCUICContextConversion.v index 5466fd492..940f7b143 100644 --- a/pcuic/theories/PCUICContextConversion.v +++ b/pcuic/theories/PCUICContextConversion.v @@ -295,6 +295,19 @@ Section ContextReduction. forward b1. eapply red1_red_ctxP_app => //. destruct b1 as [t [? ?]]. refine (existT _ {| dbody := t |} _); simpl; eauto. + - eapply forallb_All in a. eapply OnOne2_All_mix_left in X; tea. + eapply (OnOne2_exist _ (red Σ Γ')) in X. destruct X as [values' []]. + eexists; split. + eapply red_primArray_value, OnOne2_All2; tea; intuition eauto. + change (set_array_value arr values') with (set_array_value (set_array_value arr value) values'). + eapply red_primArray_value, OnOne2_All2; tea; intuition eauto. + intros; intuition eauto. + - eexists; split. + + eapply red_primArray_default; tea. + + now eapply (red_primArray_default (set_array_default arr def) x). + - eexists; split. + + eapply red_primArray_type; tea. + + now eapply (red_primArray_type (set_array_type arr ty) x). Qed. Hint Resolve red_ctx_on_free_vars : fvs. diff --git a/pcuic/theories/PCUICContextReduction.v b/pcuic/theories/PCUICContextReduction.v index a44d1be02..3b6c6814d 100644 --- a/pcuic/theories/PCUICContextReduction.v +++ b/pcuic/theories/PCUICContextReduction.v @@ -65,6 +65,8 @@ Section CtxReduction. Ltac inv_on_free_vars ::= match goal with + | [ H : is_true (on_free_vars ?P (tPrim ?p)) |- _ ] => + cbn in H; move/andP: H => [] /andP[]; intros ??? | [ H : is_true (on_free_vars ?P ?t) |- _ ] => progress cbn in H; (move/and5P: H => [] || move/and4P: H => [] || move/and3P: H => [] || move/andP: H => [] || @@ -179,6 +181,9 @@ Section CtxReduction. * relativize #|mfix0|; [erewrite on_ctx_free_vars_extend, onΓ, on_free_vars_fix_context|now len] => //. * relativize #|mfix0|; [erewrite on_ctx_free_vars_extend, onΔ, on_free_vars_fix_context|now len] => //. * now eapply red_context_app_same. + - eapply red_primArray_one_value. toAll. eapply OnOne2_All_mix_left in X; tea. solve_all. + - eapply red_primArray_default; eauto. + - eapply red_primArray_type; eauto. Qed. Lemma red_red_ctx P Γ Δ t u : diff --git a/pcuic/theories/PCUICConversion.v b/pcuic/theories/PCUICConversion.v index 6e7c65626..00074eb04 100644 --- a/pcuic/theories/PCUICConversion.v +++ b/pcuic/theories/PCUICConversion.v @@ -130,6 +130,14 @@ Section CumulSpecIsCumulAlgo. repeat split; eauto ; try reflexivity. * apply All2_reflexivity. repeat eapply Prod_reflexivity; intro x; reflexivity. * repeat split; reflexivity. + - eapply cumul_Prim. constructor; eauto; try reflexivity. + eapply OnOne2_All2; tea; intuition eauto. + + apply X0. + + reflexivity. + - eapply cumul_Prim. constructor; eauto; trea. + eapply All2_reflexivity; intro; reflexivity. + - eapply cumul_Prim; constructor; eauto; trea. + eapply All2_reflexivity; intro; reflexivity. Defined. Proposition convSpec_cumulSpec (Γ : context) (M N : term) : @@ -180,6 +188,8 @@ Section CumulSpecIsCumulAlgo. eapply All2_All_mix_left in X. 2: tea. eapply All2_impl; try exact X. cbv beta. intuition; try_with_nil. - apply cumul_mkApps; eauto. eapply cumul_CoFix. unfold tFixProp in *. eapply All2_All_mix_left in X. 2: tea. eapply All2_impl; try exact X. cbv beta. intuition; try_with_nil. + - eapply cumul_mkApps; eauto. eapply cumul_Prim. depelim o; cbn in X; constructor; intuition eauto; try try_with_nil. + solve_all; try_with_nil. Defined. Proposition eq_term_upto_univ_cumulSpec (Γ : context) {pb} M N : @@ -3565,6 +3575,8 @@ Proof using Type. rewrite map_length. eapply r0. rewrite nth_error_app_context_ge; rewrite fix_context_length; try lia. enough (#|m| + i - #|m| = i) as ->; tas; lia. + - destruct p as [? []]; cbn in X; cbn; trea. + eapply red_primArray_congr; cbn; intuition eauto; solve_all. Qed. Lemma closed_red_rel_all {Γ i body t} : @@ -3659,6 +3671,27 @@ Section MoreCongruenceLemmas. * symmetry. eassumption. Defined. +Lemma ws_cumul_pb_Prim {pb Γ p p'} : + onPrim (ws_cumul_pb Conv Σ Γ) (eq_universe Σ) p p' -> is_closed_context Γ -> + Σ;;; Γ ⊢ tPrim p ≤[pb] tPrim p'. +Proof. + intros Hp HΓ. + depelim Hp. 1-2:constructor; eauto; trea. + apply ws_cumul_pb_terms_alt in a0 as [args0 [args0' [Hargs0 Hargs0' Hargs0args0']]]. + pose proof (Hargs0_c := closed_red_terms_open_right Hargs0). + pose proof (Hargs0'_c := closed_red_terms_open_right Hargs0'). + eapply ws_cumul_pb_alt_closed in w as [def [def' []]]. + eapply ws_cumul_pb_alt_closed in w0 as [ty [ty' []]]. + eapply ws_cumul_pb_alt. + exists (tPrim (primArray; primArrayModel {| array_level := array_level a; array_default := def; array_type := ty; array_value := args0 |})). + exists (tPrim (primArray; primArrayModel {| array_level := array_level a'; array_default := def'; array_type := ty'; array_value := args0' |})). + split; eauto; pcuic; cbn; rtoProp; intuition eauto; fvs. + + eapply closed_red_terms_open_left in Hargs0. solve_all. + + eapply closed_red_terms_open_left in Hargs0'. solve_all. + + eapply red_primArray_congr; cbn; solve_all; now eapply closed_red_red. + + eapply red_primArray_congr; cbn; solve_all; now eapply closed_red_red. + + do 2 constructor; cbn; eauto. +Qed. End MoreCongruenceLemmas. @@ -3772,6 +3805,10 @@ Proof. * rewrite on_free_vars_ctx_app; solve_all. rewrite on_free_vars_fix_context; eauto; solve_all. * rewrite -> shiftnP_add, <- fix_context_length, <- app_length in *; tea. * rewrite -> shiftnP_add, <- Hfix, <- fix_context_length, <- app_length in *; tea. + - intros Γ p p' e h; cbn => H0 H1 H2. + eapply ws_cumul_pb_Prim; eauto. + depelim h; constructor; cbn in *; rtoProp; intuition eauto. + repeat toAll. solve_all. - intros Γ i u u' args args' H X X_dep H0 H1 H2. eapply ws_cumul_pb_Ind; eauto. split; eauto. rewrite on_free_vars_mkApps in H1. rewrite on_free_vars_mkApps in H2. repeat toProp; destruct_head'_and. diff --git a/pcuic/theories/PCUICCumulativitySpec.v b/pcuic/theories/PCUICCumulativitySpec.v index 12bf4aff5..0b99ec6e7 100644 --- a/pcuic/theories/PCUICCumulativitySpec.v +++ b/pcuic/theories/PCUICCumulativitySpec.v @@ -34,10 +34,10 @@ Proof. all: repeat first [ assumption | toAll ]. Qed. -Inductive eq_prim_dep (eq_term : term -> term -> Type) (Re : Universe.t -> Universe.t -> Prop) (eq_term_dep : forall x y, eq_term x y -> Type) (Re' : forall a b, Re a b -> Type) : forall x y : prim_val, eq_prim eq_term Re x y -> Type := - | eqPrimInt_dep i : eq_prim_dep eq_term Re eq_term_dep Re' (primInt; i) (primInt; i) (eqPrimInt eq_term Re i) - | eqPrimFloat_dep f : eq_prim_dep eq_term Re eq_term_dep Re' (primFloat; f) (primFloat; f) (eqPrimFloat _ _ f) - | eqPrimArray_dep a a' : +Inductive onPrim_dep (eq_term : term -> term -> Type) (Re : Universe.t -> Universe.t -> Prop) (eq_term_dep : forall x y, eq_term x y -> Type) (Re' : forall a b, Re a b -> Type) : forall x y : prim_val, onPrim eq_term Re x y -> Type := + | onPrimInt_dep i : onPrim_dep eq_term Re eq_term_dep Re' (primInt; primIntModel i) (primInt; primIntModel i) (onPrimInt eq_term Re i) + | onPrimFloat_dep f : onPrim_dep eq_term Re eq_term_dep Re' (primFloat; primFloatModel f) (primFloat; primFloatModel f) (onPrimFloat _ _ f) + | onPrimArray_dep a a' : forall (hre : Re (Universe.make a.(array_level)) (Universe.make a'.(array_level))) (eqdef : eq_term a.(array_default) a'.(array_default)) (eqty : eq_term a.(array_type) a'.(array_type)) @@ -46,8 +46,9 @@ Inductive eq_prim_dep (eq_term : term -> term -> Type) (Re : Universe.t -> Unive eq_term_dep _ _ eqdef -> eq_term_dep _ _ eqty -> All2_dep eq_term_dep eqt -> - eq_prim_dep eq_term Re eq_term_dep Re' - (primArray; primArrayModel a) (primArray; primArrayModel a') (eqPrimArray _ _ a a' hre eqdef eqty eqt). + onPrim_dep eq_term Re eq_term_dep Re' + (primArray; primArrayModel a) (primArray; primArrayModel a') (onPrimArray _ _ a a' hre eqdef eqty eqt). +Derive Signature for onPrim_dep. Reserved Notation " Σ ;;; Γ ⊢ t ≤s[ pb ] u" (at level 50, Γ, t, u at next level, format "Σ ;;; Γ ⊢ t ≤s[ pb ] u"). @@ -163,7 +164,7 @@ Inductive cumulSpec0 {cf : checker_flags} (Σ : global_env_ext) Γ (pb : conv_pb Σ ;;; Γ ⊢ tCoFix mfix idx ≤s[pb] tCoFix mfix' idx | cumul_Prim p p' : - eq_prim (fun x y => Σ ;;; Γ ⊢ x ≤s[Conv] y) (compare_universe Conv Σ) p p' -> + onPrim (fun x y => Σ ;;; Γ ⊢ x ≤s[Conv] y) (compare_universe Conv Σ) p p' -> Σ ;;; Γ ⊢ tPrim p ≤s[pb] tPrim p' (** Reductions *) @@ -444,8 +445,8 @@ Lemma cumulSpec0_rect : P cf Σ Γ pb (tCoFix mfix idx) (tCoFix mfix' idx) (cumul_CoFix _ _ _ _ _ _ H)) -> - (forall Γ pb p p' (e : eq_prim (cumulSpec0 Σ Γ Conv) (eq_universe Σ) p p'), - eq_prim_dep (cumulSpec0 Σ Γ Conv) (eq_universe Σ) (P cf Σ Γ Conv) (fun _ _ _ => True) p p' e -> + (forall Γ pb p p' (e : onPrim (cumulSpec0 Σ Γ Conv) (eq_universe Σ) p p'), + onPrim_dep (cumulSpec0 Σ Γ Conv) (eq_universe Σ) (P cf Σ Γ Conv) (fun _ _ _ => True) p p' e -> P cf Σ Γ pb (tPrim p) (tPrim p') (cumul_Prim _ _ _ _ _ e)) -> (* cumulativity rules *) @@ -527,8 +528,8 @@ Proof. + intuition auto. + auto. - eapply X20; eauto. - clear -e aux. - induction e; constructor; auto. + clear -o aux. + induction o; constructor; auto. clear -a0 aux. revert a0. induction a0; constructor; auto. - eapply X; eauto. @@ -723,8 +724,8 @@ Lemma convSpec0_ind_all : P cf Σ Γ Conv (tCoFix mfix idx) (tCoFix mfix' idx) (cumul_CoFix _ _ _ _ _ _ H)) -> - (forall Γ p p' (e : eq_prim (cumulSpec0 Σ Γ Conv) (eq_universe Σ) p p'), - eq_prim_dep (cumulSpec0 Σ Γ Conv) (eq_universe Σ) (P cf Σ Γ Conv) (fun _ _ _ => True) p p' e -> + (forall Γ p p' (e : onPrim (cumulSpec0 Σ Γ Conv) (eq_universe Σ) p p'), + onPrim_dep (cumulSpec0 Σ Γ Conv) (eq_universe Σ) (P cf Σ Γ Conv) (fun _ _ _ => True) p p' e -> P cf Σ Γ Conv (tPrim p) (tPrim p') (cumul_Prim _ _ _ _ _ e)) -> (* cumulativity rules *) diff --git a/pcuic/theories/PCUICEquality.v b/pcuic/theories/PCUICEquality.v index 5396bc44e..d4be6f3d4 100644 --- a/pcuic/theories/PCUICEquality.v +++ b/pcuic/theories/PCUICEquality.v @@ -275,16 +275,16 @@ Definition eq_predicate (eq_term : term -> term -> Type) Re p p' := ((eq_context_gen eq eq p.(pcontext) p'.(pcontext)) * eq_term p.(preturn) p'.(preturn))). -Inductive eq_prim (eq_term : term -> term -> Type) Re : prim_val -> prim_val -> Type := - | eqPrimInt i : eq_prim eq_term Re (primInt; i) (primInt; i) - | eqPrimFloat f : eq_prim eq_term Re (primFloat; f) (primFloat; f) - | eqPrimArray a a' : +Inductive onPrim (eq_term : term -> term -> Type) Re : prim_val -> prim_val -> Type := + | onPrimInt i : onPrim eq_term Re (primInt; primIntModel i) (primInt; primIntModel i) + | onPrimFloat f : onPrim eq_term Re (primFloat; primFloatModel f) (primFloat; primFloatModel f) + | onPrimArray a a' : Re (Universe.make a.(array_level)) (Universe.make a'.(array_level)) -> eq_term a.(array_default) a'.(array_default) -> eq_term a.(array_type) a'.(array_type) -> All2 eq_term a.(array_value) a'.(array_value) -> - eq_prim eq_term Re (primArray; primArrayModel a) (primArray; primArrayModel a'). -Derive Signature for eq_prim. + onPrim eq_term Re (primArray; primArrayModel a) (primArray; primArrayModel a'). +Derive Signature for onPrim. (** ** Syntactic ws_cumul_pb up-to universes We don't look at printing annotations *) @@ -379,7 +379,7 @@ Inductive eq_term_upto_univ_napp Σ (Re Rle : Universe.t -> Universe.t -> Prop) Σ ⊢ tCoFix mfix idx <==[ Rle , napp ] tCoFix mfix' idx | eq_Prim i i' : - eq_prim (eq_term_upto_univ_napp Σ Re Re 0) Re i i' -> + onPrim (eq_term_upto_univ_napp Σ Re Re 0) Re i i' -> eq_term_upto_univ_napp Σ Re Rle napp (tPrim i) (tPrim i') where " Σ ⊢ t <==[ Rle , napp ] u " := (eq_term_upto_univ_napp Σ _ Rle napp t u) : type_scope. @@ -508,6 +508,15 @@ Qed. Polymorphic Instance creflexive_eq A : CRelationClasses.Reflexive (@eq A). Proof. intro x. constructor. Qed. +Lemma onPrim_map_prop R R' Re p p' P f : tPrimProp P p -> + onPrim R Re p p' -> + (forall x y, P x -> R x y -> R' (f x) (f y)) -> + onPrim R' Re (map_prim f p) (map_prim f p'). +Proof. + destruct p as [? []]; cbn; intros h e; depelim e; intros hf; constructor; cbn; intuition eauto. + solve_all. +Qed. + #[global] Polymorphic Instance eq_predicate_refl Re Ru : CRelationClasses.Reflexive Re -> @@ -623,7 +632,7 @@ Proof. + destruct r as [[h1 h2] [[[h3 h4] h5] h6]]. eapply h1 in h3 ; auto. constructor; auto. - econstructor. - depelim e; cbn in X; constructor; intuition eauto. + depelim o; cbn in X; constructor; intuition eauto. eapply All2_All_mix_left in a0 as h; eauto. cbn in h. eapply All2_sym; solve_all. Qed. @@ -758,7 +767,7 @@ Proof. intuition eauto. transitivity (rarg y); auto. - dependent destruction e2; constructor. - depelim e; intuition eauto. depelim e1; constructor; cbn in X; intuition eauto. + depelim o; intuition eauto. depelim o0; constructor; cbn in X; intuition eauto. eapply All2_All_mix_left in b0 as h; eauto. clear b0 a0. clear -he hle a2 h. revert h a2. generalize (array_value a) (array_value a') (array_value a'0). clear -he hle. @@ -981,7 +990,7 @@ Proof. eapply All2_impl'; tea. eapply All_impl; eauto. cbn. intros x [? ?] y [[[? ?] ?] ?]. repeat split; eauto. - - intros h; depelim h. depelim e; constructor; cbn in X; constructor; intuition eauto. + - intros h; depelim h. depelim o; constructor; cbn in X; constructor; intuition eauto. solve_all. Qed. @@ -1014,7 +1023,7 @@ Proof. eapply All2_impl'; tea. eapply All_impl; eauto. cbn. intros x [? ?] y [[[? ?] ?] ?]. repeat split; eauto. - - intros h; depelim h. constructor. depelim e; cbn in X; constructor; intuition eauto. + - intros h; depelim h. constructor. depelim o; cbn in X; constructor; intuition eauto. solve_all. Qed. @@ -1071,11 +1080,8 @@ Proof. - cbn. constructor. pose proof (All2_length a). solve_all. rewrite H. eauto. - - cbn. constructor. depelim e; cbn in X; intuition eauto. - * depelim i; cbn in X; constructor; eauto. - * depelim f; cbn in X; constructor; eauto. - * constructor; cbn; eauto. - solve_all. + - cbn. constructor. depelim o; cbn in X; try constructor; cbn; intuition eauto. + solve_all. Qed. Lemma lift_compare_term `{checker_flags} pb Σ ϕ n k t t' : @@ -1154,11 +1160,8 @@ Proof. - cbn. constructor ; try sih ; eauto. pose proof (All2_length a). solve_all. now rewrite H. - - cbn; constructor. depelim e. - * cbn. depelim i; cbn in X; constructor; intuition eauto. - * cbn. depelim f; cbn in X; constructor; intuition eauto. - * cbn. depelim a; cbn in X |- *; constructor; cbn; intuition eauto. - solve_all. + - cbn; constructor. depelim o; cbn in X |- *; constructor; cbn; intuition eauto. + solve_all. Qed. Lemma eq_term_upto_univ_subst Σ Re Rle : @@ -1712,7 +1715,7 @@ Proof. now eapply eq_term_upto_univ_sym. now eapply eq_term_upto_univ_sym. now symmetry. - - depelim e; constructor; eauto. + - depelim o; constructor; eauto. now eapply eq_term_upto_univ_sym. now eapply eq_term_upto_univ_sym. eapply All2_sym; solve_all. diff --git a/pcuic/theories/PCUICParallelReduction.v b/pcuic/theories/PCUICParallelReduction.v index 3c907c2c5..74cc41989 100644 --- a/pcuic/theories/PCUICParallelReduction.v +++ b/pcuic/theories/PCUICParallelReduction.v @@ -232,7 +232,6 @@ Section ParallelReduction. | tSort _ | tInd _ _ | tConstruct _ _ _ => true - | tPrim _ => true | _ => false end. @@ -400,6 +399,10 @@ Section ParallelReduction. pred1_ctx Γ Γ' -> All2 (pred1 Γ Γ') l l' -> pred1 Γ Γ' (tEvar ev l) (tEvar ev l') + | prim_pred p p' : + pred1_ctx Γ Γ' -> + PCUICEquality.onPrim (pred1 Γ Γ') eq p p' -> pred1 Γ Γ' (tPrim p) (tPrim p') + | pred_atom_refl t : pred1_ctx Γ Γ' -> pred_atom t -> @@ -609,6 +612,13 @@ Section ParallelReduction. pred1_ctx Γ Γ' -> Pctx Γ Γ' -> All2 (P' Γ Γ') l l' -> P Γ Γ' (tEvar ev l) (tEvar ev l')) -> + + (forall (Γ Γ' : context) (p p' : prim_val), + pred1_ctx Γ Γ' -> Pctx Γ Γ' -> + PCUICEquality.onPrim (pred1 Γ Γ') eq p p' -> + PCUICEquality.onPrim (P Γ Γ') eq p p' -> + P Γ Γ' (tPrim p) (tPrim p')) -> + (forall (Γ Γ' : context) (t : term), pred1_ctx Γ Γ' -> Pctx Γ Γ' -> @@ -619,7 +629,7 @@ Section ParallelReduction. intros P Pctx Pctxover P' Hctx Hctxover. intros. assert (forall (Γ Γ' : context) (t t0 : term), pred1 Γ Γ' t t0 -> P Γ Γ' t t0). intros. - rename X20 into pr. revert Γ Γ' t t0 pr. + rename X21 into pr. revert Γ Γ' t t0 pr. fix aux 5. intros Γ Γ' t t'. move aux at top. destruct 1; match goal with @@ -634,13 +644,13 @@ Section ParallelReduction. | H : _ |- _ => eapply H; eauto end. - simpl. apply X1; auto. - clear X1 X2 X3 X4 X5 X6 X7 X8 X9 X10 X11 X12 X13 X14 X15 X16 X17 X18 X19. + clear X1 X2 X3 X4 X5 X6 X7 X8 X9 X10 X11 X12 X13 X14 X15 X16 X17 X18 X19 X20. apply Hctx. apply (on_contexts_impl a). intros. eapply X1. apply (on_contexts_impl a). intros. eapply (aux _ _ _ _ X1). - simpl. apply X2; auto. - apply Hctx, (on_contexts_impl a). exact a. intros. apply (aux _ _ _ _ X20). - - apply Hctx, (on_contexts_impl a). exact a. intros. apply (aux _ _ _ _ X20). + apply Hctx, (on_contexts_impl a). exact a. intros. apply (aux _ _ _ _ X21). + - apply Hctx, (on_contexts_impl a). exact a. intros. apply (aux _ _ _ _ X21). - eapply (All2_All2_prop (P:=pred1) (Q:=P') a0 ((extendP aux) Γ Γ')). - eapply (All2_All2_prop a1 (extendP aux Γ Γ')). - eapply (All2_branch_prop @@ -707,7 +717,7 @@ Section ParallelReduction. - eapply X8; eauto. apply (Hctx _ _ a), (on_contexts_impl a aux). - apply (Hctx _ _ a), (on_contexts_impl a aux). - - eapply (All2_All2_prop (P:=pred1) (Q:=P) a0). intros. apply (aux _ _ _ _ X20). + - eapply (All2_All2_prop (P:=pred1) (Q:=P) a0). intros. apply (aux _ _ _ _ X21). - apply (Hctx _ _ a), (on_contexts_impl a aux). - apply (All2_All2_prop (P:=pred1) (Q:=P') a0 (extendP aux _ _)). - apply (Hctxover _ _ _ _ a (on_contexts_impl a aux) @@ -743,13 +753,15 @@ Section ParallelReduction. - eapply (Hctx _ _ a), (on_contexts_impl a aux). - eapply (All2_All2_prop (P:=pred1) (Q:=P') a0 (extendP aux Γ Γ')). - eapply (Hctx _ _ a), (on_contexts_impl a aux). + - destruct o; constructor; eauto. eapply (All2_All2_prop a1 (aux Γ Γ')). + - eapply (Hctx _ _ a), (on_contexts_impl a aux). - split => //. intros. eapply Hctxover; eauto. - { eapply (on_contexts_impl X21 X20). } + { eapply (on_contexts_impl X22 X21). } eapply Hctx; eauto. - { eapply (on_contexts_impl X21 X20). } - eapply (on_contexts_impl X22 (extend_over X20 Γ Γ')). + { eapply (on_contexts_impl X22 X21). } + eapply (on_contexts_impl X23 (extend_over X21 Γ Γ')). Defined. Lemma pred1_pred1_ctx {Γ Δ t u} : pred1 Γ Δ t u -> pred1_ctx Γ Δ. @@ -757,7 +769,7 @@ Section ParallelReduction. intros H; revert Γ Δ t u H. refine (fst (pred1_ind_all_ctx _ (fun Γ Γ' => pred1_ctx Γ Γ') (fun Γ Γ' Δ Δ' => True) - _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _)); intros *. + _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _)); intros *. all:try intros **; rename_all_hyps; try solve [specialize (forall_Γ _ X3); eauto]; eauto; try solve [eexists; split; constructor; eauto]. @@ -839,7 +851,10 @@ Section ParallelReduction. eapply a; tas. eapply b. eapply on_contexts_app; auto. now eapply onctx_rel_pred1_refl. - Qed. + - constructor; eauto. + destruct p as [? []]; cbn in X; intuition eauto; econstructor; eauto. + solve_all. + Qed. Lemma pred1_ctx_refl Γ : pred1_ctx Γ Γ. Proof using. @@ -1079,7 +1094,7 @@ Qed. Proof using cf. set (Pctx := fun (Γ Γ' : context) => pred1_ctx Σ Γ Γ'). - refine (pred1_ind_all_ctx Σ _ Pctx _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _); intros *; intros; + refine (pred1_ind_all_ctx Σ _ Pctx _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _); intros *; intros; subst Pctx; rename_all_hyps; try subst Γ Γ'; lazymatch goal with @@ -1411,6 +1426,9 @@ Qed. repeat (constructor; eauto). 1-2:now eapply urenaming_vass. - econstructor; tea. solve_all. + - econstructor; tea. + depelim X1; depelim X2; cbn; constructor; cbn; eauto. + all:cbn in H; rtoProp; intuition eauto. solve_all. - destruct t => //; constructor; auto. Qed. @@ -1994,7 +2012,7 @@ Section ParallelSubstitution. Proof. intros Pover. refine (pred1_ind_all_ctx Σ _ (fun Γ Γ' => pred1_ctx Σ Γ Γ') - _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _); subst Pover; + _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _); subst Pover; try (intros until Δ; intros Δ' σ τ ons onis Hrel); trivial; repeat inv_on_free_vars. (* induction redst using ; sigma; intros Δ Δ' σ τ Hσ Hτ Hrel. *) @@ -2453,6 +2471,9 @@ Section ParallelSubstitution. - sigma. simpl. constructor; auto with pcuic. solve_all. + (* Primitive *) + - cbn in ons, onis. depelim X1; cbn; depelim X2; cbn in *; constructor; + eauto; try eapply Hrel; constructor; rtoProp; cbn; intuition eauto. solve_all. - rewrite !pred_atom_inst; auto. eapply pred1_refl_gen; auto with pcuic. Qed. diff --git a/pcuic/theories/PCUICParallelReductionConfluence.v b/pcuic/theories/PCUICParallelReductionConfluence.v index b0ea24b6b..87d4c9c5c 100644 --- a/pcuic/theories/PCUICParallelReductionConfluence.v +++ b/pcuic/theories/PCUICParallelReductionConfluence.v @@ -498,6 +498,32 @@ Section Rho. Qed. Hint Rewrite @rho_predicate_map_predicate : rho. + Equations? map_prim_wf (p : PCUICPrimitive.prim_val term) (rho : context -> forall x, depth x < depth (tPrim p) -> term) (Γ : context) : prim_val := + | (primInt; primIntModel i), _, _ := (primInt; primIntModel i); + | (primFloat; primFloatModel f), _, _ := (primFloat; primFloatModel f); + | (primArray; primArrayModel a), rho, Γ := + let default := rho Γ a.(array_default) _ in + let ty := rho Γ a.(array_type) _ in + let value := map_terms rho Γ a.(array_value) _ in + let a' := {| + array_level := array_level a; + array_type := ty; + array_default := default; + array_value := value + |} + in (primArray; primArrayModel a'). + Proof. all:cbn; clear; try lia. Qed. + + Lemma rho_prim_wf_rho_prim p rho Γ : + map_prim_wf p (fun ctx x Hx => rho ctx x) Γ = + map_prim (rho Γ) p. + Proof. + funelim (map_prim_wf _ _ _); cbn => //. + do 3 f_equal. destruct a; unfold map_array_model; cbn; f_equal. + apply map_In_spec. + Qed. + Hint Rewrite @rho_prim_wf_rho_prim : rho. + Definition inspect {A} (x : A) : { y : A | x = y } := exist x eq_refl. Arguments Nat.max : simpl never. @@ -635,6 +661,7 @@ Section Rho. rho Γ (tCoFix mfix idx) => let mfixctx := fold_fix_context_wf mfix (fun Γ x Hx => rho Γ x) Γ [] in tCoFix (map_fix_rho (t:=tCoFix mfix idx) rho Γ mfixctx mfix _) idx; + rho Γ (tPrim p) => tPrim (map_prim_wf p rho Γ); rho Γ x => x. Proof. all:try abstract lia. @@ -1906,6 +1933,9 @@ Section Rho. eapply shift_renaming; auto. clear. generalize #|m|. induction m using rev_ind. simpl. constructor. intros. rewrite map_app !fold_fix_context_app. simpl. constructor. simpl. reflexivity. apply IHm; eauto. + + - (* Prim *) + simpl; simp rho. cbn. f_equal. cbn in ont. solve_all. Qed. Lemma rho_lift0 Γ Δ P t : @@ -2388,6 +2418,7 @@ Section Rho. - eapply All2_prod_inv in X0 as [? X0]. eapply (All2_apply P), (All2_apply_arrow H0) in X0. solve_all. + - depelim X1; cbn in H1 |- *; solve_all. Qed. Lemma pred1_on_free_vars {P Γ Γ' t u} : @@ -3278,7 +3309,7 @@ Section Rho. set Pctx := fun (Γ Δ : context) => on_ctx_free_vars xpredT Γ -> pred1_ctx Σ Δ (rho_ctx Γ). - refine (pred1_ind_all_ctx Σ _ Pctx _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _); + refine (pred1_ind_all_ctx Σ _ Pctx _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _); subst Pctx; intros *. all:try intros **; rename_all_hyps; try solve [specialize (forall_Γ _ X3); eauto]; eauto; @@ -4067,6 +4098,9 @@ Section Rho. - simpl in *. simp rho. constructor. eauto. eapply All2_All_mix_left in X1; tea. eapply All2_sym, All2_map_left, All2_impl; tea => /=; intuition auto. + - cbn in *. simp rho. constructor; eauto. + depelim X2; constructor; cbn in *; rtoProp; intuition eauto. solve_all. + eapply All2_sym. solve_all. - destruct t; noconf H; simpl; constructor; eauto. Qed. diff --git a/pcuic/theories/PCUICReduction.v b/pcuic/theories/PCUICReduction.v index 92f19547c..d0dda059b 100644 --- a/pcuic/theories/PCUICReduction.v +++ b/pcuic/theories/PCUICReduction.v @@ -1923,6 +1923,93 @@ Section ReductionCongruence. - eapply red_cofix_ty. assumption. Qed. + Lemma red_primArray_one_value (arr : array_model term) (value : list term) : + OnOne2 (red Σ Γ) (array_value arr) value -> + Σ;;; Γ |- tPrim (primArray; primArrayModel arr) ⇝* tPrim (primArray; primArrayModel (set_array_value arr value)). + Proof. + intros h. + apply OnOne2_on_Trel_eq_unit in h. + apply OnOne2_on_Trel_eq_red_redl in h. + dependent induction h. + - assert (arr.(array_value) = value). + { eapply map_inj ; eauto. + intros y z e. cbn in e. inversion e. eauto. + } subst. + destruct arr; reflexivity. + - set (f := fun x : term => (x, tt)) in *. + set (g := (fun '(x, _) => x) : term × unit -> term). + assert (el : forall l, l = map f (map g l)). + { clear. intros l. induction l. + - reflexivity. + - cbn. destruct a, u. cbn. f_equal. assumption. + } + assert (el' : forall l, l = map g (map f l)). + { clear. intros l. induction l. + - reflexivity. + - cbn. f_equal. assumption. + } + eapply trans_red. + + eapply IHh; tas. symmetry. apply el. + + change (set_array_value arr value) with + (set_array_value (set_array_value arr (map g l1)) value). + econstructor. rewrite (el' value). + eapply OnOne2_map. + eapply OnOne2_impl ; eauto. + intros [? []] [? []] [h1 h2]. + unfold on_Trel in h1, h2. cbn in *. + unfold on_Trel. cbn. assumption. + Qed. + + Lemma red_primArray_value (arr : array_model term) (value : list term) : + All2 (red Σ Γ) (array_value arr) value -> + Σ;;; Γ |- tPrim (primArray; primArrayModel arr) ⇝* tPrim (primArray; primArrayModel (set_array_value arr value)). + Proof using Type. + intros h. + apply All2_many_OnOne2 in h. + induction h. + - destruct arr; reflexivity. + - eapply red_trans. + + eapply IHh. + + assert (set_array_value arr z = set_array_value (set_array_value arr y) z) as ->. + { now destruct arr. } + eapply red_primArray_one_value; eassumption. + Qed. + + Lemma red_primArray_default (arr : array_model term) (value : term) : + red Σ Γ (array_default arr) value -> + Σ;;; Γ |- tPrim (primArray; primArrayModel arr) ⇝* + tPrim (primArray; primArrayModel (set_array_default arr value)). + Proof using Type. + intros h. + destruct arr; cbn in *. + rst_induction h; eauto with pcuic. + Qed. + + Lemma red_primArray_type (arr : array_model term) (value : term) : + red Σ Γ (array_type arr) value -> + Σ;;; Γ |- tPrim (primArray; primArrayModel arr) ⇝* + tPrim (primArray; primArrayModel (set_array_type arr value)). + Proof using Type. + intros h. + destruct arr; cbn in *. + rst_induction h; eauto with pcuic. + Qed. + + Lemma red_primArray_congr (arr arr' : array_model term) : + array_level arr = array_level arr' -> + All2 (red Σ Γ) (array_value arr) (array_value arr') -> + red Σ Γ (array_default arr) (array_default arr') -> + red Σ Γ (array_type arr) (array_type arr') -> + Σ;;; Γ |- tPrim (primArray; primArrayModel arr) ⇝* + tPrim (primArray; primArrayModel arr'). + Proof using Type. + destruct arr, arr'; cbn in *. + intros <- h h0 h1. + eapply red_trans; [eapply red_primArray_value; tea|]. + eapply red_trans; [eapply red_primArray_default; tea|]. + now eapply red_trans; [eapply red_primArray_type; tea|]. + Qed. + Lemma red_prod_l : forall na A B A', red Σ Γ A A' -> diff --git a/pcuic/theories/PCUICSubstitution.v b/pcuic/theories/PCUICSubstitution.v index f96baba39..71e820a71 100644 --- a/pcuic/theories/PCUICSubstitution.v +++ b/pcuic/theories/PCUICSubstitution.v @@ -1608,6 +1608,12 @@ Qed. rewrite onΓ. apply on_free_vars_fix_context. rewrite /test_def. solve_all. now len in b2. } now rewrite -Nat.add_assoc addnP_shiftnP_k. + - destruct pr as [? []] => //. + eapply red_primArray_congr => //. + * cbn. solve_all. cbn in X1. destruct X1 as [? []]. + eapply All_All2; tea; simpl; solve_all. eapply b1; tea; solve_all. + * cbn in X, X0. intuition eauto. + * cbn in X, X0; intuition eauto. Qed. Lemma untyped_substitution_red {Γ Δ Γ' s M N} : diff --git a/pcuic/theories/PCUICTyping.v b/pcuic/theories/PCUICTyping.v index a5844cdb5..456d91eef 100644 --- a/pcuic/theories/PCUICTyping.v +++ b/pcuic/theories/PCUICTyping.v @@ -176,8 +176,8 @@ Variant case_branch_typing `{checker_flags} wf_local_fun typing Σ Γ (ci:case_i Variant primitive_typing_hyps `{checker_flags} (typing : forall (Σ : global_env_ext) (Γ : context), term -> term -> Type) Σ Γ : prim_val term -> Type := -| prim_int_hyps i : primitive_typing_hyps typing Σ Γ (primInt; i) -| prim_float_hyps f : primitive_typing_hyps typing Σ Γ (primFloat; f) +| prim_int_hyps i : primitive_typing_hyps typing Σ Γ (primInt; primIntModel i) +| prim_float_hyps f : primitive_typing_hyps typing Σ Γ (primFloat; primFloatModel f) | prim_array_hyps a (wfl : wf_universe Σ (Universe.make a.(array_level))) (hty : typing Σ Γ a.(array_type) (tSort (Universe.make a.(array_level)))) diff --git a/pcuic/theories/PCUICWfUniverses.v b/pcuic/theories/PCUICWfUniverses.v index 7ccce33e9..bc413d4da 100644 --- a/pcuic/theories/PCUICWfUniverses.v +++ b/pcuic/theories/PCUICWfUniverses.v @@ -14,6 +14,28 @@ Require Import ssreflect ssrbool. From MetaCoq.PCUIC Require Import PCUICInduction. + +Lemma test_primu_mapu pu pt fu ft p : + test_primu pu pt (mapu_prim fu ft p) = + test_primu (pu ∘ fu) (pt ∘ ft) p. +Proof. + destruct p as [? []] => //=. cbn. + now rewrite forallb_map. +Qed. + +Lemma test_primu_primProp P pu pt pu' pt' p : + tPrimProp P p -> + (forall x, pu x = pu' x) -> + (forall x, P x -> pt x = pt' x) -> + test_primu pu pt p = test_primu pu' pt' p. +Proof. + destruct p as [? []] => //=. cbn. + intuition eauto. do 3 f_equal; eauto. + solve_all. +Qed. + + + Section CheckerFlags. Context {cf:checker_flags}. @@ -297,6 +319,7 @@ Section CheckerFlags. | tConst _ u | tInd _ u | tConstruct _ _ u => forallb fu (map Universe.make u) | tEvar _ args => forallb (on_universes fu fc) args + | tPrim p => test_primu (fun x => fu (Universe.make x)) (on_universes fu fc) p | _ => true end. @@ -355,6 +378,7 @@ Qed. - rewrite forallb_map. eapply All_forallb_eq_forallb; eauto. simpl; intros []. simpl. intros. cbn. now rewrite H. + - rewrite test_primu_mapu; eapply test_primu_primProp; tea; eauto. Qed. Corollary wf_universes_lift n k t : wf_universes (lift n k t) = wf_universes t. @@ -389,6 +413,7 @@ Qed. - rewrite forallb_map. eapply All_forallb_eq_forallb; eauto. simpl; intros []. simpl. intros. cbn. now rewrite H. + - rewrite test_primu_mapu; eapply test_primu_primProp; tea; eauto. Qed. Corollary wf_universes_subst s k t : @@ -424,7 +449,7 @@ Qed. induction t using term_forall_list_ind; simpl in *; auto; try to_prop; try apply /andP; to_wfu; intuition eauto 4. - all:cbn in * ; autorewrite with map; repeat (f_equal; solve_all). + all:cbn -[Universe.make] in * ; autorewrite with map; repeat (f_equal; solve_all). - to_wfu. destruct Σ as [Σ univs']. simpl in *. eapply (wf_universe_subst_instance_univ (Σ, univs)); auto. @@ -468,6 +493,8 @@ Qed. move/andP: a => [] tctx wfu. split; auto. simpl. solve_all. now len. + - rewrite -subst_instance_univ_make. to_wfu. + eapply (wf_universe_subst_instance_univ (Σ.1, univs)) => //. Qed. Lemma weaken_wf_universe Σ Σ' t : wf Σ' -> extends Σ.1 Σ' -> @@ -507,6 +534,21 @@ Qed. intros. now eapply weaken_wf_universe_level. Qed. + Arguments Universe.make : simpl never. +Lemma test_primu_test_primu_tPrimProp {P : term -> Type} {pu put} {pu' : Level.t -> bool} {put' : term -> bool} p : + tPrimProp P p -> test_primu pu put p -> + (forall u, pu u -> pu' u) -> + (forall t, P t -> put t -> put' t) -> + test_primu pu' put' p. +Proof. + intros hp ht hf hg. + destruct p as [? []]; cbn => //. + destruct a; destruct hp; cbn in *. + rtoProp. destruct p0. intuition eauto. + eapply forallb_All in H2. eapply All_prod in a; tea. + eapply All_forallb, All_impl; tea. intuition eauto. apply hg; intuition auto. +Qed. + Lemma weaken_wf_universes {Σ : global_env_ext} Σ' t : wf Σ -> wf Σ' -> extends Σ.1 Σ' -> wf_universes Σ t -> wf_universes (Σ', Σ.2) t. @@ -532,6 +574,8 @@ Qed. intuition. - red in X; solve_all. - red in X. solve_all. + - eapply test_primu_test_primu_tPrimProp; tea; cbn; eauto. + intros; to_wfu; now eapply weaken_wf_universe. Qed. Lemma wf_universes_weaken_full : weaken_env_prop_full cumulSpec0 (lift_typing typing) (fun Σ Γ t T => @@ -782,6 +826,7 @@ Qed. now rewrite b. - rewrite /test_def /map_def /=. now rewrite a b. - rewrite /test_def /map_def /=. now rewrite a b. + - rewrite test_primu_mapu; eapply test_primu_primProp; tea; eauto. Qed. Ltac try_hyp := @@ -920,6 +965,12 @@ Qed. now eapply wf_universe_level_closed. Qed. + Lemma wf_universe_make Σ u : wf_universe Σ (Universe.make u) -> wf_universe_level Σ u. + Proof. + rewrite /wf_universe /= => hl; rewrite /wf_universe_level. + apply (hl (u, 0)). lsets. + Qed. + Lemma wf_universes_closedu {Σ : global_env} {wfΣ : wf Σ} {univs t} : on_udecl_prop Σ univs -> wf_universes (Σ, univs) t -> closedu #|polymorphic_instance univs| t. @@ -947,6 +998,9 @@ Qed. - unfold test_branch in *; solve_all. - unfold test_def in *; solve_all. - unfold test_def in *; solve_all. + - eapply test_primu_test_primu_tPrimProp; tea; cbn; eauto. + intros. to_wfu. eapply wf_universe_level_closed; tea. + now apply wf_universe_make. Qed. Lemma wf_ctx_universes_closed {Σ} {wfΣ : wf Σ} {univs ctx} : @@ -1200,6 +1254,12 @@ Qed. 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 _]]]. + + - apply/andP; split; eauto. + destruct X2; cbn => //. + rtoProp; intuition eauto. solve_all. + destruct X2; cbn => //. + simp prim_type. cbn. rtoProp; intuition. Qed. Lemma typing_wf_universes {Σ : global_env_ext} {Γ t T} : diff --git a/pcuic/theories/Syntax/PCUICDepth.v b/pcuic/theories/Syntax/PCUICDepth.v index 992607df5..71d941a11 100644 --- a/pcuic/theories/Syntax/PCUICDepth.v +++ b/pcuic/theories/Syntax/PCUICDepth.v @@ -38,6 +38,14 @@ Definition predicate_depth_gen (depth : term -> nat) (p : PCUICAst.predicate ter let pctxd := context_depth_gen depth p.(pcontext) in max (pard + pctxd) (depth p.(preturn)). +Definition prim_depth_gen (depth : term -> nat) (p : prim_val) := + match p.π2 with + | primArrayModel arr => + let vald := list_depth_gen depth arr.(array_value) in + max (depth arr.(array_default)) (max (depth arr.(array_type)) vald) + | _ => 1 + end. + Fixpoint depth t : nat := match t with | tRel i => 1 @@ -51,12 +59,14 @@ Fixpoint depth t : nat := | tProj p c => S (depth c) | tFix mfix idx => S (mfixpoint_depth_gen depth mfix) | tCoFix mfix idx => S (mfixpoint_depth_gen depth mfix) + | tPrim p => S (prim_depth_gen depth p) | _ => 1 end. Notation context_depth := (context_depth_gen depth). Notation list_depth := (list_depth_gen depth). Notation mfixpoint_depth := (mfixpoint_depth_gen depth). +Notation prim_depth := (prim_depth_gen depth). Lemma depth_mkApps f l : max (depth f) (list_depth l) <= depth (mkApps f l). Proof. @@ -134,6 +144,8 @@ Proof. apply map_depth_hom. intros. simpl. destruct x. simpl. unfold def_depth_gen. simpl. f_equal; apply aux. + - f_equal. destruct prim as [? []]; cbn; auto. + rewrite !aux map_depth_hom //. Qed. Lemma All_depth {s l} k : @@ -187,6 +199,9 @@ Proof. specialize (l0 k). specialize (l1 (n' + k)). simpl in l1. specialize (IHX k n'). lia. + - destruct p as [? []]; cbn in X; cbn; eauto; try lia. + destruct X as [? []]. specialize (l k); specialize (l0 k). + eapply (All_depth k) in a0. lia. Qed. Lemma depth_subst_instance u t : depth (subst_instance u t) = depth t. @@ -215,6 +230,8 @@ Proof. apply map_depth_hom. intros. simpl. destruct x. simpl. unfold def_depth_gen. simpl. f_equal; apply aux. + - f_equal. destruct prim as [? []]; cbn; auto. + rewrite !map_depth_hom // !aux //. Qed. Lemma depth_subst_decl s k d : @@ -334,7 +351,7 @@ Lemma term_forall_ctx_list_ind : (forall Γ (m : mfixpoint term) (n : nat), All_local_env (PCUICInduction.on_local_decl (fun Γ' t => P (Γ ,,, Γ') t)) (fix_context m) -> tFixProp (P Γ) (P (Γ ,,, fix_context m)) m -> P Γ (tCoFix m n)) -> - (forall Γ p, P Γ (tPrim p)) -> + (forall Γ p, tPrimProp (P Γ) p -> P Γ (tPrim p)) -> forall Γ (t : term), P Γ t. Proof. intros ????????????????? Γ t. @@ -417,6 +434,12 @@ Proof. - eapply X13; try (apply aux; red; simpl; lia). apply auxΓ => //. simpl. specialize (H mfix). lia. red. apply All_pair. split; apply auxl; simpl; auto. + + - eapply X14. + destruct prim as [? []]; cbn => //. + split; [apply aux; red; cbn; lia|]. + split; [apply aux; red; cbn; lia|]. + apply auxl; cbn. change (fun x => depth x) with depth; lia. Defined. Definition CasePredProp_depth (P : term -> Type) (p : predicate term) := @@ -455,7 +478,7 @@ Lemma term_ind_depth_app : (forall (m : mfixpoint term) (n : nat), onctx P (fix_context m) -> tFixProp P P m -> P (tCoFix m n)) -> - (forall p, P (tPrim p)) -> + (forall p, tPrimProp P p -> P (tPrim p)) -> forall (t : term), P t. Proof. intros ????????????????? t. @@ -537,4 +560,8 @@ Proof. - eapply X13; try (apply aux; red; simpl; lia). apply auxΓ => //. simpl. specialize (H mfix). lia. red. apply All_pair. split; apply auxl; simpl; auto. + + - eapply X14; cbn. destruct prim as [? []]; cbn => //. + do 2 (split; [eapply aux; cbn; lia|]). + eapply auxl. change (fun x => depth x) with depth. cbn. lia. Defined. diff --git a/pcuic/theories/Syntax/PCUICNamelessDef.v b/pcuic/theories/Syntax/PCUICNamelessDef.v index ce7dd06b0..6f7482eff 100644 --- a/pcuic/theories/Syntax/PCUICNamelessDef.v +++ b/pcuic/theories/Syntax/PCUICNamelessDef.v @@ -58,7 +58,7 @@ Fixpoint nameless (t : term) : bool := | tCoFix mfix idx => forallb (fun d => banon d.(dname)) mfix && forallb (test_def nameless nameless) mfix - | tPrim _ => true + | tPrim p => test_prim nameless p end. Notation nameless_ctx := (forallb (nameless_decl nameless)). @@ -106,7 +106,7 @@ Fixpoint nl (t : term) : term := | tProj p c => tProj p (nl c) | tFix mfix idx => tFix (map (map_def_anon nl nl) mfix) idx | tCoFix mfix idx => tCoFix (map (map_def_anon nl nl) mfix) idx - | tPrim p => tPrim p + | tPrim p => tPrim (map_prim nl p) end. Definition nlctx (Γ : context) : context := diff --git a/pcuic/theories/Syntax/PCUICOnFreeVars.v b/pcuic/theories/Syntax/PCUICOnFreeVars.v index 95a4f2d4b..ac499f54f 100644 --- a/pcuic/theories/Syntax/PCUICOnFreeVars.v +++ b/pcuic/theories/Syntax/PCUICOnFreeVars.v @@ -1384,7 +1384,7 @@ Lemma term_on_free_vars_ind : (forall p (m : mfixpoint term) (i : nat), tFixProp (on_free_vars p) (on_free_vars (shiftnP #|fix_context m| p)) m -> tFixProp (P p) (P (shiftnP #|fix_context m| p)) m -> P p (tCoFix m i)) -> - (forall p pr, P p (tPrim pr)) -> + (forall p pr, tPrimProp (on_free_vars p) pr -> tPrimProp (P p) pr -> P p (tPrim pr)) -> forall p (t : term), on_free_vars p t -> P p t. Proof. intros until t. revert p t. @@ -1475,6 +1475,20 @@ Proof. move=> n /= /andP[] /andP[] clb clty clmfix; constructor. * split => //; apply auxt => //. * now apply auxm. + + - destruct prim as [? []]; cbn => //. cbn in clt. + rtoProp. split => //. split => //. + revert H. generalize (array_value a). + fix auxm 1; destruct l; constructor. + * now move/andP: H. + * apply auxm. now move/andP: H. + + - destruct prim as [? []]; cbn => //. cbn in clt. + rtoProp. split => //; eauto. split; eauto. + move: (array_value a) H. + fix auxm 1; destruct array_value; constructor; eauto. + * eapply auxt. now move/andP: H. + * eapply auxm; now move/andP: H. Defined. Lemma alpha_eq_on_free_vars P (Γ Δ : context) : diff --git a/pcuic/theories/Typing/PCUICInstTyp.v b/pcuic/theories/Typing/PCUICInstTyp.v index 71bcbadc4..a5ba61b4d 100644 --- a/pcuic/theories/Typing/PCUICInstTyp.v +++ b/pcuic/theories/Typing/PCUICInstTyp.v @@ -26,7 +26,7 @@ Definition well_subst_usubst {cf} (Σ:global_env_ext) (wfΣ : wf Σ) Γ σ Δ : Σ ;;; Δ ⊢ σ : Γ -> usubst Γ σ Δ. Proof. - intuition. + intuition eauto with *. Defined. Definition well_subst_closed_subst {cf} (Σ:global_env_ext) (wfΣ : wf Σ) Γ σ Δ : @@ -326,6 +326,7 @@ Proof. rewrite <- shiftnP_add. rewrite fix_context_length. rewrite (All2_length X). eauto. + - cbn. eapply cumul_Prim. depelim X; cbn in HfreeA, HfreeB; rtoProp; constructor; cbn; eauto. solve_all. - cbn. repeat rewrite inst_mkApps. eapply cumul_Ind. * repeat rewrite map_length; eauto. * repeat toAll. @@ -353,6 +354,11 @@ Proof. apply inst_cumulSpec. Qed. +Lemma inst_prim_type σ p pty : (prim_type p pty).[σ] = prim_type (map_prim (inst σ) p) pty. +Proof. + destruct p as [? []] => //. +Qed. + Lemma type_inst {cf : checker_flags} : env_prop (fun Σ Γ t A => forall Δ σ, @@ -574,8 +580,10 @@ Proof. * now apply inst_wf_cofixpoint. * reflexivity. - - intros Σ wfΣ Γ wfΓ p pty cdecl _ hp hdecl pinv Δ σ hΔ hσ. - cbn. econstructor; tea. + - intros Σ wfΣ Γ wfΓ p pty cdecl _ hp hdecl pinv hty hind Δ σ hΔ hσ. + cbn. rewrite inst_prim_type. econstructor; tea. + 1-2:now rewrite prim_val_tag_map. + depelim hind; constructor; cbn; eauto. solve_all. - intros Σ wfΣ Γ wfΓ t A B X hwf ht iht hB ihB hcum Δ σ hΔ hσ. eapply type_Cumul. diff --git a/pcuic/theories/Typing/PCUICRenameTyp.v b/pcuic/theories/Typing/PCUICRenameTyp.v index 5e43b24b6..ce8cd6246 100644 --- a/pcuic/theories/Typing/PCUICRenameTyp.v +++ b/pcuic/theories/Typing/PCUICRenameTyp.v @@ -177,7 +177,7 @@ Proof. - rename Hren into hf. unfold urenaming in hf. destruct (nth_error Γ i) eqn:hnth; noconf pf. assert (hav : sP i). - { unfold sP, shiftnP in *. cbn in *. rewrite orb_false_r in HfreeB. intuition. } + { unfold sP, shiftnP in *. cbn in *. rewrite orb_false_r in HfreeB. intuition auto with *. } clear HfreeB. specialize hf with (1 := hav) (2 := hnth). destruct hf as [decl' [e' [eqann [hr hbo]]]]. @@ -419,6 +419,7 @@ Proof. rewrite <- fix_context_length. eapply urename_on_free_vars_shift; eauto. rewrite fix_context_length; eauto. + - eapply cumul_Prim. depelim X; cbn; cbn in HfreeB, HΓ; rtoProp; constructor; cbn; eauto. solve_all. - repeat rewrite rename_mkApps. eapply cumul_Ind. * repeat rewrite map_length; eauto. * inv_on_free_vars. @@ -789,6 +790,11 @@ Qed. Hint Rewrite <- Upn_ren : sigma. +Lemma rename_prim_type f p pty : rename f (prim_type p pty) = prim_type (map_prim (rename f) p) pty. +Proof. + destruct p as [? []]; cbn; simp prim_type => //. +Qed. + (** For an unconditional renaming defined on all variables in the source context *) Lemma typing_rename_prop : env_prop (fun Σ Γ t A => @@ -1046,8 +1052,12 @@ Proof. * now eapply rename_wf_cofixpoint. + reflexivity. - - intros Σ wfΣ Γ wfΓ p pty cdecl _ hp hdecl pinv P Δ f hf. - cbn. econstructor; tea. apply hf. + - intros Σ wfΣ Γ wfΓ p pty cdecl _ hp hdecl pinv ptyp IH P Δ f hf. + cbn. rewrite rename_prim_type /=. econstructor; tea. + * apply hf. + * rewrite prim_val_tag_map //. + * rewrite prim_val_tag_map //. + * depelim IH; eauto; cbn; constructor; cbn; eauto. solve_all. - intros Σ wfΣ Γ wfΓ t A B X hwf ht iht htB ihB cum P Δ f hf. eapply type_Cumul. diff --git a/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v b/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v index f61e294b4..6e5d1da16 100644 --- a/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v +++ b/pcuic/theories/Typing/PCUICUnivSubstitutionTyp.v @@ -95,14 +95,14 @@ Proof. - rewrite subst_instance_two. solve [econstructor; eauto]. - rewrite !subst_instance_mkApps. eapply cumul_proj. now rewrite nth_error_map Hargs. - - eapply cumul_Trans; intuition. + - eapply cumul_Trans; intuition eauto. * rewrite on_free_vars_ctx_subst_instance; eauto. * rewrite on_free_vars_subst_instance. unfold is_open_term. replace #|Γ@[u]| with #|Γ|; eauto. rewrite map_length; eauto. - eapply cumul_Evar. eapply All2_map. repeat toAll. eapply All2_impl. 1: tea. cbn; intros. eapply X0.2; eauto. - - eapply cumul_Case; try solve [intuition; eauto]. + - eapply cumul_Case; try solve [intuition eauto; eauto]. * cbv [cumul_predicate] in *; destruct_head'_prod. repeat split; eauto; cbn. + apply All2_map. repeat toAll. eapply All2_impl. 1: tea. cbn; intros. destruct_head'_prod; eauto. + apply precompose_subst_instance. eapply R_universe_instance_impl. @@ -111,21 +111,18 @@ Proof. + rewrite -> subst_instance_app, inst_case_predicate_context_subst_instance in *. eauto. * eapply All2_map. repeat toAll. eapply All2_impl. 1: tea. cbn; intros. - repeat split; eauto; intuition. + repeat split; intuition eauto. rewrite -> subst_instance_app, inst_case_branch_context_subst_instance in *; eauto. - eapply cumul_Fix. apply All2_map. repeat toAll. eapply All2_impl. 1: tea. - cbn; intros; intuition. + cbn; intros; intuition eauto. rewrite -> subst_instance_app, fix_context_subst_instance in *; eauto. - eapply cumul_CoFix. apply All2_map. repeat toAll. eapply All2_impl. 1: tea. - cbn; intros; intuition. + cbn; intros; intuition eauto. rewrite -> subst_instance_app, fix_context_subst_instance in *; eauto. - - eapply cumul_Prim. depelim e0. - * depelim i; depelim X. cbn in H. cbn. constructor. - * depelim f; depelim X. cbn in H. constructor. - * depelim X. cbn in H. noconf H; cbn in H. constructor; cbn -[Universe.make]; eauto. - + rewrite -!subst_instance_univ_make. - eapply eq_universe_subst_instance; tea. - + solve_all. + - eapply cumul_Prim. depelim e0; depelim X; cbn in H; cbn; noconf H; cbn in H; constructor; cbn -[Universe.make]; eauto. + + rewrite -!subst_instance_univ_make. + eapply eq_universe_subst_instance; tea. + + solve_all. - repeat rewrite subst_instance_mkApps. eapply cumul_Ind. * apply precompose_subst_instance_global. rewrite map_length. eapply R_global_instance_impl_same_napp; try eapply H; eauto. @@ -206,7 +203,7 @@ Lemma subst_instance_prim_val_tag (p : PCUICPrimitive.prim_val term) u : prim_val_tag (mapu_prim (subst_instance_level u) (subst_instance u) p) = prim_val_tag p. Proof. - destruct p as [? []]; simp prim_type => //=. + destruct p as [? []] => //=. Qed. Hint Resolve subst_instance_cstrs_two From c86a3a8fa8c743be1e5a62193b45d7f898c68286 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Mon, 6 Nov 2023 17:45:08 +0100 Subject: [PATCH 03/17] Alpha-conv and validity done --- pcuic/theories/PCUICAlpha.v | 21 +++- pcuic/theories/PCUICEtaExpand.v | 101 +++++++++++++++++- pcuic/theories/PCUICInversion.v | 6 +- pcuic/theories/PCUICSpine.v | 2 + pcuic/theories/PCUICValidity.v | 17 +-- .../Typing/PCUICContextConversionTyp.v | 3 +- 6 files changed, 135 insertions(+), 15 deletions(-) diff --git a/pcuic/theories/PCUICAlpha.v b/pcuic/theories/PCUICAlpha.v index 328df55be..b82cc459b 100644 --- a/pcuic/theories/PCUICAlpha.v +++ b/pcuic/theories/PCUICAlpha.v @@ -995,8 +995,25 @@ Section Alpha. 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 p prim_ty cdecl IH prim decl pinv Δ v e e'. - depelim e. econstructor; tea. now apply IH. + - intros p prim_ty cdecl IH prim decl pinv pty ptyIH Δ v e e'. + depelim e. depelim o. 1-2:econstructor; eauto; constructor. + pose proof (validity (type_Prim Σ Γ _ _ _ wfΓ prim decl pinv pty)) as [s Hs]. + eapply type_Cumul. econstructor; eauto. + * depelim ptyIH. constructor; eauto. now rewrite -e. rewrite -e; eauto. + specialize (hty _ _ e1 e'); eauto. eapply type_Cumul; tea. eapply hdef; eauto. + eapply upto_names_impl_eq_term in e1. + eapply eq_term_upto_univ_cumulSpec. + eapply eq_term_leq_term. eapply e1. + solve_all. + specialize (hty _ _ e1 e'); eauto. eapply type_Cumul; tea. eapply a0; eauto. + eapply upto_names_impl_eq_term in e1. + eapply eq_term_upto_univ_cumulSpec. + eapply eq_term_leq_term. eapply e1. + * eapply eq_context_conversion in Hs; eauto. + * simp prim_type. eapply Universe.make_inj in e. rewrite e. + eapply eq_term_upto_univ_cumulSpec. + eapply upto_names_impl_leq_term. + constructor. constructor. reflexivity. now symmetry. - intros t A B X wf ht iht har ihar hcu Δ v e e'. eapply (type_ws_cumul_pb (pb:=Cumul)). diff --git a/pcuic/theories/PCUICEtaExpand.v b/pcuic/theories/PCUICEtaExpand.v index 90f162d70..1385aacf6 100644 --- a/pcuic/theories/PCUICEtaExpand.v +++ b/pcuic/theories/PCUICEtaExpand.v @@ -47,6 +47,40 @@ Definition isConstruct t := Definition isRel t := match t with tRel _ => true | _ => false end. +Inductive onPrim (P : term -> Prop) : prim_val -> Prop := + | onPrimInt i : onPrim P (primInt; primIntModel i) + | onPrimFloat f : onPrim P (primFloat; primFloatModel f) + | onPrimArray a : + P a.(array_default) -> + P a.(array_type) -> + All P a.(array_value) -> + onPrim P (primArray; primArrayModel a). +Derive Signature for onPrim. + +Section map_All. + Context (P P' : term -> Type). + Context (ont : forall x, P x -> P' x). + + Fixpoint map_All {x} (a : All P x) : All P' x := + match a with + | All_nil => All_nil + | All_cons _ _ p a => All_cons (ont _ p) (map_All a) + end. +End map_All. + +Section map_onPrim. + Context (P P' : term -> Prop). + Context (ont : forall x, P x -> P' x). + + Definition map_onPrim {x} (o : onPrim P x) : onPrim P' x := + match o with + | onPrimInt i => onPrimInt _ i + | onPrimFloat f => onPrimFloat _ f + | onPrimArray a def ty v => + onPrimArray _ _ (ont _ def) (ont _ ty) (map_All _ _ ont v) + end. +End map_onPrim. + Section expanded. Variable Σ : global_env. @@ -85,7 +119,7 @@ Inductive expanded (Γ : list nat) : term -> Prop := #|args| >= (ind_npars mind + context_assumptions (cstr_args cdecl)) -> Forall (expanded Γ) args -> expanded Γ (mkApps (tConstruct ind c u) args) -| expanded_tPrim p : expanded Γ (tPrim p). +| expanded_tPrim p : onPrim (expanded Γ) p -> expanded Γ (tPrim p). End expanded. Derive Signature for expanded. @@ -158,7 +192,7 @@ Lemma expanded_ind : declared_constructor Σ (ind, c) mind idecl cdecl -> #|args| >= ind_npars mind + context_assumptions (cstr_args cdecl) -> Forall (expanded Σ Γ) args -> Forall (P Γ) args -> P Γ (mkApps (tConstruct ind c u) args)) -> - (forall Γ p, P Γ (tPrim p)) -> + (forall Γ p, onPrim (expanded Σ Γ) p -> onPrim (P Γ) p -> P Γ (tPrim p)) -> forall (Γ : list nat) (t : term), expanded Σ Γ t -> P Γ t. Proof. intros Σ P HRel HVar HEvar HSort HProd HLamdba HLetIn HApp HConst HInd HCase HProj HFix HCoFix HConstruct HPrim. @@ -184,6 +218,7 @@ Proof. generalize mfix at 1 3. intros mfix0 H. induction H; econstructor; cbn in *; eauto; split. - eapply HConstruct; eauto. clear - H1 f. induction H1; econstructor; eauto. + - eapply HPrim; eauto. now eapply (map_onPrim _ _ (f Γ)). Qed. From MetaCoq.PCUIC Require Import PCUICInductiveInversion PCUICLiftSubst PCUICSigmaCalculus. @@ -299,6 +334,7 @@ Proof. - rewrite lift_mkApps. cbn. eapply expanded_tConstruct_app; tea. now len. solve_all. + - cbn; constructor. depelim H0; constructor; cbn; eauto. solve_all. Qed. Lemma expanded_subst Σ a k b Γ Δ : @@ -366,6 +402,8 @@ Proof. - rewrite subst_mkApps. cbn. eapply expanded_tConstruct_app; tea. now len. solve_all. + - cbn; constructor; eauto. + depelim H1; constructor; cbn; eauto; solve_all. Qed. Lemma expanded_let_expansion Σ (Δ : context) Γ t : @@ -435,6 +473,7 @@ Proof. now len. - econstructor; eauto. solve_all. - eapply expanded_tConstruct_app; tea. now len. solve_all. + - constructor. depelim H0; constructor; cbn; eauto; solve_all. Qed. Lemma expanded_weakening Σ Γ t : expanded Σ Γ t -> forall Γ', expanded Σ (Γ ++ Γ') t. @@ -443,7 +482,8 @@ Proof. 1:{ intros. eapply expanded_tRel; tea. rewrite nth_error_app_lt. now eapply nth_error_Some_length in H. assumption. solve_all. } - all:intros; try solve [econstructor; eauto 1; solve_all; try now rewrite app_assoc]. + all:intros; try solve [econstructor; cbn; eauto 1; solve_all; try now rewrite app_assoc]. + constructor. depelim H0; constructor; cbn; eauto; solve_all. Qed. @@ -543,6 +583,11 @@ Proof. intros exp; depind exp; solve_discr => //; eauto. Qed. +Lemma expanded_tPrim_inv Σ Γ p : + expanded Σ Γ (tPrim p) -> onPrim (expanded Σ Γ) p. +Proof. + intros exp; depind exp; solve_discr => //; eauto. +Qed. Import PCUICOnOne. Module Red1Apps. @@ -659,6 +704,22 @@ Module Red1Apps. | cofix_red_body mfix0 mfix1 idx : OnOne2 (on_Trel_eq (fun t u => Σ ;;; Γ ,,, fix_context mfix0 |- t ⇝ u) dbody (fun x => (dname x, dtype x, rarg x))) mfix0 mfix1 -> Σ ;;; Γ |- tCoFix mfix0 idx ⇝ tCoFix mfix1 idx + + | array_red_val arr value : + OnOne2 (fun t u => Σ ;;; Γ |- t ⇝ u) arr.(array_value) value -> + Σ ;;; Γ |- tPrim (primArray; primArrayModel arr) ⇝ + tPrim (primArray; primArrayModel (set_array_value arr value)) + + | array_red_def arr def : + Σ ;;; Γ |- arr.(array_default) ⇝ def -> + Σ ;;; Γ |- tPrim (primArray; primArrayModel arr) ⇝ + tPrim (primArray; primArrayModel (set_array_default arr def)) + + | array_red_type arr ty : + Σ ;;; Γ |- arr.(array_type) ⇝ ty -> + Σ ;;; Γ |- tPrim (primArray; primArrayModel arr) ⇝ + tPrim (primArray; primArrayModel (set_array_type arr ty)) + where " Σ ;;; Γ |- t ⇝ u " := (red1 Σ Γ t u). Lemma red1_ind_all : @@ -784,9 +845,28 @@ Module Red1Apps. (fun x => (dname x, dtype x, rarg x))) mfix0 mfix1 -> P Γ (tCoFix mfix0 idx) (tCoFix mfix1 idx)) -> + (forall (Γ : context) (arr : array_model term) + (value : list term), + OnOne2 (Trel_conj (red1 Σ Γ) (P Γ)) (array_value arr) value -> + P Γ (tPrim (primArray; primArrayModel arr)) + (tPrim (primArray; primArrayModel (set_array_value arr value)))) -> + + (forall (Γ : context) (arr : array_model term) + (def : term), Σ;;; Γ |- array_default arr ⇝ def -> + P Γ (array_default arr) def -> + P Γ (tPrim (primArray; primArrayModel arr)) + (tPrim (primArray; primArrayModel (set_array_default arr def)))) -> + + (forall (Γ : context) (arr : array_model term) + (ty : term), Σ;;; Γ |- array_type arr ⇝ ty -> + P Γ (array_type arr) ty -> + P Γ (tPrim (primArray; primArrayModel arr)) + (tPrim (primArray; primArrayModel (set_array_type arr ty)))) -> + + forall (Γ : context) (t t0 : term), red1 Σ Γ t t0 -> P Γ t t0. Proof. - intros. rename X29 into Xlast. revert Γ t t0 Xlast. + intros. rename X32 into Xlast. revert Γ t t0 Xlast. fix aux 4. intros Γ t T. move aux at top. destruct 1; match goal with @@ -861,6 +941,10 @@ Module Red1Apps. revert o. generalize (fix_context mfix0). intros c new. revert mfix0 mfix1 new; fix auxl 3; intros l l' Hl; destruct Hl; constructor; try split; auto; intuition. + + - revert value o. generalize (array_value arr). + fix auxl 3; intros l l' Hl; destruct Hl; + constructor; try split; auto; intuition. Defined. Lemma red_tApp Σ Γ t v u : @@ -1049,6 +1133,8 @@ Proof. - rewrite subst_mkApps. cbn. eapply expanded_tConstruct_app; tea. now len. solve_all. + - cbn; econstructor; eauto. + depelim H0; constructor; cbn; eauto; solve_all. Qed. Lemma expanded_unfold_fix Σ Γ' mfix idx narg fn : @@ -1354,6 +1440,13 @@ Proof. rewrite skipn_app skipn_all2; len. replace (S n - #|mfix0|) with (S (n - #|mfix0|)) by lia. eapply wfΓ. rewrite H4 /= //. noconf H4. + - constructor. eapply expanded_tPrim_inv in exp. + depelim exp; eauto. constructor; cbn; eauto. + eapply OnOne2_impl_All_r; tea; cbn; intuition eauto. + - constructor. eapply expanded_tPrim_inv in exp. + depelim exp; eauto. constructor; cbn; eauto. + - constructor. eapply expanded_tPrim_inv in exp. + depelim exp; eauto. constructor; cbn; eauto. Qed. Lemma expanded_red {cf : checker_flags} {Σ : global_env_ext} Γ Γ' t v : wf Σ -> diff --git a/pcuic/theories/PCUICInversion.v b/pcuic/theories/PCUICInversion.v index 393f90f07..56cbd157d 100644 --- a/pcuic/theories/PCUICInversion.v +++ b/pcuic/theories/PCUICInversion.v @@ -343,12 +343,14 @@ Section Inversion. [× wf_local Σ Γ, primitive_constant Σ (prim_val_tag p) = Some prim_ty, declared_constant Σ prim_ty cdecl, - primitive_invariants cdecl & - Σ ;;; Γ ⊢ tConst prim_ty [] ≤ T]. + primitive_invariants (prim_val_tag p) cdecl, + primitive_typing_hyps typing Σ Γ p & + Σ ;;; Γ ⊢ prim_type p prim_ty ≤ T]. Proof. intros Γ p T h. depind h. - exists prim_ty, cdecl; split => //. eapply ws_cumul_pb_refl; fvs. + depelim p1 => //=; simp prim_type => //=. fvs. - destruct IHh1 as [prim_ty [cdecl []]]. exists prim_ty, cdecl. split => //. transitivity A; tea. eapply cumulSpec_cumulAlgo_curry; tea; fvs. diff --git a/pcuic/theories/PCUICSpine.v b/pcuic/theories/PCUICSpine.v index 95180d206..1c808d6b6 100644 --- a/pcuic/theories/PCUICSpine.v +++ b/pcuic/theories/PCUICSpine.v @@ -1578,6 +1578,8 @@ Proof. rewrite Nat.add_shuffle3. now rewrite app_context_assoc in b. + * cbn. destruct p as [? []]; cbn => //=. + eapply red_primArray_congr; cbn in *; intuition eauto. solve_all. - specialize (X (Γ ,,, Δ) (Some t) wf). simpl in X. apply X. reflexivity. Qed. diff --git a/pcuic/theories/PCUICValidity.v b/pcuic/theories/PCUICValidity.v index c87315ad8..1bd79bf71 100644 --- a/pcuic/theories/PCUICValidity.v +++ b/pcuic/theories/PCUICValidity.v @@ -399,12 +399,17 @@ Section Validity. eapply nth_error_all in X0 as [s Hs]; pcuic. - (* Primitive *) - destruct X0 as [s [hty hbod huniv]]. - exists s@[[]]. - change (tSort s@[[]]) with (tSort s)@[[]]. - rewrite -hty. - refine (type_Const _ _ _ [] _ wfΓ H0 _). - rewrite huniv //. + depelim X1; depelim X2; simp prim_type; cbn in *. + 1-2:destruct X0 as [s [hty hbod huniv]]; exists s@[[]]; change (tSort s@[[]]) with (tSort s)@[[]]; + rewrite -hty; refine (type_Const _ _ _ [] _ wfΓ H0 _); rewrite huniv //. + set (s := (Universe.make (array_level a))). + destruct X0 as [hty' hbod huniv]. + exists s. + eapply (type_App _ _ _ _ _ (tSort s)); tea; cycle 1. + + eapply (type_Const _ _ _ [array_level a]) in H0; tea. rewrite hty' in H0. cbn in H0. exact H0. + red. rewrite huniv. simpl. rtoProp; intuition eauto. eapply LevelSet.mem_spec. eapply (wfl (array_level a, 0)). lsets. + cbn. red. destruct check_univs => //. red. red. intros v H c. csets. + + econstructor. econstructor; eauto. econstructor; eauto. - (* Conv *) now exists s. diff --git a/pcuic/theories/Typing/PCUICContextConversionTyp.v b/pcuic/theories/Typing/PCUICContextConversionTyp.v index 9f289f114..19cbb7855 100644 --- a/pcuic/theories/Typing/PCUICContextConversionTyp.v +++ b/pcuic/theories/Typing/PCUICContextConversionTyp.v @@ -291,7 +291,8 @@ Proof. apply (All_impl X0); simpl. intros d' Ht. apply infer_typing_sort_impl with id Ht; now intros [_ IH']. - + - econstructor; tea. + depelim X2; constructor; eauto. solve_all. - econstructor; eauto. pose proof (wf_local_closed_context wfΓ). pose proof (type_closed (forall_Γ' _ X5 X6)). eapply (@closedn_on_free_vars xpred0) in H0. pose proof (subject_closed (forall_Γ'0 _ X5 X6)). eapply (@closedn_on_free_vars xpred0) in H1. From 26779c7ca53e5c69c807fa86ee3353728bd87355 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Tue, 7 Nov 2023 13:37:24 +0100 Subject: [PATCH 04/17] Proved subject reduction for arrays --- pcuic/theories/PCUICSR.v | 66 +++++++++++++++++++++++++++++++++++----- 1 file changed, 58 insertions(+), 8 deletions(-) diff --git a/pcuic/theories/PCUICSR.v b/pcuic/theories/PCUICSR.v index fc183f62f..cfd6a1994 100644 --- a/pcuic/theories/PCUICSR.v +++ b/pcuic/theories/PCUICSR.v @@ -1012,10 +1012,28 @@ Lemma closed_red1_ind (Σ : global_env_ext) (P0 : context -> term -> term -> Typ (fun x => (dname x, dtype x, rarg x))) mfix0 mfix1 -> P Γ (tCoFix mfix0 idx) (tCoFix mfix1 idx)) -> + (forall (Γ : context) (arr : array_model term) + (value : list term), + OnOne2 (Trel_conj (closed_red1 Σ Γ) (P Γ)) (array_value arr) value -> + P Γ (tPrim (primArray; primArrayModel arr)) + (tPrim (primArray; primArrayModel (set_array_value arr value)))) -> + +(forall (Γ : context) (arr : array_model term) + (def : term), closed_red1 Σ Γ (array_default arr) def -> +P Γ (array_default arr) def -> +P Γ (tPrim (primArray; primArrayModel arr)) + (tPrim (primArray; primArrayModel (set_array_default arr def)))) -> + +(forall (Γ : context) (arr : array_model term) + (ty : term), closed_red1 Σ Γ (array_type arr) ty -> +P Γ (array_type arr) ty -> +P Γ (tPrim (primArray; primArrayModel arr)) + (tPrim (primArray; primArrayModel (set_array_type arr ty)))) -> + forall (Γ : context) (t t0 : term), closed_red1 Σ Γ t t0 -> P0 Γ t t0. Proof. intros. - destruct X27 as [clΓ clt r]. + destruct X30 as [clΓ clt r]. move: clΓ clt. Ltac t := eauto; try split; eauto; @@ -1030,7 +1048,7 @@ Proof. - eapply X13. 2-3:t. inv_on_free_vars. eapply forallb_All in p0. - eapply OnOne2_All_mix_left in X27; tea. + eapply OnOne2_All_mix_left in X30; tea. eapply OnOne2_impl; tea; repeat (intuition auto; t). - forward_keep IHr. { rewrite on_free_vars_ctx_app clΓ. @@ -1042,7 +1060,7 @@ Proof. - eapply X16 => //. inv_on_free_vars. eapply forallb_All in p4. - eapply OnOne2_All_mix_left in X27; tea. cbn in X27. + eapply OnOne2_All_mix_left in X30; tea. cbn in X30. eapply OnOne2_impl; tea; cbn; intros ?? [[[]] ?]. forward_keep p5. { rewrite on_free_vars_ctx_app clΓ /=. @@ -1053,15 +1071,15 @@ Proof. intuition auto. split; auto. - eapply X22 => //. cbn in clt. eapply forallb_All in clt. - eapply OnOne2_All_mix_left in X27; tea. + eapply OnOne2_All_mix_left in X30; tea. eapply OnOne2_impl; tea; cbn; intuition auto; t. - eapply X23 => //. cbn in clt. eapply forallb_All in clt. - eapply OnOne2_All_mix_left in X27; tea. + eapply OnOne2_All_mix_left in X30; tea. eapply OnOne2_impl; tea; cbn; intuition auto; t. - eapply X24 => //. cbn in clt. eapply forallb_All in clt. - eapply OnOne2_All_mix_left in X27; tea. + eapply OnOne2_All_mix_left in X30; tea. eapply OnOne2_impl; tea; cbn; intros ?? [[[]] ?]. forward_keep p. { rewrite on_free_vars_ctx_app clΓ /=. @@ -1072,11 +1090,11 @@ Proof. intuition auto. split; auto. - eapply X25 => //. cbn in clt. eapply forallb_All in clt. - eapply OnOne2_All_mix_left in X27; tea. + eapply OnOne2_All_mix_left in X30; tea. eapply OnOne2_impl; tea; cbn; intuition auto; t. - eapply X26 => //. cbn in clt. eapply forallb_All in clt. - eapply OnOne2_All_mix_left in X27; tea. + eapply OnOne2_All_mix_left in X30; tea. eapply OnOne2_impl; tea; cbn; intros ?? [[[]] ?]. forward_keep p. { rewrite on_free_vars_ctx_app clΓ /=. @@ -1085,6 +1103,11 @@ Proof. rewrite app_length fix_context_length -shiftnP_add //. now inv_on_free_vars. } intuition auto. split; auto. + - eapply X27 => //. + cbn in clt; rtoProp. eapply forallb_All in H. + eapply OnOne2_All_mix_left in X30; tea. + eapply OnOne2_impl; tea; cbn; intros ?? [[]]. split; eauto. + split; eauto. Qed. Definition closed_red1_ind' := @@ -3081,6 +3104,33 @@ Proof. * apply conv_cumul, conv_sym. destruct disj as [<-|[_ eq]]. reflexivity. noconf eq. rewrite H4; reflexivity. + - eapply OnOne2_prod_inv in X3 as [X3 _]. + destruct X. depelim X1. depelim X2. + Transparent prim_type. + eapply (type_Prim _ _ (primArray; primArrayModel (set_array_value arr value))); tea. + constructor; cbn; eauto. + solve_all. + eapply OnOne2_All_All; tea; cbn; intuition eauto. + + - destruct X. depelim X1. depelim X2. + Transparent prim_type. + eapply (type_Prim _ _ (primArray; primArrayModel (set_array_default arr def))); tea. + constructor; cbn; eauto. + + - pose proof (type_Prim _ _ _ _ _ wfΓ heq_primitive_constant H0 X0 X1). eapply validity in X4. + destruct X. depelim X1. depelim X2. + eapply (type_ws_cumul_pb (pb:=Conv)); tea. + eapply (type_Prim _ _ (primArray; primArrayModel (set_array_type arr ty))); tea. + constructor; cbn; eauto. + * eapply type_ws_cumul_pb; tea. eexists; eauto. + now eapply (red_ws_cumul_pb (pb:=Conv)), closed_red1_red. + * solve_all. + eapply type_ws_cumul_pb; tea. eexists; eauto. + now eapply (red_ws_cumul_pb (pb:=Conv)), closed_red1_red. + * simp prim_type. eapply ws_cumul_pb_App. + eapply ws_cumul_pb_refl; fvs. cbn. + now symmetry; eapply red_ws_cumul_pb, closed_red1_red. + - (* Conversion *) specialize (forall_u _ Hu). eapply type_Cumul; tea. From a7af7452e0d26efe5ccfd3b9252ea54d5228fa5e Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Tue, 7 Nov 2023 14:04:45 +0100 Subject: [PATCH 05/17] Adapted all of PCUIC to primitive arrays (and in general term containers in primitives) --- pcuic/theories/Bidirectional/BDFromPCUIC.v | 8 + .../theories/Bidirectional/BDStrengthening.v | 39 +++-- pcuic/theories/Bidirectional/BDToPCUIC.v | 6 +- pcuic/theories/Bidirectional/BDUnique.v | 12 +- pcuic/theories/Conversion/PCUICInstConv.v | 2 +- pcuic/theories/Conversion/PCUICRenameConv.v | 4 +- pcuic/theories/PCUICAst.v | 43 +---- pcuic/theories/PCUICClassification.v | 152 ++++++++++++++++-- pcuic/theories/PCUICConvCumInversion.v | 50 +++++- pcuic/theories/PCUICConversion.v | 2 +- pcuic/theories/PCUICCumulProp.v | 26 ++- pcuic/theories/PCUICCumulativitySpec.v | 26 +-- pcuic/theories/PCUICEquality.v | 22 +-- pcuic/theories/PCUICEtaExpand.v | 10 -- pcuic/theories/PCUICExpandLets.v | 2 +- pcuic/theories/PCUICExpandLetsCorrectness.v | 34 +++- pcuic/theories/PCUICFirstorder.v | 4 +- pcuic/theories/PCUICNormal.v | 21 ++- pcuic/theories/PCUICParallelReduction.v | 6 +- pcuic/theories/PCUICPrincipality.v | 35 +++- pcuic/theories/PCUICProgress.v | 54 ++++--- pcuic/theories/PCUICWfUniverses.v | 10 +- pcuic/theories/utils/PCUICPrimitive.v | 89 ++++++++++ 23 files changed, 467 insertions(+), 190 deletions(-) diff --git a/pcuic/theories/Bidirectional/BDFromPCUIC.v b/pcuic/theories/Bidirectional/BDFromPCUIC.v index f7f33b21c..8f806cd2d 100644 --- a/pcuic/theories/Bidirectional/BDFromPCUIC.v +++ b/pcuic/theories/Bidirectional/BDFromPCUIC.v @@ -377,7 +377,15 @@ Proof. - intros p prim_ty cdecl wfΓ' hp hdecl pinv. eexists. split; [econstructor; tea|]. + destruct X0. 1-2:constructor; eauto. + eapply conv_check in hty; tea. + eapply conv_check in hdef; tea. + constructor; eauto. solve_all. + now eapply conv_check in X0. eapply ws_cumul_pb_refl; fvs. + depelim X; cbn => //. + simp prim_type. cbn. + now eapply subject_is_open_term. - intros ? ? ? ? ? ? (?&?&?) ? (?&?&?) ?. eexists. diff --git a/pcuic/theories/Bidirectional/BDStrengthening.v b/pcuic/theories/Bidirectional/BDStrengthening.v index 3882a6364..8f7027ee2 100644 --- a/pcuic/theories/Bidirectional/BDStrengthening.v +++ b/pcuic/theories/Bidirectional/BDStrengthening.v @@ -1,5 +1,6 @@ From Coq Require Import Bool List Arith Lia. From Coq Require String. +From Equations Require Import Equations. From MetaCoq.Utils Require Import utils monad_utils. From MetaCoq.Common Require Import config. From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICGlobalEnv @@ -70,6 +71,7 @@ Proof. f_equal. 1: auto. by rewrite map_length ebod shiftnP_shiftn. + - solve_all. Qed. Lemma Alli_impl_le {A P Q} {l : list A} {n} : @@ -470,7 +472,8 @@ Section OnFreeVars. cbn in Hmfix. by move: Hmfix => /andP []. - - easy. + - intros. red. intros. destruct X1; cbn => //. + simp prim_type. cbn in *. solve_all. - easy. - intros ? ? ? ? ? ? _ HT Hred. @@ -865,7 +868,11 @@ Proof using wfΣ. + by apply rename_wf_cofixpoint. - intros. red. intros P Δ f hf ht. - cbn. econstructor; tea. + cbn. rewrite rename_prim_type. econstructor; tea; rewrite ?prim_val_tag_map //. + destruct X1; cbn in *; constructor; cbn; eauto. + * eapply hty; tea; solve_all. + * eapply hdef; tea; solve_all. + * solve_all. - intros. red. intros P Δ f hf ht. econstructor ; eauto. @@ -994,17 +1001,17 @@ Proof. move: (tyt) => /typing_wf_local wfΓ. move: (tyt) => /typing_infering [T' [inf cum]]. unshelve epose proof (infering_on_free_vars _ _ _ _ _ _ _ _ inf); eauto. - pose proof (inf' := inf). apply infering_typing in inf; eauto. - exists T'. split; eauto. + pose proof (inf' := inf). apply infering_typing in inf; eauto. + exists T'. split; eauto. eapply ws_cumul_pb_forget_cumul; eauto. eapply bidirectional_renaming in inf'; eauto. eapply infering_typing; eauto. Qed. Lemma strengthening `{cf: checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} Γ Γ' Γ'' t T : Σ ;;; Γ ,,, Γ' ,,, lift_context #|Γ'| 0 Γ'' |- lift #|Γ'| #|Γ''| t : T -> - ∑ T', [× on_free_vars (strengthenP #|Γ''| #|Γ'| xpredT) T', + ∑ T', [× on_free_vars (strengthenP #|Γ''| #|Γ'| xpredT) T', (Σ ;;; Γ ,,, Γ' ,,, lift_context #|Γ'| 0 Γ'' |- lift #|Γ'| #|Γ''| t : T') , - (Σ ;;; Γ ,,, Γ' ,,, lift_context #|Γ'| 0 Γ'' |- T' <= T) & + (Σ ;;; Γ ,,, Γ' ,,, lift_context #|Γ'| 0 Γ'' |- T' <= T) & (Σ ;;; Γ ,,, Γ'' |- t : unlift #|Γ'| #|Γ''| T')]. Proof. intros Hty. @@ -1037,9 +1044,9 @@ Proof. } eapply typing_renaming_cond_P in Hty as [T' [? ?]]; eauto. - exists T'. split; eauto. - - erewrite <- (lift_unlift_term t); eauto. - - split; eauto. eapply urenaming_strengthen. + exists T'. split; eauto. + - erewrite <- (lift_unlift_term t); eauto. + - split; eauto. eapply urenaming_strengthen. - move: wfΓ'' => /wf_local_closed_context. rewrite on_free_vars_ctx_app => /andP [? ?]. apply on_ctx_free_vars_strengthenP. @@ -1056,9 +1063,9 @@ Qed. Lemma strengthening_type `{cf: checker_flags} {Σ : global_env_ext} {wfΣ : PCUICTyping.wf Σ} Γ Γ' Γ'' t s : - Σ ;;; Γ ,,, Γ' ,,, lift_context #|Γ'| 0 Γ'' |- lift #|Γ'| #|Γ''| t : tSort s -> + Σ ;;; Γ ,,, Γ' ,,, lift_context #|Γ'| 0 Γ'' |- lift #|Γ'| #|Γ''| t : tSort s -> ∑ s', (Σ ;;; Γ ,,, Γ'' |- t : tSort s') * (compare_universe Cumul Σ s' s). -Proof. +Proof. intros H; pose proof (H' := H); eapply strengthening in H. destruct H as [T' [? HT Hcumul HT']]. pose proof (Hcumul' := Hcumul). pose proof (type_closed HT). @@ -1066,21 +1073,21 @@ Proof. pose proof (type_closed HT'). eapply closedn_on_free_vars in H0. pose proof (wf_local_closed_context (typing_wf_local H')). - eapply into_ws_cumul_pb in Hcumul; eauto. + eapply into_ws_cumul_pb in Hcumul; eauto. eapply PCUICConversion.ws_cumul_pb_Sort_r_inv in Hcumul as [s' [Hred Hcumul]]. - eapply closed_red_red in Hred. pose proof (HT'':=HT'). eapply PCUICValidity.validity in HT' as [? ?]. + eapply closed_red_red in Hred. pose proof (HT'':=HT'). eapply PCUICValidity.validity in HT' as [? ?]. exists s'; split; eauto. assert (PCUICReduction.red Σ (Γ,,, Γ'') (unlift #|Γ'| #|Γ''| T') (tSort s')). { assert (tSort s' = unlift #|Γ'| #|Γ''| (tSort s')) by reflexivity. rewrite H2; clear H2. eapply red_rename. eauto. eapply (urenaming_strengthen _ _ _ _); eauto. - 2-3: eauto. eapply typing_wf_local in HT''. + 2-3: eauto. eapply typing_wf_local in HT''. eapply wf_local_closed_context in HT''. revert HT''. rewrite on_free_vars_ctx_app => /andP [? ?]. apply on_ctx_free_vars_strengthenP. all: eapply on_free_vars_ctx_on_ctx_free_vars_xpredT; eauto. - } + } eapply type_Cumul; eauto. eapply PCUICSR.subject_reduction; eauto. assert (liftSort : unlift #|Γ'| #|Γ''| (tSort s') = tSort s') by reflexivity. rewrite <- liftSort; clear liftSort. eapply PCUICConversion.cumulAlgo_cumulSpec. eapply PCUICConversion.ws_cumul_pb_red_r_inv; eauto. - eapply ws_cumul_pb_refl; eauto. eapply PCUICInversion.typing_closed_ctx; eauto. + eapply ws_cumul_pb_refl; eauto. eapply PCUICInversion.typing_closed_ctx; eauto. Qed. \ No newline at end of file diff --git a/pcuic/theories/Bidirectional/BDToPCUIC.v b/pcuic/theories/Bidirectional/BDToPCUIC.v index 7028cdcc8..1e6637ba3 100644 --- a/pcuic/theories/Bidirectional/BDToPCUIC.v +++ b/pcuic/theories/Bidirectional/BDToPCUIC.v @@ -416,7 +416,11 @@ Section BDToPCUICTyping. all: auto. - red; intros. - now econstructor. + econstructor; eauto. + depelim X1; constructor; eauto. + eapply hty; eauto. eexists. econstructor; eauto. + eapply hdef; eauto. eexists; eauto. eapply hty; eauto. eexists; econstructor; eauto. + solve_all. eapply X7; eauto. eexists. eapply hty; eauto. eexists; econstructor; eauto. - red ; intros. now eapply type_reduction. diff --git a/pcuic/theories/Bidirectional/BDUnique.v b/pcuic/theories/Bidirectional/BDUnique.v index 20547e3ee..de25a576c 100644 --- a/pcuic/theories/Bidirectional/BDUnique.v +++ b/pcuic/theories/Bidirectional/BDUnique.v @@ -236,12 +236,12 @@ Proof using wfΣ. * fvs. * now eapply type_is_open_term, infering_typing. - - inversion X1; subst. - rewrite H in H2; noconf H2. - unshelve eapply declared_constant_to_gen in H0, H3; eauto. - have eq := (declared_constant_inj _ _ H0 H3); subst cdecl0. - exists (tConst prim_ty []). - split; eapply closed_red_refl; fvs. + - depelim X1; depelim X0. + * depelim X3. rewrite e in H; noconf H. eexists. split; eapply closed_red_refl; fvs. + * depelim X3. rewrite e in H; noconf H. eexists. split; eapply closed_red_refl; fvs. + * depelim X3. rewrite e in H; noconf H. eexists. split; eapply closed_red_refl; fvs. + all:simp prim_type; cbn. cbn in hty. + all:eapply type_is_open_term, checking_typing; tea; eexists; eapply checking_typing; tea; eexists; econstructor; tea. - inversion X3 ; subst. eapply X0 in X4 as [T'' []]; subst ; tea. diff --git a/pcuic/theories/Conversion/PCUICInstConv.v b/pcuic/theories/Conversion/PCUICInstConv.v index 840dc29e7..fc6ce53b2 100644 --- a/pcuic/theories/Conversion/PCUICInstConv.v +++ b/pcuic/theories/Conversion/PCUICInstConv.v @@ -1914,7 +1914,7 @@ Proof using Type. generalize #|m|. intro k. eapply All2_map. simpl. solve_all. - simpl. constructor. - eapply onPrim_map_prop; tea. cbn; intuition eauto. + eapply onPrims_map_prop; tea. cbn; intuition eauto. Qed. Lemma inst_conv {Σ : global_env_ext} {wfΣ : wf Σ} {Γ Δ σ A B} : diff --git a/pcuic/theories/Conversion/PCUICRenameConv.v b/pcuic/theories/Conversion/PCUICRenameConv.v index 0035f456c..d2c947fdf 100644 --- a/pcuic/theories/Conversion/PCUICRenameConv.v +++ b/pcuic/theories/Conversion/PCUICRenameConv.v @@ -2,7 +2,7 @@ From Coq Require Import Morphisms. From MetaCoq.Utils Require Import utils. From MetaCoq.Common Require Import config. -From MetaCoq.PCUIC Require Import PCUICAst PCUICOnOne PCUICAstUtils PCUICCases PCUICInduction +From MetaCoq.PCUIC Require Import PCUICPrimitive PCUICAst PCUICOnOne PCUICAstUtils PCUICCases PCUICInduction PCUICLiftSubst PCUICUnivSubst PCUICCumulativity PCUICReduction PCUICGlobalEnv PCUICClosed PCUICEquality PCUICRenameDef PCUICWeakeningEnvConv PCUICSigmaCalculus PCUICClosed PCUICOnFreeVars PCUICGuardCondition @@ -263,7 +263,7 @@ Proof using Type. * unfold map_def. intuition eauto. * eauto. - simpl. constructor. - eapply onPrim_map_prop; tea. cbn; intuition eauto. + eapply onPrims_map_prop; tea. cbn; intuition eauto. Qed. Lemma urenaming_impl : diff --git a/pcuic/theories/PCUICAst.v b/pcuic/theories/PCUICAst.v index a9c3036cf..4260e3b4f 100644 --- a/pcuic/theories/PCUICAst.v +++ b/pcuic/theories/PCUICAst.v @@ -238,38 +238,6 @@ Definition isLambda t := Lemma isLambda_inv t : isLambda t -> exists na ty bod, t = tLambda na ty bod. Proof. destruct t => //; eauto. Qed. -Set Equations Transparent. - -Definition mapu_array_model {term term'} (fl : Level.t -> Level.t) (f : term -> term') - (ar : array_model term) : array_model term' := - {| array_level := fl ar.(array_level); - array_value := map f ar.(array_value); - array_default := f ar.(array_default); - array_type := f ar.(array_type) |}. - -Equations mapu_prim {term term'} (f : Level.t -> Level.t) (g : term -> term') - (p : PCUICPrimitive.prim_val term) : PCUICPrimitive.prim_val term' := -| _, _, (primInt; primIntModel i) => (primInt; primIntModel i) -| _, _, (primFloat; primFloatModel fl) => (primFloat; primFloatModel fl) -| f, g, (primArray; primArrayModel ar) => - (primArray; primArrayModel (mapu_array_model f g ar)). - -Notation map_array_model := (mapu_array_model id). -Notation map_prim := (mapu_prim id). - -Equations test_prim (p : term -> bool) (p : prim_val) : bool := -| p, (primInt; _) => true -| p, (primFloat; _) => true -| p, (primArray; primArrayModel ar) => - List.forallb p ar.(array_value) && p ar.(array_default) && p ar.(array_type). - -Equations test_primu (p : Level.t -> bool) (t : term -> bool) (p : prim_val) : bool := -| _, _, (primInt; _) => true -| _, _, (primFloat; _) => true -| p, pt, (primArray; primArrayModel ar) => - p ar.(array_level) && forallb pt ar.(array_value) && - pt ar.(array_default) && pt ar.(array_type). - (** Basic operations on the AST: lifting, substitution and tests for variable occurrences. *) Fixpoint lift n k t : term := @@ -982,13 +950,6 @@ Proof. destruct p as [? []] => //. Qed. -Definition tPrimProp {term} (P : term -> Type) (p : PCUICPrimitive.prim_val term) : Type := - match p.π2 return Type with - | primIntModel f => unit - | primFloatModel f => unit - | primArrayModel a => P a.(array_type) × P a.(array_default) × All P a.(array_value) - end. - Lemma mapu_array_model_proper {term term'} (l l' : Level.t -> Level.t) (f g : term -> term') a : l =1 l' -> f =1 g -> mapu_array_model l f a = mapu_array_model l' g a. @@ -1027,7 +988,7 @@ Proof. destruct a => //=. rewrite /mapu_array_model /=. rewrite map_id //. Qed. -Lemma test_prim_primProp {P p} : test_prim P p -> tPrimProp P p. +Lemma test_prim_primProp {term} {P p} : @test_prim term P p -> tPrimProp P p. Proof. destruct p as [? []]; cbn => //. move/andP => [] /andP[]. intuition auto. @@ -1095,7 +1056,7 @@ Proof. eapply All_forallb, All_map, All_impl; tea. intuition eauto. apply hg; intuition auto. Qed. -Lemma test_prim_mapu {put l f p} : +Lemma test_prim_mapu {term term'} {put : term' -> bool} {l} {f : term -> term'} {p} : test_prim put (mapu_prim l f p) = test_prim (fun x => put (f x)) p. Proof. destruct p as [? []]; cbn; eauto. diff --git a/pcuic/theories/PCUICClassification.v b/pcuic/theories/PCUICClassification.v index 55908fed7..c1c4d9b09 100644 --- a/pcuic/theories/PCUICClassification.v +++ b/pcuic/theories/PCUICClassification.v @@ -812,35 +812,154 @@ Section classification. congruence. Qed. - Lemma ws_cumul_pb_Axiom_l_inv {pb Γ cst u cdecl T} : + Ltac bool := rtoProp; intuition eauto. + + Lemma invert_red1_axiom_app {Γ cst u cdecl args T} : + declared_constant Σ cst cdecl -> + cst_body cdecl = None -> + Σ ;;; Γ ⊢ mkApps (tConst cst u) args ⇝1 T -> + ∑ args', (T = mkApps (tConst cst u) args') × All2 (fun arg arg' => Σ ;;; Γ ⊢ arg ⇝ arg') args args'. + Proof using wfΣ. + intros hdecl hbod. revert T. + induction args using rev_ind. + - simpl. intros; eapply invert_red_axiom in hbod; tea. 2:{ eapply into_closed_red. apply red1_red. apply X. fvs. fvs. } subst. exists []. split => //. + - rewrite mkApps_app /=; intros. + depelim X; solve_discr. + depind clrel_rel; solve_discr. + * destruct args using rev_case; solve_discr. noconf H. rewrite mkApps_app in H; noconf H. solve_discr. + * specialize (IHargs N1); forward IHargs. constructor; fvs. cbn in clrel_src. rtoProp; intuition auto. + destruct IHargs as [args' []]. exists (args' ++ [M2]). subst N1. rewrite mkApps_app /=; split => //. + eapply All2_app => //. constructor; eauto. eapply closed_red_refl; fvs. cbn in clrel_src; bool. + * exists (args ++ [N2]). rewrite mkApps_app /=. split => //. + cbn in clrel_src; bool. rewrite on_free_vars_mkApps in H; bool. toAll. + eapply All2_app => //. + eapply All_All2; tea; cbn; bool. eapply closed_red_refl; fvs. eapply All2_tip, into_closed_red; tea. now constructor. + Qed. + + Lemma closed_red_ind (P : context -> term -> term -> Type) : + (forall Γ t, is_closed_context Γ -> is_open_term Γ t -> P Γ t t) -> + (forall Γ t u, Σ ;;; Γ ⊢ t ⇝1 u -> P Γ t u) -> + (forall Γ t u v, Σ ;;; Γ ⊢ t ⇝ u -> Σ ;;; Γ ⊢ u ⇝ v -> P Γ t u -> P Γ u v -> P Γ t v) -> + forall Γ t u, Σ ;;; Γ ⊢ t ⇝ u -> P Γ t u. + Proof. + intros prefl pred ptrans. + destruct 1 as [ctx src rel]. + eapply clos_rt_rt1n_iff in rel. + induction rel. eapply prefl; tea. + eapply ptrans; tea. eapply into_closed_red; tea. now eapply red1_red. + split; fvs. now eapply clos_rt_rt1n_iff. eapply pred. split; fvs. + eapply IHrel. fvs. + Qed. + + Lemma invert_red_axiom_app {Γ cst u cdecl args T} : + declared_constant Σ cst cdecl -> + cst_body cdecl = None -> + Σ ;;; Γ ⊢ mkApps (tConst cst u) args ⇝ T -> + ∑ args', (T = mkApps (tConst cst u) args') × All2 (fun arg arg' => Σ ;;; Γ ⊢ arg ⇝ arg') args args'. + Proof using wfΣ. + intros hdecl hbod red. + remember (mkApps (tConst cst u) args) as U. revert red args HeqU. + revert Γ U T. + apply: closed_red_ind; intros; subst. + - exists args. split => //. rewrite on_free_vars_mkApps /= in H0. toAll. eapply All_All2; tea; bool. now eapply closed_red_refl. + - eapply invert_red1_axiom_app in X; tea. + -specialize (X1 _ eq_refl) as [args' []]. + subst u0. specialize (X2 _ eq_refl) as [args'' []]. subst v. + exists args''. split => //. eapply All2_trans; tea. + intros x y z red1 red2. now etransitivity. + Qed. + + Lemma All2_All2_All2 {A B C} {P Q R} (l : list A) (l' : list B) (l'' : list C) : + All2 P l l' -> All2 Q l' l'' -> + (forall x y z, P x y -> Q y z -> R x z) -> + All2 R l l''. + Proof. + induction 1 in l'' |- *; intros a; depelim a; constructor; eauto. + Qed. + + Lemma ws_cumul_pb_Axiom_l_inv {pb Γ cst u args cdecl T} : + declared_constant Σ cst cdecl -> + cst_body cdecl = None -> + Σ ;;; Γ ⊢ mkApps (tConst cst u) args ≤[pb] T -> + ∑ u' args', + [× Σ ;;; Γ ⊢ T ⇝ mkApps (tConst cst u') args', + All2 (fun args args' => Σ ;;; Γ ⊢ args ≤[Conv] args') args args' & + PCUICEquality.R_universe_instance (eq_universe Σ) u u']. + Proof using wfΣ. + intros hdecl hb H. + eapply ws_cumul_pb_red in H as (v & v' & [tv tv' eqp]). + epose proof (invert_red_axiom_app hdecl hb tv) as [args' [-> ?]]. + eapply compare_term_mkApps_l_inv in eqp as [t' [l' []]]; subst v'. + depelim e. + exists u', l'. split => //. + eapply closed_red_open_right in tv'. rewrite on_free_vars_mkApps /= in tv'. + solve_all. + eapply All2_All2_All2; tea. cbn. + intros x y z red [eq op]. + eapply ws_cumul_pb_red. exists y, z; split => //. eapply closed_red_refl; fvs. + Qed. + + Lemma ws_cumul_pb_Axiom_r_inv {pb Γ cst u args cdecl T} : declared_constant Σ cst cdecl -> cst_body cdecl = None -> - Σ ;;; Γ ⊢ tConst cst u ≤[pb] T -> - ∑ u', Σ ;;; Γ ⊢ T ⇝ tConst cst u' × PCUICEquality.R_universe_instance (eq_universe Σ) u u'. + Σ ;;; Γ ⊢ T ≤[pb] mkApps (tConst cst u) args -> + ∑ u' args', + [× Σ ;;; Γ ⊢ T ⇝ mkApps (tConst cst u') args', + All2 (fun args args' => Σ ;;; Γ ⊢ args ≤[Conv] args') args' args & + PCUICEquality.R_universe_instance (eq_universe Σ) u' u]. Proof using wfΣ. intros hdecl hb H. eapply ws_cumul_pb_red in H as (v & v' & [tv tv' eqp]). - epose proof (invert_red_axiom hdecl hb tv). subst v. - depelim eqp. - exists u'. split => //. + epose proof (invert_red_axiom_app hdecl hb tv') as [args' [-> ?]]. + eapply compare_term_mkApps_r_inv in eqp as [t' [l' []]]; subst v. + depelim e. + exists u0, l'. split => //. + eapply closed_red_open_right in tv. rewrite on_free_vars_mkApps /= in tv. + solve_all. + eapply All2_sym in a0. eapply All2_sym. + eapply All2_All2_All2; tea. cbn. + intros x y z red [eq op]. + eapply ws_cumul_pb_red. exists z, y; split => //. eapply closed_red_refl; fvs. Qed. - Lemma invert_cumul_axiom_ind {Γ cst cdecl u ind u' args} : + Lemma invert_cumul_axiom_ind {Γ cst cdecl u ind u' args args'} : declared_constant Σ cst cdecl -> cst_body cdecl = None -> - Σ ;;; Γ ⊢ tConst cst u ≤ mkApps (tInd ind u') args -> False. + Σ ;;; Γ ⊢ mkApps (tConst cst u) args ≤ mkApps (tInd ind u') args' -> False. Proof using wfΣ. - intros hd hb ht; eapply ws_cumul_pb_Axiom_l_inv in ht as (u'' & hred & hcmp); eauto. + intros hd hb ht; eapply ws_cumul_pb_Axiom_l_inv in ht as (u'' & args'' & [hred hcmp _]); eauto. eapply invert_red_mkApps_tInd in hred as (? & []); auto. solve_discr. Qed. - Lemma invert_cumul_axiom_prod {Γ cst cdecl u na dom codom} : + Lemma invert_cumul_axiom_prod {Γ cst cdecl u args na dom codom} : declared_constant Σ cst cdecl -> cst_body cdecl = None -> - Σ ;;; Γ ⊢ tConst cst u ≤ tProd na dom codom -> False. + Σ ;;; Γ ⊢ mkApps (tConst cst u) args ≤ tProd na dom codom -> False. + Proof using wfΣ. + intros hd hb ht; eapply ws_cumul_pb_Axiom_l_inv in ht as (u'' & args' & [hred hcmp _]); eauto. + eapply invert_red_prod in hred as (? & ? & []); auto. solve_discr. + Qed. + + Lemma invert_cumul_prim_type_ind {Γ cst cdecl p ind u' args'} : + declared_constant Σ cst cdecl -> + primitive_invariants (prim_val_tag p) cdecl -> + Σ ;;; Γ ⊢ prim_type p cst ≤ mkApps (tInd ind u') args' -> False. + Proof using wfΣ. + intros hd hb ht. + destruct p as [? []]; simp prim_type in ht. + 1-2:eapply (invert_cumul_axiom_ind (args := [])) in ht; tea; now destruct hb as [s []]. + eapply (invert_cumul_axiom_ind (args := [_])) in ht; tea. apply hb. + Qed. + + Lemma invert_cumul_prim_type_prod {Γ cst cdecl p na dom codom} : + declared_constant Σ cst cdecl -> + primitive_invariants (prim_val_tag p) cdecl -> + Σ ;;; Γ ⊢ prim_type p cst ≤ tProd na dom codom -> False. Proof using wfΣ. - intros hd hb ht; eapply ws_cumul_pb_Axiom_l_inv in ht as (u'' & hred & hcmp); eauto. - eapply invert_red_prod in hred as (? & ? & []); auto. discriminate. + intros hd hb ht. + destruct p as [? []]; simp prim_type in ht. + 1-2:eapply (invert_cumul_axiom_prod (args := [])) in ht; tea; now destruct hb as [s []]. + eapply (invert_cumul_axiom_prod (args := [_])) in ht; tea. apply hb. Qed. Lemma whnf_classification' t i u args : @@ -865,9 +984,10 @@ Section classification. unfold unfold_fix. destruct nth_error; [|easy]. destruct o as [[? [? ?]]|]; [|easy]. inversion H; eauto. - now rewrite head_mkApps /head /=. - - eapply inversion_Prim in typed as [prim_ty [cdecl [? ? ? [? hp]]]]; eauto. - eapply invert_cumul_axiom_ind in w; eauto. - apply hp. + - eapply inversion_Prim in typed as [prim_ty [cdecl [? ? ? ? hp]]]; eauto. + destruct p as [? []]; simp prim_type in w. + 1-2:eapply (invert_cumul_axiom_ind (args := [])) in w; eauto; destruct p0 as [s []]; eauto. + eapply (invert_cumul_axiom_ind (args := [array_type a0])) in w; eauto. destruct p0 as [s []]; eauto. Qed. Lemma whnf_classification t i u args : diff --git a/pcuic/theories/PCUICConvCumInversion.v b/pcuic/theories/PCUICConvCumInversion.v index f97ac0625..53b5e80af 100644 --- a/pcuic/theories/PCUICConvCumInversion.v +++ b/pcuic/theories/PCUICConvCumInversion.v @@ -284,18 +284,19 @@ Section fixed. - apply whnf_red_isIndConstructApp in w as ?. destruct hd. all: cbn. - 1-9, 12-15: apply conv_cum_alt; eauto. + 1-9, 12-16: apply conv_cum_alt; eauto. 1-13: constructor. 1-13: exists x1, x. 1-13: split; eauto with pcuic. 1-13: (eapply eq_term_upto_univ_napp_nonind; [exact e|try exact H1]). 1-13: cbn in *; congruence. - all: depelim w; subst. - all: depelim e. - all: depelim w0; subst. - all: apply All2_length in a. - all: try (constructor; constructor; rewrite a; auto). - all: destruct leq; cbn; repeat constructor => //. + 1-3: depelim w; subst; depelim e; depelim w0; subst; + apply All2_length in a; + try (constructor; constructor; rewrite a; auto). + + sq. exists (tPrim i'), (tPrim i'0). split => //. all:eauto with pcuic. + eapply (eq_term_upto_univ_napp_nonind _ _ 0); eauto. now constructor. + - eapply alt_into_ws_cumul_pb_terms => //. clear -a1 a a0. induction a1 in args, args', x2, a, x3, a0, a1 |- *; depelim a; depelim a0; [now constructor|]. @@ -574,5 +575,40 @@ Section fixed. exists c'0, c'1; split; eauto. all:eapply into_closed_red; eauto. Qed. +(* + Lemma conv_cum_tPrim_inv leq Γ p p' : + conv_cum leq Σ Γ (tPrim p) (tPrim p') -> + whnf RedFlags.default Σ Γ (tPrim p) -> + whnf RedFlags.default Σ Γ (tPrim p') -> + ∥ onPrims (fun t t' => Σ ;;; Γ ⊢ t = t') (eq_universe Σ) p p' ∥. + Proof using wfΣ. + intros conv whl whr. + apply conv_cum_alt in conv as [(?&?&[r1 r2 ?])]; auto. sq. + have clΓ := clrel_ctx r1. + have cll := clrel_src r1. + have clr := clrel_src r2. + eapply whnf_red_inv in r1; eauto. + eapply whnf_red_inv in r2; eauto. + depelim r1. + depelim r2. + depelim c. + depelim o; depelim o0; depelim o1; cbn in cll, clr; rtoProp; constructor; eauto. + * now rewrite e e0. + * eapply ws_cumul_pb_alt; do 2 eexists; split; tea. + * eapply ws_cumul_pb_alt; do 2 eexists; split; tea. + * solve_all. + 2:{ intros x y h. epose proof (conv_cum_alt (leq:=Conv)). unfold conv_cum in H. cbn in H. } + eapply alt_into_ws_cumul_pb_terms; tea; solve_all. + eapply + + eapply All2_trans. + + 2-3: eapply All2_impl; tea. + split; [easy|]. + apply ws_cumul_pb_alt_closed; eauto. + exists c'0, c'1; split; eauto. + all:eapply into_closed_red; eauto. + Qed. + *) End fixed. diff --git a/pcuic/theories/PCUICConversion.v b/pcuic/theories/PCUICConversion.v index 00074eb04..fb88ad900 100644 --- a/pcuic/theories/PCUICConversion.v +++ b/pcuic/theories/PCUICConversion.v @@ -3672,7 +3672,7 @@ Section MoreCongruenceLemmas. Defined. Lemma ws_cumul_pb_Prim {pb Γ p p'} : - onPrim (ws_cumul_pb Conv Σ Γ) (eq_universe Σ) p p' -> is_closed_context Γ -> + onPrims (ws_cumul_pb Conv Σ Γ) (eq_universe Σ) p p' -> is_closed_context Γ -> Σ;;; Γ ⊢ tPrim p ≤[pb] tPrim p'. Proof. intros Hp HΓ. diff --git a/pcuic/theories/PCUICCumulProp.v b/pcuic/theories/PCUICCumulProp.v index 9af3c720a..7ceee9cd8 100644 --- a/pcuic/theories/PCUICCumulProp.v +++ b/pcuic/theories/PCUICCumulProp.v @@ -820,6 +820,17 @@ Proof using Type. eauto. Qed. +Lemma cumul_prop_subst_instance_level Σ univs u u' (l : Level.t) : + wf Σ.1 -> + consistent_instance_ext Σ univs u -> + consistent_instance_ext Σ univs u' -> + eq_univ_prop (subst_instance u (Universe.make l)) + (subst_instance u' (Universe.make l)). +Proof using Type. + intros wfΣ cu cu'. red. + unfold subst_instance; cbn. intuition. +Qed. + Lemma cumul_prop_subst_instance_instance Σ univs u u' (i : Instance.t) : wf Σ.1 -> consistent_instance_ext Σ univs u -> @@ -866,6 +877,10 @@ Proof using Type. eapply All2_map. eapply All_All2; tea. cbn. intuition auto. rewrite /id. reflexivity. + - cbn. constructor; splits; simpl. + destruct p as [? []]; constructor; cbn in *; intuition eauto. + unshelve eapply cumul_prop_subst_instance_level; tea. exact (array_level a). + solve_all. Qed. Lemma R_eq_univ_prop_consistent_instances Σ univs u u' : @@ -885,7 +900,7 @@ Proof using Type. eapply All2_impl. eapply All_All_All2; eauto. lia. simpl; intros. - intuition. + intuition auto with *. Qed. Lemma untyped_subslet_inds Γ ind u u' mdecl : @@ -1324,9 +1339,12 @@ Proof using Hcf Hcf'. eapply eq_term_eq_term_prop_impl; eauto. now symmetry in a. - - depelim X2. - eapply inversion_Prim in X1 as [prim_ty' [cdecl' []]]; tea. - rewrite H in e. noconf e. eapply cumul_cumul_prop; eauto. pcuic. + - depelim X4. + eapply inversion_Prim in X3 as [prim_ty' [cdecl' []]]; tea. depelim o. + 1-2:rewrite H in e; noconf e; eapply cumul_cumul_prop; eauto; pcuic. + cbn in H, e2. rewrite H in e2. noconf e2. eapply cumul_cumul_prop; eauto; pcuic. + move: w; simp prim_type. intro. etransitivity; tea. constructor; fvs. cbn. fvs. + depelim X1. fvs. eapply eq_term_leq_term. symmetry; repeat constructor; eauto. Qed. End no_prop_leq_type. diff --git a/pcuic/theories/PCUICCumulativitySpec.v b/pcuic/theories/PCUICCumulativitySpec.v index 0b99ec6e7..431eb9dc3 100644 --- a/pcuic/theories/PCUICCumulativitySpec.v +++ b/pcuic/theories/PCUICCumulativitySpec.v @@ -34,22 +34,6 @@ Proof. all: repeat first [ assumption | toAll ]. Qed. -Inductive onPrim_dep (eq_term : term -> term -> Type) (Re : Universe.t -> Universe.t -> Prop) (eq_term_dep : forall x y, eq_term x y -> Type) (Re' : forall a b, Re a b -> Type) : forall x y : prim_val, onPrim eq_term Re x y -> Type := - | onPrimInt_dep i : onPrim_dep eq_term Re eq_term_dep Re' (primInt; primIntModel i) (primInt; primIntModel i) (onPrimInt eq_term Re i) - | onPrimFloat_dep f : onPrim_dep eq_term Re eq_term_dep Re' (primFloat; primFloatModel f) (primFloat; primFloatModel f) (onPrimFloat _ _ f) - | onPrimArray_dep a a' : - forall (hre : Re (Universe.make a.(array_level)) (Universe.make a'.(array_level))) - (eqdef : eq_term a.(array_default) a'.(array_default)) - (eqty : eq_term a.(array_type) a'.(array_type)) - (eqt : All2 eq_term a.(array_value) a'.(array_value)), - Re' _ _ hre -> - eq_term_dep _ _ eqdef -> - eq_term_dep _ _ eqty -> - All2_dep eq_term_dep eqt -> - onPrim_dep eq_term Re eq_term_dep Re' - (primArray; primArrayModel a) (primArray; primArrayModel a') (onPrimArray _ _ a a' hre eqdef eqty eqt). -Derive Signature for onPrim_dep. - Reserved Notation " Σ ;;; Γ ⊢ t ≤s[ pb ] u" (at level 50, Γ, t, u at next level, format "Σ ;;; Γ ⊢ t ≤s[ pb ] u"). @@ -164,7 +148,7 @@ Inductive cumulSpec0 {cf : checker_flags} (Σ : global_env_ext) Γ (pb : conv_pb Σ ;;; Γ ⊢ tCoFix mfix idx ≤s[pb] tCoFix mfix' idx | cumul_Prim p p' : - onPrim (fun x y => Σ ;;; Γ ⊢ x ≤s[Conv] y) (compare_universe Conv Σ) p p' -> + onPrims (fun x y => Σ ;;; Γ ⊢ x ≤s[Conv] y) (compare_universe Conv Σ) p p' -> Σ ;;; Γ ⊢ tPrim p ≤s[pb] tPrim p' (** Reductions *) @@ -445,8 +429,8 @@ Lemma cumulSpec0_rect : P cf Σ Γ pb (tCoFix mfix idx) (tCoFix mfix' idx) (cumul_CoFix _ _ _ _ _ _ H)) -> - (forall Γ pb p p' (e : onPrim (cumulSpec0 Σ Γ Conv) (eq_universe Σ) p p'), - onPrim_dep (cumulSpec0 Σ Γ Conv) (eq_universe Σ) (P cf Σ Γ Conv) (fun _ _ _ => True) p p' e -> + (forall Γ pb p p' (e : onPrims (cumulSpec0 Σ Γ Conv) (eq_universe Σ) p p'), + onPrims_dep (cumulSpec0 Σ Γ Conv) (eq_universe Σ) (P cf Σ Γ Conv) (fun _ _ _ => True) p p' e -> P cf Σ Γ pb (tPrim p) (tPrim p') (cumul_Prim _ _ _ _ _ e)) -> (* cumulativity rules *) @@ -724,8 +708,8 @@ Lemma convSpec0_ind_all : P cf Σ Γ Conv (tCoFix mfix idx) (tCoFix mfix' idx) (cumul_CoFix _ _ _ _ _ _ H)) -> - (forall Γ p p' (e : onPrim (cumulSpec0 Σ Γ Conv) (eq_universe Σ) p p'), - onPrim_dep (cumulSpec0 Σ Γ Conv) (eq_universe Σ) (P cf Σ Γ Conv) (fun _ _ _ => True) p p' e -> + (forall Γ p p' (e : onPrims (cumulSpec0 Σ Γ Conv) (eq_universe Σ) p p'), + onPrims_dep (cumulSpec0 Σ Γ Conv) (eq_universe Σ) (P cf Σ Γ Conv) (fun _ _ _ => True) p p' e -> P cf Σ Γ Conv (tPrim p) (tPrim p') (cumul_Prim _ _ _ _ _ e)) -> (* cumulativity rules *) diff --git a/pcuic/theories/PCUICEquality.v b/pcuic/theories/PCUICEquality.v index d4be6f3d4..b9db09530 100644 --- a/pcuic/theories/PCUICEquality.v +++ b/pcuic/theories/PCUICEquality.v @@ -275,17 +275,6 @@ Definition eq_predicate (eq_term : term -> term -> Type) Re p p' := ((eq_context_gen eq eq p.(pcontext) p'.(pcontext)) * eq_term p.(preturn) p'.(preturn))). -Inductive onPrim (eq_term : term -> term -> Type) Re : prim_val -> prim_val -> Type := - | onPrimInt i : onPrim eq_term Re (primInt; primIntModel i) (primInt; primIntModel i) - | onPrimFloat f : onPrim eq_term Re (primFloat; primFloatModel f) (primFloat; primFloatModel f) - | onPrimArray a a' : - Re (Universe.make a.(array_level)) (Universe.make a'.(array_level)) -> - eq_term a.(array_default) a'.(array_default) -> - eq_term a.(array_type) a'.(array_type) -> - All2 eq_term a.(array_value) a'.(array_value) -> - onPrim eq_term Re (primArray; primArrayModel a) (primArray; primArrayModel a'). -Derive Signature for onPrim. - (** ** Syntactic ws_cumul_pb up-to universes We don't look at printing annotations *) @@ -379,7 +368,7 @@ Inductive eq_term_upto_univ_napp Σ (Re Rle : Universe.t -> Universe.t -> Prop) Σ ⊢ tCoFix mfix idx <==[ Rle , napp ] tCoFix mfix' idx | eq_Prim i i' : - onPrim (eq_term_upto_univ_napp Σ Re Re 0) Re i i' -> + onPrims (eq_term_upto_univ_napp Σ Re Re 0) Re i i' -> eq_term_upto_univ_napp Σ Re Rle napp (tPrim i) (tPrim i') where " Σ ⊢ t <==[ Rle , napp ] u " := (eq_term_upto_univ_napp Σ _ Rle napp t u) : type_scope. @@ -508,15 +497,6 @@ Qed. Polymorphic Instance creflexive_eq A : CRelationClasses.Reflexive (@eq A). Proof. intro x. constructor. Qed. -Lemma onPrim_map_prop R R' Re p p' P f : tPrimProp P p -> - onPrim R Re p p' -> - (forall x y, P x -> R x y -> R' (f x) (f y)) -> - onPrim R' Re (map_prim f p) (map_prim f p'). -Proof. - destruct p as [? []]; cbn; intros h e; depelim e; intros hf; constructor; cbn; intuition eauto. - solve_all. -Qed. - #[global] Polymorphic Instance eq_predicate_refl Re Ru : CRelationClasses.Reflexive Re -> diff --git a/pcuic/theories/PCUICEtaExpand.v b/pcuic/theories/PCUICEtaExpand.v index 1385aacf6..dcaf8fbfb 100644 --- a/pcuic/theories/PCUICEtaExpand.v +++ b/pcuic/theories/PCUICEtaExpand.v @@ -47,16 +47,6 @@ Definition isConstruct t := Definition isRel t := match t with tRel _ => true | _ => false end. -Inductive onPrim (P : term -> Prop) : prim_val -> Prop := - | onPrimInt i : onPrim P (primInt; primIntModel i) - | onPrimFloat f : onPrim P (primFloat; primFloatModel f) - | onPrimArray a : - P a.(array_default) -> - P a.(array_type) -> - All P a.(array_value) -> - onPrim P (primArray; primArrayModel a). -Derive Signature for onPrim. - Section map_All. Context (P P' : term -> Type). Context (ont : forall x, P x -> P' x). diff --git a/pcuic/theories/PCUICExpandLets.v b/pcuic/theories/PCUICExpandLets.v index c7cbbbc08..76aa43586 100644 --- a/pcuic/theories/PCUICExpandLets.v +++ b/pcuic/theories/PCUICExpandLets.v @@ -41,7 +41,7 @@ Fixpoint trans (t : term) : term := | tCoFix mfix idx => let mfix' := List.map (map_def trans trans) mfix in tCoFix mfix' idx - | tPrim i => tPrim i + | tPrim i => tPrim (map_prim trans i) end. Notation trans_decl := (map_decl trans). diff --git a/pcuic/theories/PCUICExpandLetsCorrectness.v b/pcuic/theories/PCUICExpandLetsCorrectness.v index 3a2dfb985..6eee52285 100644 --- a/pcuic/theories/PCUICExpandLetsCorrectness.v +++ b/pcuic/theories/PCUICExpandLetsCorrectness.v @@ -245,6 +245,7 @@ Proof. intros i. eapply closedP_mon. lia. - f_equal. solve_all. - f_equal; red in X, X0; solve_all. + - f_equal; solve_all. Qed. Definition on_fst {A B C} (f:A->C) (p:A×B) := (f p.1, p.2). @@ -446,6 +447,7 @@ Proof. cbn. rtoProp; intuition auto. unfold map_def; cbn. f_equal. rewrite a //. solve_all. rewrite b //. solve_all. now len. + - f_equal. solve_all. eapply b. solve_all. Qed. Lemma trans_subst_ctx (Γ : context) xs k t : @@ -498,6 +500,7 @@ Proof. - f_equal. unfold tFixProp in X. rewrite !map_map_compose. autorewrite with map. solve_all_one. + - f_equal. solve_all. Qed. Lemma trans_subst_instance_ctx Γ u : @@ -929,6 +932,7 @@ Section wtsub. All (fun d => wt Γ d.(dtype) × wt (Γ ,,, fix_context mfix) d.(dbody)) mfix | tEvar _ l => False | tRel i => wf_local Σ Γ + | tPrim p => primitive_typing_hyps (fun Σ Γ t T => wt Γ t) Σ Γ p | _ => unit end. Import PCUICGeneration PCUICInversion. @@ -997,6 +1001,8 @@ Section wtsub. eapply All_prod. eapply (All_impl a). intros ? h; exact h. eapply (All_impl a0). intros ? h; eexists; tea. + - eapply inversion_Prim in h as (?&?&[]). + depelim p0; constructor; eauto. 1-2:now eexists. solve_all. now eexists. eauto. Qed. End wtsub. @@ -2195,6 +2201,13 @@ Proof. rewrite -(trans_fix_context (shiftnP #|Γ| xpred0) _ idx) //. now rewrite trans_local_app in X0. cbn; congruence. intros. repeat split; reflexivity. + + - cbn. eapply red_primArray_value. cbn. eapply All2_map. depelim wt. + eapply OnOne2_All_mix_left in X; tea. + eapply OnOne2_All2; tea; cbn; intuition eauto. + + - cbn; depelim wt. eapply red_primArray_default; cbn; eauto. + - cbn; depelim wt. eapply red_primArray_type; cbn; eauto. Qed. Lemma trans_R_global_instance {Σ : global_env} Re Rle gref napp u u' : @@ -2336,6 +2349,10 @@ Proof. cbn. eauto using subrelation_refl. - constructor. solve_all; eauto using subrelation_refl. - constructor; solve_all; eauto using subrelation_refl. + - constructor; depelim X0; cbn in X; constructor; cbn; intuition eauto. + * eapply a2; eauto using subrelation_refl. + * eapply a1; eauto using subrelation_refl. + * solve_all. eauto using subrelation_refl. Qed. Lemma trans_compare_term {cf} {Σ : global_env} {pb ϕ T U} : @@ -3580,6 +3597,11 @@ Proof. split;auto;eassumption. Qed. +Lemma trans_prim_ty p prim_ty : trans (prim_type p prim_ty) = prim_type (map_prim trans p) prim_ty. +Proof. + destruct p as [? []]; simp prim_type; cbn; reflexivity. +Qed. + Theorem pcuic_expand_lets {cf} (Σ : SE.global_env_ext) Γ t T : wf Σ -> typing Σ Γ t T -> @@ -3899,11 +3921,12 @@ Proof. eapply (subject_is_open_term (Σ := Σ)); tea. len in IHdb. eauto. + rewrite trans_wf_cofixpoint //. - - cbn. econstructor. - 3:eapply trans_declared_constant. all:eauto. - destruct X0 as [s []]; exists s; split => //. - * cbn. rewrite H1 => //. - * cbn. now rewrite H2. + - cbn. rewrite trans_prim_ty. econstructor; eauto. rewrite prim_val_tag_map //. + * now eapply trans_declared_constant in H0. + * destruct p as [? []]; cbn in X0 |- *. + 1-2:destruct X0 as [s []]; exists s; split => //; rewrite ?H1 ?H2 //=. + destruct X0; split => //. rewrite H1 //. rewrite H2 //. + * depelim X1; depelim X2; constructor; intuition eauto. cbn. solve_all. - eapply (type_ws_cumul_pb (pb:=Cumul)). + eauto. + now exists s. @@ -5471,6 +5494,7 @@ Proof. 4:eauto. len. cbn. now rewrite context_assumptions_smash_context context_assumptions_map /=. solve_all. + - depelim H0; constructor; cbn; eauto. solve_all. Qed. Lemma expanded_trans_local {cf:checker_flags} diff --git a/pcuic/theories/PCUICFirstorder.v b/pcuic/theories/PCUICFirstorder.v index 549af512f..d8e9da6f9 100644 --- a/pcuic/theories/PCUICFirstorder.v +++ b/pcuic/theories/PCUICFirstorder.v @@ -671,8 +671,8 @@ Proof using Type. eapply andb_true_iff in Hfo as [Hfo _]. rewrite /check_recursivity_kind E in Hty. now eapply (negb_False _ Hfo). - + eapply inversion_Prim in Hty as [prim_ty [cdecl [wf hp hdecl [s []] cum]]]; eauto. - now eapply invert_cumul_axiom_ind in cum; tea. + + eapply inversion_Prim in Hty as [prim_ty [cdecl [wf hp hdecl inv ty cum]]]; eauto. + now eapply invert_cumul_prim_type_ind in cum; tea. - destruct t; inv Hhead. + exfalso. now eapply invert_ind_ind in Hty. + apply inversion_mkApps in Hty as Hcon; auto. diff --git a/pcuic/theories/PCUICNormal.v b/pcuic/theories/PCUICNormal.v index c33225406..cd0180bb1 100644 --- a/pcuic/theories/PCUICNormal.v +++ b/pcuic/theories/PCUICNormal.v @@ -933,7 +933,7 @@ Proof. apply nth_error_None in H. lia. - eapply red1_mkApps_tCoFix_inv in r as [[(?&->&?)|(?&->&?)]|(?&->&?)]; eauto. - - depelim r. solve_discr. + - depelim r; solve_discr; eapply whnf_prim. Qed. Lemma whnf_pres {cf:checker_flags} Σ {wfΣ : wf Σ} Γ t t' : @@ -1007,7 +1007,9 @@ Inductive whnf_red Σ Γ : term -> term -> Type := red Σ (Γ,,, fix_context mfix) (dbody d) (dbody d')) mfix mfix' -> whnf_red Σ Γ (tCoFix mfix idx) (tCoFix mfix' idx) -| whnf_red_tPrim i : whnf_red Σ Γ (tPrim i) (tPrim i). +| whnf_red_tPrim i i' : + onPrims (red Σ Γ) eq i i' -> + whnf_red Σ Γ (tPrim i) (tPrim i'). Derive Signature for whnf_red. @@ -1048,6 +1050,9 @@ Proof. cbn. intros ? ? (->&->&r1&r2). eauto. + - depelim o. 1-2: reflexivity. + eapply red_primArray_congr; eauto. + now eapply Universe.make_inj in e. Qed. #[global] @@ -1150,6 +1155,7 @@ Proof. 2: apply All2_same; auto. constructor. apply All2_same; auto. + - constructor. destruct p as [? []]; constructor; eauto. eapply All2_same; auto. Qed. Ltac inv_on_free_vars := @@ -1403,6 +1409,12 @@ Proof. eapply context_pres_let_bodies_red; eauto. apply fix_context_pres_let_bodies. now apply All2_length in a. + - constructor. depelim o; depelim o0; constructor; eauto. + * rewrite -x //. + * etransitivity; tea. + * etransitivity; tea. + * eapply All2_trans; eauto. + tc. Qed. Lemma whne_red1_inv {cf:checker_flags} Σ {wfΣ : wf Σ} Γ t t' : @@ -1509,7 +1521,10 @@ Proof. cbn. intros ? ? (?&[= -> -> ->]). auto. - - depelim r; solve_discr. + - depelim r; solve_discr; constructor; eauto. + * constructor; eauto. eapply OnOne2_All2; eauto. + * constructor; eauto. cbn. eapply All2_same; eauto. + * constructor; eauto. cbn. eapply All2_same; eauto. Qed. Lemma whnf_red_inv {cf:checker_flags} {Σ : global_env_ext} Γ t t' : diff --git a/pcuic/theories/PCUICParallelReduction.v b/pcuic/theories/PCUICParallelReduction.v index 74cc41989..a2073ad3f 100644 --- a/pcuic/theories/PCUICParallelReduction.v +++ b/pcuic/theories/PCUICParallelReduction.v @@ -401,7 +401,7 @@ Section ParallelReduction. | prim_pred p p' : pred1_ctx Γ Γ' -> - PCUICEquality.onPrim (pred1 Γ Γ') eq p p' -> pred1 Γ Γ' (tPrim p) (tPrim p') + onPrims (pred1 Γ Γ') eq p p' -> pred1 Γ Γ' (tPrim p) (tPrim p') | pred_atom_refl t : pred1_ctx Γ Γ' -> @@ -615,8 +615,8 @@ Section ParallelReduction. (forall (Γ Γ' : context) (p p' : prim_val), pred1_ctx Γ Γ' -> Pctx Γ Γ' -> - PCUICEquality.onPrim (pred1 Γ Γ') eq p p' -> - PCUICEquality.onPrim (P Γ Γ') eq p p' -> + onPrims (pred1 Γ Γ') eq p p' -> + onPrims (P Γ Γ') eq p p' -> P Γ Γ' (tPrim p) (tPrim p')) -> (forall (Γ Γ' : context) (t : term), diff --git a/pcuic/theories/PCUICPrincipality.v b/pcuic/theories/PCUICPrincipality.v index e09d17d4f..b7ab280db 100644 --- a/pcuic/theories/PCUICPrincipality.v +++ b/pcuic/theories/PCUICPrincipality.v @@ -343,10 +343,10 @@ Section Principality. repeat split; eauto. eapply type_CoFix; eauto. - apply inversion_Prim in hA as [prim_ty [cdecl []]] => //; pcuic. - exists (tConst prim_ty []). + exists (prim_type p prim_ty). intros B hB. apply inversion_Prim in hB as [prim_ty' [cdecl' []]] => //; pcuic. - econstructor; tea. + econstructor; tea; fvs. Qed. (** A weaker version that is often convenient to use. *) @@ -463,6 +463,18 @@ Proof. intros. now eapply eq_term_empty_eq_term. Qed. +Lemma All2_tip_l {A B} {R : A -> B -> Type} x y : All2 R [x] y -> ∑ y', (y = [y']) * R x y'. +Proof. + intros a; depelim a. depelim a. exists y0; split => //. +Qed. + +Lemma Forall2_tip_l {A B} {R : A -> B -> Prop} x y : Forall2 R [x] y -> exists y', (y = [y']) * R x y'. +Proof. + intros a; depelim a. depelim a. exists y0; split => //. +Qed. + +From MetaCoq.PCUIC Require Import PCUICClassification. + Lemma typing_leq_term {cf:checker_flags} (Σ : global_env_ext) Γ t t' T T' : wf Σ.1 -> on_udecl Σ.1 Σ.2 -> @@ -739,8 +751,25 @@ Proof. constructor. apply eq_term_empty_leq_term in eqty. now eapply leq_term_empty_leq_term. - - depelim X2. + - epose proof (type_Prim _ _ _ _ _ X H H0 X0 X1). eapply validity in X5. + depelim X2; depelim X4; depelim o. + 1-2:econstructor; tea. + depelim X1. destruct X5 as [s ?]. econstructor; tea. + eapply inversion_Prim in X3 as [prim_ty' [cdecl' []]]; eauto. + rewrite e2 in H. noconf H. + unshelve eapply declared_constant_to_gen in H0, d; tea. + pose proof (declared_constant_inj _ _ H0 d). subst cdecl'. + simp prim_type in w |- *. + eapply (ws_cumul_pb_Axiom_l_inv (args := [_])) in w as [u' [args' []]]; tea. 2:eapply declared_constant_from_gen, H0. 2:eapply p. + eapply cumulAlgo_cumulSpec. etransitivity. now eapply red_ws_cumul_pb. + eapply All2_tip_l in a3 as [y' [-> Heq]]. red in r. eapply Forall2_map_inv in r. + eapply Forall2_tip_l in r. cbn. eapply ws_cumul_pb_eq_le. + eapply (ws_cumul_pb_mkApps (args := [_]) (args' := [_])). + * constructor; fvs. constructor. red. eapply Forall2_map. destruct r as [? [-> eq]]. constructor. symmetry. + etransitivity; tea. now symmetry. constructor. + * constructor; [|constructor]. symmetry. etransitivity; tea. constructor; fvs. symmetry. + now eapply eq_term_empty_eq_term. - eapply type_Cumul'. eapply X1; eauto. now exists s. diff --git a/pcuic/theories/PCUICProgress.v b/pcuic/theories/PCUICProgress.v index 46374b542..e737ac41d 100644 --- a/pcuic/theories/PCUICProgress.v +++ b/pcuic/theories/PCUICProgress.v @@ -256,12 +256,14 @@ forall (P : global_env_ext -> context -> term -> term -> Type) wf_cofixpoint Σ.1 mfix -> P Σ Γ (tCoFix mfix n) decl.(dtype)) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (p : prim_val) prim_ty cdecl, - PΓ Σ Γ -> - primitive_constant Σ.1 (prim_val_tag p) = Some prim_ty -> - declared_constant Σ.1 prim_ty cdecl -> - primitive_invariants cdecl -> - P Σ Γ (tPrim p) (tConst prim_ty [])) -> + (forall (Σ : global_env_ext) (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (p : prim_val) prim_ty cdecl, + PΓ Σ Γ -> + primitive_constant Σ (prim_val_tag p) = Some prim_ty -> + declared_constant Σ prim_ty cdecl -> + primitive_invariants (prim_val_tag p) cdecl -> + primitive_typing_hyps typing Σ Γ p -> + primitive_typing_hyps P Σ Γ p -> + P Σ Γ (tPrim p) (prim_type p prim_ty)) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (t A B : term) s, PΓ Σ Γ -> @@ -437,12 +439,16 @@ Lemma typing_ind_env `{cf : checker_flags} : wf_cofixpoint Σ.1 mfix -> P Σ Γ (tCoFix mfix n) decl.(dtype)) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (p : prim_val) prim_ty cdecl, - PΓ Σ Γ -> - primitive_constant Σ.1 (prim_val_tag p) = Some prim_ty -> - declared_constant Σ.1 prim_ty cdecl -> - primitive_invariants cdecl -> - P Σ Γ (tPrim p) (tConst prim_ty [])) -> + + (forall (Σ : global_env_ext) (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (p : prim_val) prim_ty cdecl, + PΓ Σ Γ -> + primitive_constant Σ (prim_val_tag p) = Some prim_ty -> + declared_constant Σ prim_ty cdecl -> + primitive_invariants (prim_val_tag p) cdecl -> + primitive_typing_hyps typing Σ Γ p -> + primitive_typing_hyps P Σ Γ p -> + P Σ Γ (tPrim p) (prim_type p prim_ty)) -> + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (t A B : term) s, PΓ Σ Γ -> @@ -554,17 +560,29 @@ Proof. now eapply ws_cumul_pb_Sort_Prod_inv in w. Qed. -Lemma typing_spine_axiom {cf : checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} Γ cst u cdecl args T : +Lemma typing_spine_axiom {cf : checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} Γ cst u cdecl args args' T : declared_constant Σ cst cdecl -> cdecl.(cst_body) = None -> - typing_spine Σ Γ (tConst cst u) args T -> args = []. + typing_spine Σ Γ (mkApps (tConst cst u) args) args' T -> args' = []. Proof. intros hdecl hb. - induction args => //. + induction args' => //. intros sp. depelim sp. now eapply invert_cumul_axiom_prod in w. Qed. +Lemma typing_spine_prim_type {cf : checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} Γ prim prim_ty cdecl args T : + declared_constant Σ prim_ty cdecl -> + primitive_invariants (prim_val_tag prim) cdecl -> + typing_spine Σ Γ (prim_type prim prim_ty) args T -> args = []. +Proof. + intros hdecl hb. + induction args => //. + destruct prim as [? []]; cbn in *; intros sp; destruct hb; simp prim_type in *. + 1-2:destruct a0; eapply (typing_spine_axiom _ _ _ _ []) in sp; tea. + eapply (typing_spine_axiom _ _ _ _ [array_type a0]) in sp; tea. +Qed. + Lemma typing_value_head_napp {cf : checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} fn args hd T : negb (isApp fn) -> Σ ;;; [] |- mkApps fn (args ++ [hd]) : T -> @@ -621,9 +639,9 @@ Proof. now constructor. * (* primitive *) cbn. - eapply inversion_Prim in hfn as [prim_ty [cdecl [hwf hp hdecl [s []]]]]; tea. - eapply typing_spine_strengthen in hcum. 3:tea. 2:{ eapply validity; econstructor; eauto. now exists s. } - now eapply typing_spine_axiom, app_tip_nil in hcum. + eapply inversion_Prim in hfn as [prim_ty [cdecl [hwf hp hdecl hinv hty hcum']]]; tea. + eapply typing_spine_strengthen in hcum. 3:tea. 2:{ eapply validity; econstructor; eauto. } + now eapply typing_spine_prim_type, app_tip_nil in hcum. Qed. Lemma typing_value_head {cf : checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} fn args hd T : diff --git a/pcuic/theories/PCUICWfUniverses.v b/pcuic/theories/PCUICWfUniverses.v index bc413d4da..ed40f49a1 100644 --- a/pcuic/theories/PCUICWfUniverses.v +++ b/pcuic/theories/PCUICWfUniverses.v @@ -14,8 +14,7 @@ Require Import ssreflect ssrbool. From MetaCoq.PCUIC Require Import PCUICInduction. - -Lemma test_primu_mapu pu pt fu ft p : +Lemma test_primu_mapu {term term'} pu pt fu (ft : term -> term') p : test_primu pu pt (mapu_prim fu ft p) = test_primu (pu ∘ fu) (pt ∘ ft) p. Proof. @@ -23,7 +22,7 @@ Proof. now rewrite forallb_map. Qed. -Lemma test_primu_primProp P pu pt pu' pt' p : +Lemma test_primu_primProp {term} P pu (pt : term -> bool) pu' (pt' : term -> bool) p : tPrimProp P p -> (forall x, pu x = pu' x) -> (forall x, P x -> pt x = pt' x) -> @@ -34,12 +33,9 @@ Proof. solve_all. Qed. - - Section CheckerFlags. Context {cf:checker_flags}. - Lemma wf_universe_type0 Σ : wf_universe Σ Universe.type0. Proof using Type. simpl. @@ -325,8 +321,6 @@ Section CheckerFlags. Definition wf_universes t := on_universes wf_universeb closedu t. - - Lemma wf_universeb_instance_forall u : forallb wf_universeb (map Universe.make u) = wf_universeb_instance Σ u. Proof using Type. diff --git a/pcuic/theories/utils/PCUICPrimitive.v b/pcuic/theories/utils/PCUICPrimitive.v index b3502ce76..04b856b97 100644 --- a/pcuic/theories/utils/PCUICPrimitive.v +++ b/pcuic/theories/utils/PCUICPrimitive.v @@ -152,3 +152,92 @@ Definition string_of_prim {term} (soft : term -> string) (p : prim_val term) : s | primFloatModel f => "(float: " ^ string_of_float f ^ ")" | primArrayModel a => "(array:" ^ string_of_list soft a.(array_value) ^ ")" end. + +(** Predicates *) + +Inductive onPrim {term} (P : term -> Prop) : prim_val term -> Prop := + | onPrimInt i : onPrim P (primInt; primIntModel i) + | onPrimFloat f : onPrim P (primFloat; primFloatModel f) + | onPrimArray a : + P a.(array_default) -> + P a.(array_type) -> + All P a.(array_value) -> + onPrim P (primArray; primArrayModel a). +Derive Signature for onPrim. + +Inductive onPrims {term} (eq_term : term -> term -> Type) Re : prim_val term -> prim_val term -> Type := + | onPrimsInt i : onPrims eq_term Re (primInt; primIntModel i) (primInt; primIntModel i) + | onPrimsFloat f : onPrims eq_term Re (primFloat; primFloatModel f) (primFloat; primFloatModel f) + | onPrimsArray a a' : + Re (Universe.make a.(array_level)) (Universe.make a'.(array_level)) -> + eq_term a.(array_default) a'.(array_default) -> + eq_term a.(array_type) a'.(array_type) -> + All2 eq_term a.(array_value) a'.(array_value) -> + onPrims eq_term Re (primArray; primArrayModel a) (primArray; primArrayModel a'). +Derive Signature NoConfusion for onPrims. + +Definition tPrimProp {term} (P : term -> Type) (p : PCUICPrimitive.prim_val term) : Type := + match p.π2 return Type with + | primIntModel f => unit + | primFloatModel f => unit + | primArrayModel a => P a.(array_type) × P a.(array_default) × All P a.(array_value) + end. + + +Inductive onPrims_dep {term} (eq_term : term -> term -> Type) (Re : Universe.t -> Universe.t -> Prop) (eq_term_dep : forall x y, eq_term x y -> Type) (Re' : forall a b, Re a b -> Type) : forall x y : prim_val term, onPrims eq_term Re x y -> Type := + | onPrimsInt_dep i : onPrims_dep eq_term Re eq_term_dep Re' (primInt; primIntModel i) (primInt; primIntModel i) (onPrimsInt eq_term Re i) + | onPrimsFloat_dep f : onPrims_dep eq_term Re eq_term_dep Re' (primFloat; primFloatModel f) (primFloat; primFloatModel f) (onPrimsFloat _ _ f) + | onPrimsArray_dep a a' : + forall (hre : Re (Universe.make a.(array_level)) (Universe.make a'.(array_level))) + (eqdef : eq_term a.(array_default) a'.(array_default)) + (eqty : eq_term a.(array_type) a'.(array_type)) + (eqt : All2 eq_term a.(array_value) a'.(array_value)), + Re' _ _ hre -> + eq_term_dep _ _ eqdef -> + eq_term_dep _ _ eqty -> + All2_dep eq_term_dep eqt -> + onPrims_dep eq_term Re eq_term_dep Re' + (primArray; primArrayModel a) (primArray; primArrayModel a') (onPrimsArray _ _ a a' hre eqdef eqty eqt). +Derive Signature for onPrims_dep. + +Set Equations Transparent. + +Definition mapu_array_model {term term'} (fl : Level.t -> Level.t) (f : term -> term') + (ar : array_model term) : array_model term' := + {| array_level := fl ar.(array_level); + array_value := map f ar.(array_value); + array_default := f ar.(array_default); + array_type := f ar.(array_type) |}. + +Equations mapu_prim {term term'} (f : Level.t -> Level.t) (g : term -> term') + (p : PCUICPrimitive.prim_val term) : PCUICPrimitive.prim_val term' := +| _, _, (primInt; primIntModel i) => (primInt; primIntModel i) +| _, _, (primFloat; primFloatModel fl) => (primFloat; primFloatModel fl) +| f, g, (primArray; primArrayModel ar) => + (primArray; primArrayModel (mapu_array_model f g ar)). + +Notation map_array_model := (mapu_array_model id). +Notation map_prim := (mapu_prim id). + +Equations test_prim {term} (p : term -> bool) (p : prim_val term) : bool := +| p, (primInt; _) => true +| p, (primFloat; _) => true +| p, (primArray; primArrayModel ar) => + List.forallb p ar.(array_value) && p ar.(array_default) && p ar.(array_type). + +Equations test_primu {term} (p : Level.t -> bool) (t : term -> bool) (p : prim_val term) : bool := +| _, _, (primInt; _) => true +| _, _, (primFloat; _) => true +| p, pt, (primArray; primArrayModel ar) => + p ar.(array_level) && forallb pt ar.(array_value) && + pt ar.(array_default) && pt ar.(array_type). + +Lemma onPrims_map_prop {term term'} R R' Re p p' P f : @tPrimProp term P p -> + onPrims R Re p p' -> + (forall x y, P x -> R x y -> R' (f x) (f y)) -> + onPrims (term:=term') R' Re (map_prim f p) (map_prim f p'). +Proof. + destruct p as [? []]; cbn; intros h e; depelim e; intros hf; constructor; cbn; intuition eauto. + eapply All2_All_mix_left in a1; tea. + eapply All2_map, All2_impl; tea; cbn; intuition eauto. +Qed. From e93c39cae5a7f8d0267286d97ade963f9f2d6dc3 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Fri, 17 Nov 2023 17:49:00 +0100 Subject: [PATCH 06/17] Done with template <-> PCUIC. Working on safe checker --- .vscode/metacoq.code-workspace | 71 ++++++++++++++----- pcuic/theories/Syntax/PCUICPosition.v | 57 ++++++++++++++- pcuic/theories/Syntax/PCUICReflect.v | 2 +- pcuic/theories/utils/PCUICPrimitive.v | 22 +++--- safechecker/theories/PCUICEqualityDec.v | 41 ++++++++++- safechecker/theories/PCUICErrors.v | 53 ++++++++++++++ safechecker/theories/PCUICSafeConversion.v | 27 +++++++ safechecker/theories/PCUICSafeReduce.v | 12 ++-- safechecker/theories/PCUICSafeRetyping.v | 2 +- template-coq/theories/TermEquality.v | 7 +- template-coq/theories/Typing.v | 31 +++++++- template-coq/theories/TypingWf.v | 3 +- template-coq/theories/WcbvEval.v | 1 + template-pcuic/theories/PCUICToTemplate.v | 5 +- .../theories/PCUICToTemplateCorrectness.v | 29 +++++--- .../theories/TemplateMonadToPCUIC.v | 7 +- template-pcuic/theories/TemplateToPCUIC.v | 5 ++ .../theories/TemplateToPCUICCorrectness.v | 44 +++++++++++- .../theories/TemplateToPCUICExpanded.v | 3 + .../theories/TemplateToPCUICWcbvEval.v | 1 + 20 files changed, 359 insertions(+), 64 deletions(-) diff --git a/.vscode/metacoq.code-workspace b/.vscode/metacoq.code-workspace index 6a0a8cce3..d6ac88a4c 100644 --- a/.vscode/metacoq.code-workspace +++ b/.vscode/metacoq.code-workspace @@ -11,31 +11,64 @@ - "-R", "safechecker-plugin/theories", "MetaCoq.SafeCheckerPlugin", + "-R", "safechecker-plugin/theories,MetaCoq.SafeCheckerPlugin", - "-R", "utils/theories", "MetaCoq.Utils", - "-R", "common/theories", "MetaCoq.Common", - "-R", "template-pcuic/theories", "MetaCoq.TemplatePCUIC", + "-R", "utils/theories,MetaCoq.Utils", + "-R", "common/theories,MetaCoq.Common", + "-R", "template-pcuic/theories,MetaCoq.TemplatePCUIC", - "-I", "template-coq", + "-I template-coq", // "-bt", get backtraces from Coq on errors - "-I", "template-coq/build", - "-R", "template-coq/theories", "MetaCoq.Template", - "-R", "examples", "MetaCoq.Examples", - "-R", "checker/theories", "MetaCoq.Checker", - "-R", "pcuic/theories", "MetaCoq.PCUIC", - "-I", "safechecker/src", - "-R", "safechecker/theories", "MetaCoq.SafeChecker", - "-I", "erasure/src", - "-R", "erasure/theories", "MetaCoq.Erasure", - "-R", "erasure-plugin/theories", "MetaCoq.ErasurePlugin", - "-R", "translations", "MetaCoq.Translations", - "-R", "test-suite/plugin-demo/theories", "MetaCoq.ExtractedPluginDemo", - "-R", "test-suite", "MetaCoq.TestSuite" + "-I template-coq/build", + "-R", "template-coq/theories,MetaCoq.Template", + "-R", "examples,MetaCoq.Examples", + "-R", "checker/theories,MetaCoq.Checker", + "-R", "pcuic/theories,MetaCoq.PCUIC", + "-I safechecker/src", + "-R", "safechecker/theories,MetaCoq.SafeChecker", + "-I erasure/src", + "-R", "erasure/theories,MetaCoq.Erasure", + "-R", "erasure-plugin/theories,MetaCoq.ErasurePlugin", + "-R", "translations,MetaCoq.Translations", + "-R", "test-suite/plugin-demo/theories,MetaCoq.ExtractedPluginDemo", + "-R", "test-suite,MetaCoq.TestSuite" ], // When enabled, will trim trailing whitespace when saving a file. "files.trimTrailingWhitespace": true, - "coqtop.binPath": "_opam/bin" + "coqtop.binPath": "_opam/bin", + "files.exclude": { + "**/*.vo": true, + "**/*.vok": true, + "**/*.vos": true, + "**/*.aux": true, + "**/*.glob": true, + "**/.git": true, + "**/.svn": true, + "**/.hg": true, + "**/CVS": true, + "**/.DS_Store": true, + "**/Thumbs.db": true + }, + "coq-lsp.args": [ + "-R", "safechecker-plugin/theories,MetaCoq.SafeCheckerPlugin", + "-R", "utils/theories,MetaCoq.Utils", + "-R", "common/theories,MetaCoq.Common", + "-R", "template-pcuic/theories,MetaCoq.TemplatePCUIC", + + "-I", "template-coq", + //"-bt", // get backtraces from Coq on errors + //"-Itemplate-coq/build", + "-R", "template-coq/theories,MetaCoq.Template", + "-R", "examples,MetaCoq.Examples", + "-R", "pcuic/theories,MetaCoq.PCUIC", + "-R", "safechecker/theories,MetaCoq.SafeChecker", + "-R", "erasure/theories,MetaCoq.Erasure", + "-R", "erasure-plugin/theories,MetaCoq.ErasurePlugin", + "-R", "translations,MetaCoq.Translations", + "-R", "test-suite/plugin-demo/theories,MetaCoq.ExtractedPluginDemo", + "-R", "test-suite,MetaCoq.TestSuite" + ], + "coq-lsp.path": "_opam/bin/coq-lsp" } } diff --git a/pcuic/theories/Syntax/PCUICPosition.v b/pcuic/theories/Syntax/PCUICPosition.v index 18ee68f23..439689ef7 100644 --- a/pcuic/theories/Syntax/PCUICPosition.v +++ b/pcuic/theories/Syntax/PCUICPosition.v @@ -47,7 +47,10 @@ Inductive choice := | prod_r | let_bd | let_ty -| let_in. +| let_in +| array_def +| array_ty +| array_val (n : nat). Derive NoConfusion NoConfusionHom EqDec for choice. @@ -115,6 +118,13 @@ Fixpoint validpos t (p : position) {struct p} := | let_bd, tLetIn na b B t => validpos b p | let_ty, tLetIn na b B t => validpos B p | let_in, tLetIn na b B t => validpos t p + | array_def, tPrim (primArray; primArrayModel a) => validpos (array_default a) p + | array_ty, tPrim (primArray; primArrayModel a) => validpos (array_type a) p + | array_val n, tPrim (primArray; primArrayModel a) => + match nth_error (array_value a) n with + | Some d => validpos d p + | None => false + end | _, _ => false end end. @@ -253,6 +263,11 @@ Proof. * 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. + + depelim e. depelim o; eauto. + + depelim e; depelim o; eauto. + + depelim e. depelim o; eauto. + cbn in vp. destruct (nth_error (array_value a) n) eqn:e'; try discriminate. + cbn. eapply All2_nth_error_Some in a0 as [t' [-> ?]]; tea. eapply ih; tea. Qed. Lemma eq_term_valid_pos : @@ -690,6 +705,34 @@ Proof. unshelve eapply Acc_cofix_mfix_bd with (1 := e1) (p := exist q _). -- simpl. rewrite e1 in e. assumption. -- eapply ihm. + - destruct q as [q e]. destruct q as [| c q]. + + constructor. intros [p' e'] h. + unfold posR in h. cbn in h. + dependent destruction h. + destruct p as [? []]; + destruct c. all: noconf e'. + * simpl in e'. + case_eq (nth_error m n0). + 2:{ intro h. pose proof e' as hh. rewrite h in hh. discriminate. } + intros [na ty bo ra] e1. + eapply All_nth_error in X as ihm. 2: exact e1. + simpl in ihm. + unshelve eapply Acc_cofix_mfix_ty with (1 := e1) (p := exist p _). + -- simpl. rewrite e1 in e'. assumption. + -- eapply ihm. + * simpl in e'. + case_eq (nth_error m n0). + 2:{ intro h. pose proof e' as hh. rewrite h in hh. discriminate. } + intros [na ty bo ra] e1. + eapply All_nth_error in X as ihm. 2: exact e1. + simpl in ihm. + unshelve eapply Acc_cofix_mfix_bd with (1 := e1) (p := exist p _). + -- simpl. rewrite e1 in e'. assumption. + -- eapply ihm. + + + destruct p as [? []]; cbn in X, e; intuition eauto; cbn in e. + + destruct x eqn:hp; try discriminate; cbn in e. constructor. Qed. Fixpoint atpos t (p : position) {struct p} : term := @@ -933,7 +976,11 @@ Variant stack_entry : Type := | Lambda_bd (na : aname) (A : term) | LetIn_bd (na : aname) (B t : term) | LetIn_ty (na : aname) (b t : term) -| LetIn_in (na : aname) (b B : term). +| LetIn_in (na : aname) (b B : term) +| PrimArray_ty (l : Level.t) (l : list term) (def : term) +| PrimArray_def (l : Level.t) (l : list term) (ty : term) +(* Hole in one of the values *) +| PrimArray_val (l : Level.t) (bef : list term) (after : list term) (def : term) (ty : term). Definition stack := list stack_entry. @@ -1029,6 +1076,9 @@ Definition fill_hole (t : term) (se : stack_entry) : term := | LetIn_bd na B u => tLetIn na t B u | LetIn_ty na b u => tLetIn na b t u | LetIn_in na b B => tLetIn na b B t + | PrimArray_def l v ty => tPrim (primArray; primArrayModel {| array_level := l; array_value := v; array_default := t; array_type := ty |}) + | PrimArray_ty l v def => tPrim (primArray; primArrayModel {| array_level := l; array_value := v; array_default := def; array_type := t |}) + | PrimArray_val l bef after def ty => tPrim (primArray; primArrayModel {| array_level := l; array_value := bef ++ (t :: after); array_default := def; array_type := ty |}) end. (* Not using fold_left here to get the right unfolding behavior *) @@ -1307,6 +1357,9 @@ Definition closedn_stack_entry k se := | LetIn_bd na B u => closedn k B && closedn (S k) u | LetIn_ty na b u => closedn k b && closedn (S k) u | LetIn_in na b B => closedn k b && closedn k B + | PrimArray_def l v ty => forallb (closedn k) v && closedn k ty + | PrimArray_ty l v def => forallb (closedn k) v && closedn k def + | PrimArray_val l bef after def ty => forallb (closedn k) bef && forallb (closedn k) after && closedn k def && closedn k ty end. Fixpoint closedn_stack k π := diff --git a/pcuic/theories/Syntax/PCUICReflect.v b/pcuic/theories/Syntax/PCUICReflect.v index 92ae085e0..47fafc878 100644 --- a/pcuic/theories/Syntax/PCUICReflect.v +++ b/pcuic/theories/Syntax/PCUICReflect.v @@ -115,7 +115,7 @@ Fixpoint eqb_term (u v : term) : bool := eqb x.(rarg) y.(rarg) && eqb x.(dname) y.(dname)) mfix mfix' - | tPrim p, tPrim p' => @eqb_prim_val _ eqb_term p p' + | tPrim p, tPrim p' => @eqb_prim_val _ eqb eqb_term p p' | _, _ => false end. diff --git a/pcuic/theories/utils/PCUICPrimitive.v b/pcuic/theories/utils/PCUICPrimitive.v index 04b856b97..ac205a86e 100644 --- a/pcuic/theories/utils/PCUICPrimitive.v +++ b/pcuic/theories/utils/PCUICPrimitive.v @@ -78,15 +78,15 @@ Instance reflect_eq_spec_float : ReflectEq SpecFloat.spec_float := EqDec_Reflect Import ReflectEq. -Definition eqb_array {term} {eqt : term -> term -> bool} (x y : array_model term) : bool := - eqb x.(array_level) y.(array_level) && +Definition eqb_array {term} {equ : Level.t -> Level.t -> bool} {eqt : term -> term -> bool} (x y : array_model term) : bool := + equ x.(array_level) y.(array_level) && forallb2 eqt x.(array_value) y.(array_value) && eqt x.(array_default) y.(array_default) && eqt x.(array_type) y.(array_type). #[program,global] Instance reflect_eq_array {term} {req : ReflectEq term}: ReflectEq (array_model term) := - { eqb := eqb_array (eqt := eqb) }. + { eqb := eqb_array (equ := eqb) (eqt := eqb) }. Next Obligation. intros term req [] []; cbn. unfold eqb_array. cbn. change (forallb2 eqb) with (eqb (A := list term)). @@ -97,14 +97,14 @@ Next Obligation. all:constructor; congruence. Qed. -Equations eqb_prim_model {term} {req : term -> term -> bool} {t : prim_tag} (x y : prim_model term t) : bool := +Equations eqb_prim_model {term} {equ : Level.t -> Level.t -> bool} {req : term -> term -> bool} {t : prim_tag} (x y : prim_model term t) : bool := | primIntModel x, primIntModel y := ReflectEq.eqb x y | primFloatModel x, primFloatModel y := ReflectEq.eqb x y - | primArrayModel x, primArrayModel y := eqb_array (eqt:=req) x y. + | primArrayModel x, primArrayModel y := eqb_array (equ := equ) (eqt:=req) x y. #[global, program] Instance prim_model_reflecteq {term} {req : ReflectEq term} {p : prim_tag} : ReflectEq (prim_model term p) := - {| ReflectEq.eqb := eqb_prim_model (req := eqb) |}. + {| ReflectEq.eqb := eqb_prim_model (equ := eqb) (req := eqb) |}. Next Obligation. intros. depelim x; depelim y; simp eqb_prim_model. case: ReflectEq.eqb_spec; constructor; subst; auto. congruence. @@ -116,15 +116,15 @@ Qed. #[global] Instance prim_model_eqdec {term} {req : ReflectEq term} : forall p : prim_tag, EqDec (prim_model term p) := _. -Equations eqb_prim_val {term} {req : term -> term -> bool} (x y : prim_val term) : bool := - | (primInt; i), (primInt; i') := eqb_prim_model (req := req) i i' - | (primFloat; f), (primFloat; f') := eqb_prim_model (req := req) f f' - | (primArray; a), (primArray; a') := eqb_prim_model (req := req) a a' +Equations eqb_prim_val {term} {equ : Level.t -> Level.t -> bool} {req : term -> term -> bool} (x y : prim_val term) : bool := + | (primInt; i), (primInt; i') := eqb_prim_model (equ := equ) (req := req) i i' + | (primFloat; f), (primFloat; f') := eqb_prim_model (equ := equ) (req := req) f f' + | (primArray; a), (primArray; a') := eqb_prim_model (equ := equ) (req := req) a a' | x, y := false. #[global, program] Instance prim_val_reflect_eq {term} {req : ReflectEq term} : ReflectEq (prim_val term) := - {| ReflectEq.eqb := eqb_prim_val (req := eqb) |}. + {| ReflectEq.eqb := eqb_prim_val (equ := eqb) (req := eqb) |}. Next Obligation. intros. funelim (eqb_prim_val x y); simp eqb_prim_val. change (eqb_prim_model i i') with (eqb i i'). diff --git a/safechecker/theories/PCUICEqualityDec.v b/safechecker/theories/PCUICEqualityDec.v index b79ef3c25..1d6c7ccc6 100644 --- a/safechecker/theories/PCUICEqualityDec.v +++ b/safechecker/theories/PCUICEqualityDec.v @@ -2,7 +2,7 @@ From Coq Require Import ProofIrrelevance ssreflect ssrbool. From MetaCoq.Utils Require Import utils. From MetaCoq.Common Require Import config uGraph. -From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICTactics +From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICPrimitive PCUICTactics PCUICReflect PCUICLiftSubst PCUICUnivSubst PCUICTyping PCUICGlobalEnv PCUICCumulativity PCUICEquality PCUICWfUniverses PCUICInduction. @@ -158,7 +158,8 @@ Fixpoint eqb_term_upto_univ_napp eqb_binder_annot x.(dname) y.(dname) ) mfix mfix' - | tPrim p, tPrim p' => eqb p p' + | tPrim p, tPrim p' => eqb_prim_val (equ := fun l l' => equ (Universe.make l) (Universe.make l')) + (req := eqb_term_upto_univ_napp equ equ gen_compare_global_instance 0) p p' | _, _ => false end. @@ -396,6 +397,8 @@ Proof. eapply All_impl ; tea. move => ? [? ?]. now apply /andP. + - destruct p as [? []]; cbn in *; rtoProp; intuition eauto. + solve_all. Qed. Definition reflect_eq_predicate {Σ equ lequ} @@ -530,6 +533,8 @@ Proof. destruct gr; auto; now repeat rewrite hlookup. Qed. +Transparent eqb_prim_val eqb_prim_model. + Lemma reflect_eq_term_upto_univ Σ equ lequ (p : Universe.t -> bool) (q : nat -> term -> bool) (Re Rle : Universe.t -> Universe.t -> Prop) @@ -754,7 +759,34 @@ Proof. constructor. constructor. constructor ; try easy. now inversion e3. - - cbn - [eqb]. eqspecs. do 2 constructor. + - destruct p0 as [? []], prim as [? []]; + cbn in X, ht, ht'; + rewrite /eqb_prim_val /eqb_prim_model; cbn -[eqb]; try eqspecs; + try repeat constructor. + 1-8:intros e; depelim e; try depelim o; intuition eauto. + unfold eqb_array. eqspecs; rewrite ?andb_true_r ?andb_false_r. + rtoProp; intuition auto. specialize (a1 equ Re 0 he _ H4 H0). + specialize (a2 equ Re 0 he _ H5 H1). + case: a1; rewrite ?andb_true_r ?andb_false_r; [intros eq|constructor; intros e; depelim e; depelim o]; eauto. + case: a2; rewrite ?andb_true_r ?andb_false_r; [intros eq'|constructor; intros e; depelim e; depelim o]; eauto. + equspec equ he; eauto; cbn. + 2:constructor; intros e; depelim e; depelim o; eauto. + repeat toAll. + revert b0 H2. + destruct a as [l d ty v], a0 as [l' d' ty' v']; cbn in *. + intros a. + induction a in v' |- *; intros. + * depelim H2; cbn; constructor; eauto; try repeat constructor; cbn; eauto. + intros heq; depelim heq; cbn in *; depelim o; eauto. depelim a0. + * intuition auto. depelim H2; cbn; try constructor; eauto. + ++ intros heq; depelim heq; depelim o; eauto. depelim a2. + ++ specialize (IHa _ H2). case: IHa; intros htl. + +++ rewrite andb_true_r. specialize (b equ Re 0 he _ a0 i). + case: b; repeat constructor; eauto; depelim htl; depelim o; eauto. + intros heq; depelim heq; depelim o; cbn in *; eauto. eapply f. now depelim a3. + +++ rewrite andb_false_r; constructor. + intros heq; depelim heq; depelim o; cbn in *; eauto. eapply htl. + repeat constructor; eauto. cbn. now depelim a2. Qed. Lemma eqb_term_upto_univ_impl (equ lequ : _ -> _ -> bool) @@ -944,6 +976,9 @@ Proof. - eapply forallb_All in wt; eapply All_mix in X; try apply wt; clear wt. eapply All_All2; eauto; simpl; intuition eauto; apply andb_and in a as [? ?]; eauto. + - destruct p as [? []]; cbn -[Universe.make] in X, wt; rtoProp; intuition eauto; constructor; eauto. + + eapply hRe. now move/wf_universe_reflect: H. + + solve_all. eapply All_All2; eauto; simpl; intuition eauto. Defined. Lemma eqb_term_upto_univ_refl Σ (eqb leqb : Universe.t -> Universe.t -> bool) (Re : Universe.t -> Universe.t -> Prop) diff --git a/safechecker/theories/PCUICErrors.v b/safechecker/theories/PCUICErrors.v index aefec8e10..361b6688a 100644 --- a/safechecker/theories/PCUICErrors.v +++ b/safechecker/theories/PCUICErrors.v @@ -89,6 +89,26 @@ Inductive ConversionError := (Γ : context) (mfix : mfixpoint term) (Γ' : context) (mfix' : mfixpoint term) +| DistinctPrimTags + (Γ : context) (p : prim_val) + (Γ' : context) (p' : prim_val) + +| DistinctPrimValues + (Γ : context) (p : prim_val) + (Γ' : context) (p' : prim_val) + +| ArrayNotConvertibleValues + (Γ : context) (a : array_model term) (Γ' : context) (a' : array_model term) + (e : ConversionError) + +| ArrayNotConvertibleDefault + (Γ : context) (a : array_model term) (Γ' : context) (a' : array_model term) + (e : ConversionError) + +| ArrayNotConvertibleTypes + (Γ : context) (a : array_model term) (Γ' : context) (a' : array_model term) + (e : ConversionError) + | StackHeadError (leq : conv_pb) (Γ1 : context) @@ -271,6 +291,39 @@ Fixpoint string_of_conv_error Σ (e : ConversionError) : string := nl ^ "and" ^ nl ^ print_term Σ Γ' (tCoFix mfix' idx) ^ nl ^ "have a different number of mutually defined functions." + | ArrayNotConvertibleValues Γ a Γ' a' e => + "The two arrays " ^ nl ^ + print_term Σ Γ (tPrim (primArray; primArrayModel a)) ^ + nl ^ "and" ^ nl ^ + print_term Σ Γ (tPrim (primArray; primArrayModel a')) ^ + nl ^ " have non-convertible values: " ^ + nl ^ string_of_conv_error Σ e + | ArrayNotConvertibleTypes Γ a Γ' a' e => + "The two arrays " ^ nl ^ + print_term Σ Γ (tPrim (primArray; primArrayModel a)) ^ + nl ^ "and" ^ nl ^ + print_term Σ Γ (tPrim (primArray; primArrayModel a')) ^ + nl ^ " have non-convertible types: " ^ + nl ^ string_of_conv_error Σ e + | ArrayNotConvertibleDefault Γ a Γ' a' e => + "The two arrays " ^ nl ^ + print_term Σ Γ (tPrim (primArray; primArrayModel a)) ^ + nl ^ "and" ^ nl ^ + print_term Σ Γ (tPrim (primArray; primArrayModel a')) ^ + nl ^ " have non-convertible default values: " ^ + nl ^ string_of_conv_error Σ e + | DistinctPrimTags Γ p Γ' p' => + "The two primitive values " ^ nl ^ + print_term Σ Γ (tPrim p) ^ + nl ^ "and" ^ nl ^ + print_term Σ Γ (tPrim p') ^ + nl ^ " have distinct tags" + | DistinctPrimValues Γ p Γ' p' => + "The two primitive values " ^ nl ^ + print_term Σ Γ (tPrim p) ^ + nl ^ "and" ^ nl ^ + print_term Σ Γ (tPrim p') ^ + nl ^ " have distinct values" | StackHeadError leq Γ1 t1 args1 u1 l1 Γ2 t2 u2 l2 e => "TODO stackheaderror" ^ nl ^ string_of_conv_error Σ e diff --git a/safechecker/theories/PCUICSafeConversion.v b/safechecker/theories/PCUICSafeConversion.v index 57544eeac..c49846fcd 100644 --- a/safechecker/theories/PCUICSafeConversion.v +++ b/safechecker/theories/PCUICSafeConversion.v @@ -732,6 +732,7 @@ Section Conversion. prog_discr (tProj _ _) (tProj _ _) := False ; prog_discr (tFix _ _) (tFix _ _) := False ; prog_discr (tCoFix _ _) (tCoFix _ _) := False ; + prog_discr (tPrim _) (tPrim _) := False ; prog_discr _ _ := True. (* Note that the arity of this should be the same for all s as otherwise @@ -1298,6 +1299,9 @@ Section Conversion. | prog_view_CoFix mfix idx mfix' idx' : prog_view (tCoFix mfix idx) (tCoFix mfix' idx') + | prog_view_Prim p p' : + prog_view (tPrim p) (tPrim p') + | prog_view_other : forall u v, prog_discr u v -> prog_view u v. @@ -1326,6 +1330,9 @@ Section Conversion. prog_viewc (tCoFix mfix idx) (tCoFix mfix' idx') := prog_view_CoFix mfix idx mfix' idx' ; + prog_viewc (tPrim p) (tPrim p') := + prog_view_Prim p p' ; + prog_viewc u v := prog_view_other u v I. Lemma welltyped_wf_local Σ Γ t : @@ -3476,6 +3483,26 @@ Qed. ) } ; + | prog_view_Prim p p' with inspect (eqb (prim_val_tag p) (prim_val_tag p')) := { + | @exist false tag_uneq := + no (DistinctPrimTag (Γ ,,, stack_context π1) p (Γ ,,, stack_context π2) p') + | @exist true eqtag with p, p' := { + | (primInt; primIntModel i) | (primInt, primIntModel i') with inspect (eqb i i') := + { | @exist true eqi := yes + | @exist false neqi := no (DistingPrimValues (Γ ,,, stack_context π1) p (Γ ,,, stack_context π2) p') } + | (primFloat; primFloatModel f) | (primFloat; primFloatModel f') with inspect (eqb f f') := + { | @exist true eqf := yes + | @exist false neqf := no (DistingPrimValues (Γ ,,, stack_context π1) p (Γ ,,, stack_context π2) p') } + | (primArray; primArrayModel a) | (primArray; primArrayModel a') with + with isconv_red_raw Conv (array_type a) (Prod_l na B1 :: π1) (array_type a') (Prod_l na' B2 :: π2) aux : + inspect (eqb f f') := + { | @exist true eqf := yes + | @exist false neqf := no (DistingPrimValues (Γ ,,, stack_context π1) p (Γ ,,, stack_context π2) p') } + + + } + } ; + | prog_view_other t1 t2 h := isconv_fallback leq t1 π1 t2 π2 aux }. diff --git a/safechecker/theories/PCUICSafeReduce.v b/safechecker/theories/PCUICSafeReduce.v index 65f8c71b2..a61cabcc5 100644 --- a/safechecker/theories/PCUICSafeReduce.v +++ b/safechecker/theories/PCUICSafeReduce.v @@ -1319,8 +1319,8 @@ Corollary R_Acc_aux : unfold is_true in typ. unfold PCUICAst.PCUICEnvironment.fst_ctx in *. congruence. - - eapply inversion_Prim in typ as (prim_ty & cdecl & [? ? ? [? []]]); tea. - now eapply invert_cumul_axiom_ind in w; tea. + - eapply inversion_Prim in typ as (prim_ty & cdecl & [? ? ? ?]); tea. + now eapply (invert_cumul_prim_type_ind) in w; tea. Qed. Definition isCoFix_app t := @@ -1353,8 +1353,8 @@ Corollary R_Acc_aux : unfold unfold_fix. destruct o as [[? [-> ?]] | ->]; eauto. - unfold isCoFix_app in cof. now rewrite decompose_app_mkApps in cof. - - eapply inversion_Prim in typ as [prim_ty [cdecl [? ? ? [? []]]]]; tea. - now eapply invert_cumul_axiom_ind in w; tea. + - eapply inversion_Prim in typ as [prim_ty [cdecl [? ? ? ?]]]; tea. + now eapply invert_cumul_prim_type_ind in w; tea. Qed. Lemma whnf_fix_arg_whne mfix idx body Σ Γ t before args aftr ty : @@ -1532,8 +1532,8 @@ Corollary R_Acc_aux : simpl in h. rewrite stack_context_appstack in h. destruct h as [T h]. apply inversion_App in h as (?&?&?&?&?); auto. - apply inversion_Prim in t0 as (prim_ty & cdecl & [? ? ? [s []]]); auto. - eapply PCUICClassification.invert_cumul_axiom_prod; eauto. + apply inversion_Prim in t0 as (prim_ty & cdecl & [? ? ? ?]); auto. + eapply PCUICClassification.invert_cumul_prim_type_prod; eauto. - unfold zipp. case_eq (decompose_stack π). intros l ρ e. constructor. constructor. eapply whne_mkApps. eapply whne_rel_nozeta. assumption. diff --git a/safechecker/theories/PCUICSafeRetyping.v b/safechecker/theories/PCUICSafeRetyping.v index eb31d9b85..385dabc26 100644 --- a/safechecker/theories/PCUICSafeRetyping.v +++ b/safechecker/theories/PCUICSafeRetyping.v @@ -366,7 +366,7 @@ Qed. | exist None _ => ! }; infer Γ wfΓ (tPrim p) wt with inspect (abstract_primitive_constant X p.π1) := - { | exist (Some prim_ty) eqp => ret (tConst prim_ty []) + { | exist (Some prim_ty) eqp => ret (prim_type p prim_ty) | exist None _ => ! }. Next Obligation. diff --git a/template-coq/theories/TermEquality.v b/template-coq/theories/TermEquality.v index a11f62bb4..653533f9a 100644 --- a/template-coq/theories/TermEquality.v +++ b/template-coq/theories/TermEquality.v @@ -241,9 +241,10 @@ Inductive eq_term_upto_univ_napp Σ (Re Rle : Universe.t -> Universe.t -> Prop) | eq_Int i : eq_term_upto_univ_napp Σ Re Rle napp (tInt i) (tInt i) | eq_Float f : eq_term_upto_univ_napp Σ Re Rle napp (tFloat f) (tFloat f) | eq_Array u u' arr arr' def def' ty ty' : - All2 (eq_term_upto_univ_napp Σ Re Rle 0) arr arr' -> - eq_term_upto_univ_napp Σ Re Rle 0 def def' -> - eq_term_upto_univ_napp Σ Re Rle 0 ty ty' -> + Re (Universe.make u) (Universe.make u') -> + All2 (eq_term_upto_univ_napp Σ Re Re 0) arr arr' -> + eq_term_upto_univ_napp Σ Re Re 0 def def' -> + eq_term_upto_univ_napp Σ Re Re 0 ty ty' -> eq_term_upto_univ_napp Σ Re Rle napp (tArray u arr def ty) (tArray u' arr' def' ty'). Notation eq_term_upto_univ Σ Re Rle := (eq_term_upto_univ_napp Σ Re Rle 0). diff --git a/template-coq/theories/Typing.v b/template-coq/theories/Typing.v index cc451c39b..7fa6c4555 100644 --- a/template-coq/theories/Typing.v +++ b/template-coq/theories/Typing.v @@ -276,7 +276,19 @@ Inductive red1 (Σ : global_env) (Γ : context) : term -> term -> Type := | cofix_red_body mfix0 mfix1 idx : OnOne2 (on_Trel_eq (red1 Σ (Γ ,,, fix_context mfix0)) dbody (fun x => (dname x, dtype x, rarg x))) mfix0 mfix1 -> - red1 Σ Γ (tCoFix mfix0 idx) (tCoFix mfix1 idx). + red1 Σ Γ (tCoFix mfix0 idx) (tCoFix mfix1 idx) + +| array_red_val l v v' d ty : + OnOne2 (fun t u => red1 Σ Γ t u) v v' -> + red1 Σ Γ (tArray l v d ty) (tArray l v' d ty) + +| array_red_def l v d d' ty : + red1 Σ Γ d d' -> + red1 Σ Γ (tArray l v d ty) (tArray l v d' ty) + +| array_red_type l v d ty ty' : + red1 Σ Γ ty ty' -> + red1 Σ Γ (tArray l v d ty) (tArray l v d ty'). Lemma red1_ind_all : forall (Σ : global_env) (P : context -> term -> term -> Type), @@ -398,9 +410,19 @@ Lemma red1_ind_all : (fun x => (dname x, dtype x, rarg x))) mfix0 mfix1 -> P Γ (tCoFix mfix0 idx) (tCoFix mfix1 idx)) -> + (forall Γ l v v' d ty, + OnOne2 (fun t u => Trel_conj (red1 Σ Γ) (P Γ) t u) v v' -> + P Γ (tArray l v d ty) (tArray l v' d ty)) -> + + (forall Γ l v d d' ty, + red1 Σ Γ d d' -> P Γ d d' -> P Γ (tArray l v d ty) (tArray l v d' ty)) -> + + (forall Γ l v d ty ty', + red1 Σ Γ ty ty' -> P Γ ty ty' -> P Γ (tArray l v d ty) (tArray l v d ty')) -> + forall (Γ : context) (t t0 : term), red1 Σ Γ t t0 -> P Γ t t0. Proof. - intros. rename X30 into Xlast. revert Γ t t0 Xlast. + intros. rename X33 into Xlast. revert Γ t t0 Xlast. fix aux 4. intros Γ t T. move aux at top. destruct 1; @@ -458,8 +480,11 @@ Proof. revert o. generalize (fix_context mfix0). intros c H28. revert mfix0 mfix1 H28; fix auxl 3; intros l l' Hl; destruct Hl; constructor; try split; auto; intuition. -Defined. + - eapply X30. + revert v v' o. fix auxl 3; intros ? ? Hl; destruct Hl; + constructor; try split; auto; intuition. +Defined. (** *** Reduction diff --git a/template-coq/theories/TypingWf.v b/template-coq/theories/TypingWf.v index a4f97e740..affb93cc9 100644 --- a/template-coq/theories/TypingWf.v +++ b/template-coq/theories/TypingWf.v @@ -847,6 +847,7 @@ Section WfRed. cbn. unfold wf_decl. simpl. intros ? [? ? ? ?] ?. simpl in *. intuition eauto with wf. + - constructor; eauto. eapply (OnOne2_All_All X); tea; intuition eauto. Qed. @@ -1143,7 +1144,7 @@ Section TypingWf. Proof using Type. intros wfx; revert y. induction wfx using term_wf_forall_list_ind; intros [] wfy; - eapply wf_inv in wfy; simpl in wfy; simpl; + eapply wf_inv in wfy; simpl in wfy; simpl; try intros [= ?]; try intuition congruence. Qed. diff --git a/template-coq/theories/WcbvEval.v b/template-coq/theories/WcbvEval.v index a41410626..72be49adb 100644 --- a/template-coq/theories/WcbvEval.v +++ b/template-coq/theories/WcbvEval.v @@ -55,6 +55,7 @@ Fixpoint csubst t k u := let mfix' := List.map (map_def (csubst t k) (csubst t k')) mfix in tCoFix mfix' idx | tCast c kd c' => tCast (csubst t k c) kd (csubst t k c') + | tArray l v d ty => tArray l (List.map (csubst t k) v) (csubst t k d) (csubst t k ty) | x => x end. diff --git a/template-pcuic/theories/PCUICToTemplate.v b/template-pcuic/theories/PCUICToTemplate.v index 4656b28d5..a1f4d1933 100644 --- a/template-pcuic/theories/PCUICToTemplate.v +++ b/template-pcuic/theories/PCUICToTemplate.v @@ -13,10 +13,11 @@ Definition uint63_from_model (i : uint63_model) : Uint63.int := Definition float64_from_model (f : float64_model) : PrimFloat.float := FloatOps.SF2Prim (proj1_sig f). -Definition trans_prim (t : prim_val) : Ast.term := +Definition trans_prim (trans : PCUICAst.term -> Ast.term) (t : prim_val) : Ast.term := match t.π2 with | primIntModel i => Ast.tInt i | primFloatModel f => Ast.tFloat f + | primArrayModel a => Ast.tArray (array_level a) (map trans (array_value a)) (trans (array_default a)) (trans (array_type a)) end. Definition trans_predicate (t : PCUICAst.predicate Ast.term) : predicate Ast.term := @@ -53,7 +54,7 @@ Fixpoint trans (t : PCUICAst.term) : Ast.term := | PCUICAst.tCoFix mfix idx => let mfix' := List.map (map_def trans trans) mfix in tCoFix mfix' idx - | PCUICAst.tPrim i => trans_prim i + | PCUICAst.tPrim i => trans_prim trans i end. Notation trans_decl := (map_decl trans). diff --git a/template-pcuic/theories/PCUICToTemplateCorrectness.v b/template-pcuic/theories/PCUICToTemplateCorrectness.v index 03a23608f..0f80efe77 100644 --- a/template-pcuic/theories/PCUICToTemplateCorrectness.v +++ b/template-pcuic/theories/PCUICToTemplateCorrectness.v @@ -106,7 +106,7 @@ Proof. rewrite b. now rewrite forget_types_length map_context_length. - f_equal; auto; red in X; solve_list. - f_equal; auto; red in X; solve_list. - - destruct p as [? []]; eauto. + - destruct p as [? []]; cbn in X; cbn; f_equal; intuition eauto; solve_all. Qed. Definition on_fst {A B C} (f:A->C) (p:A×B) := (f p.1, p.2). @@ -284,7 +284,7 @@ Proof. cbn in *. now rewrite e e0. + apply IHX. - - destruct p as [? []]; eauto. + - destruct p as [? []]; cbn in X |- *; f_equal; intuition eauto; solve_all. Qed. Lemma trans_subst10 u B: @@ -334,7 +334,7 @@ Proof. destruct p. now rewrite e e0. + apply IHX. - - destruct p as [? []]; eauto. + - destruct p as [? []]; cbn in X |- *; f_equal; intuition eauto; solve_all. Qed. Lemma trans_subst_instance_ctx Γ u : @@ -683,6 +683,7 @@ Section wtsub. | tFix mfix idx | tCoFix mfix idx => All (fun d => wt Γ d.(dtype) × wt (Γ ,,, fix_context mfix) d.(dbody)) mfix | tEvar _ l => False + | tPrim p => tPrimProp (wt Γ) p | _ => unit end. Import PCUICGeneration PCUICInversion. @@ -749,6 +750,9 @@ Section wtsub. eapply All_prod. eapply (All_impl a). intros ? h; exact h. eapply (All_impl a0). intros ? h; eexists; tea. + - eapply inversion_Prim in h as (?&?&[]); eauto. + destruct prim as [? []]; cbn in *; eauto; try exact tt. + depelim p0. repeat split. now eexists. now eexists. solve_all. now eexists. Qed. End wtsub. @@ -1028,7 +1032,7 @@ Proof. cbn; eauto. cbn in p0. destruct p0. eauto. - cbn. red in X. solve_all. - cbn. red in X. solve_all. - - destruct p as [? []]; constructor. + - destruct p as [? []]; cbn in X, H |- *; constructor; solve_all; eauto. Qed. #[global] Hint Resolve trans_wf : wf. @@ -1058,6 +1062,7 @@ Proof. - constructor; solve_all. - eapply Typing.cofix_red_body; solve_all. eapply b0, All2_app => //. reflexivity. + - eapply Typing.array_red_val. solve_all. Qed. Lemma map_map2 {A B C D} (f : A -> B) (g : C -> D -> A) l l' : @@ -1381,6 +1386,9 @@ Proof. intuition auto. rewrite /trans_local map_app in X. now rewrite -trans_fix_context. + - cbn. apply TT.array_red_val. + eapply OnOne2_All_mix_left in X; tea. + eapply OnOne2_map. solve_all. Qed. Lemma trans_R_global_instance Σ Re Rle gref napp u u' : @@ -1449,7 +1457,7 @@ Proof. red in X0. solve_all_one. eapply trans_eq_context_gen_eq_binder_annot in a. now rewrite !map_context_trans. - - destruct p as [? []]; constructor. + - depelim X0; cbn in X |- *; try econstructor; intuition eauto; solve_all. Qed. Lemma trans_leq_term {cf} Σ ϕ T U : @@ -2363,11 +2371,14 @@ Proof. + fold trans;subst types. now apply trans_mfix_All2. + now rewrite trans_wf_cofixpoint. - - cbn. destruct p as [? []]; cbn; econstructor; eauto. - 1,3: eapply trans_declared_constant; tea. - all:move: X0; rewrite /Ast.Env.primitive_invariants /primitive_invariants; - intros [s []]; exists s; split => //; + - cbn. destruct p as [? []]; simp prim_type; cbn; econstructor; eauto. + 1,3,5: eapply trans_declared_constant; tea. + all:cbn in *. + all:move: X0; rewrite /Ast.Env.primitive_invariants /primitive_invariants. + 1-2:intros [s []]; exists s; split => //; destruct cdecl as [ty [?|] ?]; cbn in *; subst; auto => //. + intros []; split => //. rewrite H1 //. rewrite H2 //. + all:depelim X2; eauto. intros _. solve_all. - eapply TT.type_Conv. + eassumption. + eassumption. diff --git a/template-pcuic/theories/TemplateMonadToPCUIC.v b/template-pcuic/theories/TemplateMonadToPCUIC.v index 48726cffd..06477c287 100644 --- a/template-pcuic/theories/TemplateMonadToPCUIC.v +++ b/template-pcuic/theories/TemplateMonadToPCUIC.v @@ -163,7 +163,12 @@ Section with_tc. ret (tCoFix mfix' idx) | Ast.tInt n => ret (tPrim (primInt; primIntModel n)) | Ast.tFloat n => ret (tPrim (primFloat; primFloatModel n)) - end. + | Ast.tArray l v d ty => + v' <- monad_map@{t u t t} monad_trans' v ;; + d' <- monad_trans' d ;; + ty' <- monad_trans' ty ;; + ret (tPrim ((primArray; primArrayModel {| array_level := l; array_value := v'; array_default := d'; array_type := ty' |}))) + end. End with_helper. End with_tc. diff --git a/template-pcuic/theories/TemplateToPCUIC.v b/template-pcuic/theories/TemplateToPCUIC.v index 7568a7985..595df18f5 100644 --- a/template-pcuic/theories/TemplateToPCUIC.v +++ b/template-pcuic/theories/TemplateToPCUIC.v @@ -103,6 +103,11 @@ Section Trans. tCoFix mfix' idx | Ast.tInt n => tPrim (primInt; primIntModel n) | Ast.tFloat n => tPrim (primFloat; primFloatModel n) + | Ast.tArray l v d ty => tPrim (primArray; primArrayModel + {| array_level := l; + array_value := List.map trans v; + array_default := trans d; + array_type := trans ty |}) end. Definition trans_decl (d : Ast.Env.context_decl) := diff --git a/template-pcuic/theories/TemplateToPCUICCorrectness.v b/template-pcuic/theories/TemplateToPCUICCorrectness.v index 3eae7d881..76845a68a 100644 --- a/template-pcuic/theories/TemplateToPCUICCorrectness.v +++ b/template-pcuic/theories/TemplateToPCUICCorrectness.v @@ -239,6 +239,7 @@ Proof. red in X0. f_equal => //. rewrite /id. unfold trans_predicate. f_equal; solve_all. f_equal. solve_all. + do 4 f_equal; solve_all. Qed. Lemma trans_decl_weakening {cf} Σ {Σ' : global_env_map} t : @@ -379,6 +380,10 @@ Proof. rewrite e. now rewrite map_length. - f_equal; auto. maps. solve_all. - f_equal; auto; solve_all. + - f_equal; auto. + - f_equal; auto. + - cbn; do 4 f_equal; auto. + rewrite /map_array_model. f_equal; eauto. cbn. solve_all. Qed. Lemma trans_mkApp u a : trans (Template.Ast.mkApp u a) = tApp (trans u) (trans a). @@ -435,6 +440,10 @@ Proof. f_equal; auto; solve_all. - f_equal; auto; solve_list. - f_equal; auto; solve_list. + - f_equal; auto. + - f_equal; auto. + - cbn; f_equal; auto. rewrite /map_array_model; cbn; do 3 f_equal; eauto. + solve_all. Qed. Notation Tterm :=Template.Ast.term. @@ -470,6 +479,9 @@ Proof. rewrite /map_branch /trans_branch /= /id. now intros; f_equal. - rewrite /id /map_predicate /=. f_equal; solve_all. + - f_equal; auto. + - f_equal; auto. + - cbn. rewrite /mapu_array_model; cbn; do 4 f_equal; auto; solve_all. Qed. @@ -958,6 +970,10 @@ Section Trans_Global. intros [? ? ? ?] [? ? ? ?] [[[? ?] [[ih1 ih2] [? ?]]] [? ?]]. simpl in *. intuition eauto. now eapply ih1. now eapply ih2. + - constructor; eauto. intuition auto; constructor; cbn; eauto. + eapply (IHt1 Re); eauto. reflexivity. + eapply (IHt2 Re); eauto. reflexivity. + solve_all. eapply a; eauto. reflexivity. Qed. Lemma trans_leq_term ϕ T U : @@ -1698,6 +1714,8 @@ Section Trans_Global. cbn. rewrite /trans_decl /vass /=. now rewrite trans_lift. + simpl; congruence. + - simpl. eapply array_red_val. eapply (OnOne2_All_mix_left a) in X. + apply OnOne2_map; solve_all. Qed. End Trans_Global. @@ -2451,18 +2469,39 @@ Proof. now eapply TypingWf.typing_wf in Hs'. -- destruct decl; reflexivity. - - cbn. econstructor; cbn; eauto. + - cbn. replace (tConst prim_ty []) with (prim_type (primInt; primIntModel p) prim_ty) by now simp prim_type. + econstructor; cbn; eauto. + rewrite trans_env_retroknowledge //. + now apply forall_decls_declared_constant. + move: X0; rewrite /Ast.Env.primitive_invariants /primitive_invariants. intros [s []]; exists s; split => //; destruct cdecl as [ty [?|] ?]; cbn in *; subst; auto => //. - - cbn. econstructor; cbn; eauto. + + constructor. + - cbn. replace (tConst prim_ty []) with (prim_type (primFloat; primFloatModel p) prim_ty) by now simp prim_type. + econstructor; cbn; eauto. + rewrite trans_env_retroknowledge //. + now apply forall_decls_declared_constant. + move: X0; rewrite /Ast.Env.primitive_invariants /primitive_invariants. intros [s []]; exists s; split => //; destruct cdecl as [ty [?|] ?]; cbn in *; subst; auto => //. + + constructor. + - cbn. set (a := {| array_level := _ |}). + replace (tApp (tConst prim_ty [u]) (trans (trans_global_env Σ.1) ty)) with (prim_type (primArray; primArrayModel a) prim_ty) by now simp prim_type. + econstructor; cbn; eauto. + + rewrite trans_env_retroknowledge //. + + now apply forall_decls_declared_constant. + + move: X0; rewrite /Ast.Env.primitive_invariants /primitive_invariants. + intros []; split => //; eauto. + * apply forall_decls_declared_constant in H0; eauto. + rewrite /trans_constant_body in H0 |- *. + now rewrite H1 H2 H3 /= in H0 |- *. + * rewrite /trans_constant_body in H0 |- *. + now rewrite H1 H2 H3 /= in H0 |- *. + + constructor; eauto. cbn [array_level a]. eapply validity in X2; eauto. + eapply PCUICWfUniverses.isType_wf_universes in X2. cbn [trans PCUICWfUniverses.wf_universes] in X2. + unfold PCUICWfUniverses.wf_universes in X2. cbn [PCUICWfUniverses.on_universes] in X2. + move: X2. case: PCUICWfUniverses.wf_universe_reflect => //; eauto. eauto. + cbn [a array_value]. solve_all. - assert (WfAst.wf Σ B). { now apply typing_wf in X2. } eapply type_Cumul; eauto. @@ -2594,6 +2633,7 @@ Proof. rewrite map2_length; len. eauto. - unfold test_def; red in X. solve_all. - unfold test_def; solve_all. + - solve_all. Qed. From MetaCoq.PCUIC Require Import PCUICOnFreeVars. diff --git a/template-pcuic/theories/TemplateToPCUICExpanded.v b/template-pcuic/theories/TemplateToPCUICExpanded.v index bd4dfd7e8..cf6f8e115 100644 --- a/template-pcuic/theories/TemplateToPCUICExpanded.v +++ b/template-pcuic/theories/TemplateToPCUICExpanded.v @@ -235,6 +235,9 @@ Proof with eauto using expanded. - wf_inv wf ?. econstructor. solve_all. - 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. + - repeat constructor. + - repeat constructor. + - wf_inv wf [[] ?]. repeat cbn; constructor. constructor; cbn; eauto. solve_all. Qed. diff --git a/template-pcuic/theories/TemplateToPCUICWcbvEval.v b/template-pcuic/theories/TemplateToPCUICWcbvEval.v index a3a1b7b07..2e95e9d3d 100644 --- a/template-pcuic/theories/TemplateToPCUICWcbvEval.v +++ b/template-pcuic/theories/TemplateToPCUICWcbvEval.v @@ -158,6 +158,7 @@ Proof. f_equal; auto; solve_all. - f_equal; auto; solve_list. - f_equal; auto; solve_list. + - cbn; f_equal; auto. do 3 f_equal. rewrite /map_array_model /=; f_equal; eauto; solve_all. Qed. Lemma trans_substl {cf} Σ a b : From 3ee01a38d7f3dddc5d8c58a1304fe8dc5523b827 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Mon, 20 Nov 2023 15:34:12 +0100 Subject: [PATCH 07/17] Update PCUICPosition to include arrays --- pcuic/theories/PCUICReduction.v | 4 ++ pcuic/theories/Syntax/PCUICPosition.v | 97 +++++++++++++++++++++++---- 2 files changed, 88 insertions(+), 13 deletions(-) diff --git a/pcuic/theories/PCUICReduction.v b/pcuic/theories/PCUICReduction.v index d0dda059b..dcf117dfe 100644 --- a/pcuic/theories/PCUICReduction.v +++ b/pcuic/theories/PCUICReduction.v @@ -2219,6 +2219,10 @@ Section Stacks. apply OnOne2_app. constructor; cbn. rewrite -app_assoc in h; auto. + - now apply array_red_type; cbn in *. + - now apply array_red_def; cbn in *. + - apply array_red_val; cbn in *. + eapply OnOne2_app. now constructor. Qed. Corollary red_context_zip : diff --git a/pcuic/theories/Syntax/PCUICPosition.v b/pcuic/theories/Syntax/PCUICPosition.v index 439689ef7..ee75e7d03 100644 --- a/pcuic/theories/Syntax/PCUICPosition.v +++ b/pcuic/theories/Syntax/PCUICPosition.v @@ -490,6 +490,49 @@ Proof. - simpl. cbn in e2. rewrite e in e2. assumption. - specialize (ih2 q). eapply ih2. all: assumption. } + assert ( + forall a (p : pos a.(array_default)) + (e : validpos (tPrim (primArray; primArrayModel a)) (array_def :: proj1_sig p)), + Acc posR p -> + Acc posR (exist (array_def :: proj1_sig p) e) + ) as Acc_array_def. + { intros a p e h. + induction h as [p ih1 ih2] in e |- *. + constructor. intros [q e2] h. + dependent destruction h. + simple refine (let q := exist p0 _ : pos a.(array_default) in _). + - simpl. cbn in e2. assumption. + - specialize (ih2 q). eapply ih2. all: assumption. + } + assert ( + forall a (p : pos a.(array_type)) + (e : validpos (tPrim (primArray; primArrayModel a)) (array_ty :: proj1_sig p)), + Acc posR p -> + Acc posR (exist (array_ty :: proj1_sig p) e) + ) as Acc_array_ty. + { intros a p e h. + induction h as [p ih1 ih2] in e |- *. + constructor. intros [q e2] h. + dependent destruction h. + simple refine (let q := exist p0 _ : pos a.(array_type) in _). + - simpl. cbn in e2. assumption. + - specialize (ih2 q). eapply ih2. all: assumption. + } + assert ( + forall a n v (p : pos v) + (e : nth_error a.(array_value) n = Some v) + (e1 : validpos (tPrim (primArray; primArrayModel a)) (array_val n :: proj1_sig p)), + Acc posR p -> + Acc posR (exist (array_val n :: proj1_sig p) e1) + ) as Acc_array_val. + { intros a n v p e e1 h. + induction h as [p ih1 ih2] in e, e1 |- *. + constructor. intros [q e2] h. + dependent destruction h. + simple refine (let q := exist p0 _ : pos v in _). + - simpl. cbn in e2. rewrite e in e2. assumption. + - specialize (ih2 q). eapply ih2. all: assumption. + } intro t. induction t using term_forall_list_ind ; intros q. all: try solve [ destruct q as [q e] ; destruct q as [| c q] ; [ @@ -711,28 +754,29 @@ Proof. dependent destruction h. destruct p as [? []]; destruct c. all: noconf e'. + * eapply Acc_array_def with (p := exist p0 e'), X. + * eapply Acc_array_ty with (p := exist p0 e'), X. * simpl in e'. - case_eq (nth_error m n0). + case_eq (nth_error (array_value a) n). 2:{ intro h. pose proof e' as hh. rewrite h in hh. discriminate. } - intros [na ty bo ra] e1. + intros v e1. destruct X as [_ [_ X]]. eapply All_nth_error in X as ihm. 2: exact e1. simpl in ihm. - unshelve eapply Acc_cofix_mfix_ty with (1 := e1) (p := exist p _). + unshelve eapply Acc_array_val with (1 := e1) (p := exist p0 _). -- simpl. rewrite e1 in e'. assumption. -- eapply ihm. - * simpl in e'. - case_eq (nth_error m n0). - 2:{ intro h. pose proof e' as hh. rewrite h in hh. discriminate. } - intros [na ty bo ra] e1. + + destruct p as [? []]; destruct c; noconf e. + * eapply Acc_array_def with (p := exist q e), X. + * eapply Acc_array_ty with (p := exist q e), X. + * simpl in e. + case_eq (nth_error (array_value a) n). + 2:{ intro h. pose proof e as hh. rewrite h in hh. discriminate. } + intros v e1. destruct X as [_ [_ X]]. eapply All_nth_error in X as ihm. 2: exact e1. simpl in ihm. - unshelve eapply Acc_cofix_mfix_bd with (1 := e1) (p := exist p _). - -- simpl. rewrite e1 in e'. assumption. + unshelve eapply Acc_array_val with (1 := e1) (p := exist q _). + -- simpl. rewrite e1 in e. assumption. -- eapply ihm. - - - destruct p as [? []]; cbn in X, e; intuition eauto; cbn in e. - + destruct x eqn:hp; try discriminate; cbn in e. constructor. Qed. Fixpoint atpos t (p : position) {struct p} : term := @@ -782,6 +826,13 @@ Fixpoint atpos t (p : position) {struct p} : term := | let_bd, tLetIn na b B t => atpos b p | let_ty, tLetIn na b B t => atpos B p | let_in, tLetIn na b B t => atpos t p + | array_def, tPrim (primArray; primArrayModel a) => atpos a.(array_default) p + | array_ty, tPrim (primArray; primArrayModel a) => atpos a.(array_type) p + | array_val n, tPrim (primArray; primArrayModel a) => + match nth_error a.(array_value) n with + | Some v => atpos v p + | None => tRel 0 + end | _, _ => tRel 0 end end. @@ -818,6 +869,14 @@ Proof. + simpl. destruct nth_error as [[na ty bo ra]|] eqn:e. * apply IHp. * rewrite hh. reflexivity. + + destruct prim as [? []]; simpl. 1-2:now rewrite hh. + now rewrite IHp. + + destruct prim as [? []]; simpl. 1-2:now rewrite hh. + now rewrite IHp. + + destruct prim as [? []]; simpl. 1-2:now rewrite hh. + destruct nth_error eqn:e. + * now rewrite IHp. + * now rewrite hh. Qed. Lemma poscat_valid : @@ -836,6 +895,7 @@ Proof. all: simpl in *; repeat match goal with + | [prim: prim_val |- _] => destruct prim as [? []] | [H: context[nth_error ?a ?b] |- _] => destruct (nth_error a b); [|discriminate] | [H: context[context_choice_term ?a ?b] |- _] => @@ -869,6 +929,7 @@ Proof. all: repeat match goal with + | [prim: prim_val |- _] => destruct prim as [? []] | |- context[nth_error ?a ?b] => destruct (nth_error a b); auto | |- context[context_choice_term ?a ?b] => @@ -1400,6 +1461,9 @@ Definition stack_entry_choice (se : stack_entry) : choice := | LetIn_bd na B t => let_bd | LetIn_ty na b t => let_ty | LetIn_in na b B => let_in + | PrimArray_def u v ty => array_def + | PrimArray_ty u v def => array_ty + | PrimArray_val l bef aft def ty => array_val #|bef| end. Definition stack_position : stack -> position := @@ -1432,6 +1496,7 @@ Proof. + rewrite nth_error_snoc; auto. - destruct brs as ((?&[])&?); cbn. all: rewrite nth_error_snoc; auto. + - rewrite nth_error_snoc; auto. Qed. Lemma stack_position_valid : @@ -1451,6 +1516,7 @@ Proof. + rewrite nth_error_snoc; auto. - destruct brs as ((?&[])&?); cbn. all: rewrite nth_error_snoc; auto. + - rewrite nth_error_snoc; auto. Qed. Definition stack_pos (t : term) (π : stack) : pos (zipc t π) := @@ -1528,6 +1594,8 @@ Section Stacks. - destruct brs as ((?&[])&?); cbn in *. + apply app_inj_length_l in H0 as (_&H0); auto; noconf H0. reflexivity. + - eapply app_inj_length_l in H0 as (_&H0); auto. + now noconf H0. Qed. Definition isStackApp (π : stack) : bool := @@ -1764,6 +1832,9 @@ Proof. - destruct brs as [[brs []] brs']; cbn. rewrite -app_tip_assoc !forallb_app; cbn. len; ring_simplify; rewrite - !andb_assoc; repeat bool_congr. + - cbn. bool_congr. + - cbn; rewrite -!andb_assoc; bool_congr. + - cbn. rewrite !forallb_app -!andb_assoc; repeat bool_congr. Qed. Lemma closedn_zip k t π : closedn k (zipc t π) = closedn_stack k π && closedn (#|stack_context π| + k) t. From 67bba980bd387059ce16998c01824f8002750872 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Tue, 21 Nov 2023 15:25:11 +0100 Subject: [PATCH 08/17] Updatet SafeConversion to arrays --- pcuic/theories/Conversion/PCUICInstConv.v | 6 +- .../theories/Conversion/PCUICOnFreeVarsConv.v | 6 +- pcuic/theories/PCUICClassification.v | 174 ------ pcuic/theories/PCUICConversion.v | 266 ++++++++- pcuic/theories/PCUICEtaExpand.v | 8 - pcuic/theories/PCUICPrincipality.v | 10 - pcuic/theories/PCUICSR.v | 33 +- pcuic/theories/PCUICSafeLemmata.v | 83 +++ safechecker/theories/PCUICErrors.v | 18 + safechecker/theories/PCUICSafeConversion.v | 503 +++++++++++++++--- safechecker/theories/PCUICSafeReduce.v | 2 +- utils/theories/All_Forall.v | 55 ++ 12 files changed, 861 insertions(+), 303 deletions(-) diff --git a/pcuic/theories/Conversion/PCUICInstConv.v b/pcuic/theories/Conversion/PCUICInstConv.v index fc6ce53b2..32c8f3e5c 100644 --- a/pcuic/theories/Conversion/PCUICInstConv.v +++ b/pcuic/theories/Conversion/PCUICInstConv.v @@ -1416,7 +1416,7 @@ Proof using Type. sigma in X. eapply X. * eapply inst_ext. rewrite ren_lift_renaming. now sigma. - - eapply usubst_Up; eauto; intuition. + - eapply usubst_Up; eauto; intuition auto with *. Qed. @@ -1465,7 +1465,7 @@ Proof using Type. eapply on_free_vars_inst. 2: now eapply h. now easy. - - eapply usubst_Up'; eauto; intuition. + - eapply usubst_Up'; eauto; intuition auto with *. Qed. Lemma well_subst_Up' {Σ : global_env_ext} {wfΣ : wf Σ} {Γ Δ σ na t A} : @@ -1491,7 +1491,7 @@ Proof using Type. sigma in X0. eapply X0. + eapply inst_ext. rewrite ren_lift_renaming. now sigma. - - eapply usubst_Up'; eauto; intuition. + - eapply usubst_Up'; eauto; intuition auto with *. Qed. diff --git a/pcuic/theories/Conversion/PCUICOnFreeVarsConv.v b/pcuic/theories/Conversion/PCUICOnFreeVarsConv.v index 6a0e1dad7..a6b9afb6c 100644 --- a/pcuic/theories/Conversion/PCUICOnFreeVarsConv.v +++ b/pcuic/theories/Conversion/PCUICOnFreeVarsConv.v @@ -120,7 +120,7 @@ Proof. * unfold shiftnP in Hn. unfold sP , shiftnP. toProp. toProp Hn. destruct Hn. + intuition. - + right. toProp. toProp H. destruct H; intuition. + + right. toProp. toProp H. destruct H; intuition auto with *. * destruct (hf eq_refl) as [decl' [Hfn _]]. clear hf Hn. unfold sP , shiftnP. rewrite orb_false_r. assert (shiftn #|Ξ| f n < #|Δ,,, rename_context f Ξ|). @@ -131,8 +131,8 @@ Proof. 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. + * toProp H. intuition auto with *. + * toProp H. destruct H; [toProp H |]; intuition auto with *. Defined. Lemma urename_is_open_term P Γ Δ f u : let sP := shiftnP #|Γ| P in diff --git a/pcuic/theories/PCUICClassification.v b/pcuic/theories/PCUICClassification.v index c1c4d9b09..0ce8e468e 100644 --- a/pcuic/theories/PCUICClassification.v +++ b/pcuic/theories/PCUICClassification.v @@ -122,10 +122,6 @@ Section Arities. End Arities. -Lemma All2_map_left' {A B C} (P : A -> B -> Type) l l' (f : C -> A) : - All2 P (map f l) l' -> All2 (fun x y => P (f x) y) l l'. -Proof. intros. rewrite - (map_id l') in X. eapply All2_map_inv; eauto. Qed. - Section Spines. Context {cf : checker_flags}. Context {Σ : global_env_ext}. @@ -792,176 +788,6 @@ Section classification. False. Proof. eauto using wh_neutral_empty_gen. Qed. - (* TODO move *) - Lemma invert_red_axiom {Γ cst u cdecl T} : - declared_constant Σ cst cdecl -> - cst_body cdecl = None -> - Σ ;;; Γ ⊢ tConst cst u ⇝ T -> - T = tConst cst u. - Proof using wfΣ. - intros hdecl hb. - unshelve eapply declared_constant_to_gen in hdecl; eauto. - generalize_eq x (tConst cst u). - move=> e [clΓ clt] red. - revert cst u hdecl hb e. - eapply clos_rt_rt1n_iff in red. - induction red; simplify_dep_elim. - - reflexivity. - - depelim r; solve_discr. - unshelve eapply declared_constant_to_gen in isdecl; eauto. - congruence. - Qed. - - Ltac bool := rtoProp; intuition eauto. - - Lemma invert_red1_axiom_app {Γ cst u cdecl args T} : - declared_constant Σ cst cdecl -> - cst_body cdecl = None -> - Σ ;;; Γ ⊢ mkApps (tConst cst u) args ⇝1 T -> - ∑ args', (T = mkApps (tConst cst u) args') × All2 (fun arg arg' => Σ ;;; Γ ⊢ arg ⇝ arg') args args'. - Proof using wfΣ. - intros hdecl hbod. revert T. - induction args using rev_ind. - - simpl. intros; eapply invert_red_axiom in hbod; tea. 2:{ eapply into_closed_red. apply red1_red. apply X. fvs. fvs. } subst. exists []. split => //. - - rewrite mkApps_app /=; intros. - depelim X; solve_discr. - depind clrel_rel; solve_discr. - * destruct args using rev_case; solve_discr. noconf H. rewrite mkApps_app in H; noconf H. solve_discr. - * specialize (IHargs N1); forward IHargs. constructor; fvs. cbn in clrel_src. rtoProp; intuition auto. - destruct IHargs as [args' []]. exists (args' ++ [M2]). subst N1. rewrite mkApps_app /=; split => //. - eapply All2_app => //. constructor; eauto. eapply closed_red_refl; fvs. cbn in clrel_src; bool. - * exists (args ++ [N2]). rewrite mkApps_app /=. split => //. - cbn in clrel_src; bool. rewrite on_free_vars_mkApps in H; bool. toAll. - eapply All2_app => //. - eapply All_All2; tea; cbn; bool. eapply closed_red_refl; fvs. eapply All2_tip, into_closed_red; tea. now constructor. - Qed. - - Lemma closed_red_ind (P : context -> term -> term -> Type) : - (forall Γ t, is_closed_context Γ -> is_open_term Γ t -> P Γ t t) -> - (forall Γ t u, Σ ;;; Γ ⊢ t ⇝1 u -> P Γ t u) -> - (forall Γ t u v, Σ ;;; Γ ⊢ t ⇝ u -> Σ ;;; Γ ⊢ u ⇝ v -> P Γ t u -> P Γ u v -> P Γ t v) -> - forall Γ t u, Σ ;;; Γ ⊢ t ⇝ u -> P Γ t u. - Proof. - intros prefl pred ptrans. - destruct 1 as [ctx src rel]. - eapply clos_rt_rt1n_iff in rel. - induction rel. eapply prefl; tea. - eapply ptrans; tea. eapply into_closed_red; tea. now eapply red1_red. - split; fvs. now eapply clos_rt_rt1n_iff. eapply pred. split; fvs. - eapply IHrel. fvs. - Qed. - - Lemma invert_red_axiom_app {Γ cst u cdecl args T} : - declared_constant Σ cst cdecl -> - cst_body cdecl = None -> - Σ ;;; Γ ⊢ mkApps (tConst cst u) args ⇝ T -> - ∑ args', (T = mkApps (tConst cst u) args') × All2 (fun arg arg' => Σ ;;; Γ ⊢ arg ⇝ arg') args args'. - Proof using wfΣ. - intros hdecl hbod red. - remember (mkApps (tConst cst u) args) as U. revert red args HeqU. - revert Γ U T. - apply: closed_red_ind; intros; subst. - - exists args. split => //. rewrite on_free_vars_mkApps /= in H0. toAll. eapply All_All2; tea; bool. now eapply closed_red_refl. - - eapply invert_red1_axiom_app in X; tea. - -specialize (X1 _ eq_refl) as [args' []]. - subst u0. specialize (X2 _ eq_refl) as [args'' []]. subst v. - exists args''. split => //. eapply All2_trans; tea. - intros x y z red1 red2. now etransitivity. - Qed. - - Lemma All2_All2_All2 {A B C} {P Q R} (l : list A) (l' : list B) (l'' : list C) : - All2 P l l' -> All2 Q l' l'' -> - (forall x y z, P x y -> Q y z -> R x z) -> - All2 R l l''. - Proof. - induction 1 in l'' |- *; intros a; depelim a; constructor; eauto. - Qed. - - Lemma ws_cumul_pb_Axiom_l_inv {pb Γ cst u args cdecl T} : - declared_constant Σ cst cdecl -> - cst_body cdecl = None -> - Σ ;;; Γ ⊢ mkApps (tConst cst u) args ≤[pb] T -> - ∑ u' args', - [× Σ ;;; Γ ⊢ T ⇝ mkApps (tConst cst u') args', - All2 (fun args args' => Σ ;;; Γ ⊢ args ≤[Conv] args') args args' & - PCUICEquality.R_universe_instance (eq_universe Σ) u u']. - Proof using wfΣ. - intros hdecl hb H. - eapply ws_cumul_pb_red in H as (v & v' & [tv tv' eqp]). - epose proof (invert_red_axiom_app hdecl hb tv) as [args' [-> ?]]. - eapply compare_term_mkApps_l_inv in eqp as [t' [l' []]]; subst v'. - depelim e. - exists u', l'. split => //. - eapply closed_red_open_right in tv'. rewrite on_free_vars_mkApps /= in tv'. - solve_all. - eapply All2_All2_All2; tea. cbn. - intros x y z red [eq op]. - eapply ws_cumul_pb_red. exists y, z; split => //. eapply closed_red_refl; fvs. - Qed. - - Lemma ws_cumul_pb_Axiom_r_inv {pb Γ cst u args cdecl T} : - declared_constant Σ cst cdecl -> - cst_body cdecl = None -> - Σ ;;; Γ ⊢ T ≤[pb] mkApps (tConst cst u) args -> - ∑ u' args', - [× Σ ;;; Γ ⊢ T ⇝ mkApps (tConst cst u') args', - All2 (fun args args' => Σ ;;; Γ ⊢ args ≤[Conv] args') args' args & - PCUICEquality.R_universe_instance (eq_universe Σ) u' u]. - Proof using wfΣ. - intros hdecl hb H. - eapply ws_cumul_pb_red in H as (v & v' & [tv tv' eqp]). - epose proof (invert_red_axiom_app hdecl hb tv') as [args' [-> ?]]. - eapply compare_term_mkApps_r_inv in eqp as [t' [l' []]]; subst v. - depelim e. - exists u0, l'. split => //. - eapply closed_red_open_right in tv. rewrite on_free_vars_mkApps /= in tv. - solve_all. - eapply All2_sym in a0. eapply All2_sym. - eapply All2_All2_All2; tea. cbn. - intros x y z red [eq op]. - eapply ws_cumul_pb_red. exists z, y; split => //. eapply closed_red_refl; fvs. - Qed. - - Lemma invert_cumul_axiom_ind {Γ cst cdecl u ind u' args args'} : - declared_constant Σ cst cdecl -> - cst_body cdecl = None -> - Σ ;;; Γ ⊢ mkApps (tConst cst u) args ≤ mkApps (tInd ind u') args' -> False. - Proof using wfΣ. - intros hd hb ht; eapply ws_cumul_pb_Axiom_l_inv in ht as (u'' & args'' & [hred hcmp _]); eauto. - eapply invert_red_mkApps_tInd in hred as (? & []); auto. solve_discr. - Qed. - - Lemma invert_cumul_axiom_prod {Γ cst cdecl u args na dom codom} : - declared_constant Σ cst cdecl -> - cst_body cdecl = None -> - Σ ;;; Γ ⊢ mkApps (tConst cst u) args ≤ tProd na dom codom -> False. - Proof using wfΣ. - intros hd hb ht; eapply ws_cumul_pb_Axiom_l_inv in ht as (u'' & args' & [hred hcmp _]); eauto. - eapply invert_red_prod in hred as (? & ? & []); auto. solve_discr. - Qed. - - Lemma invert_cumul_prim_type_ind {Γ cst cdecl p ind u' args'} : - declared_constant Σ cst cdecl -> - primitive_invariants (prim_val_tag p) cdecl -> - Σ ;;; Γ ⊢ prim_type p cst ≤ mkApps (tInd ind u') args' -> False. - Proof using wfΣ. - intros hd hb ht. - destruct p as [? []]; simp prim_type in ht. - 1-2:eapply (invert_cumul_axiom_ind (args := [])) in ht; tea; now destruct hb as [s []]. - eapply (invert_cumul_axiom_ind (args := [_])) in ht; tea. apply hb. - Qed. - - Lemma invert_cumul_prim_type_prod {Γ cst cdecl p na dom codom} : - declared_constant Σ cst cdecl -> - primitive_invariants (prim_val_tag p) cdecl -> - Σ ;;; Γ ⊢ prim_type p cst ≤ tProd na dom codom -> False. - Proof using wfΣ. - intros hd hb ht. - destruct p as [? []]; simp prim_type in ht. - 1-2:eapply (invert_cumul_axiom_prod (args := [])) in ht; tea; now destruct hb as [s []]. - eapply (invert_cumul_axiom_prod (args := [_])) in ht; tea. apply hb. - Qed. - Lemma whnf_classification' t i u args : axiom_free_value Σ [] t -> wh_normal Σ [] t -> diff --git a/pcuic/theories/PCUICConversion.v b/pcuic/theories/PCUICConversion.v index fb88ad900..23a5771aa 100644 --- a/pcuic/theories/PCUICConversion.v +++ b/pcuic/theories/PCUICConversion.v @@ -27,6 +27,9 @@ Require Import Equations.Prop.DepElim. *well-scoped* objects. *) +Notation "Σ ;;; Γ ⊢ t ⇝1 u" := (closed_red1 Σ Γ t u) (at level 50, Γ, t, u at next level, + format "Σ ;;; Γ ⊢ t ⇝1 u"). + Implicit Types (cf : checker_flags) (Σ : global_env_ext). Set Default Goal Selector "!". @@ -762,6 +765,148 @@ Section ConvCongruences. constructor. 2:cbn. all:eauto with fvs. Qed. + (* TODO move *) + Lemma invert_red_axiom {Γ cst u cdecl T} : + declared_constant Σ cst cdecl -> + cst_body cdecl = None -> + Σ ;;; Γ ⊢ tConst cst u ⇝ T -> + T = tConst cst u. + Proof using wfΣ. + intros hdecl hb. + unshelve eapply declared_constant_to_gen in hdecl; eauto. + generalize_eq x (tConst cst u). + move=> e [clΓ clt] red. + revert cst u hdecl hb e. + eapply clos_rt_rt1n_iff in red. + induction red; simplify_dep_elim. + - reflexivity. + - depelim r; solve_discr. + unshelve eapply declared_constant_to_gen in isdecl; eauto. + congruence. + Qed. + + Ltac bool := rtoProp; intuition eauto. + + Lemma invert_red1_axiom_app {Γ cst u cdecl args T} : + declared_constant Σ cst cdecl -> + cst_body cdecl = None -> + Σ ;;; Γ ⊢ mkApps (tConst cst u) args ⇝1 T -> + ∑ args', (T = mkApps (tConst cst u) args') × All2 (fun arg arg' => Σ ;;; Γ ⊢ arg ⇝ arg') args args'. + Proof using wfΣ. + intros hdecl hbod. revert T. + induction args using rev_ind. + - simpl. intros; eapply invert_red_axiom in hbod; tea. + 2:{ eapply closed_red1_red; tea. } + subst. exists []. split => //. + - rewrite mkApps_app /=; intros. + depelim X; solve_discr. + depind clrel_rel; solve_discr. + * destruct args using rev_case; solve_discr; noconf H. + rewrite mkApps_app in H; noconf H. solve_discr. + * specialize (IHargs N1); forward IHargs. + { constructor; fvs. cbn in clrel_src. rtoProp; intuition auto. } + destruct IHargs as [args' []]. exists (args' ++ [M2]). subst N1. rewrite mkApps_app /=; split => //. + eapply All2_app => //. constructor; eauto. eapply closed_red_refl; fvs. cbn in clrel_src; bool. + * exists (args ++ [N2]). rewrite mkApps_app /=. split => //. + cbn in clrel_src; bool. rewrite on_free_vars_mkApps in H; bool. toAll. + eapply All2_app => //. + + eapply All_All2; tea; cbn; bool. eapply closed_red_refl; fvs. + + constructor; [|constructor]. eapply into_closed_red; tea. now constructor. + Qed. + + Lemma closed_red_ind (P : context -> term -> term -> Type) : + (forall Γ t, is_closed_context Γ -> is_open_term Γ t -> P Γ t t) -> + (forall Γ t u, Σ ;;; Γ ⊢ t ⇝1 u -> P Γ t u) -> + (forall Γ t u v, Σ ;;; Γ ⊢ t ⇝ u -> Σ ;;; Γ ⊢ u ⇝ v -> P Γ t u -> P Γ u v -> P Γ t v) -> + forall Γ t u, Σ ;;; Γ ⊢ t ⇝ u -> P Γ t u. + Proof. + intros prefl pred ptrans. + destruct 1 as [ctx src rel]. + eapply clos_rt_rt1n_iff in rel. + induction rel. + - eapply prefl; tea. + - eapply ptrans; tea. + + eapply into_closed_red; tea. now eapply red1_red. + + split; fvs. now eapply clos_rt_rt1n_iff. + + eapply pred. split; fvs. + + eapply IHrel. fvs. + Qed. + + Lemma invert_red_axiom_app {Γ cst u cdecl args T} : + declared_constant Σ cst cdecl -> + cst_body cdecl = None -> + Σ ;;; Γ ⊢ mkApps (tConst cst u) args ⇝ T -> + ∑ args', (T = mkApps (tConst cst u) args') × All2 (fun arg arg' => Σ ;;; Γ ⊢ arg ⇝ arg') args args'. + Proof using wfΣ. + intros hdecl hbod red. + remember (mkApps (tConst cst u) args) as U. revert red args HeqU. + revert Γ U T. + apply: closed_red_ind; intros; subst. + - exists args. split => //. rewrite on_free_vars_mkApps /= in H0. toAll. eapply All_All2; tea; bool. now eapply closed_red_refl. + - eapply invert_red1_axiom_app in X; tea. + - specialize (X1 _ eq_refl) as [args' []]. + subst u0. specialize (X2 _ eq_refl) as [args'' []]. subst v. + exists args''. split => //. eapply All2_trans; tea. + intros x y z red1 red2. now etransitivity. + Qed. + + Lemma invert_red_tPrim {Γ p T} : + Σ ;;; Γ ⊢ tPrim p ⇝ T -> + (∑ i, p = (primInt; primIntModel i) /\ T = tPrim p) + + (∑ f, p = (primFloat; primFloatModel f) /\ T = tPrim p) + + ∑ a a', + [× p = (primArray; primArrayModel a), T = tPrim (primArray; primArrayModel a'), + a.(array_level) = a'.(array_level), + Σ ;;; Γ ⊢ a.(array_default) ⇝ a'.(array_default), + Σ ;;; Γ ⊢ a.(array_type) ⇝ a'.(array_type) & + All2 (fun x y => Σ ;;; Γ ⊢ x ⇝ y) a.(array_value) a'.(array_value)]. + Proof using wfΣ. + intros red; remember (tPrim p) as U. revert red p HeqU. + revert Γ U T. + apply: closed_red_ind; intros; subst. + - destruct p as [? []]. + * now left. + * left. now right. + * right. exists a, a. + cbn in H0; rtoProp; intuition auto. + split => //. solve_all; eapply All_All2; tea; cbn; intuition eauto using into_closed_red. + - right. depelim X. depelim clrel_rel; solve_discr; + cbn in clrel_src; rtoProp; intuition auto. + + exists arr, (set_array_value arr value); split => //. + toAll. cbn. + { clear -clrel_ctx H o. induction H in value, o |- *. + * depelim o. + * depelim o. + ++ constructor; eauto. + +++ eapply into_closed_red; eauto. + +++ eapply All_All2; tea; cbn; intuition eauto. + now eapply into_closed_red. + ++ constructor; eauto. now eapply into_closed_red. } + + exists arr, (set_array_default arr def); split => //. + * eapply into_closed_red; eauto. + * toAll; eapply All_All2; tea; cbn; intuition eauto using into_closed_red. + + exists arr, (set_array_type arr ty); split => //. + * eapply into_closed_red; eauto. + * toAll; eapply All_All2; tea; cbn; intuition eauto using into_closed_red. + - specialize (X1 _ eq_refl) as []. + * destruct s as [[? []]|[? []]]; subst; + specialize (X2 _ eq_refl) as []. + + left. destruct s; [left|right]; eauto. + + destruct s as [? [? []]]. subst. noconf e. + + left; destruct s; [left|right]; eauto. + + destruct s as [? [? []]]; subst; noconf e. + * right. destruct s as [a [a' []]]. + subst. specialize (X2 _ eq_refl) as [[|]|]. + + destruct s as [? []]; congruence. + + destruct s as [? []]; congruence. + + destruct s as [a2 [a2' []]]; subst. + noconf e. + exists a, a2'. split => //; try congruence. + ** etransitivity; tea. + ** etransitivity; tea. + ** eapply All2_trans; tea. + intros ?????. eapply closed_red_trans; tea. + Qed. End ConvCongruences. Notation red_terms Σ Γ := (All2 (closed_red Σ Γ)). @@ -1253,6 +1398,124 @@ Section Inversions. transitivity l0 => //. transitivity l1 => //. now symmetry. Qed. + Lemma ws_cumul_pb_Axiom_l_inv {pb Γ cst u args cdecl T} : + declared_constant Σ cst cdecl -> + cst_body cdecl = None -> + Σ ;;; Γ ⊢ mkApps (tConst cst u) args ≤[pb] T -> + ∑ u' args', + [× Σ ;;; Γ ⊢ T ⇝ mkApps (tConst cst u') args', + All2 (fun args args' => Σ ;;; Γ ⊢ args ≤[Conv] args') args args' & + PCUICEquality.R_universe_instance (eq_universe Σ) u u']. + Proof using wfΣ. + intros hdecl hb H. + eapply ws_cumul_pb_red in H as (v & v' & [tv tv' eqp]). + epose proof (invert_red_axiom_app hdecl hb tv) as [args' [-> ?]]. + eapply compare_term_mkApps_l_inv in eqp as [t' [l' []]]; subst v'. + depelim e. + exists u', l'. split => //. + eapply closed_red_open_right in tv'. rewrite on_free_vars_mkApps /= in tv'. + solve_all. + eapply All2_All2_All2; tea. cbn. + intros x y z red [eq op]. + eapply ws_cumul_pb_red. exists y, z; split => //. eapply closed_red_refl; fvs. + Qed. + + Lemma ws_cumul_pb_Axiom_r_inv {pb Γ cst u args cdecl T} : + declared_constant Σ cst cdecl -> + cst_body cdecl = None -> + Σ ;;; Γ ⊢ T ≤[pb] mkApps (tConst cst u) args -> + ∑ u' args', + [× Σ ;;; Γ ⊢ T ⇝ mkApps (tConst cst u') args', + All2 (fun args args' => Σ ;;; Γ ⊢ args ≤[Conv] args') args' args & + PCUICEquality.R_universe_instance (eq_universe Σ) u' u]. + Proof using wfΣ. + intros hdecl hb H. + eapply ws_cumul_pb_red in H as (v & v' & [tv tv' eqp]). + epose proof (invert_red_axiom_app hdecl hb tv') as [args' [-> ?]]. + eapply compare_term_mkApps_r_inv in eqp as [t' [l' []]]; subst v. + depelim e. + exists u0, l'. split => //. + eapply closed_red_open_right in tv. rewrite on_free_vars_mkApps /= in tv. + solve_all. + eapply All2_sym in a0. eapply All2_sym. + eapply All2_All2_All2; tea. cbn. + intros x y z red [eq op]. + eapply ws_cumul_pb_red. exists z, y; split => //. eapply closed_red_refl; fvs. + Qed. + + Lemma invert_cumul_axiom_ind {Γ cst cdecl u ind u' args args'} : + declared_constant Σ cst cdecl -> + cst_body cdecl = None -> + Σ ;;; Γ ⊢ mkApps (tConst cst u) args ≤ mkApps (tInd ind u') args' -> False. + Proof using wfΣ. + intros hd hb ht; eapply ws_cumul_pb_Axiom_l_inv in ht as (u'' & args'' & [hred hcmp _]); eauto. + eapply invert_red_mkApps_tInd in hred as (? & []); auto. solve_discr. + Qed. + + Lemma invert_cumul_axiom_prod {Γ cst cdecl u args na dom codom} : + declared_constant Σ cst cdecl -> + cst_body cdecl = None -> + Σ ;;; Γ ⊢ mkApps (tConst cst u) args ≤ tProd na dom codom -> False. + Proof using wfΣ. + intros hd hb ht; eapply ws_cumul_pb_Axiom_l_inv in ht as (u'' & args' & [hred hcmp _]); eauto. + eapply invert_red_prod in hred as (? & ? & []); auto. solve_discr. + Qed. + + Lemma invert_cumul_prim_type_ind {Γ cst cdecl p ind u' args'} : + declared_constant Σ cst cdecl -> + primitive_invariants (prim_val_tag p) cdecl -> + Σ ;;; Γ ⊢ prim_type p cst ≤ mkApps (tInd ind u') args' -> False. + Proof using wfΣ. + intros hd hb ht. + destruct p as [? []]; simp prim_type in ht. + 1-2:eapply (invert_cumul_axiom_ind (args := [])) in ht; tea; now destruct hb as [s []]. + eapply (invert_cumul_axiom_ind (args := [_])) in ht; tea. apply hb. + Qed. + + Lemma invert_cumul_prim_type_prod {Γ cst cdecl p na dom codom} : + declared_constant Σ cst cdecl -> + primitive_invariants (prim_val_tag p) cdecl -> + Σ ;;; Γ ⊢ prim_type p cst ≤ tProd na dom codom -> False. + Proof using wfΣ. + intros hd hb ht. + destruct p as [? []]; simp prim_type in ht. + 1-2:eapply (invert_cumul_axiom_prod (args := [])) in ht; tea; now destruct hb as [s []]. + eapply (invert_cumul_axiom_prod (args := [_])) in ht; tea. apply hb. + Qed. + + Lemma invert_cumul_Prim {Γ pb p p'} : + Σ ;;; Γ ⊢ tPrim p ≤[pb] tPrim p' -> + onPrims (ws_cumul_pb Conv Σ Γ) (eq_universe Σ) p p'. + Proof using wfΣ. + intros hd. + generalize (ws_cumul_pb_is_open_term hd); move/and3P => [] cl cl' cl''. + move/ws_cumul_pb_alt_closed: hd => [v [v' []]]. + move/invert_red_tPrim => hl. + move/invert_red_tPrim => hr. + destruct hl as [[[i []]|[f []]]|[a [a' []]]], + hr as [[[i' []]|[f' []]]|[a1 [a2' []]]]; subst; try congruence; + intros eq; depelim eq; depelim o; cbn in cl', cl''; + rtoProp; intuition eauto; constructor; eauto. + - rewrite e1 e4 //. + - etransitivity. + * eapply red_ws_cumul_pb; tea. + * symmetry. etransitivity; [eapply red_ws_cumul_pb; tea|]. + symmetry. constructor; eauto; fvs. + - etransitivity. + * eapply red_ws_cumul_pb; tea. + * symmetry. etransitivity; [eapply red_ws_cumul_pb; tea|]. + symmetry. constructor; eauto; fvs. + - etransitivity. + * eapply red_terms_ws_cumul_pb_terms; tea. + * symmetry. etransitivity; [eapply red_terms_ws_cumul_pb_terms; tea|]. + symmetry. + assert (All (is_open_term Γ) (array_value a')). + { clear -cf wfΣ a0. solve_all; now eapply closed_red_open_right in X. } + assert (All (is_open_term Γ) (array_value a2')). + { clear -cf wfΣ a2. solve_all; now eapply closed_red_open_right in X. } + solve_all. constructor; eauto; fvs. + Qed. + End Inversions. #[global] Hint Resolve closed_red_terms_open_left closed_red_terms_open_right : fvs. @@ -1420,9 +1683,6 @@ Section Inversions. - eapply ws_cumul_pb_App_r; eauto with fvs. Qed. - Notation "Σ ;;; Γ ⊢ t ⇝1 u" := (closed_red1 Σ Γ t u) (at level 50, Γ, t, u at next level, - format "Σ ;;; Γ ⊢ t ⇝1 u"). - Lemma closed_red1_mkApps_left {Γ} {t u ts} : Σ ;;; Γ ⊢ t ⇝1 u -> forallb (is_open_term Γ) ts -> diff --git a/pcuic/theories/PCUICEtaExpand.v b/pcuic/theories/PCUICEtaExpand.v index dcaf8fbfb..75e6e8d67 100644 --- a/pcuic/theories/PCUICEtaExpand.v +++ b/pcuic/theories/PCUICEtaExpand.v @@ -6,14 +6,6 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICTyping PCUICReduction PCUICProgram PCUICLiftSubst PCUICCSubst PCUICUnivSubst. (* move *) -Lemma All_fold_tip {A : Type} (P : list A -> A -> Type) {x} : All_fold P [x] -> P [] x. -Proof. - intros a; now depelim a. -Qed. - -Lemma All_tip {A} {P : A -> Type} {a : A} : P a <~> All P [a]. -Proof. split; intros. repeat constructor; auto. now depelim X. Qed. - Lemma nApp_mkApps t f args : t = mkApps f args -> ~~ isApp t -> t = f /\ args = []. Proof. diff --git a/pcuic/theories/PCUICPrincipality.v b/pcuic/theories/PCUICPrincipality.v index b7ab280db..63191d88d 100644 --- a/pcuic/theories/PCUICPrincipality.v +++ b/pcuic/theories/PCUICPrincipality.v @@ -463,16 +463,6 @@ Proof. intros. now eapply eq_term_empty_eq_term. Qed. -Lemma All2_tip_l {A B} {R : A -> B -> Type} x y : All2 R [x] y -> ∑ y', (y = [y']) * R x y'. -Proof. - intros a; depelim a. depelim a. exists y0; split => //. -Qed. - -Lemma Forall2_tip_l {A B} {R : A -> B -> Prop} x y : Forall2 R [x] y -> exists y', (y = [y']) * R x y'. -Proof. - intros a; depelim a. depelim a. exists y0; split => //. -Qed. - From MetaCoq.PCUIC Require Import PCUICClassification. Lemma typing_leq_term {cf:checker_flags} (Σ : global_env_ext) Γ t t' T T' : diff --git a/pcuic/theories/PCUICSR.v b/pcuic/theories/PCUICSR.v index cfd6a1994..791e0d37d 100644 --- a/pcuic/theories/PCUICSR.v +++ b/pcuic/theories/PCUICSR.v @@ -201,20 +201,6 @@ Ltac hide H := | ?ty => change ty with (@hidebody _ ty) in H end. -Lemma All2i_nth_error {A B} {P : nat -> A -> B -> Type} {l l' n x c k} : - All2i P k l l' -> - nth_error l n = Some x -> - nth_error l' n = Some c -> - P (k + n)%nat x c. -Proof. - induction 1 in n |- *. - * rewrite !nth_error_nil => //. - * destruct n. - + simpl. intros [= <-] [= <-]. now rewrite Nat.add_0_r. - + simpl. intros hnth hnth'. specialize (IHX _ hnth hnth'). - now rewrite Nat.add_succ_r. -Qed. - Lemma conv_context_smash_end {cf Σ} {wfΣ : wf Σ} (Γ Δ Δ' : context) : wf_local Σ (Γ ,,, Δ) -> wf_local Σ (Γ ,,, Δ') -> @@ -841,17 +827,6 @@ Proof. intros. eapply ctx_inst_merge; try rewrite ?(List.rev_involutive Δ) //; tea. Qed. -Lemma All2i_All2i_mix {A B} {P Q : nat -> A -> B -> Type} - {n} {l : list A} {l' : list B} : - All2i P n l l' -> All2i Q n l l' -> All2i (fun i x y => (P i x y * Q i x y)%type) n l l'. -Proof. - induction 2; simpl; intros; constructor. - inv X; intuition auto. - apply IHX0. inv X; intuition auto. -Qed. - -Definition conj_impl {A B} : A -> (A -> B) -> A × B := fun x f => (x, f x). - Lemma is_open_term_snoc (Γ : context) M d : on_free_vars (shiftnP 1 (shiftnP #|Γ| xpred0)) M -> is_open_term (Γ ,, d) M. Proof. rewrite /=. @@ -1292,10 +1267,6 @@ Proof. eapply All2_refl; reflexivity. Qed. -Lemma All2_tip {A} {P} (t u : A) : P t u -> All2 P [t] [u]. -Proof. now repeat constructor. Qed. -#[global] Hint Resolve All2_tip : core. - Lemma map2_set_binder_name_expand_lets nas Γ Δ : #|nas| = #|Δ| -> map2 set_binder_name nas (expand_lets_ctx Γ Δ) = @@ -2248,8 +2219,8 @@ Proof. do 2 case. move=> hnth [] wfbr wfbctxargs wfbrctx wfcbc' wfcbcty'. case => eqbctx. case. case => wfbctx _. move=> [] [] Hbody IHbody [] brty IHbrty. - eapply conj_impl. solve_all. move=> cvcbc. - apply conj_impl; [|move=> wfcb']. + eapply and_assum. solve_all. move=> cvcbc. + apply and_assum; [|move=> wfcb']. { now eapply typing_wf_local in wfcbcty'. } split => //. have declc : declared_constructor Σ (ci, cstr) mdecl idecl cdecl. diff --git a/pcuic/theories/PCUICSafeLemmata.v b/pcuic/theories/PCUICSafeLemmata.v index 97b6621ea..9d398c82a 100644 --- a/pcuic/theories/PCUICSafeLemmata.v +++ b/pcuic/theories/PCUICSafeLemmata.v @@ -490,6 +490,14 @@ Section Lemmata. cbn in *. eapply closed_context_conversion; tea. now symmetry. + - eapply inversion_Prim in typ as (?&?&[]); eauto. + depelim p0. now eexists. + - eapply inversion_Prim in typ as (?&?&[]); eauto. + depelim p0. now eexists. + - eapply inversion_Prim in typ as (?&?&[]); eauto. + depelim p0. cbn in *. + eapply All_app in hvalue as [_ hvalue]. depelim hvalue. + now eexists. Qed. Lemma cored_red : @@ -777,6 +785,81 @@ Section Lemmata. apply isProdmkApps in hh. assumption. Qed. + Fixpoint isAppPrim (t : term) : bool := + match t with + | tApp t l => isAppPrim t + | tPrim p => true + | _ => false + end. + + Lemma isAppPrim_mkApps : + forall t l, isAppPrim (mkApps t l) = isAppPrim t. + Proof using Type. + intros t l. revert t. + induction l ; intros t. + - reflexivity. + - cbn. rewrite IHl. reflexivity. + Qed. + + Definition isPrim t := + match t with + | tPrim _ => true + | _ => false + end. + + Lemma isPrimmkApps : + forall t l, + isPrim (mkApps t l) -> + l = []. + Proof using Type. + intros t l h. + revert t h. + induction l ; intros t h. + - reflexivity. + - cbn in h. specialize IHl with (1 := h). subst. + cbn in h. discriminate h. + Qed. + + Lemma isAppPrim_isPrim : + forall Γ t, + isAppPrim t -> + welltyped Σ Γ t -> + isPrim t. + Proof using hΣ. + intros Γ t hp hw. + induction t in Γ, hp, hw |- *. + all: try discriminate hp. + 2:reflexivity. + - simpl in hp. + specialize IHt1 with (1 := hp). + assert (welltyped Σ Γ t1) as h. + { destruct hw as [T h]. + apply inversion_App in h as hh ; auto. + destruct hh as [na [A' [B' [? [? ?]]]]]. + eexists. eassumption. + } + specialize IHt1 with (1 := h). + destruct t1. + all: try discriminate IHt1. + destruct hw as [T hw']. + apply inversion_App in hw' as ihw' ; auto. + destruct ihw' as [na' [A' [B' [hP [? ?]]]]]. + apply inversion_Prim in hP as [s1 [s2 [? htyp ? ? ? bot]]] ; auto. + now eapply invert_cumul_prim_type_prod in bot. + Qed. + + Lemma mkApps_Prim_nil : + forall Γ p l, + welltyped Σ Γ (mkApps (tPrim p) l) -> + l = []. + Proof using hΣ. + intros Γ p l h. + pose proof (isAppPrim_isPrim) as hh. + specialize hh with (2 := h). + rewrite isAppPrim_mkApps in hh. + specialize hh with (1 := eq_refl). + apply isPrimmkApps in hh. assumption. + Qed. (* TODO MOVE or even replace old lemma *) Lemma decompose_stack_noStackApp : diff --git a/safechecker/theories/PCUICErrors.v b/safechecker/theories/PCUICErrors.v index 361b6688a..2e9fe7bac 100644 --- a/safechecker/theories/PCUICErrors.v +++ b/safechecker/theories/PCUICErrors.v @@ -97,6 +97,12 @@ Inductive ConversionError := (Γ : context) (p : prim_val) (Γ' : context) (p' : prim_val) +| ArrayNotConvertibleLevels + (Γ : context) (a : array_model term) (Γ' : context) (a' : array_model term) + +| ArrayValuesNotSameLength + (Γ : context) (a : array_model term) (Γ' : context) (a' : array_model term) + | ArrayNotConvertibleValues (Γ : context) (a : array_model term) (Γ' : context) (a' : array_model term) (e : ConversionError) @@ -291,6 +297,12 @@ Fixpoint string_of_conv_error Σ (e : ConversionError) : string := nl ^ "and" ^ nl ^ print_term Σ Γ' (tCoFix mfix' idx) ^ nl ^ "have a different number of mutually defined functions." + | ArrayValuesNotSameLength Γ a Γ' a' => + "The two arrays " ^ nl ^ + print_term Σ Γ (tPrim (primArray; primArrayModel a)) ^ + nl ^ "and" ^ nl ^ + print_term Σ Γ (tPrim (primArray; primArrayModel a')) ^ + nl ^ " have different length" | ArrayNotConvertibleValues Γ a Γ' a' e => "The two arrays " ^ nl ^ print_term Σ Γ (tPrim (primArray; primArrayModel a)) ^ @@ -298,6 +310,12 @@ Fixpoint string_of_conv_error Σ (e : ConversionError) : string := print_term Σ Γ (tPrim (primArray; primArrayModel a')) ^ nl ^ " have non-convertible values: " ^ nl ^ string_of_conv_error Σ e + | ArrayNotConvertibleLevels Γ a Γ' a' => + "The two arrays " ^ nl ^ + print_term Σ Γ (tPrim (primArray; primArrayModel a)) ^ + nl ^ "and" ^ nl ^ + print_term Σ Γ (tPrim (primArray; primArrayModel a')) ^ + nl ^ " have non-convertible universe levels" | ArrayNotConvertibleTypes Γ a Γ' a' e => "The two arrays " ^ nl ^ print_term Σ Γ (tPrim (primArray; primArrayModel a)) ^ diff --git a/safechecker/theories/PCUICSafeConversion.v b/safechecker/theories/PCUICSafeConversion.v index c49846fcd..e2f8e5be6 100644 --- a/safechecker/theories/PCUICSafeConversion.v +++ b/safechecker/theories/PCUICSafeConversion.v @@ -1,5 +1,5 @@ (* Distributed under the terms of the MIT license. *) -From Coq Require Import ProofIrrelevance. +From Coq Require Import ProofIrrelevance ssreflect ssrbool. From MetaCoq.Utils Require Import utils. From MetaCoq.Common Require Import config uGraph. From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils @@ -590,7 +590,7 @@ Section Conversion. destruct (decompose_stack π) in isr. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. destruct (isr _ wfΣ) as [isr']. - depelim isr'; rewrite mkApps_tApp in *; try solve [solve_discr]. + depelim isr'; rewrite -> mkApps_tApp in *; try solve [solve_discr]. apply whne_mkApps_inv in w; [|easy]. destruct w as [|(?&?&?&?&?&?&?&?)]; [|discriminate]. depelim w; solve_discr; discriminate. @@ -624,7 +624,7 @@ Section Conversion. destruct decompose_stack eqn:decomp. apply decompose_stack_eq in decomp as ->. cbn; intros ->. - rewrite zipc_appstack, zipp_as_mkApps, decompose_stack_appstack. + rewrite zipc_appstack zipp_as_mkApps decompose_stack_appstack. cbn. now rewrite app_nil_r. Qed. @@ -1040,15 +1040,14 @@ Section Conversion. + unfold is_constructor. pose proof (eq_sym eq2) as eql. apply decompose_stack_at_length in eql. subst. - rewrite nth_error_app_ge by auto. + rewrite nth_error_app_ge; auto. replace (#|l| - #|l|) with 0 by lia. cbn. case_eq (decompose_stack ρ). intros l0 s1 ee. rewrite ee in hd. pose proof (decompose_stack_eq _ _ _ ee). subst. cbn in hd. subst. rewrite zipc_appstack. cbn. - unfold isConstruct_app. rewrite decompose_app_mkApps by auto. - reflexivity. + unfold isConstruct_app. rewrite decompose_app_mkApps; auto. Qed. Lemma unfold_one_fix_red_zippx : @@ -1088,14 +1087,14 @@ Section Conversion. + unfold is_constructor. pose proof (eq_sym eq2) as eql. apply decompose_stack_at_length in eql. subst. - rewrite nth_error_app_ge by auto. + rewrite -> nth_error_app_ge by auto. replace (#|l| - #|l|) with 0 by lia. cbn. case_eq (decompose_stack ρ). intros l0 s1 ee. rewrite ee in hd. pose proof (decompose_stack_eq _ _ _ ee). subst. cbn in hd. subst. rewrite zipc_appstack. cbn. - unfold isConstruct_app. rewrite decompose_app_mkApps by auto. + unfold isConstruct_app. rewrite -> decompose_app_mkApps by auto. reflexivity. Qed. @@ -1131,15 +1130,14 @@ Section Conversion. + unfold is_constructor. pose proof (eq_sym eq2) as eql. apply decompose_stack_at_length in eql. subst. - rewrite nth_error_app_ge by auto. + rewrite -> nth_error_app_ge by auto. replace (#|l| - #|l|) with 0 by lia. cbn. case_eq (decompose_stack ρ). intros l0 s1 ee. rewrite ee in hd. pose proof (decompose_stack_eq _ _ _ ee). subst. cbn in hd. subst. rewrite zipc_appstack. cbn. - unfold isConstruct_app. rewrite decompose_app_mkApps by auto. - reflexivity. + unfold isConstruct_app. rewrite decompose_app_mkApps; auto. Qed. Lemma unfold_one_fix_cored : @@ -1173,14 +1171,14 @@ Section Conversion. + unfold is_constructor. pose proof (eq_sym eq2) as eql. apply decompose_stack_at_length in eql. subst. - rewrite nth_error_app_ge by auto. + rewrite -> nth_error_app_ge by auto. replace (#|l| - #|l|) with 0 by lia. cbn. case_eq (decompose_stack ρ). intros l0 s1 ee. rewrite ee in hd. pose proof (decompose_stack_eq _ _ _ ee). subst. cbn in hd. subst. rewrite zipc_appstack. cbn. - unfold isConstruct_app. rewrite decompose_app_mkApps by auto. + unfold isConstruct_app. rewrite -> decompose_app_mkApps by auto. reflexivity. Qed. @@ -1242,7 +1240,7 @@ Section Conversion. rewrite <- (stack_context_decompose ρ), decomp in wh. change (App_l c :: θ) with (appstack [c] θ) in *. rewrite !decompose_stack_appstack. - rewrite zipp_as_mkApps, !decompose_stack_appstack in h. + rewrite zipp_as_mkApps !decompose_stack_appstack in h. destruct h as (ty&typ). cbn in *. rewrite stack_context_appstack in typ. @@ -1267,7 +1265,7 @@ Section Conversion. apply whnf_ne. econstructor. + eauto. - + rewrite nth_error_snoc by easy. + + rewrite -> nth_error_snoc by easy. eauto. + eapply whnf_fix_arg_whne; eauto. now destruct cred. @@ -1663,8 +1661,8 @@ Qed. specialize (H _ wfΣ). apply conv_cum_alt in H as [(?&?&[r1 r2 eq])]; auto. 2: pose proof (hΣ _ wfΣ); sq ; eauto. - rewrite zipp_as_mkApps in r1, r2. - erewrite <- abstract_env_lookup_correct' in eq1, eq2; eauto. + rewrite !zipp_as_mkApps in r1, r2. + erewrite <- !abstract_env_lookup_correct' in eq1, eq2; eauto. symmetry in eq1, eq2. eapply declared_constant_from_gen in eq1, eq2. generalize hΣ. intros []; eauto. @@ -1921,7 +1919,7 @@ Qed. destruct h1 as [h1]. eapply welltyped_zipc_zipp in h; auto. eapply welltyped_zipc_zipp in h'; auto. - rewrite zipp_as_mkApps in h, h'. + rewrite !zipp_as_mkApps in h, h'. destruct h as [s h]. destruct h' as [s' h']. eapply PCUICValidity.inversion_mkApps in h as [A [hcase _]]. eapply PCUICValidity.inversion_mkApps in h' as [A' [hcase' _]]. @@ -2033,7 +2031,7 @@ Qed. apply welltyped_zipc_tCase_brs_length in h' as (?&?&?&?); eauto. unshelve eapply declared_inductive_to_gen in H, H1; eauto. pose proof (PCUICInductiveInversion.declared_inductive_unique_sig H H1) as u; noconf u. - rewrite app_length in *. + rewrite !app_length in H0, H2. cbn in *. lia. Qed. @@ -2047,7 +2045,7 @@ Qed. apply welltyped_zipc_tCase_brs_length in h' as (?&?&?&?); eauto. unshelve eapply declared_inductive_to_gen in H, H1; eauto. pose proof (PCUICInductiveInversion.declared_inductive_unique_sig H H1) as u; noconf u. - rewrite app_length in *. + rewrite !app_length in H0, H2. cbn in *. lia. Qed. @@ -2591,7 +2589,7 @@ Qed. Next Obligation. apply h''; clear h''. destruct H as [H]; inversion H; constructor. - rewrite map_app, <- app_assoc; simpl; assumption. + rewrite map_app -app_assoc; simpl; assumption. Qed. Next Obligation. apply h''; clear h''. @@ -2720,7 +2718,7 @@ Qed. destruct ir as (_&[wh]); eauto. eapply eqb_term_upto_univ_impl with (p := wf_universeb Σ) (q := closedu) in eq; tea. 2-3: intros; apply iff_reflect; eapply (abstract_env_compare_universe_correct _ wfΣ Conv) ; now eapply wf_universe_iff. - 2:{ intros. rewrite wf_universeb_instance_forall in *. + 2:{ intros. rewrite !wf_universeb_instance_forall in H0, H1. apply wf_universe_instance_iff in H0. apply wf_universe_instance_iff in H1. apply compare_global_instance_correct; eauto. @@ -2817,7 +2815,7 @@ Qed. pose proof (hΣ _ wfΣ). eapply eqb_term_upto_univ_impl in eq; tea. 2-3: intros; apply iff_reflect; eapply (abstract_env_compare_universe_correct _ wfΣ Conv) ; now eapply wf_universe_iff. - 2:{ intros. rewrite wf_universeb_instance_forall in *. + 2:{ intros. rewrite !wf_universeb_instance_forall in H0, H1. apply wf_universe_instance_iff in H0. apply wf_universe_instance_iff in H1. eapply (compare_global_instance_correct wfΣ); eauto. } @@ -3073,7 +3071,7 @@ Qed. destruct hp as [hp]. eapply welltyped_zipc_zipp in h; auto. eapply welltyped_zipc_zipp in h'; auto. - rewrite zipp_as_mkApps in h, h'. + rewrite !zipp_as_mkApps in h, h'. destruct h as [s h]. destruct h' as [s' h']. eapply PCUICValidity.inversion_mkApps in h as [A [hcase _]]. eapply PCUICValidity.inversion_mkApps in h' as [A' [hcase' _]]. @@ -3267,6 +3265,119 @@ Qed. destruct H as [[]]; constructor; auto. Qed. +Notation tPrimArray a := (tPrim (primArray; primArrayModel a)). + +Equations (noeqns) isconv_array_values_aux + (Γ : context) + (a1 : array_model term) + (π1 : stack) + (h1 : wtp Γ (tPrimArray a1) π1) + (a2 : array_model term) + (π2 : stack) + (h2 : wtp Γ (tPrimArray a2) π2) + (hx : conv_stack_ctx Γ π1 π2) + (aux : Aux Term Γ (tPrimArray a1) π1 (tPrimArray a2) π2 h2) + (pre1 pre2 post1 post2 : list term) + (eq1 : a1.(array_value) = pre1 ++ post1) + (eq2 : a2.(array_value) = pre2 ++ post2) : + ConversionResult (forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ws_cumul_pb_terms Σ (Γ,,, stack_context π1) post1 post2∥) := + isconv_array_values_aux + Γ a1 π1 h1 a2 π2 h2 + hx aux pre1 pre2 [] [] eq1 eq2 => yes; + + isconv_array_values_aux + Γ a1 π1 h1 a2 π2 h2 + hx aux pre1 pre2 (t1 :: post1) (t2 :: post2) eq1 eq2 + with isconv_red_raw + Conv + t1 (PrimArray_val a1.(array_level) + pre1 post1 + a1.(array_default) a1.(array_type) :: π1) + t2 (PrimArray_val a2.(array_level) + pre2 post2 + a2.(array_default) a2.(array_type) :: π2) aux := { + + | Error ce not_conv_term => no ce; + + | Success conv_tm + with isconv_array_values_aux + Γ a1 π1 h1 a2 π2 h2 hx aux + (pre1 ++ [t1]) (pre2 ++ [t2]) post1 post2 _ _ := { + + | Error ce not_conv_rest => no ce; + + | Success conv_rest => yes + } + }; + + isconv_array_values_aux + Γ a1 π1 h1 a2 π2 h2 hx aux + pre1 pre2 post1 post2 eq eq2 => no (ArrayValuesNotSameLength + (Γ,,, stack_context π1) a1 + (Γ,,, stack_context π2) a2). + Next Obligation. + destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]; + specialize_Σ wfΣ. + destruct H as [H]. + depelim H. + Qed. + Next Obligation. + destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]; + specialize_Σ wfΣ. + destruct H as [H]. + depelim H. + Qed. + Next Obligation. + destruct a1; auto. + Qed. + Next Obligation. + destruct a2; auto. + Qed. + Next Obligation. + apply R_positionR. all: simpl. + 1: destruct a1; cbn in *; subst; reflexivity. + rewrite stack_position_cons. + rewrite <- app_nil_r. + eapply positionR_poscat. + constructor. + Qed. + Next Obligation. + rewrite <- app_assoc; auto. + Qed. + Next Obligation. + rewrite <- app_assoc; auto. + Qed. + Next Obligation. + specialize_Σ wfΣ. + destruct conv_tm, conv_rest. + unfold zipp in X; simpl in *. + constructor; constructor; auto. + Qed. + Next Obligation. + contradiction not_conv_rest. intros. + specialize_Σ wfΣ. + destruct H as [H]; depelim H. + constructor; auto. + Qed. + Next Obligation. + contradiction not_conv_term. intros Σ wfΣ. + specialize_Σ wfΣ. + destruct H as [H]; depelim H. + constructor; auto. + Qed. + + Definition isconv_array_values + (Γ : context) + (a1 : array_model term) + (π1 : stack) + (h1 : wtp Γ (tPrimArray a1) π1) + (a2 : array_model term) + (π2 : stack) + (h2 : wtp Γ (tPrimArray a2) π2) + (hx : conv_stack_ctx Γ π1 π2) + (aux : Aux Term Γ (tPrimArray a1) π1 (tPrimArray a2) π2 h2) := + isconv_array_values_aux Γ a1 π1 h1 a2 π2 h2 hx aux [] [] a1.(array_value) a2.(array_value) eq_refl eq_refl. + Lemma conv_cum_red_inv Σ (wfΣ : abstract_env_ext_rel X Σ) leq Γ t1 t2 t1' t2' : red Σ Γ t1 t1' -> red Σ Γ t2 t2' -> @@ -3483,24 +3594,32 @@ Qed. ) } ; - | prog_view_Prim p p' with inspect (eqb (prim_val_tag p) (prim_val_tag p')) := { - | @exist false tag_uneq := - no (DistinctPrimTag (Γ ,,, stack_context π1) p (Γ ,,, stack_context π2) p') - | @exist true eqtag with p, p' := { - | (primInt; primIntModel i) | (primInt, primIntModel i') with inspect (eqb i i') := + | prog_view_Prim p p' with eq_dec (prim_val_tag p) (prim_val_tag p') := { + | right tag_uneq := no (DistinctPrimTags (Γ ,,, stack_context π1) p (Γ ,,, stack_context π2) p') + | left eqtag with p, p' := { + | (primInt; primIntModel i) | (primInt; primIntModel i') with inspect (eqb i i') := { | @exist true eqi := yes - | @exist false neqi := no (DistingPrimValues (Γ ,,, stack_context π1) p (Γ ,,, stack_context π2) p') } + | @exist false neqi := no (DistinctPrimValues (Γ ,,, stack_context π1) p (Γ ,,, stack_context π2) p') } | (primFloat; primFloatModel f) | (primFloat; primFloatModel f') with inspect (eqb f f') := { | @exist true eqf := yes - | @exist false neqf := no (DistingPrimValues (Γ ,,, stack_context π1) p (Γ ,,, stack_context π2) p') } - | (primArray; primArrayModel a) | (primArray; primArrayModel a') with - with isconv_red_raw Conv (array_type a) (Prod_l na B1 :: π1) (array_type a') (Prod_l na' B2 :: π2) aux : - inspect (eqb f f') := - { | @exist true eqf := yes - | @exist false neqf := no (DistingPrimValues (Γ ,,, stack_context π1) p (Γ ,,, stack_context π2) p') } - - - } + | @exist false neqf := no (DistinctPrimValues (Γ ,,, stack_context π1) p (Γ ,,, stack_context π2) p') } + | (primArray; primArrayModel a) | (primArray; primArrayModel a') + with inspect (abstract_env_eq X (Universe.make (array_level a)) (Universe.make (array_level a'))) := + { | @exist false neql := no (ArrayNotConvertibleLevels (Γ ,,, stack_context π1) a (Γ ,,, stack_context π2) a') + | @exist true eql with isconv_red_raw Conv (array_type a) (PrimArray_ty a.(array_level) a.(array_value) a.(array_default) :: π1) + (array_type a') (PrimArray_ty a'.(array_level) a'.(array_value) a'.(array_default) :: π2) aux := { + | Success convdiscrty with isconv_red_raw Conv (array_default a) (PrimArray_def a.(array_level) a.(array_value) a.(array_type) :: π1) + (array_default a') (PrimArray_def a'.(array_level) a'.(array_value) a'.(array_type) :: π2) aux := { + | Success convdiscrdef with isconv_array_values Γ a π1 _ a' π2 _ hx aux := { + | Success convdiscrval := yes + | Error e h := no (ArrayNotConvertibleValues (Γ ,,, stack_context π1) a (Γ ,,, stack_context π2) a' e) + } + | Error e h := no (ArrayNotConvertibleDefault (Γ ,,, stack_context π1) a (Γ ,,, stack_context π2) a' e) + } ; + | Error e h := no (ArrayNotConvertibleTypes (Γ ,,, stack_context π1) a (Γ ,,, stack_context π2) a' e) + } + } + } } ; | prog_view_other t1 t2 h := @@ -3716,7 +3835,7 @@ Qed. apply welltyped_zipc_zipp in h1; auto. clear aux. eapply welltyped_zipc_zipp in h2; eauto. - rewrite !zipp_as_mkApps in *. + rewrite !zipp_as_mkApps in h1, h2 |- *. apply mkApps_Prod_nil in h1 as ->; auto. apply mkApps_Prod_nil in h2 as ->; auto. destruct h. @@ -3734,10 +3853,10 @@ Qed. apply welltyped_zipc_zipp in h1; auto. clear aux. eapply welltyped_zipc_zipp in h2; eauto. - rewrite !zipp_as_mkApps in *. + rewrite !zipp_as_mkApps in H, h1, h2 |- *. apply mkApps_Prod_nil in h1; auto. apply mkApps_Prod_nil in h2; auto. - rewrite h1, h2 in H. + rewrite h1 h2 in H. apply Prod_conv_cum_inv in H as (?&_&?); auto. Qed. Next Obligation. @@ -3750,10 +3869,10 @@ Qed. apply welltyped_zipc_zipp in h1; auto. clear aux. eapply welltyped_zipc_zipp in h2; eauto. - rewrite !zipp_as_mkApps in *. + rewrite !zipp_as_mkApps in H, h1, h2 |- *. apply mkApps_Prod_nil in h1; auto. apply mkApps_Prod_nil in h2; auto. - rewrite h1, h2 in H. + rewrite h1 h2 in H. apply Prod_conv_cum_inv in H as (?&?&?); auto. eapply (ssrbool.elimF (eqb_annot_reflect _ _)); tea. now unfold eqb_binder_annot. @@ -3766,10 +3885,10 @@ Qed. apply welltyped_zipc_zipp in h1; auto. clear aux. eapply welltyped_zipc_zipp in h2; eauto. - rewrite !zipp_as_mkApps in *. + rewrite !zipp_as_mkApps in h1, h2, H |- *. apply mkApps_Prod_nil in h1; auto. apply mkApps_Prod_nil in h2; auto. - rewrite h1, h2 in H. + rewrite h1 h2 in H. apply Prod_conv_cum_inv in H as (?&?&_); auto. Qed. (* tCase *) @@ -3881,10 +4000,10 @@ Qed. unfold reduce_term in eq4. rewrite e' in eq4. cbn in eq4. epose proof (eqb_term_upto_univ_refl Σ _ _ _ _ 0 c' _ _). - rewrite H in eq4. + erewrite H in eq4. - discriminate. - intros. apply iff_reflect. eapply abstract_env_compare_universe_correct with (conv_pb := Conv) ; eauto. - - intros. rewrite wf_universeb_instance_forall in *. + - intros. rewrite !wf_universeb_instance_forall in H0, H1. apply wf_universe_instance_iff in H0. apply wf_universe_instance_iff in H1. eapply compare_global_instance_correct; eauto. @@ -3904,10 +4023,10 @@ Qed. rewrite <- H2 in eq4. cbn in eq4. epose proof (eqb_term_upto_univ_refl Σ _ _ _ _ 0 c' _ _). - rewrite H1 in eq4. + erewrite H1 in eq4. - discriminate. - intros. apply iff_reflect. eapply abstract_env_compare_universe_correct with (conv_pb := Conv); eauto. - - intros. rewrite wf_universeb_instance_forall in *. + - intros. rewrite !wf_universeb_instance_forall in H3, H4. apply wf_universe_instance_iff in H3. apply wf_universe_instance_iff in H4. eapply compare_global_instance_correct; eauto. @@ -3985,10 +4104,10 @@ Qed. unfold reduce_term in eq3. rewrite e' in eq3. cbn in eq3. epose proof (eqb_term_upto_univ_refl Σ _ _ _ _ 0 c _ _). - rewrite H in eq3. + erewrite H in eq3. - discriminate. - intros. apply iff_reflect. eapply abstract_env_compare_universe_correct with (conv_pb := Conv); eauto. - - intros. rewrite wf_universeb_instance_forall in *. + - intros. rewrite !wf_universeb_instance_forall in H0, H1. apply wf_universe_instance_iff in H0. apply wf_universe_instance_iff in H1. eapply compare_global_instance_correct; eauto. @@ -4008,10 +4127,10 @@ Qed. rewrite <- H2 in eq3. cbn in eq3. epose proof (eqb_term_upto_univ_refl Σ _ _ _ _ 0 c _ _). - rewrite H1 in eq3. + erewrite H1 in eq3. - discriminate. - intros. apply iff_reflect. eapply abstract_env_compare_universe_correct with (conv_pb := Conv); eauto. - - intros. rewrite wf_universeb_instance_forall in *. + - intros. rewrite !wf_universeb_instance_forall in H3, H4. apply wf_universe_instance_iff in H3. apply wf_universe_instance_iff in H4. eapply compare_global_instance_correct; eauto. @@ -4136,10 +4255,10 @@ Qed. unfold reduce_term in eq4. rewrite e' in eq4. cbn in eq4. epose proof (eqb_term_upto_univ_refl Σ _ _ _ _ 0 c' _ _). - rewrite H in eq4. + erewrite H in eq4. - discriminate. - intros. apply iff_reflect. eapply abstract_env_compare_universe_correct with (conv_pb := Conv); eauto. - - intros. rewrite wf_universeb_instance_forall in *. + - intros. rewrite !wf_universeb_instance_forall in H0, H1. apply wf_universe_instance_iff in H0. apply wf_universe_instance_iff in H1. eapply compare_global_instance_correct; eauto. @@ -4158,10 +4277,10 @@ Qed. rewrite <- H2 in eq4. cbn in eq4. epose proof (eqb_term_upto_univ_refl Σ _ _ _ _ 0 c' _ _). - rewrite H1 in eq4. + erewrite H1 in eq4. - discriminate. - intros. apply iff_reflect. eapply abstract_env_compare_universe_correct with (conv_pb := Conv); eauto. - - intros. rewrite wf_universeb_instance_forall in *. + - intros. rewrite !wf_universeb_instance_forall in H3, H4. apply wf_universe_instance_iff in H3. apply wf_universe_instance_iff in H4. eapply compare_global_instance_correct; eauto. @@ -4241,10 +4360,10 @@ Qed. unfold reduce_term in eq3. rewrite e' in eq3. cbn in eq3. epose proof (eqb_term_upto_univ_refl Σ _ _ _ _ 0 c _ _). - rewrite H in eq3. + erewrite H in eq3. - discriminate. - intros. apply iff_reflect. eapply abstract_env_compare_universe_correct with (conv_pb := Conv); eauto. - - intros. rewrite wf_universeb_instance_forall in *. + - intros. rewrite !wf_universeb_instance_forall in H0, H1. apply wf_universe_instance_iff in H0. apply wf_universe_instance_iff in H1. eapply compare_global_instance_correct; eauto. @@ -4263,10 +4382,10 @@ Qed. rewrite <- H2 in eq3. cbn in eq3. epose proof (eqb_term_upto_univ_refl Σ _ _ _ _ 0 c _ _). - rewrite H1 in eq3. + erewrite H1 in eq3. - discriminate. - intros. apply iff_reflect. eapply abstract_env_compare_universe_correct with (conv_pb := Conv); eauto. - - intros. rewrite wf_universeb_instance_forall in *. + - intros. rewrite !wf_universeb_instance_forall in H3, H4. apply wf_universe_instance_iff in H3. apply wf_universe_instance_iff in H4. eapply compare_global_instance_correct; eauto. @@ -4657,6 +4776,248 @@ Qed. congruence. Qed. + (* Primitive arrays *) + Next Obligation. + rename H into wfΣ. pose proof (hΣ _ wfΣ) as [hΣ]. + eapply conv_cum_zipp; eauto. + * sq. eapply ws_cumul_pb_Prim; eauto. + + symmetry in eqi; eapply eqb_eq in eqi. subst i'. + constructor. + + specialize (hx Σ wfΣ) as []. fvs. + * specialize (hx Σ wfΣ). + specialize (h1 Σ wfΣ). + apply welltyped_zipc_zipp in h1; auto. + clear aux. + eapply welltyped_zipc_zipp in h2; eauto. + rewrite !zipp_as_mkApps in h1, h2 |- *. + apply mkApps_Prim_nil in h1 as ->; auto. + apply mkApps_Prim_nil in h2 as ->; auto. + Qed. + + Next Obligation. + Proof. + destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. + pose proof (hΣ _ wfΣ) as [hΣ]. + specialize (H _ wfΣ). + specialize (hx Σ wfΣ). + specialize (h1 Σ wfΣ). + apply welltyped_zipc_zipp in h1; eauto. + clear aux. + eapply welltyped_zipc_zipp in h2; eauto. + rewrite !zipp_as_mkApps in H, h1, h2 |- *. + apply mkApps_Prim_nil in h1; auto. + apply mkApps_Prim_nil in h2; auto. + rewrite h1 h2 in H. + destruct H as [H]. cbn in H. + eapply invert_cumul_Prim in H; depelim H. + now rewrite eqb_refl in neqi. + Qed. + + Next Obligation. + Proof. + symmetry in eqf. + apply eqb_eq in eqf; subst f'. + rename H into wfΣ. pose proof (hΣ _ wfΣ) as [hΣ]. + clear aux. + specialize (hx Σ wfΣ). + specialize (h1 Σ wfΣ). + specialize (h2 Σ wfΣ). + apply welltyped_zipc_zipp in h1; eauto. + eapply welltyped_zipc_zipp in h2; eauto. + rewrite !zipp_as_mkApps in h1, h2 |- *. + apply mkApps_Prim_nil in h1; auto. + apply mkApps_Prim_nil in h2; auto. + rewrite h1 h2; cbn. + red. sq. eapply ws_cumul_pb_Prim; eauto; fvs. + constructor. + Qed. + + Next Obligation. + Proof. + destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. + pose proof (hΣ _ wfΣ) as [hΣ]. + specialize (H _ wfΣ). + specialize (hx Σ wfΣ). + specialize (h1 Σ wfΣ). + apply welltyped_zipc_zipp in h1; eauto. + clear aux. + eapply welltyped_zipc_zipp in h2; eauto. + rewrite !zipp_as_mkApps in h1, h2, H |- *. + apply mkApps_Prim_nil in h1; auto. + apply mkApps_Prim_nil in h2; auto. + rewrite h1 h2 in H. + destruct H as [H]. cbn in H. + eapply invert_cumul_Prim in H; depelim H. + now rewrite eqb_refl in neqf. + Qed. + + Next Obligation. + Proof. + destruct a; eauto. + Qed. + + Next Obligation. + Proof. + destruct a'; eauto. + Qed. + + Next Obligation. + Proof. + eapply R_positionR. all: simpl. + - now destruct a. + - rewrite <- app_nil_r, stack_position_cons. + eapply positionR_poscat. constructor. + Qed. + + Next Obligation. + Proof. + now destruct a. + Qed. + + Next Obligation. + Proof. + now destruct a'; eauto. + Qed. + + Next Obligation. + Proof. + eapply R_positionR. all: simpl. + - now destruct a. + - rewrite <- app_nil_r, stack_position_cons. + eapply positionR_poscat. constructor. + Qed. + + Next Obligation. + Proof. + pose proof (hΣ _ H) as [hΣ]. + specialize (h1 Σ H). clear aux. + specialize (h2 Σ H). + apply welltyped_zipc_zipp in h1; eauto. + eapply welltyped_zipc_zipp in h2; eauto. + rewrite !zipp_as_mkApps in h1, h2 |- *. + pose proof (mkApps_Prim_nil _ _ _ _ _ h1); auto. + pose proof (mkApps_Prim_nil _ _ _ _ _ h2); auto. + rewrite H0 H1; cbn. + specialize (convdiscrty _ H) as [convty]. + specialize (convdiscrdef _ H) as [convdef]. + specialize (convdiscrval _ H) as [convval]. + cbn in convty, convdef. + symmetry in eql. unfold abstract_env_eq in eql. + eapply abstract_env_compare_universe_correct in eql; eauto. + 2:{ destruct h1 as [? ty]; eapply typing_wf_universes in ty; eauto. + move/andP: ty => []. rewrite H0 /=. cbn -[wf_universeb]. + rtoProp; intuition auto. + now move/wf_universe_reflect: H3. } + 2:{ destruct h2 as [? ty]; eapply typing_wf_universes in ty; eauto. + move/andP: ty => []. rewrite H1 /=. cbn -[wf_universeb]. + rtoProp; intuition auto. + now move/wf_universe_reflect: H3. } + constructor. eapply ws_cumul_pb_Prim; eauto; fvs. + constructor; eauto. + Qed. + + Next Obligation. + Proof. + destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. + pose proof (hΣ _ wfΣ) as [hΣ]. + specialize (h1 Σ wfΣ). clear aux. + specialize (h2 Σ wfΣ). + apply welltyped_zipc_zipp in h1; eauto. + eapply welltyped_zipc_zipp in h2; eauto. + rewrite !zipp_as_mkApps in H, h1, h2. + pose proof (mkApps_Prim_nil _ _ _ _ _ h1); auto. + pose proof (mkApps_Prim_nil _ _ _ _ _ h2); auto. + rewrite H0 H1 in H; cbn. + specialize (H _ wfΣ). + apply h. intros ? H2. + pose proof (abstract_env_ext_irr _ wfΣ H2). subst Σ0. + cbn in H. + destruct H as [H]; eapply invert_cumul_Prim in H; depelim H. now sq. + Qed. + + Next Obligation. + Proof. + destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. + pose proof (hΣ _ wfΣ) as [hΣ]. + specialize (h1 Σ wfΣ). clear aux. + specialize (h2 Σ wfΣ). + apply welltyped_zipc_zipp in h1; eauto. + eapply welltyped_zipc_zipp in h2; eauto. + rewrite !zipp_as_mkApps in H, h1, h2. + pose proof (mkApps_Prim_nil _ _ _ _ _ h1); auto. + pose proof (mkApps_Prim_nil _ _ _ _ _ h2); auto. + rewrite H0 H1 in H; cbn. + specialize (H _ wfΣ). + apply h. intros ? H2. + pose proof (abstract_env_ext_irr _ wfΣ H2). subst Σ0. + cbn in H. + destruct H as [H]; eapply invert_cumul_Prim in H; depelim H. now sq. + Qed. + + Next Obligation. + Proof. + destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. + pose proof (hΣ _ wfΣ) as [hΣ]. + specialize (h1 Σ wfΣ). clear aux. + specialize (h2 Σ wfΣ). + apply welltyped_zipc_zipp in h1; eauto. + eapply welltyped_zipc_zipp in h2; eauto. + rewrite !zipp_as_mkApps in H, h1, h2. + pose proof (mkApps_Prim_nil _ _ _ _ _ h1); auto. + pose proof (mkApps_Prim_nil _ _ _ _ _ h2); auto. + rewrite H0 H1 in H; cbn. + specialize (H _ wfΣ). + apply h. intros ? H2. + pose proof (abstract_env_ext_irr _ wfΣ H2). subst Σ0. + cbn in H. + destruct H as [H]; eapply invert_cumul_Prim in H; depelim H. now sq. + Qed. + + Next Obligation. + Proof. + destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. + pose proof (hΣ _ wfΣ) as [hΣ]. + specialize (h1 Σ wfΣ). clear aux. + specialize (h2 Σ wfΣ). + apply welltyped_zipc_zipp in h1; eauto. + eapply welltyped_zipc_zipp in h2; eauto. + rewrite !zipp_as_mkApps in H, h1, h2. + pose proof (mkApps_Prim_nil _ _ _ _ _ h1); auto. + pose proof (mkApps_Prim_nil _ _ _ _ _ h2); auto. + rewrite H0 H1 in H; cbn. + specialize (H _ wfΣ). + symmetry in neql. + destruct H as [H]; eapply invert_cumul_Prim in H; depelim H. + unshelve eapply (abstract_env_compare_universe_correct _ _ Conv) in e0; eauto. + - unfold abstract_env_eq in neql. now erewrite neql in e0. + - rewrite H0 in h1. destruct h1 as [? wt]. + eapply typing_wf_universes in wt; eauto. + move/andP: wt => []. cbn -[wf_universeb wf_universe]. rtoProp; intuition auto. + now move/wf_universe_reflect: H2. + - rewrite H1 in h2. destruct h2 as [? wt]. + eapply typing_wf_universes in wt; eauto. + move/andP: wt => []. cbn -[wf_universeb wf_universe]. rtoProp; intuition auto. + now move/wf_universe_reflect: H2. + Qed. + + Next Obligation. + Proof. + destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. + pose proof (hΣ _ wfΣ) as [hΣ]. + specialize (h1 Σ wfΣ). clear aux. + specialize (h2 Σ wfΣ). + apply welltyped_zipc_zipp in h1; eauto. + eapply welltyped_zipc_zipp in h2; eauto. + rewrite !zipp_as_mkApps in H, h1, h2. + pose proof (mkApps_Prim_nil _ _ _ _ _ h1); auto. + pose proof (mkApps_Prim_nil _ _ _ _ _ h2); auto. + rewrite H0 H1 in H; cbn. + specialize (H _ wfΣ). + rewrite H0 /= in h1; rewrite H1 /= in h2. + destruct H as [H]; eapply invert_cumul_Prim in H; depelim H. + all:cbn in tag_uneq; congruence. + Qed. + (* Fallback *) Next Obligation. unshelve eapply R_stateR. @@ -4745,7 +5106,7 @@ Qed. rewrite app_length in h. cbn in h. simpl. split. + rewrite mkApps_app in H. assumption. - + rewrite !stack_position_cons, !stack_position_appstack. + + rewrite !stack_position_cons !stack_position_appstack. rewrite <- !app_assoc. apply positionR_poscat. assert (h' : forall n m, positionR (repeat app_l n ++ [app_r]) (repeat app_l m)). { clear. intro n. induction n ; intro m. @@ -4854,8 +5215,7 @@ Qed. } ; | ccview_cofix mfix idx with inspect (unfold_cofix mfix idx) := { | @exist (Some (narg, fn)) eq2 with inspect (decompose_stack ρ) := { - | @exist (args, ξ) eq' := Some (tCase ci p (mkApps fn args) brs) - } ; + | @exist (args, ξ) eq' := Some (tCase ci p (mkApps fn args) brs) } ; | @exist None eq2 := False_rect _ _ } ; | ccview_other cred _ := None @@ -4882,6 +5242,7 @@ Qed. Qed. Next Obligation. + Proof. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. simpl_reduce_stack Σ wfΣ. destruct (h _ wfΣ) as (?&typ); auto. @@ -4891,8 +5252,7 @@ Qed. eapply PCUICValidity.inversion_mkApps in scrut_ty as (?&typ&?); auto. apply inversion_CoFix in typ as (?&?&?&?&?&?&?); auto. unfold unfold_cofix in eq2. - rewrite e in eq2. - congruence. + now rewrite e in eq2. Qed. Lemma unfold_one_case_cored : @@ -5293,8 +5653,11 @@ Qed. split; [constructor; eauto with pcuic|]. eauto with pcuic. - apply whnf_mkApps_tPrim_inv in wh as ->. - constructor; eexists _, []. - eauto using whnf_red with pcuic. + constructor; eexists (tPrim prim), []. + split. + * constructor. destruct prim as [? []]; constructor; auto. + eapply All2_same. intros; reflexivity. + * split; [constructor|]. eapply whnf_prim. - constructor; eexists _, (decompose_stack π).1. clear H. erewrite <- abstract_env_lookup_correct' in e; eauto. split; [econstructor|]; eauto. @@ -5625,7 +5988,7 @@ Qed. + intros. eapply iff_reflect. destruct leq. * eapply (abstract_env_compare_universe_correct _ H Conv) ; now eapply wf_universe_iff. * eapply (abstract_env_compare_universe_correct _ H Cumul) ; now eapply wf_universe_iff. - + intros. rewrite wf_universeb_instance_forall in *. + + intros. rewrite !wf_universeb_instance_forall in H0, H1. apply wf_universe_instance_iff in H0. apply wf_universe_instance_iff in H1. eapply compare_global_instance_correct; eauto. diff --git a/safechecker/theories/PCUICSafeReduce.v b/safechecker/theories/PCUICSafeReduce.v index a61cabcc5..8cb42f7c8 100644 --- a/safechecker/theories/PCUICSafeReduce.v +++ b/safechecker/theories/PCUICSafeReduce.v @@ -1533,7 +1533,7 @@ Corollary R_Acc_aux : destruct h as [T h]. apply inversion_App in h as (?&?&?&?&?); auto. apply inversion_Prim in t0 as (prim_ty & cdecl & [? ? ? ?]); auto. - eapply PCUICClassification.invert_cumul_prim_type_prod; eauto. + eapply PCUICConversion.invert_cumul_prim_type_prod; eauto. - unfold zipp. case_eq (decompose_stack π). intros l ρ e. constructor. constructor. eapply whne_mkApps. eapply whne_rel_nozeta. assumption. diff --git a/utils/theories/All_Forall.v b/utils/theories/All_Forall.v index 021ceed29..ad2966a7e 100644 --- a/utils/theories/All_Forall.v +++ b/utils/theories/All_Forall.v @@ -37,6 +37,10 @@ Derive Signature for All2. Derive NoConfusionHom for All2. #[global] Hint Constructors All2 : core. +Lemma All2_tip {A} {P} (t u : A) : P t u -> All2 P [t] [u]. +Proof. now repeat constructor. Qed. +#[global] Hint Resolve All2_tip : core. + Inductive All2_dep {A B : Type} {R : A -> B -> Type} (R' : forall a b, R a b -> Type) : forall {xs ys}, All2 R xs ys -> Type := | All2_dep_nil : All2_dep R' All2_nil | All2_dep_cons : forall (x : A) (y : B) (l : list A) (l' : list B) (r : R x y) (rs : All2 R l l'), @@ -461,6 +465,34 @@ Lemma All2_map_inv {A B C D} {R : C -> D -> Type} {f : A -> C} {g : B -> D} {l l All2 R (map f l) (map g l') -> All2 (fun x y => R (f x) (g y)) l l'. Proof. apply All2_map_equiv. Qed. +Lemma All2_tip_l {A B} {R : A -> B -> Type} x y : All2 R [x] y -> ∑ y', (y = [y']) * R x y'. +Proof. + intros a; depelim a. depelim a. exists y0; split => //. +Qed. + +Lemma All2i_All2i_mix {A B} {P Q : nat -> A -> B -> Type} + {n} {l : list A} {l' : list B} : + All2i P n l l' -> All2i Q n l l' -> All2i (fun i x y => (P i x y * Q i x y)%type) n l l'. +Proof. + induction 2; simpl; intros; constructor. + inv X; intuition auto. + apply IHX0. inv X; intuition auto. +Qed. + +Lemma All2i_nth_error {A B} {P : nat -> A -> B -> Type} {l l' n x c k} : + All2i P k l l' -> + nth_error l n = Some x -> + nth_error l' n = Some c -> + P (k + n)%nat x c. +Proof. + induction 1 in n |- *. + * rewrite !nth_error_nil => //. + * destruct n. + + simpl. intros [= <-] [= <-]. now rewrite Nat.add_0_r. + + simpl. intros hnth hnth'. specialize (IHX _ hnth hnth'). + now rewrite Nat.add_succ_r. +Qed. + (* Lemma All2_List_Forall_mix_left {A : Type} {P : A -> Prop} {Q : A -> A -> Prop} *) (* {l l' : list A} : *) (* Forall P l -> All2 Q l l' -> All2 (fun x y => P x /\ Q x y) l l'. *) @@ -524,6 +556,14 @@ Proof. apply IHX0. inv X; intuition auto. Qed. +Lemma All2_All2_All2 {A B C} {P Q R} (l : list A) (l' : list B) (l'' : list C) : + All2 P l l' -> All2 Q l' l'' -> + (forall x y z, P x y -> Q y z -> R x z) -> + All2 R l l''. +Proof. + induction 1 in l'' |- *; intros a; depelim a; constructor; eauto. +Qed. + Lemma Forall_All {A : Type} (P : A -> Prop) l : Forall P l -> All P l. Proof. @@ -1550,6 +1590,10 @@ Lemma All2_map_left {A B C} {P : A -> C -> Type} {l l'} {f : B -> A} : All2 (fun x y => P (f x) y) l l' -> All2 P (map f l) l'. Proof. apply All2_map_left_equiv. Qed. +Lemma All2_map_left' {A B C} (P : A -> B -> Type) l l' (f : C -> A) : + All2 P (map f l) l' -> All2 (fun x y => P (f x) y) l l'. +Proof. intros. rewrite - (map_id l') in X. eapply All2_map_inv; eauto. Qed. + Lemma All2_map_right {A B C} {P : A -> C -> Type} {l l'} {f : B -> C} : All2 (fun x y => P x (f y)) l l' -> All2 P l (map f l'). Proof. apply All2_map_right_equiv. Qed. @@ -2072,6 +2116,11 @@ Proof. + eapply IHl. inversion h. assumption. Qed. +Lemma Forall2_tip_l {A B} {R : A -> B -> Prop} x y : Forall2 R [x] y -> exists y', (y = [y']) /\ R x y'. +Proof. + intros a; depelim a. depelim a. exists y0; split => //. +Qed. + Lemma All2_length {A B} {P : A -> B -> Type} {l l'} : All2 P l l' -> #|l| = #|l'|. Proof. induction 1; simpl; auto. Qed. @@ -3311,6 +3360,12 @@ Inductive All_fold {A} (P : list A -> A -> Type) : list A -> Type := | All_fold_cons {d Γ} : All_fold P Γ -> P Γ d -> All_fold P (d :: Γ). Derive Signature NoConfusionHom for All_fold. + +Lemma All_fold_tip {A : Type} (P : list A -> A -> Type) {x} : All_fold P [x] -> P [] x. +Proof. + intros a; now depelim a. +Qed. + Inductive All2_fold {A} (P : list A -> list A -> A -> A -> Type) : list A -> list A -> Type := | All2_fold_nil : All2_fold P nil nil From 51cb57dd479aa8dd5c627c9e79b043765cba6532 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Tue, 21 Nov 2023 17:00:33 +0100 Subject: [PATCH 09/17] Typechecking of arrays --- safechecker/theories/PCUICTypeChecker.v | 207 +++++++++++++++++------- 1 file changed, 150 insertions(+), 57 deletions(-) diff --git a/safechecker/theories/PCUICTypeChecker.v b/safechecker/theories/PCUICTypeChecker.v index f754a93e6..749dab06d 100644 --- a/safechecker/theories/PCUICTypeChecker.v +++ b/safechecker/theories/PCUICTypeChecker.v @@ -1310,6 +1310,126 @@ Section Typecheck. End check_mfix. + 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_primitive_invariants (p : prim_tag) (decl : constant_body) : typing_result_comp (primitive_invariants p decl) := + | primInt | decl := + check_eq_true (eqb decl.(cst_body) None) (Msg "primitive type is registered to a defined constant") ;; + check_eq_true (eqb decl.(cst_universes) Monomorphic_ctx) (Msg "primitive type is registered to a polymorphic constant") ;; + check_eq_true (isSort decl.(cst_type)) (Msg "primitive type is registered to an axiom whose type is not a sort") ;; + ret _ + | primFloat | decl := + check_eq_true (eqb decl.(cst_body) None) (Msg "primitive type is registered to a defined constant") ;; + check_eq_true (eqb decl.(cst_universes) Monomorphic_ctx) (Msg "primitive type is registered to a polymorphic constant") ;; + check_eq_true (isSort decl.(cst_type)) (Msg "primitive type is registered to an axiom whose type is not a sort") ;; + ret _ + | primArray | decl := + let s := Universe.make (Level.lvar 0) in + check_eq_true (eqb decl.(cst_body) None) (Msg "primitive type is registered to a defined constant") ;; + check_eq_true (eqb decl.(cst_universes) (Polymorphic_ctx array_uctx)) (Msg "primitive type is registered to a polymorphic constant") ;; + check_eq_true (eqb decl.(cst_type) (tImpl (tSort s) (tSort s))) (Msg "primitive type for arrays is registered to an axiom whose type is not of shape Type -> Type") ;; + ret _. + Proof. + all:try eapply eqb_eq in i. + all:try apply eqb_eq in i0. + - destruct (cst_type decl) => //. now exists u. + - destruct X1 as []. rewrite H in absurd; eauto. + - destruct X1 as []. apply absurd. now rewrite H1; apply eqb_refl. + - destruct X1 as []. apply absurd; rewrite H0; apply eqb_refl. + - destruct (cst_type decl) => //. now exists u. + - destruct X1 as []. rewrite H in absurd; eauto. + - destruct X1 as []. apply absurd. now rewrite H1; apply eqb_refl. + - destruct X1 as []. apply absurd; rewrite H0; apply eqb_refl. + - destruct (cst_type decl) => //. split => //. now apply eqb_eq in i1. + - destruct H as []. rewrite H in absurd; eauto. + - destruct H as []. apply absurd. now rewrite H1; apply eqb_refl. + - destruct H as []. apply absurd; rewrite H0; apply eqb_refl. + Qed. + + Section make_All. + Context {A} {B : A -> Type} (f : forall x : A, typing_result_comp (B x)). + + Equations? make_All (l : list A) : typing_result_comp (All B l) := + | [] := ret (All_nil) + | hd :: tl := + hd' <- f hd ;; + tl' <- make_All tl ;; + ret (All_cons hd' tl'). + Proof. + all:depelim X0; eauto. + Qed. + + End make_All. + + Section check_primitive. + 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 Obligation Tactic := idtac. + Equations? check_primitive_typing (p : prim_val) : typing_result_comp (forall Σ (wfΣ : abstract_env_ext_rel X Σ), ∥ primitive_typing_hyps checking Σ Γ p ∥) := + | (primInt; primIntModel i) := ret _ + | (primFloat; primFloatModel f) := ret _ + | (primArray; primArrayModel a) := + check_eq_true (abstract_env_ext_wf_universeb X (Universe.make a.(array_level))) (Msg "primitive array level is not well-formed") ;; + check_type <- bdcheck infer Γ wfΓ a.(array_type) (tSort (Universe.make a.(array_level))) _ ;; + check_default <- bdcheck infer Γ wfΓ a.(array_default) a.(array_type) _ ;; + check_values <- make_All (fun x => bdcheck infer Γ wfΓ x a.(array_type) _) a.(array_value) ;; + ret _. + Proof. + all:intros. + all:try pose proof (abstract_env_ext_wf _ wfΣ) as [wfΣ']. + all:try specialize (wfΓ _ wfΣ). + 1-2:do 2 constructor. + - eauto. + - sq. erewrite <- abstract_env_ext_wf_universeb_correct in i; tea. + eexists; eapply type_Sort; eauto. + now move/@wf_universe_reflect: i. + - specialize (check_type _ wfΣ) as []. + sq; eapply checking_typing in X0; eauto. now eexists. + erewrite <- abstract_env_ext_wf_universeb_correct in i; tea. + eexists; eapply type_Sort; eauto. + now move/@wf_universe_reflect: i. + - specialize (check_type _ wfΣ) as []. + sq; eapply checking_typing in X0; eauto. now eexists. + erewrite <- abstract_env_ext_wf_universeb_correct in i; tea. + eexists; eapply type_Sort; eauto. + now move/@wf_universe_reflect: i. + - specialize (check_type _ wfΣ) as []. + specialize (check_default _ wfΣ) as []. + assert (∥ Σ;;; Γ |- array_type a : tSort (Universe.make (array_level a)) ∥) as []. + { sq. eapply checking_typing in X0; eauto. + erewrite <- abstract_env_ext_wf_universeb_correct in i; tea. + eexists; eapply type_Sort; eauto. now move/@wf_universe_reflect: i. } + assert (∥ All (fun x : term => Σ;;; Γ |- x ◃ array_type a) (array_value a) ∥). + { induction check_values. + - repeat constructor. + - destruct IHcheck_values. specialize (p _ wfΣ) as []. + destruct wfΓ as []. + constructor. constructor; eauto. } + sq; constructor; eauto. + erewrite <- abstract_env_ext_wf_universeb_correct in i; tea. + now move/@wf_universe_reflect: i. + - destruct (abstract_env_ext_exists X) as [[Σ hΣ]]. + specialize (H _ hΣ) as [tyh]. + depelim tyh. eapply absurd. solve_all. + pose proof (abstract_env_ext_irr _ wfΣ hΣ). now subst Σ0. + - destruct (abstract_env_ext_exists X) as [[Σ hΣ]]. + specialize (H _ hΣ) as [tyh]. + specialize (check_type _ hΣ) as []. + depelim tyh. eapply absurd. intros. + pose proof (abstract_env_ext_irr _ wfΣ hΣ). now subst Σ0. + - destruct (abstract_env_ext_exists X) as [[Σ hΣ]]. + specialize (H _ hΣ) as [tyh]. + depelim tyh. eapply absurd. intros. + pose proof (abstract_env_ext_irr _ wfΣ hΣ). now subst Σ0. + - destruct (abstract_env_ext_exists X) as [[Σ hΣ]]. + specialize (H _ hΣ) as [tyh]. + erewrite <- abstract_env_ext_wf_universeb_correct in absurd; tea. + eapply absurd. depelim tyh. + now move/wf_universe_reflect: wfl. + Qed. + End check_primitive. Lemma chop_firstn_skipn {A} n (l : list A) : chop n l = (firstn n l, skipn n l). Proof using Type. @@ -1323,9 +1443,6 @@ Section Typecheck. now exists t. 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 (Γ : 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 ∥ }) by struct t := @@ -1472,10 +1589,9 @@ Section Typecheck. | exist (Some prim_ty) eqp with inspect (abstract_env_lookup X prim_ty) := { | exist (Some (ConstantDecl d)) HH => let ty := d.(cst_type) in - check_eq_true (eqb d.(cst_body) None) (Msg "primitive type is registered to a defined constant") ;; - check_eq_true (eqb d.(cst_universes) Monomorphic_ctx) (Msg "primitive type is registered to a polymorphic constant") ;; - check_eq_true (isSort d.(cst_type)) (Msg "primitive type is registered to an axiom whose type is not a sort") ;; - ret (tConst prim_ty []; _) + inv <- check_primitive_invariants (prim_val_tag p) d ;; + ty <- check_primitive_typing infer Γ HΓ p ;; + ret (prim_type p prim_ty; _) | _ => raise (UndeclaredConstant prim_ty) } }. @@ -2561,84 +2677,61 @@ Section Typecheck. congruence. Qed. + (* Primitives *) Next Obligation. - eapply eqb_eq in i. eapply eqb_eq in i0. + specialize (ty _ wfΣ) as []. sq. econstructor; eauto. + erewrite abstract_primitive_constant_correct in eqp; eauto. rewrite -(abstract_env_lookup_correct' _ (Σ := Σ)) // in HH. - split. econstructor. rewrite eqp. - erewrite abstract_primitive_constant_correct; try reflexivity. eassumption. red. apply declared_constant_from_gen. unfold declared_constant_gen. now rewrite -HH. - destruct (cst_type d) eqn:hty => //. - exists u. split => //. - Qed. - Next Obligation. - destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. - cbn in *. specialize_Σ wfΣ ; sq. - depelim X1. - eapply eqb_eq in i. eapply eqb_eq in i0. - rewrite -(abstract_env_lookup_correct' _ (Σ := Σ)) // in HH. - rewrite (abstract_primitive_constant_correct _ _ _ wfΣ) in eqp. - rewrite e1 in eqp. noconf eqp. - symmetry in HH. - destruct (hΣ _ wfΣ). unshelve eapply declared_constant_to_gen in d0; eauto. - rewrite d0 in HH; noconf HH. - destruct p1 as [s' []]. rewrite H in absurd. now apply absurd. Qed. Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. - cbn in *. specialize_Σ wfΣ ; sq. - depelim X1. - eapply eqb_eq in i. - rewrite -(abstract_env_lookup_correct' _ (Σ := Σ)) // in HH. - rewrite (abstract_primitive_constant_correct _ _ _ wfΣ) in eqp. - rewrite e1 in eqp. noconf eqp. - destruct (hΣ _ wfΣ). unshelve eapply declared_constant_to_gen in d0; eauto. - symmetry in HH. rewrite d0 in HH; noconf HH. - destruct p1 as [s' []]. apply absurd. case: eqb_spec => //. + cbn in *. specialize_Σ wfΣ. + eapply absurd. intros ? h. + pose proof (abstract_env_ext_irr _ wfΣ h). subst Σ0. sq. + now depelim X1. Qed. - - Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. - cbn in *. specialize_Σ wfΣ ; sq. - depelim X1. + pose proof (heΣ _ wfΣ) as [heΣ]. + cbn in *. specialize_Σ wfΣ. destruct X1 as []. + depelim X1. eapply absurd; eauto. rewrite -(abstract_env_lookup_correct' _ (Σ := Σ)) // in HH. - rewrite (abstract_primitive_constant_correct _ _ _ wfΣ) in eqp. + erewrite abstract_primitive_constant_correct in eqp; eauto. rewrite e1 in eqp. noconf eqp. - destruct (hΣ _ wfΣ). unshelve eapply declared_constant_to_gen in d0; eauto. - symmetry in HH. rewrite d0 in HH; noconf HH. - destruct p1 as [s' []]. apply absurd. case: eqb_spec => //. + unshelve eapply declared_constant_to_gen in d0; eauto. 3:eapply heΣ0. + unfold declared_constant_gen in d0. rewrite d0 in HH. now noconf HH. Qed. - Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. - cbn in *. specialize_Σ wfΣ ; sq. + pose proof (heΣ _ wfΣ) as [heΣ]. + cbn in *. specialize_Σ wfΣ. destruct X1 as []. depelim X1. rewrite -(abstract_env_lookup_correct' _ (Σ := Σ)) // in e0. - rewrite (abstract_primitive_constant_correct _ _ _ wfΣ) in eqp. + erewrite abstract_primitive_constant_correct in eqp; eauto. rewrite e1 in eqp. noconf eqp. - destruct (hΣ _ wfΣ). unshelve eapply declared_constant_to_gen in d; eauto. - symmetry in e0. rewrite d in e0; noconf e0. + unshelve eapply declared_constant_to_gen in d; eauto. 3:eapply heΣ0. + unfold declared_constant_gen in d. rewrite d in e0. now noconf e0. Qed. - Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. - cbn in *. specialize_Σ wfΣ ; sq. + pose proof (heΣ _ wfΣ) as [heΣ]. + cbn in *. specialize_Σ wfΣ. destruct X1 as []. depelim X1. rewrite -(abstract_env_lookup_correct' _ (Σ := Σ)) // in e0. - rewrite (abstract_primitive_constant_correct _ _ _ wfΣ) in eqp. + erewrite abstract_primitive_constant_correct in eqp; eauto. rewrite e1 in eqp. noconf eqp. - destruct (hΣ _ wfΣ). unshelve eapply declared_constant_to_gen in d; eauto. - symmetry in e0. rewrite /declared_constant_gen in d. - rewrite e0 in d; noconf d. + unshelve eapply declared_constant_to_gen in d; eauto. 3:eapply heΣ0. + unfold declared_constant_gen in d. rewrite d in e0. now noconf e0. Qed. - Next Obligation. destruct (abstract_env_ext_exists X) as [[Σ wfΣ]]. - cbn in *. specialize_Σ wfΣ ; sq. + pose proof (heΣ _ wfΣ) as [heΣ]. + cbn in *. specialize_Σ wfΣ. destruct X1 as []. depelim X1. - rewrite (abstract_primitive_constant_correct _ _ _ wfΣ) in e0. - unfold prim_val_tag in e1. congruence. + erewrite abstract_primitive_constant_correct in e0; eauto. + rewrite e1 in e0. noconf e0. Qed. Definition check_isType := infer_isType infer. From 69c1f0583615141690186478da6d1dd04e3301af Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Fri, 24 Nov 2023 16:34:36 +0100 Subject: [PATCH 10/17] In erasure, must update solve_all to automate proofs --- erasure/_CoqProject.in | 1 + erasure/theories/EAst.v | 2 +- erasure/theories/EAstUtils.v | 2 +- erasure/theories/EInduction.v | 37 ++++++- erasure/theories/ELiftSubst.v | 8 +- erasure/theories/EPretty.v | 7 +- erasure/theories/EPrimitive.v | 159 ++++++++++++++++++++++++++++ erasure/theories/EReflect.v | 13 ++- erasure/theories/ESubstitution.v | 10 ++ erasure/theories/Extract.v | 51 ++++++--- erasure/theories/Typed/Certifying.v | 1 + 11 files changed, 262 insertions(+), 29 deletions(-) create mode 100644 erasure/theories/EPrimitive.v diff --git a/erasure/_CoqProject.in b/erasure/_CoqProject.in index 5397184fa..a4a94c374 100644 --- a/erasure/_CoqProject.in +++ b/erasure/_CoqProject.in @@ -1,5 +1,6 @@ -R theories MetaCoq.Erasure +theories/EPrimitive.v theories/EAst.v theories/EAstUtils.v theories/EInduction.v diff --git a/erasure/theories/EAst.v b/erasure/theories/EAst.v index 164dd8e3c..cd1eec17f 100644 --- a/erasure/theories/EAst.v +++ b/erasure/theories/EAst.v @@ -1,7 +1,7 @@ (* Distributed under the terms of the MIT license. *) From MetaCoq.Utils Require Import utils. From MetaCoq.Common Require Import BasicAst Universes. -From MetaCoq.PCUIC Require Import PCUICPrimitive. +From MetaCoq.Erasure Require Import EPrimitive. (** * Extracted terms These are the terms produced by extraction: compared to kernel terms, diff --git a/erasure/theories/EAstUtils.v b/erasure/theories/EAstUtils.v index 58c38f7a7..ec168e08e 100644 --- a/erasure/theories/EAstUtils.v +++ b/erasure/theories/EAstUtils.v @@ -373,7 +373,7 @@ Fixpoint string_of_term (t : term) : string := ^ string_of_term c ^ ")" | tFix l n => "Fix(" ^ (string_of_list (string_of_def string_of_term) l) ^ "," ^ string_of_nat n ^ ")" | tCoFix l n => "CoFix(" ^ (string_of_list (string_of_def string_of_term) l) ^ "," ^ string_of_nat n ^ ")" - | tPrim p => "Prim(" ^ PCUICPrimitive.string_of_prim string_of_term p ^ ")" + | tPrim p => "Prim(" ^ EPrimitive.string_of_prim string_of_term p ^ ")" end. (** Compute all the global environment dependencies of the term *) diff --git a/erasure/theories/EInduction.v b/erasure/theories/EInduction.v index 87d535a5e..093b0e56d 100644 --- a/erasure/theories/EInduction.v +++ b/erasure/theories/EInduction.v @@ -2,8 +2,8 @@ Require Import List ssreflect ssrbool. From MetaCoq.Utils Require Import utils. From MetaCoq.Common Require Import BasicAst. -From MetaCoq.Erasure Require Import EAst EAstUtils. From MetaCoq.PCUIC Require Import PCUICSize. +From MetaCoq.Erasure Require Import EPrimitive EAst EAstUtils. From Equations Require Import Equations. Set Equations Transparent. @@ -15,6 +15,13 @@ Set Equations Transparent. (** Custom induction principle on syntax, dealing with the various lists appearing in terms. *) +Definition primProp P (p : prim_val term) : Type := + match p.π2 with + | primIntModel _ => unit + | primFloatModel _ => unit + | primArrayModel a => P a.(array_default) × All P a.(array_value) + end. + Lemma term_forall_list_ind : forall P : term -> Type, (P tBox) -> @@ -34,7 +41,7 @@ Lemma term_forall_list_ind : (forall (s : projection) (t : term), P t -> P (tProj s t)) -> (forall (m : mfixpoint term) (n : nat), All (fun x => P (dbody x)) m -> P (tFix m n)) -> (forall (m : mfixpoint term) (n : nat), All (fun x => P (dbody x)) m -> P (tCoFix m n)) -> - (forall p, P (tPrim p)) -> + (forall p, primProp P p -> P (tPrim p)) -> forall t : term, P t. Proof. intros until t. revert t. @@ -66,6 +73,12 @@ Proof. fix auxm 1. destruct mfix; constructor; [|apply auxm]. apply auxt. + + destruct p as [? []]; cbn => //. + destruct a as [def v]; cbn. + split. eapply auxt. + revert v; fix auxv 1. + destruct v; constructor; [apply auxt|apply auxv]. Defined. Ltac applyhyp := @@ -89,7 +102,14 @@ Ltac inv H := | @ahyp _ ?X => inversion_clear X end. -Fixpoint size t : nat := +Definition prim_size (f : term -> nat) (p : prim_val term) : nat := + match p.π2 with + | primIntModel _ => 0 + | primFloatModel _ => 0 + | primArrayModel a => f a.(array_default) + list_size f a.(array_value) + end. + +Fixpoint size (t : term) : nat := match t with | tRel i => 1 | tEvar ev args => S (list_size size args) @@ -101,6 +121,7 @@ Fixpoint size t : nat := | tFix mfix idx => S (list_size (fun x => size (dbody x)) mfix) | tCoFix mfix idx => S (list_size (fun x => size (dbody x)) mfix) | tConstruct _ _ ignore_args => S (list_size size ignore_args) + | tPrim p => S (prim_size size p) | _ => 1 end. @@ -217,7 +238,7 @@ Section MkApps_rec. (pproj : forall (s : projection) (t : term), P t -> P (tProj s t)) (pfix : forall (m : mfixpoint term) (n : nat), All (fun x => P (dbody x)) m -> P (tFix m n)) (pcofix : forall (m : mfixpoint term) (n : nat), All (fun x => P (dbody x)) m -> P (tCoFix m n)) - (pprim : forall p, P (tPrim p)). + (pprim : forall p, primProp P p -> P (tPrim p)). Definition inspect {A} (x : A) : { y : A | x = y } := exist _ x eq_refl. @@ -243,7 +264,7 @@ Section MkApps_rec. | tProj p c => pproj p c (rec c) | tFix mfix idx => pfix mfix idx (All_rec P dbody mfix (fun x H => rec x)) | tCoFix mfix idx => pcofix mfix idx (All_rec P dbody mfix (fun x H => rec x)) - | tPrim p => pprim p. + | tPrim p => pprim p _. Proof. all:unfold MR; cbn; auto with arith. 4:lia. - clear -napp nonnil da rec. @@ -255,6 +276,12 @@ Section MkApps_rec. rewrite da in H0. cbn in H0. rewrite <- H0. unfold id in H. change (fun x => size x) with size in H. abstract lia. - clear -da. abstract (eapply decompose_app_inv in da; now symmetry). + - destruct p as [? []]; cbn => //. + destruct a as [def v]; cbn. + split. + * eapply rec; red; cbn. lia. + * refine (All_rec P id v (fun x H => rec x _)); red; cbn. + unfold id in H. change (fun x => size x) with size in H. lia. Qed. End MkApps_rec. diff --git a/erasure/theories/ELiftSubst.v b/erasure/theories/ELiftSubst.v index f37390211..d78515df6 100644 --- a/erasure/theories/ELiftSubst.v +++ b/erasure/theories/ELiftSubst.v @@ -1,7 +1,7 @@ (* Distributed under the terms of the MIT license. *) From MetaCoq.Utils Require Import utils. From MetaCoq.Common Require Import BasicAst. -From MetaCoq.Erasure Require Import EAst EAstUtils EInduction. +From MetaCoq.Erasure Require Import EPrimitive EAst EAstUtils EInduction. Require Import ssreflect. (** * Lifting and substitution for the AST @@ -36,7 +36,7 @@ Fixpoint lift n k t : term := | tVar _ => t | tConst _ => t | tConstruct ind i args => tConstruct ind i (map (lift n k) args) - | tPrim _ => t + | tPrim p => tPrim (map_prim (lift n k) p) end. Notation lift0 n := (lift n 0). @@ -71,6 +71,7 @@ Fixpoint subst s k u := let mfix' := List.map (map_def (subst s k')) mfix in tCoFix mfix' idx | tConstruct ind i args => tConstruct ind i (map (subst s k) args) + | tPrim p => tPrim (map_prim (subst s k) p) | x => x end. @@ -98,6 +99,7 @@ Fixpoint closedn k (t : term) : bool := let k' := List.length mfix + k in List.forallb (test_def (closedn k')) mfix | tConstruct ind i args => forallb (closedn k) args + | tPrim p => test_prim (closedn k) p | _ => true end. @@ -230,6 +232,8 @@ Proof. - now elim (leb k n). - destruct x; cbn. now rewrite H0. + - destruct p as [? []]; cbn in X |- *; intuition eauto. do 2 f_equal. + rewrite /map_array_model; destruct a; cbn => //=. f_equal; eauto. solve_all. Qed. Lemma lift0_p : forall M, lift0 0 M = M. diff --git a/erasure/theories/EPretty.v b/erasure/theories/EPretty.v index 3ce0b3a7c..9afba35b0 100644 --- a/erasure/theories/EPretty.v +++ b/erasure/theories/EPretty.v @@ -2,8 +2,7 @@ From Coq Require Import Program. From MetaCoq.Utils Require Import utils. From MetaCoq.Common Require Import BasicAst. -From MetaCoq.Erasure Require Import EAst EAstUtils EGlobalEnv. -From MetaCoq.PCUIC Require Import PCUICPrimitive. +From MetaCoq.Erasure Require Import EPrimitive EAst EAstUtils EGlobalEnv. (** * Pretty printing *) @@ -85,11 +84,11 @@ Module PrintTermTree. let ctx' := List.map (fun d => {| decl_name := dname d; decl_body := None |}) defs in print_list (print_def (print_term (ctx' ++ Γ)%list true false)) (nl ^ " with ") defs. - Definition print_prim {term} (soft : term -> Tree.t) (p : @prim_val EAst.term) : Tree.t := + Definition print_prim (soft : EAst.term -> Tree.t) (p : @prim_val EAst.term) : Tree.t := match p.π2 return Tree.t with | primIntModel f => "(int: " ^ Primitive.string_of_prim_int f ^ ")" | primFloatModel f => "(float: " ^ Primitive.string_of_float f ^ ")" - (* | primArrayModel a => "(array:" ^ ")" *) + | primArrayModel a => "(array:" ^ soft a.(array_default) ^ " , " ^ string_of_list soft a.(array_value) ^ ")" end. Fixpoint print_term (Γ : context) (top : bool) (inapp : bool) (t : term) {struct t} : Tree.t := diff --git a/erasure/theories/EPrimitive.v b/erasure/theories/EPrimitive.v new file mode 100644 index 000000000..a4715ffe4 --- /dev/null +++ b/erasure/theories/EPrimitive.v @@ -0,0 +1,159 @@ +(* Distributed under the terms of the MIT license. *) +From MetaCoq.Utils Require Import utils. +From MetaCoq.Common Require Import Universes Primitive Reflect + Environment EnvironmentTyping. +(* From MetaCoq.Erasure Require Import BasicAst. *) +From Equations Require Import Equations. +From Coq Require Import ssreflect. +From Coq Require Import Uint63 SpecFloat. + +Implicit Type term : Set. + +Record array_model {term : Set} : Set := + { array_default : term; + array_value : list term }. +Derive NoConfusion for array_model. + +Arguments array_model : clear implicits. +#[global] +Instance array_model_eqdec {term} (e : EqDec term) : EqDec (array_model term). +Proof. eqdec_proof. Qed. + +(* We use this inductive definition instead of the more direct case analysis + [prim_model_of] so that [prim_model] can be used in the inductive definition + of terms, otherwise it results in a non-strictly positive definition. + *) +Inductive prim_model (term : Set) : prim_tag -> Set := +| primIntModel (i : PrimInt63.int) : prim_model term primInt +| primFloatModel (f : PrimFloat.float) : prim_model term primFloat +| primArrayModel (a : array_model term) : prim_model term primArray. + +Arguments primIntModel {term}. +Arguments primFloatModel {term}. +Arguments primArrayModel {term}. + +Derive Signature NoConfusion for prim_model. + +Definition prim_model_of (term : Set) (p : prim_tag) : Set := + match p with + | primInt => PrimInt63.int + | primFloat => PrimFloat.float + | primArray => array_model term + end. + +Definition prim_val term : Set := ∑ t : prim_tag, prim_model term t. +Definition prim_val_tag {term} (s : prim_val term) := s.π1. +Definition prim_val_model {term} (s : prim_val term) : prim_model term (prim_val_tag s) := s.π2. + +Definition prim_model_val {term} (p : prim_val term) : prim_model_of term (prim_val_tag p) := + match prim_val_model p in prim_model _ t return prim_model_of term t with + | primIntModel i => i + | primFloatModel f => f + | primArrayModel a => a + end. + +Lemma exist_irrel_eq {A} (P : A -> bool) (x y : sig P) : + proj1_sig x = proj1_sig y -> x = y. +Proof. + destruct x as [x p], y as [y q]; simpl; intros ->. + now destruct (uip p q). +Qed. + +#[global] +Instance reflect_eq_Z : ReflectEq Z := EqDec_ReflectEq _. + +Import ReflectEq. + +Local Obligation Tactic := idtac. + +Definition eqb_array {term} {eqt : term -> term -> bool} (x y : array_model term) := + forallb2 eqt x.(array_value) y.(array_value) && + eqt x.(array_default) y.(array_default). + +#[program,global] +Instance reflect_eq_array {term} {req : ReflectEq term}: ReflectEq (array_model term) := + { eqb := eqb_array (eqt := eqb) }. +Next Obligation. + intros term req [] []; cbn. unfold eqb_array. cbn. + change (forallb2 eqb) with (eqb (A := list term)). + case: eqb_spec => //=. intros ->. + case: eqb_spec => //=. intros ->. constructor; auto. + all:constructor; congruence. +Qed. + +Equations eqb_prim_model {term} {req : term -> term -> bool} {t : prim_tag} (x y : prim_model term t) : bool := + | primIntModel x, primIntModel y := ReflectEq.eqb x y + | primFloatModel x, primFloatModel y := ReflectEq.eqb x y + | primArrayModel x, primArrayModel y := eqb_array (eqt:=req) x y. + +#[global, program] +Instance prim_model_reflecteq {term} {req : ReflectEq term} {p : prim_tag} : ReflectEq (prim_model term p) := + {| ReflectEq.eqb := eqb_prim_model (req := eqb) |}. +Next Obligation. + intros. depelim x; depelim y; simp eqb_prim_model. + case: ReflectEq.eqb_spec; constructor; subst; auto. congruence. + case: ReflectEq.eqb_spec; constructor; subst; auto. congruence. + change (eqb_array a a0) with (eqb a a0). + case: ReflectEq.eqb_spec; constructor; subst; auto. congruence. +Qed. + +#[global] +Instance prim_model_eqdec {term} {req : ReflectEq term} : forall p : prim_tag, EqDec (prim_model term p) := _. + +Equations eqb_prim_val {term} {req : term -> term -> bool} (x y : prim_val term) : bool := + | (primInt; i), (primInt; i') := eqb_prim_model (req := req) i i' + | (primFloat; f), (primFloat; f') := eqb_prim_model (req := req) f f' + | (primArray; a), (primArray; a') := eqb_prim_model (req := req) a a' + | x, y := false. + +#[global, program] +Instance prim_val_reflect_eq {term} {req : ReflectEq term} : ReflectEq (prim_val term) := + {| ReflectEq.eqb := eqb_prim_val (req := eqb) |}. +Next Obligation. + intros. funelim (eqb_prim_val x y); simp eqb_prim_val. + change (eqb_prim_model i i') with (eqb i i'). + case: ReflectEq.eqb_spec; constructor; subst; auto. intros H; noconf H. cbn in n. auto. + constructor. intros H; noconf H; auto. + constructor. intros H; noconf H; auto. + constructor. intros H; noconf H; auto. + change (eqb_prim_model f f') with (eqb f f'). + case: ReflectEq.eqb_spec; constructor; subst; auto. intros H; noconf H; auto. + constructor. intros H; noconf H; auto. + constructor. intros H; noconf H; auto. + constructor. intros H; noconf H; auto. + change (eqb_prim_model a a') with (eqb a a'). + case: ReflectEq.eqb_spec; constructor; subst; auto. intros H; noconf H; auto. +Qed. + +#[global] +Instance prim_tag_model_eqdec term {req : ReflectEq term} : EqDec (prim_val term) := _. + +(** Printing *) + +Definition string_of_prim {term} (soft : term -> string) (p : prim_val term) : string := + match p.π2 return string with + | primIntModel f => "(int: " ^ string_of_prim_int f ^ ")" + | primFloatModel f => "(float: " ^ string_of_float f ^ ")" + | primArrayModel a => "(array:" ^ soft a.(array_default) ^ " , " ^ string_of_list soft a.(array_value) ^ ")" + end. + +Definition map_array_model {term term'} (f : term -> term') (a : array_model term) : array_model term' := + {| array_default := f a.(array_default); + array_value := map f a.(array_value) |}. + +Definition map_prim {term term'} (f : term -> term') (p : prim_val term) : prim_val term' := + match p.π2 return prim_val term' with + | primIntModel f => (primInt; primIntModel f) + | primFloatModel f => (primFloat; primFloatModel f) + | primArrayModel a => (primArray; primArrayModel (map_array_model f a)) + end. + +Definition test_array_model {term} (f : term -> bool) (a : array_model term) : bool := + f a.(array_default) && forallb f a.(array_value). + +Definition test_prim {term} (f : term -> bool) (p : prim_val term) : bool := + match p.π2 return bool with + | primIntModel f => true + | primFloatModel f => true + | primArrayModel a => test_array_model f a + end. diff --git a/erasure/theories/EReflect.v b/erasure/theories/EReflect.v index 69f9ebaba..58f99e2f4 100644 --- a/erasure/theories/EReflect.v +++ b/erasure/theories/EReflect.v @@ -1,7 +1,7 @@ From Coq Require Import ssreflect ssrbool. From MetaCoq.Utils Require Import utils. From MetaCoq.Common Require Import BasicAst Reflect. -From MetaCoq.Erasure Require Import EAst EInduction. +From MetaCoq.Erasure Require Import EPrimitive EAst EInduction. From Equations Require Import Equations. Local Ltac finish := @@ -112,8 +112,15 @@ Proof. subst. inversion e0. subst. destruct (eq_dec rarg rarg0) ; nodec. subst. left. reflexivity. - - destruct (eq_dec p prim); nodec. - left; subst. reflexivity. + - destruct p as [? []]; destruct p0 as [? []]; cbn ; nodec. + + destruct (eq_dec i i0); nodec; subst; eauto. + + destruct (eq_dec f f0); nodec; subst; eauto. + + destruct a as [? ?], a0 as [? ?]; cbn. + destruct X as [hd hv]. cbn in *. + destruct (hd array_default0); nodec; subst; eauto. + revert array_value0; induction hv; intros []; eauto; nodec. + destruct (p t); subst; nodec. + destruct (IHhv l0); nodec. noconf e; eauto. Defined. #[global] diff --git a/erasure/theories/ESubstitution.v b/erasure/theories/ESubstitution.v index 5ff50816b..af58c8a04 100644 --- a/erasure/theories/ESubstitution.v +++ b/erasure/theories/ESubstitution.v @@ -111,6 +111,13 @@ Proof. eapply All2_impl. exact X1. intros ? ? [[] [? []]]. split; eauto. + - econstructor. + induction H2; constructor. + induction H; constructor; depelim X2; eauto. + depelim X1. + eapply All2_All_mix_left in X3; eauto. + eapply All2_impl. exact X3. + cbn. intros ? ? [? ?]. eauto. Qed. Lemma erases_extends' (Σ:global_env_ext) Γ t T: @@ -574,6 +581,9 @@ Proof. eapply is_type_subst; eauto. - cbn. depelim H1. * cbn; constructor. + depelim H1. depelim H1; repeat constructor. + depelim X2. + + eapply (@substitution _ Σ _ Γ Γ' s Δ) in H1. * constructor. eapply is_type_subst in X3; tea. - eapply H; eauto. Qed. diff --git a/erasure/theories/Extract.v b/erasure/theories/Extract.v index af9063ea3..67b4c5ff7 100644 --- a/erasure/theories/Extract.v +++ b/erasure/theories/Extract.v @@ -6,9 +6,8 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICPrimitive PCUICTyp PCUICElimination PCUICWcbvEval PCUICFirstorder PCUICWellScopedCumulativity PCUICFirstorder PCUICNormalization PCUICReduction PCUICConversion PCUICPrincipality PCUICNormal. - -From MetaCoq.Erasure Require EAst EGlobalEnv. - +From MetaCoq.Erasure Require EPrimitive EAst EGlobalEnv. +Import EPrimitive. Module E := EAst. Local Existing Instance extraction_checker_flags. @@ -68,14 +67,20 @@ Reserved Notation "Σ ;;; Γ |- s ⇝ℇ t" (at level 50, Γ, s, t at next level Definition erase_context (Γ : context) : list name := map (fun d => d.(decl_name).(binder_name)) Γ. -Definition erase_prim_model {t : prim_tag} (e : @prim_model term t) : @prim_model E.term t := - match e in @prim_model _ x return prim_model E.term x with - | primIntModel i => primIntModel i - | primFloatModel f => primFloatModel f - end. +Inductive erase_prim_model (erase : term -> E.term -> Prop) : forall {t : prim_tag}, + @PCUICPrimitive.prim_model term t -> @prim_model E.term t -> Prop := +| erase_primInt i : @erase_prim_model erase primInt (PCUICPrimitive.primIntModel i) (primIntModel i) +| erase_primFloat f : @erase_prim_model erase primFloat (PCUICPrimitive.primFloatModel f) (primFloatModel f) +| erase_primArray a ed ev : + erase a.(PCUICPrimitive.array_default) ed -> + All2 erase a.(PCUICPrimitive.array_value) ev -> + @erase_prim_model erase primArray + (PCUICPrimitive.primArrayModel a) (primArrayModel {| array_default := ed; array_value := ev |}). -Definition erase_prim_val (p : prim_val term) : prim_val E.term := - (p.π1; erase_prim_model p.π2). +Inductive erase_prim_val (erase : term -> E.term -> Prop) : + PCUICPrimitive.prim_val term -> prim_val E.term -> Prop := +| erase_prim t m m' : @erase_prim_model erase t m m' -> + erase_prim_val erase (t; m) (t; m'). Inductive erases (Σ : global_env_ext) (Γ : context) : term -> E.term -> Prop := erases_tRel : forall i : nat, Σ;;; Γ |- tRel i ⇝ℇ E.tRel i @@ -126,7 +131,9 @@ Inductive erases (Σ : global_env_ext) (Γ : context) : term -> E.term -> Prop : × Σ;;; Γ ,,, fix_context mfix |- dbody d ⇝ℇ E.dbody d') mfix mfix' -> Σ;;; Γ |- tCoFix mfix n ⇝ℇ E.tCoFix mfix' n - | erases_tPrim : forall p, Σ;;; Γ |- tPrim p ⇝ℇ E.tPrim (erase_prim_val p) + | erases_tPrim : forall p p', + erase_prim_val (erases Σ Γ) p p' -> + Σ;;; Γ |- tPrim p ⇝ℇ E.tPrim p' | erases_box : forall t : term, isErasable Σ Γ t -> Σ;;; Γ |- t ⇝ℇ E.tBox where "Σ ;;; Γ |- s ⇝ℇ t" := (erases Σ Γ s t). @@ -197,7 +204,10 @@ Lemma erases_forall_list_ind (dbody d) (EAst.dbody d') ) mfix mfix' -> P Γ (tCoFix mfix n) (E.tCoFix mfix' n)) - (Hprim : forall Γ p, P Γ (tPrim p) (E.tPrim (erase_prim_val p))) + (Hprim : forall Γ p p', + erase_prim_val (erases Σ Γ) p p' -> + erase_prim_val (P Γ) p p' -> + P Γ (tPrim p) (E.tPrim p')) (Hbox : forall Γ t, isErasable Σ Γ t -> P Γ t E.tBox) : forall Γ t t0, Σ;;; Γ |- t ⇝ℇ t0 -> @@ -238,6 +248,14 @@ Proof. fix f' 3. intros mfix mfix' []; [now constructor|]. constructor; [now apply f|now apply f']. + - eapply Hprim; tea. + induction H. constructor. + induction H; constructor. + * now eapply f. + * destruct a; cbn in *. + revert array_value ev X. + fix f' 3; intros mfix mfix' []; [now constructor|]. + constructor; [now apply f|now apply f']. Defined. Definition erases_constant_body (Σ : global_env_ext) (cb : constant_body) (cb' : E.constant_body) := @@ -333,7 +351,14 @@ Inductive erases_deps (Σ : global_env) (Σ' : E.global_declarations) : E.term - | erases_deps_tCoFix defs i : Forall (fun d => erases_deps Σ Σ' (E.dbody d)) defs -> erases_deps Σ Σ' (E.tCoFix defs i) -| erases_deps_tPrim p : erases_deps Σ Σ' (E.tPrim p). +| erases_deps_tPrimInt i : + erases_deps Σ Σ' (E.tPrim (primInt; primIntModel i)) +| erases_deps_tPrimFloat f : + erases_deps Σ Σ' (E.tPrim (primFloat; primFloatModel f)) +| erases_deps_tPrimArray a : + erases_deps Σ Σ' a.(array_default) -> + Forall (erases_deps Σ Σ') a.(array_value) -> + erases_deps Σ Σ' (E.tPrim (primArray; primArrayModel a)). Definition option_is_none {A} (o : option A) := match o with diff --git a/erasure/theories/Typed/Certifying.v b/erasure/theories/Typed/Certifying.v index 565a2f3a8..e2c693743 100644 --- a/erasure/theories/Typed/Certifying.v +++ b/erasure/theories/Typed/Certifying.v @@ -51,6 +51,7 @@ Definition change_modpath (mpath : modpath) (suffix : string) (to_rename : kerna | tCoFix mfix idx => tCoFix (map (map_def go go) mfix) idx | tInt n => tInt n | tFloat n => tFloat n + | tArray l v def ty => tArray l (map go v) (go def) (go ty) end. Fixpoint map_constants_global_decls (k : kername -> kername) (f : constant_body -> constant_body) (Σ : global_declarations) : global_declarations := From 8ecbd0e3764556dbe0e9ae32d54619a2d74cc079 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Thu, 23 Nov 2023 15:21:49 +0100 Subject: [PATCH 11/17] WIP in erasure of arrays --- erasure/theories/ECSubst.v | 4 +- erasure/theories/EEtaExpandedFix.v | 84 ++++++++++++++-- erasure/theories/EInduction.v | 11 +- erasure/theories/ELiftSubst.v | 70 +++++++------ erasure/theories/EPrimitive.v | 155 +++++++++++++++++++++++++++++ erasure/theories/EReflect.v | 2 +- erasure/theories/EWellformed.v | 4 +- erasure/theories/Typed/ClosedAux.v | 5 +- 8 files changed, 283 insertions(+), 52 deletions(-) diff --git a/erasure/theories/ECSubst.v b/erasure/theories/ECSubst.v index 8500b0404..e05deb65d 100644 --- a/erasure/theories/ECSubst.v +++ b/erasure/theories/ECSubst.v @@ -1,7 +1,7 @@ (* Distributed under the terms of the MIT license. *) From Coq Require Import Program. From MetaCoq.Utils Require Import utils. -From MetaCoq.Erasure Require Import EAst EAstUtils EInduction ELiftSubst. +From MetaCoq.Erasure Require Import EPrimitive EAst EAstUtils EInduction ELiftSubst. Require Import ssreflect ssrbool. From Equations Require Import Equations. @@ -37,6 +37,7 @@ Fixpoint csubst t k u := let mfix' := List.map (map_def (csubst t k')) mfix in tCoFix mfix' idx | tConstruct ind n args => tConstruct ind n (map (csubst t k) args) + | tPrim p => tPrim (map_prim (csubst t k) p) | x => x end. @@ -168,7 +169,6 @@ Proof. - move/andP => []. intros. f_equal; solve_all; eauto. - move/andP => []. intros. f_equal; solve_all; eauto. - move/andP => []. intros. f_equal; solve_all; eauto. - destruct x0; cbn in *. f_equal; auto. Qed. Lemma subst_csubst_comm l t k b : diff --git a/erasure/theories/EEtaExpandedFix.v b/erasure/theories/EEtaExpandedFix.v index 951bdf971..257deaee4 100644 --- a/erasure/theories/EEtaExpandedFix.v +++ b/erasure/theories/EEtaExpandedFix.v @@ -2,7 +2,7 @@ From Coq Require Import Utf8 Program. From MetaCoq.Utils Require Import utils. From MetaCoq.Common Require Import config Kernames BasicAst EnvMap. -From MetaCoq.Erasure Require Import EAst EGlobalEnv EAstUtils EEnvMap EExtends EWellformed. +From MetaCoq.Erasure Require Import EPrimitive EAst EGlobalEnv EAstUtils EEnvMap EExtends EWellformed. From MetaCoq.Erasure Require Import EWcbvEvalInd EProgram EWcbvEval. From MetaCoq.Erasure Require Import EInduction ELiftSubst ESpineView ECSubst EProgram. @@ -60,11 +60,28 @@ Inductive expanded (Γ : list nat): term -> Prop := #|args| >= ind_npars mind + cdecl.(cstr_nargs) -> Forall (expanded Γ) args -> expanded Γ (mkApps (tConstruct ind idx []) args) -| expanded_tPrim p : expanded Γ (tPrim p) +| expanded_tPrim p : primProp (expanded Γ) p -> expanded Γ (tPrim p) | expanded_tBox : expanded Γ tBox. End expanded. +Section make_All. + Context {A} {B : A -> Type} (f : forall x : A, B x). + + Equations make_All (l : list A) : All B l := + | [] := All_nil + | hd :: tl := All_cons (f hd) (make_All tl). +End make_All. + +Section make_All_All. + Context {A} {B : A -> Type} {C : A -> Type} (f : forall x : A, B x -> C x). + + Equations make_All_All {l : list A} (a : All B l) : All C l := + | All_nil := All_nil + | All_cons bhd btl := All_cons (f _ bhd) (make_All_All btl). + +End make_All_All. + Lemma expanded_ind : ∀ (Σ : global_declarations) (P : list nat → term → Prop), (∀ (Γ : list nat) (n : nat) (args : list term) (m : nat), @@ -138,7 +155,7 @@ Lemma expanded_ind : → Forall (expanded Σ Γ) args → Forall (P Γ) args → P Γ (mkApps (tConstruct ind idx []) args)) - → (∀ Γ p, P Γ (tPrim p)) + → (∀ Γ p, primProp (expanded Σ Γ) p -> primProp (P Γ) p -> P Γ (tPrim p)) → (∀ Γ : list nat, P Γ tBox) → ∀ (Γ : list nat) (t : term), expanded Σ Γ t → P Γ t. Proof. @@ -158,6 +175,9 @@ Proof. generalize mfix at 1 3. intros mfix0 H. induction H; econstructor; cbn in *; eauto; split. - eapply HConstruct; eauto. clear - H1 f. induction H1; econstructor; eauto. + - eapply HPrim; eauto. + depelim X; constructor. intuition eauto. + eapply (make_All_All (f Γ) b). Qed. Definition expanded_constant_decl Σ (cb : constant_body) : Prop := @@ -267,6 +287,40 @@ Section isEtaExp. | None => false end. + Section PrimIn. + Context {term : Set}. + + Equations InPrim (x : term) (p : prim_val term) : Prop := + | x | (primInt; primIntModel i) := False + | x | (primFloat; primFloatModel _) := False + | x | (primArray; primArrayModel a) := + x = a.(array_default) \/ In x a.(array_value). + + Lemma InPrim_primProp p P : (forall x, InPrim x p -> P x) -> primProp P p. + Proof. + intros hin. + destruct p as [? []]; constructor. + split. eapply hin; eauto. now left. + cbn in hin. + induction (array_value a); constructor. + eapply hin. now right; left. eapply IHl. + intros. eapply hin. intuition eauto. now right; right. + Qed. + + Equations test_primIn (p : prim_val term) (f : forall x : term, InPrim x p -> bool) : bool := + | (primInt; primIntModel i) | _ := true + | (primFloat; primFloatModel _) | _ := true + | (primArray; primArrayModel a) | f := + f a.(array_default) (or_introl eq_refl) && forallb_InP a.(array_value) (fun x H => f x (or_intror H)). + + Lemma test_primIn_spec p (f : term -> bool) : + test_primIn p (fun x H => f x) = test_prim f p. + Proof. + funelim (test_primIn p (fun x H => f x)); cbn => //. + rewrite forallb_InP_spec //. + Qed. + End PrimIn. + Import TermSpineView. Definition is_nil {A} (l : list A) := match l with nil => true | _ => false end. @@ -292,7 +346,7 @@ Section isEtaExp. | tBox => true | tVar _ => true | tConst _ => true - | tPrim _ => true + | tPrim p => test_primIn p (fun x H => isEtaExp Γ x) | tConstruct ind i block_args => isEtaExp_app ind i 0 && is_nil block_args }. Proof using Σ. all:try lia. @@ -317,6 +371,9 @@ Section isEtaExp. change (fun x => size x) with size in H. pose proof (size_mkApps_l napp nnil). lia. - eapply (In_size snd size) in H. cbn in H; lia. + - destruct p as [? []]; cbn in *; eauto. destruct H; subst; try lia. + eapply (In_size id size) in H. unfold id in H. + change (fun x => size x) with size in H. lia. Qed. Lemma isEtaExp_app_mon ind c i i' : i <= i' -> isEtaExp_app ind c i -> isEtaExp_app ind c i'. @@ -328,6 +385,7 @@ Section isEtaExp. Qed. Hint Rewrite @forallb_InP_spec : isEtaExp. + Hint Rewrite @test_primIn_spec : isEtaExp. Lemma isEtaExp_mkApps_nonnil Γ f v : ~~ isApp f -> v <> [] -> @@ -431,6 +489,7 @@ Section isEtaExp. - destruct nth_error eqn:E; try easy. erewrite nth_error_app_left; eauto. - rewrite app_assoc. eapply a, b. reflexivity. - rewrite app_assoc. eapply a, b. reflexivity. + - eapply InPrim_primProp in H. solve_all. - rewrite isEtaExp_mkApps => //. cbn [expanded_head_viewc]. rtoProp; repeat solve_all. - rewrite isEtaExp_mkApps => //. cbn [expanded_head_viewc]. rtoProp; repeat solve_all. rewrite app_assoc. rtoProp; intuition auto. @@ -449,6 +508,7 @@ Section isEtaExp. - destruct block_args; cbn in *; eauto. - eapply a in b. 2: f_equal. revert b. now len. - eapply a in b. 2: f_equal. revert b. now len. + - eapply InPrim_primProp in H; solve_all. - cbn. destruct block_args; cbn in *; eauto. - cbn. solve_all. rtoProp; intuition auto. eapply a in H0. 2: reflexivity. revert H0. now len. - destruct nth_error eqn:Hn; cbn in H1; try easy. @@ -481,6 +541,8 @@ Section isEtaExp. solve_all. rewrite app_assoc. eapply a0; cbn; eauto. now len. cbn. now rewrite app_assoc. - rewrite app_assoc. eapply a0; len; eauto. now rewrite app_assoc. + - eapply InPrim_primProp in H. solve_all. + eapply primProp_map. solve_all. - fold csubst. move/andP: H1 => [] etaexp h. rewrite csubst_mkApps /=. rewrite isEtaExp_Constructor. solve_all. @@ -525,7 +587,7 @@ Section isEtaExp. #|Γ| = k -> nth_error mfix idx = Some d -> closed (EAst.tFix mfix idx) -> - forallb (fun x => isEtaExp (rev_map (S ∘ rarg) mfix) x.(dbody)) mfix -> + forallb (fun x => isLambda x.(dbody) && isEtaExp (rev_map (S ∘ rarg) mfix) x.(dbody)) mfix -> isEtaExp (Γ ++ [1 + d.(EAst.rarg)] ++ Δ) b -> isEtaExp (Γ ++ Δ) (ECSubst.csubst (EAst.tFix mfix idx) k b). Proof using Type*. intros Hk Hnth Hcl. @@ -555,6 +617,9 @@ Section isEtaExp. { cbn in Hcl. solve_all. rewrite Nat.add_0_r in a0. eauto. } now rewrite app_assoc. solve_all. + - eapply InPrim_primProp in H. solve_all. + eapply primProp_map. eapply primProp_impl; tea. intuition eauto. + destruct H. eapply i; eauto. solve_all. - solve_all. fold csubst. move/andP: H1 => [] etaexp h. rewrite csubst_mkApps /=. rewrite isEtaExp_Constructor. solve_all. @@ -632,13 +697,13 @@ Section isEtaExp. Lemma isEtaExp_fixsubstl Δ mfix t : forallb (fun x => - (* isLambda x.(dbody) && *) + isLambda x.(dbody) && isEtaExp (rev_map (S ∘ rarg) mfix) x.(dbody)) mfix -> isEtaExp ((rev_map (S ∘ rarg) mfix) ++ Δ) t -> isEtaExp Δ (substl (fix_subst mfix) t). Proof using Type*. intros Hall Heta. - assert (Hcl : closed (EAst.tFix mfix 0) ). { cbn. solve_all. rtoProp; intuition auto. eapply isEtaExp_closed in H. revert H. now len. } + assert (Hcl : closed (EAst.tFix mfix 0) ). { cbn. solve_all. rtoProp; intuition auto. eapply isEtaExp_closed in H0. revert H0. now len. } revert Hcl Hall Heta. intros Hcl Hall Heta. cbn in Hcl. solve_all. @@ -809,6 +874,8 @@ Proof. eapply In_All in H0. solve_all. - econstructor. rewrite forallb_InP_spec in H0. eapply forallb_Forall in H0. eapply In_All in H. solve_all. + - econstructor. eapply InPrim_primProp in H. rewrite test_primIn_spec in H0. + solve_all. - rtoProp. eapply In_All in H. rewrite forallb_InP_spec in H2. eapply forallb_Forall in H2. eapply isEtaExp_app_expanded in H0 as (? & ? & ? & ? & ?). @@ -848,6 +915,7 @@ Proof. - rewrite isEtaExp_Constructor. rtoProp. repeat split. 2: eapply forallb_Forall; solve_all. eapply expanded_isEtaExp_app_; eauto. + - rewrite test_primIn_spec. solve_all. Qed. Definition isEtaExp_constant_decl Σ cb := @@ -987,6 +1055,8 @@ Proof. - eapply isEtaExp_app_extends; tea. - eapply In_All in H0. solve_all. - eapply In_All in H; solve_all. + - eapply InPrim_primProp in H. + rewrite !test_primIn_spec in H0 |- *. solve_all. - eapply In_All in H; solve_all. rewrite isEtaExp_Constructor //. rtoProp; intuition auto. eapply isEtaExp_app_extends; tea. diff --git a/erasure/theories/EInduction.v b/erasure/theories/EInduction.v index 093b0e56d..19f8929ad 100644 --- a/erasure/theories/EInduction.v +++ b/erasure/theories/EInduction.v @@ -15,13 +15,6 @@ Set Equations Transparent. (** Custom induction principle on syntax, dealing with the various lists appearing in terms. *) -Definition primProp P (p : prim_val term) : Type := - match p.π2 with - | primIntModel _ => unit - | primFloatModel _ => unit - | primArrayModel a => P a.(array_default) × All P a.(array_value) - end. - Lemma term_forall_list_ind : forall P : term -> Type, (P tBox) -> @@ -74,7 +67,7 @@ Proof. destruct mfix; constructor; [|apply auxm]. apply auxt. - destruct p as [? []]; cbn => //. + destruct p as [? []]; cbn => //; constructor. destruct a as [def v]; cbn. split. eapply auxt. revert v; fix auxv 1. @@ -276,7 +269,7 @@ Section MkApps_rec. rewrite da in H0. cbn in H0. rewrite <- H0. unfold id in H. change (fun x => size x) with size in H. abstract lia. - clear -da. abstract (eapply decompose_app_inv in da; now symmetry). - - destruct p as [? []]; cbn => //. + - destruct p as [? []]; cbn => //; constructor. destruct a as [def v]; cbn. split. * eapply rec; red; cbn. lia. diff --git a/erasure/theories/ELiftSubst.v b/erasure/theories/ELiftSubst.v index d78515df6..83ad52ef7 100644 --- a/erasure/theories/ELiftSubst.v +++ b/erasure/theories/ELiftSubst.v @@ -160,6 +160,14 @@ Hint Extern 0 (_ = _) => progress f_equal : all. #[global] Hint Unfold on_snd snd : all. +#[global] Hint Extern 10 => + match goal with + | [ H : _ × _ |- _ ] => destruct H + end : all. + +#[global] Hint Extern 10 => + progress cbn beta match fix cofix iota zeta in * : all. + Lemma on_snd_eq_id_spec {A B} (f : B -> B) (x : A * B) : f (snd x) = snd x <-> on_snd f x = x. @@ -203,12 +211,24 @@ Ltac change_Sk := |- context [S (?x + ?y)] => progress change (S (x + y)) with (S x + y) end. - -Ltac solve_all := +Ltac primProp := + repeat match goal with + | [ H : is_true (test_prim _ _) |- _ ] => eapply test_prim_impl_primProp in H + | [ H : test_prim _ _ = true |- _ ] => eapply test_prim_impl_primProp in H + | [ |- is_true (test_prim _ _) ] => eapply primProp_impl_test_prim + | [ |- test_prim _ _ = true ] => eapply primProp_impl_test_prim + | [ H : primProp ?P ?p, H' : primProp ?Q ?p |- _ ] => + let H'' := fresh in pose proof (H'' := primProp_conj H H'); clear H H' + end. + +Tactic Notation "solve_all_k" int(k) := unfold tFixProp in *; + primProp; repeat toAll; try All_map; try close_Forall; change_Sk; auto with all; - intuition eauto 4 with all. + intuition eauto k with all. + +Ltac solve_all := solve_all_k 5. Ltac nth_leb_simpl := match goal with @@ -231,9 +251,6 @@ Proof. try (f_equal; auto; solve_all). - now elim (leb k n). - - destruct x; cbn. now rewrite H0. - - destruct p as [? []]; cbn in X |- *; intuition eauto. do 2 f_equal. - rewrite /map_array_model; destruct a; cbn => //=. f_equal; eauto. solve_all. Qed. Lemma lift0_p : forall M, lift0 0 M = M. @@ -252,7 +269,7 @@ Proof. intros M. elim M using term_forall_list_ind; intros; simpl; - rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; + rewrite -> ?map_map_compose, ?map_prim_compose, ?compose_on_snd, ?compose_map_def, ?map_length; try (rewrite -> H, ?H0, ?H1; auto); try (f_equal; auto; solve_all). - elim (leb_spec k n); intros. @@ -272,7 +289,7 @@ Proof. intros M. elim M using term_forall_list_ind; intros; simpl; - rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length, ?Nat.add_assoc; + rewrite -> ?map_map_compose, ?map_prim_compose, ?compose_on_snd, ?compose_map_def, ?map_length, ?Nat.add_assoc; try solve [f_equal; auto; solve_all]; repeat nth_leb_simpl. f_equal; auto. solve_all. f_equal. rewrite Nat.add_assoc. @@ -320,7 +337,7 @@ Lemma simpl_subst_rec : Proof. intros M. induction M using term_forall_list_ind; intros; simpl; - rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; + rewrite -> ?map_map_compose, ?map_prim_compose, ?compose_on_snd, ?compose_map_def, ?map_length; try solve [f_equal; auto; solve_all]; repeat nth_leb_simpl. Qed. @@ -342,7 +359,7 @@ Proof. intros M. elim M using term_forall_list_ind; intros; simpl; try easy; - rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length, ?Nat.add_assoc; + rewrite -> ?map_map_compose, ?map_prim_compose, ?compose_on_snd, ?compose_map_def, ?map_length, ?Nat.add_assoc; try solve [f_equal; auto; solve_all]. - repeat nth_leb_simpl. @@ -367,7 +384,7 @@ Proof. |- context [tRel _] => idtac | |- _ => cbn -[plus] end; try easy; - rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length, ?Nat.add_assoc; + rewrite -> ?map_map_compose, ?map_prim_compose, ?compose_on_snd, ?compose_map_def, ?map_length, ?Nat.add_assoc; try solve [f_equal; auto; solve_all]. - unfold subst at 1. unfold lift at 4. @@ -423,7 +440,7 @@ Proof. |- context [tRel _] => idtac | |- _ => simpl end; try easy; - rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length, ?Nat.add_assoc; + rewrite -> ?map_map_compose, ?map_prim_compose, ?compose_on_snd, ?compose_map_def, ?map_length, ?Nat.add_assoc; try solve [f_equal; auto; solve_all]. - unfold subst at 2. @@ -450,22 +467,19 @@ Lemma lift_closed n k t : closedn k t -> lift n k t = t. Proof. revert k. elim t using term_forall_list_ind; intros; try easy; - rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; + rewrite -> ?map_map_compose, ?map_prim_compose, ?compose_on_snd, ?compose_map_def, ?map_length; unfold test_def in *; simpl closed in *; try solve [simpl lift; simpl closed; f_equal; auto; rtoProp; solve_all]; try easy. + - rewrite lift_rel_lt; auto. revert H. elim (Nat.ltb_spec n0 k); intros; try easy. - - cbn. f_equal; auto. - rtoProp; solve_all. - rtoProp; solve_all. - destruct x; f_equal; cbn in *. eauto. Qed. Lemma closed_upwards {k t} k' : closedn k t -> k' >= k -> closedn k' t. Proof. revert k k'. elim t using term_forall_list_ind; intros; try lia; - rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; + rewrite -> ?map_map_compose, ?map_prim_compose, ?compose_on_snd, ?compose_map_def, ?map_length; simpl closed in *; unfold test_snd, test_def in *; try solve [(try f_equal; simpl; repeat (rtoProp; solve_all); eauto)]. @@ -483,24 +497,18 @@ Proof. assert (n - k > 0) by lia. assert (exists n', n - k = S n'). exists (pred (n - k)). lia. destruct H2. rewrite H2. simpl. now rewrite Nat.sub_0_r. - - f_equal; eauto; solve_all. destruct x; cbn in *; eauto. - now rewrite H. Qed. Lemma subst_closed n k t : closedn k t -> subst n k t = t. Proof. revert k. elim t using term_forall_list_ind; intros; try easy; - rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; + rewrite -> ?map_map_compose, ?map_prim_compose, ?compose_on_snd, ?compose_map_def, ?map_length; unfold test_def in *; simpl closed in *; try solve [simpl subst; simpl closed; f_equal; auto; rtoProp; solve_all]; try easy. - cbn. revert H. elim (Nat.ltb_spec n0 k); intros; try easy. elim (Nat.leb_spec k n0); intros; try easy. - - cbn. f_equal; auto. - rtoProp; solve_all. - rtoProp; solve_all. - destruct x; f_equal; cbn in *. now apply a0. Qed. (* Lemma lift_to_extended_list_k Γ k : forall k', *) @@ -526,7 +534,7 @@ Lemma subst_app_decomp l l' k t : Proof. induction t in k |- * using term_forall_list_ind; simpl; auto; rewrite ?subst_mkApps; try change_Sk; - try (f_equal; rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; + try (f_equal; rewrite -> ?map_map_compose, ?map_prim_compose, ?compose_on_snd, ?compose_map_def, ?map_length; eauto; solve_all). - repeat nth_leb_simpl. @@ -542,7 +550,7 @@ Lemma subst_app_simpl l l' k t : Proof. induction t in k |- * using term_forall_list_ind; simpl; eauto; rewrite ?subst_mkApps; try change_Sk; - try (f_equal; rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length, ?Nat.add_assoc; + try (f_equal; rewrite -> ?map_map_compose, ?map_prim_compose, ?compose_on_snd, ?compose_map_def, ?map_length, ?Nat.add_assoc; eauto; solve_all; eauto). - repeat nth_leb_simpl. @@ -583,13 +591,12 @@ Lemma closedn_subst_eq s k k' t : closedn (k + k' + #|s|) t = closedn (k + k') (subst s k' t). Proof. - intros Hs. solve_all. revert Hs. + intros Hs. toAll. revert Hs. induction t in k' |- * using EInduction.term_forall_list_ind; intros; simpl in *; autorewrite with map => //; simpl closed in *; try change_Sk; - unfold test_def in *; simpl in *; - solve_all. + unfold test_def in *; simpl in *. - elim (Nat.leb_spec k' n); intros. simpl. destruct nth_error eqn:Heq. @@ -611,18 +618,21 @@ Proof. - specialize (IHt2 (S k')). rewrite <- Nat.add_succ_comm in IHt2. rewrite IHt1 // IHt2 //. + - solve_all. - eapply All_forallb_eq_forallb; eauto. - rewrite IHt //. f_equal. eapply All_forallb_eq_forallb; tea. cbn. intros. specialize (H (#|x.1| + k')). rewrite Nat.add_assoc (Nat.add_comm k) in H. now rewrite !Nat.add_assoc. + - solve_all. - eapply All_forallb_eq_forallb; tea. cbn. intros. specialize (H (#|m| + k')). now rewrite !Nat.add_assoc !(Nat.add_comm k) in H |- *. - eapply All_forallb_eq_forallb; tea. cbn. intros. specialize (H (#|m| + k')). now rewrite !Nat.add_assoc !(Nat.add_comm k) in H |- *. + - solve_all. Qed. Lemma closedn_subst s k t : diff --git a/erasure/theories/EPrimitive.v b/erasure/theories/EPrimitive.v index a4715ffe4..866b3c414 100644 --- a/erasure/theories/EPrimitive.v +++ b/erasure/theories/EPrimitive.v @@ -157,3 +157,158 @@ Definition test_prim {term} (f : term -> bool) (p : prim_val term) : bool := | primFloatModel f => true | primArrayModel a => test_array_model f a end. + +Inductive primProp {term} P : prim_val term -> Type := + | primPropInt i : primProp P (primInt; primIntModel i) + | primPropFloat f : primProp P (primFloat; primFloatModel f) + | primPropArray a : P a.(array_default) × All P a.(array_value) -> + primProp P (primArray; primArrayModel a). +Derive Signature NoConfusion for primProp. + +Section primtheory. + Context {term term' term''} (g : term' -> term'') (f : term -> term') (p : prim_val term). + + Lemma map_prim_compose : + map_prim g (map_prim f p) = map_prim (g ∘ f) p. + Proof. + destruct p as [? []]; cbn; auto. + do 2 f_equal. destruct a; rewrite /map_array_model //= map_map_compose //. + Qed. +End primtheory. +#[global] Hint Rewrite @map_prim_compose : all. + +Lemma reflectT_forallb {A} {P : A -> Type} {p l} : + (forall x, reflectT (P x) (p x)) -> + reflectT (All P l) (forallb p l). +Proof. + move=> hr; induction l; cbn; try constructor; eauto. + case: (hr a) => /= // pa. + - case: IHl; constructor; eauto. + intros ha; apply f. now depelim ha. + - constructor. intros ha; now depelim ha. +Qed. + +Section primtheory. + Context {term} {p : prim_val term}. + + Lemma map_prim_id f : + (forall x, f x = x) -> + map_prim f p = p. + Proof. + destruct p as [? []]; cbn; auto. + intros hf. do 2 f_equal. destruct a; rewrite /map_array_model //=; f_equal; eauto. + rewrite map_id_f //. + Qed. + + Lemma map_prim_id_prop f P : + primProp P p -> + (forall x, P x -> f x = x) -> + map_prim f p = p. + Proof. + destruct p as [? []]; cbn; auto. + intros hp hf. do 2 f_equal. destruct a; rewrite /map_array_model //=; f_equal; eauto; depelim hp. + * eapply hf; intuition eauto. + * eapply All_map_id. destruct p0 as [_ hp]. + eapply All_impl; tea. + Qed. + + Lemma map_prim_eq {term'} {f g : term -> term'} : + (forall x, f x = g x) -> + map_prim f p = map_prim g p. + Proof. + destruct p as [? []]; cbn; auto. + intros hf. do 2 f_equal. destruct a; rewrite /map_array_model //=; f_equal; eauto. + now eapply map_ext. + Qed. + + Lemma map_prim_eq_prop {term' P} {f g : term -> term'} : + primProp P p -> + (forall x, P x -> f x = g x) -> + map_prim f p = map_prim g p. + Proof. + intros hp; depelim hp; subst; cbn; auto. intros hf. + do 2 f_equal. destruct a; rewrite /map_array_model //=; f_equal; eauto. + * eapply hf; intuition eauto. + * destruct p0 as [_ hp]. + eapply All_map_eq. + eapply All_impl; tea. + Qed. + + Lemma test_prim_primProp P pr : + (forall t, reflectT (P t) (pr t)) -> + reflectT (primProp P p) (test_prim pr p). + Proof. + intros hr. + destruct p as [? []]; cbn; repeat constructor => //. + destruct a as []; cbn. + rewrite /test_array_model /=. + case: (hr array_default0) => //= hd. + - case: (reflectT_forallb hr) => hv; repeat constructor; eauto. + intros hp; depelim hp; intuition eauto. + - constructor. intros hp; depelim hp; intuition eauto. + Qed. + + Lemma test_prim_impl_primProp pr : + test_prim pr p -> primProp pr p. + Proof. + case: (test_prim_primProp (fun x => pr x) pr) => //. + intros t; destruct (pr t); constructor; eauto. + Qed. + + Lemma primProp_impl_test_prim (pr : term -> bool) : + primProp pr p -> test_prim pr p. + Proof. + case: (test_prim_primProp (fun x => pr x) pr) => //. + intros t; destruct (pr t); constructor; eauto. + Qed. + + Lemma primProp_conj {P Q} : primProp P p -> primProp Q p -> primProp (fun x => P x × Q x) p. + Proof. + destruct p as [? []]; cbn; repeat constructor; intuition eauto. + now depelim X. now depelim X0. + depelim X; depelim X0. + now eapply All_mix. + Qed. + + Lemma primProp_impl {P Q} : primProp P p -> (forall x, P x -> Q x) -> primProp Q p. + Proof. + intros hp hpq; destruct p as [? []]; cbn in *; intuition eauto; constructor. + depelim hp; intuition eauto. + eapply All_impl; tea. + Qed. + + Lemma primProp_map {term' P} {f : term -> term'} : + primProp (P ∘ f) p -> + primProp P (map_prim f p). + Proof. + destruct p as [? []]; cbn in *; intuition eauto; constructor. + depelim X; intuition eauto. + eapply All_map, All_impl; tea. cbn; eauto. + Qed. + + Lemma test_prim_map {term' pr} {f : term -> term'} : + test_prim pr (map_prim f p) = test_prim (pr ∘ f) p. + Proof. + destruct p as [? []]; cbn in *; intuition auto. + destruct a as []; rewrite /test_array_model //=. f_equal. + rewrite forallb_map //. + Qed. + + Lemma test_prim_eq_prop P pr pr' : + primProp P p -> + (forall x, P x -> pr x = pr' x) -> + test_prim pr p = test_prim pr' p. + Proof. + destruct p as [? []]; cbn in *; intuition auto. + destruct a as []; rewrite /test_array_model //=. + depelim X; f_equal; intuition eauto. + eapply All_forallb_eq_forallb; tea. + Qed. + +End primtheory. + +#[global] Hint Rewrite @test_prim_map : map. +#[global] Hint Resolve map_prim_id map_prim_id_prop : all. +#[global] Hint Resolve map_prim_eq map_prim_eq_prop : all. +#[global] Hint Resolve primProp_impl primProp_map : all. +#[global] Hint Resolve test_prim_eq_prop : all. diff --git a/erasure/theories/EReflect.v b/erasure/theories/EReflect.v index 58f99e2f4..9fbc550b8 100644 --- a/erasure/theories/EReflect.v +++ b/erasure/theories/EReflect.v @@ -116,7 +116,7 @@ Proof. + destruct (eq_dec i i0); nodec; subst; eauto. + destruct (eq_dec f f0); nodec; subst; eauto. + destruct a as [? ?], a0 as [? ?]; cbn. - destruct X as [hd hv]. cbn in *. + depelim X. destruct p as [hd hv]. cbn in *. destruct (hd array_default0); nodec; subst; eauto. revert array_value0; induction hv; intros []; eauto; nodec. destruct (p t); subst; nodec. diff --git a/erasure/theories/EWellformed.v b/erasure/theories/EWellformed.v index 693297e44..034bd6334 100644 --- a/erasure/theories/EWellformed.v +++ b/erasure/theories/EWellformed.v @@ -2,7 +2,7 @@ From Coq Require Import Utf8 Program. From MetaCoq.Utils Require Import utils. From MetaCoq.Common Require Import config Kernames. -From MetaCoq.Erasure Require Import EAst EAstUtils ELiftSubst ECSubst EGlobalEnv. +From MetaCoq.Erasure Require Import EPrimitive EAst EAstUtils ELiftSubst ECSubst EGlobalEnv. From MetaCoq.PCUIC Require Import PCUICTactics. Local Open Scope string_scope. @@ -119,7 +119,7 @@ Section wf. | _ => true end && forallb (wellformed k) block_args else is_nil block_args | tVar _ => has_tVar - | tPrim _ => has_tPrim + | tPrim p => has_tPrim && test_prim (wellformed k) p end. End wf. diff --git a/erasure/theories/Typed/ClosedAux.v b/erasure/theories/Typed/ClosedAux.v index 89fa05df0..da4836d84 100644 --- a/erasure/theories/Typed/ClosedAux.v +++ b/erasure/theories/Typed/ClosedAux.v @@ -2,7 +2,7 @@ From Coq Require Import List. From Coq Require Import ssrbool. From MetaCoq.Erasure.Typed Require Import Utils. From MetaCoq.Utils Require Import utils. -From MetaCoq.Erasure Require Import EAst. +From MetaCoq.Erasure Require Import EPrimitive EAst. From MetaCoq.Erasure Require Import EAstUtils. From MetaCoq.Erasure Require Import ECSubst. From MetaCoq.Erasure Require Import EInduction. @@ -123,6 +123,7 @@ Proof. repeat (f_equal; try lia). rewrite <- (proj2 clos); repeat (f_equal; try lia). + - solve_all_k 6. Qed. Lemma closedn_subst s k k' t : @@ -192,6 +193,8 @@ Proof. * repeat (f_equal; try lia). * rewrite <- (proj2 all'). repeat (f_equal; try lia). + - solve_all. eapply primProp_map. + eapply primProp_impl; tea. cbn. intuition eauto. eapply a; tea. solve_all. Qed. Lemma closedn_subst0 s k t : From 335484751ad42f42246fcef24565ee4058e92bdb Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Fri, 24 Nov 2023 16:36:10 +0100 Subject: [PATCH 12/17] Adapt erasure for primitive arrays --- .vscode/metacoq.code-workspace | 37 +++++----- erasure-plugin/_PluginProject.in | 2 + .../src/metacoq_erasure_plugin.mlpack | 1 + erasure/theories/EAstUtils.v | 34 ++++++--- erasure/theories/EConstructorsAsBlocks.v | 7 +- erasure/theories/EDeps.v | 22 +++++- erasure/theories/EEtaExpanded.v | 25 +++++-- erasure/theories/EEtaExpandedFix.v | 74 +++---------------- erasure/theories/EImplementBox.v | 26 +++---- erasure/theories/EInduction.v | 8 ++ erasure/theories/EInlineProjections.v | 16 ++-- erasure/theories/ELiftSubst.v | 3 +- erasure/theories/EOptimizePropDiscr.v | 13 ++-- erasure/theories/EPrimitive.v | 66 ++++++++++++++++- erasure/theories/ERemoveParams.v | 24 ++++-- erasure/theories/ESubstitution.v | 18 +++-- .../EWcbvEvalCstrsAsBlocksFixLambdaInd.v | 10 ++- erasure/theories/EWcbvEvalCstrsAsBlocksInd.v | 6 +- erasure/theories/EWcbvEvalEtaInd.v | 6 +- erasure/theories/EWellformed.v | 7 -- erasure/theories/ErasureCorrectness.v | 7 ++ erasure/theories/ErasureFunction.v | 55 ++++++++------ erasure/theories/ErasureFunctionProperties.v | 59 ++++++++++----- erasure/theories/ErasureProperties.v | 16 +++- erasure/theories/Extract.v | 8 +- erasure/theories/Typed/Optimize.v | 9 ++- erasure/theories/Typed/OptimizeCorrectness.v | 41 ++++++++-- erasure/theories/Typed/TypeAnnotations.v | 2 +- pcuic/theories/utils/PCUICPrimitive.v | 2 +- utils/theories/All_Forall.v | 16 ++++ 30 files changed, 401 insertions(+), 219 deletions(-) diff --git a/.vscode/metacoq.code-workspace b/.vscode/metacoq.code-workspace index d6ac88a4c..0c15d9de8 100644 --- a/.vscode/metacoq.code-workspace +++ b/.vscode/metacoq.code-workspace @@ -11,27 +11,28 @@ - "-R", "safechecker-plugin/theories,MetaCoq.SafeCheckerPlugin", - "-R", "utils/theories,MetaCoq.Utils", - "-R", "common/theories,MetaCoq.Common", - "-R", "template-pcuic/theories,MetaCoq.TemplatePCUIC", + "-R", "safechecker-plugin/theories","MetaCoq.SafeCheckerPlugin", - "-I template-coq", + "-R", "utils/theories","MetaCoq.Utils", + "-R", "common/theories","MetaCoq.Common", + "-R", "template-pcuic/theories","MetaCoq.TemplatePCUIC", + + "-I", "template-coq", // "-bt", get backtraces from Coq on errors - "-I template-coq/build", - "-R", "template-coq/theories,MetaCoq.Template", - "-R", "examples,MetaCoq.Examples", - "-R", "checker/theories,MetaCoq.Checker", - "-R", "pcuic/theories,MetaCoq.PCUIC", - "-I safechecker/src", - "-R", "safechecker/theories,MetaCoq.SafeChecker", - "-I erasure/src", - "-R", "erasure/theories,MetaCoq.Erasure", - "-R", "erasure-plugin/theories,MetaCoq.ErasurePlugin", - "-R", "translations,MetaCoq.Translations", - "-R", "test-suite/plugin-demo/theories,MetaCoq.ExtractedPluginDemo", - "-R", "test-suite,MetaCoq.TestSuite" + "-I", "template-coq/build", + "-R", "template-coq/theories", "MetaCoq.Template", + "-R", "examples", "MetaCoq.Examples", + "-R", "checker/theories" "MetaCoq.Checker", + "-R", "pcuic/theories","MetaCoq.PCUIC", + "-I", "safechecker/src", + "-R", "safechecker/theories","MetaCoq.SafeChecker", + "-I", "erasure/src", + "-R", "erasure/theories","MetaCoq.Erasure", + "-R", "erasure-plugin/theories","MetaCoq.ErasurePlugin", + "-R", "translations","MetaCoq.Translations", + "-R", "test-suite/plugin-demo/theories","MetaCoq.ExtractedPluginDemo", + "-R", "test-suite","MetaCoq.TestSuite" ], // When enabled, will trim trailing whitespace when saving a file. diff --git a/erasure-plugin/_PluginProject.in b/erasure-plugin/_PluginProject.in index dc5b48e22..217e83bb3 100644 --- a/erasure-plugin/_PluginProject.in +++ b/erasure-plugin/_PluginProject.in @@ -50,6 +50,8 @@ src/pCUICSafeReduce.ml src/pCUICSafeRetyping.ml src/pCUICSafeRetyping.mli +src/ePrimitive.mli +src/ePrimitive.ml src/eAst.ml src/eAst.mli src/eAstUtils.ml diff --git a/erasure-plugin/src/metacoq_erasure_plugin.mlpack b/erasure-plugin/src/metacoq_erasure_plugin.mlpack index 29a6ec0d0..8bd95c56b 100644 --- a/erasure-plugin/src/metacoq_erasure_plugin.mlpack +++ b/erasure-plugin/src/metacoq_erasure_plugin.mlpack @@ -39,6 +39,7 @@ PCUICUnivSubst PCUICSafeReduce PCUICSafeRetyping +EPrimitive EAst EAstUtils ELiftSubst diff --git a/erasure/theories/EAstUtils.v b/erasure/theories/EAstUtils.v index ec168e08e..ec9068a60 100644 --- a/erasure/theories/EAstUtils.v +++ b/erasure/theories/EAstUtils.v @@ -2,7 +2,7 @@ From Equations Require Import Equations. From MetaCoq.Utils Require Import utils. From MetaCoq.Common Require Import BasicAst Kernames. -From MetaCoq.Erasure Require Import EAst. +From MetaCoq.Erasure Require Import EPrimitive EAst. Require Import ssreflect ssrbool. Global Hint Resolve app_tip_nil : core. @@ -378,24 +378,36 @@ Fixpoint string_of_term (t : term) : string := (** Compute all the global environment dependencies of the term *) -Fixpoint term_global_deps (t : EAst.term) := +Section PrimDeps. + Context (deps : term -> KernameSet.t). + + Equations prim_global_deps (p : prim_val term) : KernameSet.t := + | (primInt; primIntModel i) => KernameSet.empty + | (primFloat; primFloatModel f) => KernameSet.empty + | (primArray; primArrayModel a) => + List.fold_left (fun acc x => KernameSet.union (deps x) acc) a.(array_value) (deps a.(array_default)). + +End PrimDeps. + +Fixpoint term_global_deps (t : term) := match t with - | EAst.tConst kn => KernameSet.singleton kn - | EAst.tConstruct {| inductive_mind := kn |} _ args => + | tConst kn => KernameSet.singleton kn + | tConstruct {| inductive_mind := kn |} _ args => List.fold_left (fun acc x => KernameSet.union (term_global_deps x) acc) args (KernameSet.singleton kn) - | EAst.tLambda _ x => term_global_deps x - | EAst.tApp x y - | EAst.tLetIn _ x y => KernameSet.union (term_global_deps x) (term_global_deps y) - | EAst.tCase (ind, _) x brs => + | tLambda _ x => term_global_deps x + | tApp x y + | tLetIn _ x y => KernameSet.union (term_global_deps x) (term_global_deps y) + | tCase (ind, _) x brs => KernameSet.union (KernameSet.singleton (inductive_mind ind)) (List.fold_left (fun acc x => KernameSet.union (term_global_deps (snd x)) acc) brs (term_global_deps x)) - | EAst.tFix mfix _ | EAst.tCoFix mfix _ => - List.fold_left (fun acc x => KernameSet.union (term_global_deps (EAst.dbody x)) acc) mfix + | tFix mfix _ | tCoFix mfix _ => + List.fold_left (fun acc x => KernameSet.union (term_global_deps (dbody x)) acc) mfix KernameSet.empty - | EAst.tProj p c => + | tProj p c => KernameSet.union (KernameSet.singleton (inductive_mind p.(proj_ind))) (term_global_deps c) + | tPrim p => prim_global_deps term_global_deps p | _ => KernameSet.empty end. \ No newline at end of file diff --git a/erasure/theories/EConstructorsAsBlocks.v b/erasure/theories/EConstructorsAsBlocks.v index 78dab5e91..58fe52125 100644 --- a/erasure/theories/EConstructorsAsBlocks.v +++ b/erasure/theories/EConstructorsAsBlocks.v @@ -2,7 +2,7 @@ From Coq Require Import Utf8 Program. From MetaCoq.Utils Require Import utils. From MetaCoq.Common Require Import config Kernames BasicAst EnvMap. -From MetaCoq.Erasure Require Import EAst EAstUtils EInduction EArities +From MetaCoq.Erasure Require Import EPrimitive EAst EAstUtils EInduction EArities ELiftSubst ESpineView EGlobalEnv EWellformed EEnvMap EWcbvEval EEtaExpanded ECSubst EWcbvEvalEtaInd EProgram. @@ -62,7 +62,7 @@ Section transform_blocks. | tVar n => EAst.tVar n | tConst n => EAst.tConst n | tConstruct ind i block_args => EAst.tConstruct ind i [] - | tPrim p => EAst.tPrim p }. + | tPrim p => EAst.tPrim (map_primIn p (fun x H => transform_blocks x)) }. Proof. all:try lia. all:try apply (In_size); tea. @@ -82,6 +82,7 @@ Section transform_blocks. change (fun x => size x) with size in H. pose proof (size_mkApps_l napp nnil). lia. - eapply (In_size snd size) in H. cbn in *. lia. + - now eapply InPrim_size in H. Qed. End Def. @@ -116,6 +117,7 @@ Section transform_blocks. unfold test_def in *; simpl closed in *; try solve [simpl; subst; simpl closed; f_equal; auto; rtoProp; solve_all; solve_all]; try easy. + - solve_all_k 6. - rewrite !closedn_mkApps in H1 *. rtoProp; intuition auto. solve_all. - destruct (chop nargs v) eqn:E. @@ -714,6 +716,7 @@ Proof. cbn -[transform_blocks isEtaExp] in *. rtoProp. eauto. - unfold wf_fix in *. len. solve_all. rtoProp; intuition auto. solve_all. + - solve_all_k 7. - rewrite !wellformed_mkApps in Hw |- * => //. rtoProp. eapply isEtaExp_mkApps in H1. rewrite decompose_app_mkApps in H1; eauto. destruct construct_viewc; eauto. cbn in d. eauto. diff --git a/erasure/theories/EDeps.v b/erasure/theories/EDeps.v index 8189ffdfe..dd272ca21 100644 --- a/erasure/theories/EDeps.v +++ b/erasure/theories/EDeps.v @@ -1,9 +1,9 @@ From Coq Require Import Arith List. From Equations Require Import Equations. From MetaCoq.PCUIC Require Import - PCUICAst PCUICAstUtils PCUICTyping PCUICInversion PCUICWeakeningEnv PCUICWeakeningEnvTyp. + PCUICPrimitive PCUICAst PCUICAstUtils PCUICTyping PCUICInversion PCUICWeakeningEnv PCUICWeakeningEnvTyp. Set Warnings "-notation-overridden". -From MetaCoq.Erasure Require Import EAst EAstUtils ECSubst EInduction +From MetaCoq.Erasure Require Import EPrimitive EAst EAstUtils ECSubst EInduction ELiftSubst EGlobalEnv EWcbvEval Extract ESubstitution. From MetaCoq.Erasure Require EExtends. Set Warnings "+notation-overridden". @@ -84,6 +84,8 @@ Proof. depelim X. constructor; [|easy]. now apply e. + - depelim X; depelim er; constructor; cbn. solve_all. + destruct p. solve_all. Qed. Lemma erases_deps_subst Σ Σ' s k t : @@ -134,6 +136,7 @@ Proof. depelim X. constructor; [|easy]. now apply e. + - depelim X; depelim er; constructor; cbn; intuition auto; solve_all. Qed. Lemma erases_deps_subst1 Σ Σ' t k u : @@ -191,6 +194,7 @@ Proof. depelim X. constructor; [|easy]. now apply e. + - depelim X; depelim er; constructor; cbn; intuition auto; solve_all. Qed. Lemma erases_deps_substl Σ Σ' s t : @@ -402,7 +406,7 @@ Lemma erases_deps_forall_ind Σ Σ' Forall (fun d : Extract.E.def Extract.E.term => erases_deps Σ Σ' (Extract.E.dbody d)) defs -> Forall (fun d => P (E.dbody d)) defs -> P (Extract.E.tCoFix defs i)) - (Hprim : forall p, P (Extract.E.tPrim p)): + (Hprim : forall p, primProp (erases_deps Σ Σ') p -> primProp P p -> P (Extract.E.tPrim p)): forall t, erases_deps Σ Σ' t -> P t. Proof. fix f 2. @@ -441,6 +445,15 @@ Proof. fix f' 2. intros defs []; [now constructor|]. constructor; [now apply f|now apply f']. + - eapply Hprim; tea; constructor. + - eapply Hprim; tea; constructor. + - eapply Hprim; tea; constructor. + intuition auto; solve_all. + split. auto. destruct a as [d v]. cbn in *. + eapply Forall_All. + revert v H. + fix aux 2. + intros ? [];constructor; auto. Defined. (* Lemma fresh_global_erase {cf : checker_flags} Σ Σ' kn : @@ -559,6 +572,7 @@ Proof. rewrite H in declm. eapply PCUICWeakeningEnv.lookup_env_Some_fresh in kn_fresh. eauto. eauto. + - depelim X0; intuition auto; constructor; auto. Qed. Derive Signature for erases_global_decls. @@ -681,6 +695,8 @@ Proof. depelim all_deps. destruct p as (?&?&?). now constructor; eauto. + - eapply inversion_Prim in wt as [prim_ty [decl []]]; eauto. + depelim H0; depelim p1; depelim X; cbn in *; try constructor; cbn; intuition eauto. solve_all. Qed. Lemma Forall2_nth_error_left {A B} {P} {l : list A} {l' : list B} : Forall2 P l l' -> diff --git a/erasure/theories/EEtaExpanded.v b/erasure/theories/EEtaExpanded.v index 18cccf6e2..96377e3bd 100644 --- a/erasure/theories/EEtaExpanded.v +++ b/erasure/theories/EEtaExpanded.v @@ -5,7 +5,7 @@ From Coq Require Import Utf8 Program. From MetaCoq.Utils Require Import utils. From MetaCoq.Common Require Import config Kernames EnvMap BasicAst. -From MetaCoq.Erasure Require Import EAst EAstUtils EInduction EGlobalEnv EExtends EWellformed ELiftSubst ESpineView ECSubst EWcbvEval EWcbvEvalInd EProgram. +From MetaCoq.Erasure Require Import EPrimitive EAst EAstUtils EInduction EGlobalEnv EExtends EWellformed ELiftSubst ESpineView ECSubst EWcbvEval EWcbvEvalInd EProgram. Local Open Scope string_scope. Set Asymmetric Patterns. @@ -75,7 +75,7 @@ Section isEtaExp. | tBox => true | tVar _ => true | tConst _ => true - | tPrim _ => true + | tPrim p => test_primIn p (fun x H => isEtaExp x) | tConstruct ind i block_args => isEtaExp_app ind i 0 && is_nil block_args }. Proof. all:try lia. @@ -92,11 +92,15 @@ Section isEtaExp. change (fun x => size x) with size in H. pose proof (size_mkApps_l napp nnil). lia. - eapply (In_size snd size) in H. cbn in H; lia. + - destruct p as [? []]; cbn in *; intuition eauto. + subst. lia. + eapply (In_size id size) in H0. unfold id in H0. + change (fun x => size x) with size in H0. lia. Qed. End isEtaExp. -Global Hint Rewrite @forallb_InP_spec : isEtaExp. +Global Hint Rewrite @test_primIn_spec @forallb_InP_spec : isEtaExp. Tactic Notation "simp_eta" "in" hyp(H) := simp isEtaExp in H; rewrite -?isEtaExp_equation_1 in H. Ltac simp_eta := simp isEtaExp; rewrite -?isEtaExp_equation_1. @@ -276,8 +280,8 @@ Section WeakEtaExp. eapply a0 => //. - move/andP: H0 => [] etaexp h. rewrite csubst_mkApps /=. - rewrite isEtaExp_Constructor. solve_all. - rewrite map_length. rtoProp; solve_all. solve_all. + rewrite isEtaExp_Constructor; solve_all. + rtoProp; solve_all. solve_all. now destruct block_args. - rewrite csubst_mkApps /=. move/andP: H1 => [] eu ev. @@ -417,6 +421,7 @@ Proof. - eapply In_All in H; solve_all. move/andP: b => [] -> /=. eauto. - eapply In_All in H; solve_all. + - solve_all. - eapply In_All in H; solve_all. rewrite isEtaExp_Constructor //. rtoProp; intuition auto. eapply isEtaExp_app_extends; tea. @@ -471,7 +476,7 @@ Inductive expanded : term -> Prop := #|args| >= cstr_arity mind cdecl -> Forall expanded args -> expanded (mkApps (tConstruct ind idx []) args) -| expanded_tPrim p : expanded (tPrim p) +| expanded_tPrim p : primProp expanded p -> expanded (tPrim p) | expanded_tBox : expanded tBox. End expanded. @@ -510,7 +515,7 @@ forall (Σ : global_declarations) (P : term -> Prop), (args : list term), declared_constructor Σ (ind, idx) mind idecl cdecl -> #|args| >= cstr_arity mind cdecl -> Forall (expanded Σ) args -> Forall P args -> P (mkApps (tConstruct ind idx []) args)) -> -(forall p, P (tPrim p)) -> +(forall p, primProp (expanded Σ) p -> primProp P p -> P (tPrim p)) -> (P tBox) -> forall t : term, expanded Σ t -> P t. Proof. @@ -523,6 +528,9 @@ Proof. - eapply H8; eauto. induction H13; econstructor; cbn in *; intuition eauto. - eapply H9; eauto. induction H13; econstructor; cbn in *; eauto. - eapply H10; eauto. clear - H15 f. induction H15; econstructor; cbn in *; eauto. + - eapply H11; eauto. + depelim X; constructor. destruct p; split; eauto. + eapply (make_All_All f a0). Qed. Local Hint Constructors expanded : core. @@ -595,6 +603,7 @@ Proof. intuition auto. - econstructor. rewrite forallb_InP_spec in H0. eapply forallb_Forall in H0. eapply In_All in H. solve_all. + - econstructor. rewrite test_primIn_spec in H0. solve_all. - rtoProp. eapply In_All in H. rewrite forallb_InP_spec in H2. eapply forallb_Forall in H2. eapply isEtaExp_app_expanded in H0 as (? & ? & ? & ? & ?). @@ -621,6 +630,7 @@ Proof. - rewrite isEtaExp_Constructor. rtoProp; repeat split. 2: eapply forallb_Forall. 2: solve_all. eapply expanded_isEtaExp_app_; eauto. + - solve_all. Qed. Lemma expanded_global_env_isEtaExp_env {Σ} : expanded_global_env Σ -> isEtaExp_env Σ. @@ -695,6 +705,7 @@ Proof. - simp_eta. rtoProp; intuition auto. eapply In_All in H0; solve_all. - eapply In_All in H. simp_eta; rtoProp; intuition auto. solve_all. + - simp_eta. solve_all_k 7. primProp. solve_all. - eapply In_All in H. simp_eta; rtoProp; intuition auto. rewrite EEtaExpanded.isEtaExp_Constructor. rtoProp; repeat split. eauto. solve_all. destruct block_args; cbn in *; eauto. diff --git a/erasure/theories/EEtaExpandedFix.v b/erasure/theories/EEtaExpandedFix.v index 257deaee4..3d50dab63 100644 --- a/erasure/theories/EEtaExpandedFix.v +++ b/erasure/theories/EEtaExpandedFix.v @@ -65,23 +65,6 @@ Inductive expanded (Γ : list nat): term -> Prop := End expanded. -Section make_All. - Context {A} {B : A -> Type} (f : forall x : A, B x). - - Equations make_All (l : list A) : All B l := - | [] := All_nil - | hd :: tl := All_cons (f hd) (make_All tl). -End make_All. - -Section make_All_All. - Context {A} {B : A -> Type} {C : A -> Type} (f : forall x : A, B x -> C x). - - Equations make_All_All {l : list A} (a : All B l) : All C l := - | All_nil := All_nil - | All_cons bhd btl := All_cons (f _ bhd) (make_All_All btl). - -End make_All_All. - Lemma expanded_ind : ∀ (Σ : global_declarations) (P : list nat → term → Prop), (∀ (Γ : list nat) (n : nat) (args : list term) (m : nat), @@ -287,40 +270,6 @@ Section isEtaExp. | None => false end. - Section PrimIn. - Context {term : Set}. - - Equations InPrim (x : term) (p : prim_val term) : Prop := - | x | (primInt; primIntModel i) := False - | x | (primFloat; primFloatModel _) := False - | x | (primArray; primArrayModel a) := - x = a.(array_default) \/ In x a.(array_value). - - Lemma InPrim_primProp p P : (forall x, InPrim x p -> P x) -> primProp P p. - Proof. - intros hin. - destruct p as [? []]; constructor. - split. eapply hin; eauto. now left. - cbn in hin. - induction (array_value a); constructor. - eapply hin. now right; left. eapply IHl. - intros. eapply hin. intuition eauto. now right; right. - Qed. - - Equations test_primIn (p : prim_val term) (f : forall x : term, InPrim x p -> bool) : bool := - | (primInt; primIntModel i) | _ := true - | (primFloat; primFloatModel _) | _ := true - | (primArray; primArrayModel a) | f := - f a.(array_default) (or_introl eq_refl) && forallb_InP a.(array_value) (fun x H => f x (or_intror H)). - - Lemma test_primIn_spec p (f : term -> bool) : - test_primIn p (fun x H => f x) = test_prim f p. - Proof. - funelim (test_primIn p (fun x H => f x)); cbn => //. - rewrite forallb_InP_spec //. - Qed. - End PrimIn. - Import TermSpineView. Definition is_nil {A} (l : list A) := match l with nil => true | _ => false end. @@ -489,7 +438,6 @@ Section isEtaExp. - destruct nth_error eqn:E; try easy. erewrite nth_error_app_left; eauto. - rewrite app_assoc. eapply a, b. reflexivity. - rewrite app_assoc. eapply a, b. reflexivity. - - eapply InPrim_primProp in H. solve_all. - rewrite isEtaExp_mkApps => //. cbn [expanded_head_viewc]. rtoProp; repeat solve_all. - rewrite isEtaExp_mkApps => //. cbn [expanded_head_viewc]. rtoProp; repeat solve_all. rewrite app_assoc. rtoProp; intuition auto. @@ -508,7 +456,6 @@ Section isEtaExp. - destruct block_args; cbn in *; eauto. - eapply a in b. 2: f_equal. revert b. now len. - eapply a in b. 2: f_equal. revert b. now len. - - eapply InPrim_primProp in H; solve_all. - cbn. destruct block_args; cbn in *; eauto. - cbn. solve_all. rtoProp; intuition auto. eapply a in H0. 2: reflexivity. revert H0. now len. - destruct nth_error eqn:Hn; cbn in H1; try easy. @@ -541,12 +488,14 @@ Section isEtaExp. solve_all. rewrite app_assoc. eapply a0; cbn; eauto. now len. cbn. now rewrite app_assoc. - rewrite app_assoc. eapply a0; len; eauto. now rewrite app_assoc. - - eapply InPrim_primProp in H. solve_all. - eapply primProp_map. solve_all. + - solve_all_k 6. - fold csubst. move/andP: H1 => [] etaexp h. rewrite csubst_mkApps /=. rewrite isEtaExp_Constructor. solve_all. - rewrite map_length. rtoProp; solve_all. solve_all. destruct block_args; cbn in *; eauto. + rtoProp; solve_all. destruct block_args. + 2:{ cbn in *; eauto. } + solve_all. + destruct block_args; cbn in *; eauto. - rewrite csubst_mkApps /=. move/andP : H2 => [] /andP [] eu ef ev. rewrite isEtaExp_mkApps //. @@ -600,9 +549,7 @@ Section isEtaExp. erewrite option_default_ext; eauto. f_equal. destruct i; cbn; lia. + now rewrite !nth_error_app1 in H0 |- *; try lia. - - intros. eapply forallb_All in H1; eapply All_mix in H; tea. - eapply All_forallb, All_map, All_impl; tea; cbv beta. - intros x Hx. eapply Hx; eauto. apply Hx. + - solve_all. eapply a; trea. solve_all. - eapply H with (Γ := 0 :: Γ0); cbn -[isEtaExp]; eauto. - solve_all. move/andP: H2 => [] etab etab'. simp_eta. apply/andP. split; eauto. @@ -617,14 +564,11 @@ Section isEtaExp. { cbn in Hcl. solve_all. rewrite Nat.add_0_r in a0. eauto. } now rewrite app_assoc. solve_all. - - eapply InPrim_primProp in H. solve_all. - eapply primProp_map. eapply primProp_impl; tea. intuition eauto. + - solve_all. eapply primProp_impl; tea. intuition eauto. destruct H. eapply i; eauto. solve_all. - solve_all. fold csubst. move/andP: H1 => [] etaexp h. rewrite csubst_mkApps /=. - rewrite isEtaExp_Constructor. solve_all. - rewrite map_length. rtoProp; solve_all. - rewrite forallb_map. + rewrite isEtaExp_Constructor. solve_all. rtoProp; solve_all. eapply All_forallb. clear Heq0 Heq. eapply All_impl; tea; cbv beta. intros x Hx. @@ -874,7 +818,7 @@ Proof. eapply In_All in H0. solve_all. - econstructor. rewrite forallb_InP_spec in H0. eapply forallb_Forall in H0. eapply In_All in H. solve_all. - - econstructor. eapply InPrim_primProp in H. rewrite test_primIn_spec in H0. + - econstructor. rewrite test_primIn_spec in H0. solve_all. - rtoProp. eapply In_All in H. rewrite forallb_InP_spec in H2. eapply forallb_Forall in H2. diff --git a/erasure/theories/EImplementBox.v b/erasure/theories/EImplementBox.v index 001c2dc9c..132539142 100644 --- a/erasure/theories/EImplementBox.v +++ b/erasure/theories/EImplementBox.v @@ -2,7 +2,7 @@ From Coq Require Import Utf8 Program. From MetaCoq.Utils Require Import utils. From MetaCoq.Common Require Import config Kernames BasicAst EnvMap. -From MetaCoq.Erasure Require Import EAst EAstUtils EInduction EArities +From MetaCoq.Erasure Require Import EPrimitive EAst EAstUtils EInduction EArities ELiftSubst ESpineView EGlobalEnv EWellformed EEnvMap EWcbvEval EEtaExpanded ECSubst EWcbvEvalEtaInd EProgram. @@ -54,7 +54,7 @@ Section implement_box. | tVar n => EAst.tVar n | tConst n => EAst.tConst n | tConstruct ind i block_args => EAst.tConstruct ind i (map_InP block_args (fun d H => implement_box d)) - | tPrim p => EAst.tPrim p. + | tPrim p => EAst.tPrim (map_primIn p (fun x H => implement_box x)). Proof. all:try lia. all:try apply (In_size); tea. @@ -62,6 +62,7 @@ Section implement_box. - now apply (In_size id size). - now eapply (In_size id size). - eapply (In_size snd size) in H. cbn in *. lia. + - now eapply InPrim_size in H. Qed. End Def. @@ -93,11 +94,12 @@ Section implement_box. rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; unfold test_def in *; simpl closed in *; - try solve [simpl; subst; simpl closed; f_equal; auto; rtoProp; solve_all; solve_all]; try easy. + (*try solve [simpl; subst; simpl closed; f_equal; auto; rtoProp; solve_all; solve_all]; try easy. rtoProp. split. eauto. solve_all. replace (#|x.1| + S k) with (#|x.1| + k + 1) by lia. - eapply closedn_lift. eauto. + eapply closedn_lift. eauto.*) + try solve [simpl; subst; simpl closed; f_equal; auto; rtoProp; solve_all; solve_all_k 6]; try easy. Qed. Hint Rewrite @forallb_InP_spec : isEtaExp. @@ -175,7 +177,7 @@ Section implement_box. - cbn. do 2 f_equal. 1: eauto. rewrite !map_map. solve_all. eapply In_All. intros ? ?. unfold map_def. cbn. f_equal. - setoid_rewrite -> closed_subst at 2. + setoid_rewrite -> closed_subst at 2. replace (#|x.1| + S k) with ((#|x.1| + k) + 1) by lia. rewrite <- commut_lift_subst_rec. 2: lia. rewrite <- closed_subst. @@ -183,11 +185,10 @@ Section implement_box. f_equal. eapply H; eauto. - cbn. f_equal. rewrite !map_map. solve_all. - eapply In_All. intros ? ?. unfold map_def. cbn. f_equal. erewrite H; eauto. - f_equal. now rewrite map_length. + eapply In_All. intros ? ?. unfold map_def. cbn. f_equal. rewrite H; eauto. - cbn. f_equal. rewrite !map_map. solve_all. - eapply In_All. intros ? ?. unfold map_def. cbn. f_equal. erewrite H; eauto. - f_equal. now rewrite map_length. + eapply In_All. intros ? ?. unfold map_def. cbn. f_equal. rewrite H; eauto. + - cbn. solve_all. Qed. Lemma implement_box_substl s t : @@ -252,7 +253,7 @@ Section implement_box. unfold cunfold_cofix. rewrite nth_error_map. destruct nth_error eqn:heq. - intros [= <- <-] => /=. f_equal. + intros [= <- <-] => /=. f_equal. rewrite implement_box_substl //. 2:congruence. f_equal. f_equal. apply implement_box_cofix_subst. Qed. @@ -468,11 +469,10 @@ Proof. eapply wellformed_lift. eauto. - rewrite lookup_constructor_implement_box. intuition auto. - unfold wf_fix in *. rtoProp. solve_all. solve_all. now eapply isLambda_implement_box. - - unfold wf_fix in *. rtoProp. solve_all. - len. solve_all. len. destruct x. - cbn -[implement_box isEtaExp] in *. rtoProp. eauto. + - unfold wf_fix in *. rtoProp. solve_all. len. solve_all. - unfold wf_fix in *. len. solve_all. rtoProp; intuition auto. solve_all. + - solve_all_k 6. Qed. Lemma transform_wellformed_decl' {efl : EEnvFlags} {Σ : global_declarations} d : diff --git a/erasure/theories/EInduction.v b/erasure/theories/EInduction.v index 19f8929ad..208332ffc 100644 --- a/erasure/theories/EInduction.v +++ b/erasure/theories/EInduction.v @@ -124,6 +124,14 @@ Proof. rewrite IHl. simpl. lia. Qed. +Lemma InPrim_size x p : InPrim x p -> size x < S (prim_size size p). +Proof. + destruct p as [? []]; cbn => //. + intros [->|H]; try lia. + eapply (In_size id size) in H; unfold id in H. + change (fun x => size x) with size in H. lia. +Qed. + Lemma decompose_app_rec_size t l : let da := decompose_app_rec t l in size da.1 + list_size size da.2 = size t + list_size size l. diff --git a/erasure/theories/EInlineProjections.v b/erasure/theories/EInlineProjections.v index 0eaa0b364..e9cee9812 100644 --- a/erasure/theories/EInlineProjections.v +++ b/erasure/theories/EInlineProjections.v @@ -2,7 +2,7 @@ From Coq Require Import Utf8 Program. From MetaCoq.Utils Require Import utils. From MetaCoq.Common Require Import config Kernames BasicAst. -From MetaCoq.Erasure Require Import EAst EAstUtils EExtends +From MetaCoq.Erasure Require Import EPrimitive EAst EAstUtils EExtends ELiftSubst ECSubst EGlobalEnv EWellformed EWcbvEval Extract EEnvMap EArities EProgram. @@ -123,7 +123,7 @@ Section optimize. | tVar _ => t | tConst _ => t | tConstruct ind n args => tConstruct ind n (map optimize args) - | tPrim _ => t + | tPrim p => tPrim (map_prim optimize p) end. Lemma optimize_mkApps f l : optimize (mkApps f l) = mkApps (optimize f) (map optimize l). @@ -163,9 +163,10 @@ Section optimize. rewrite hl /= andb_true_r. rewrite IHt //=; len. apply Nat.ltb_lt. lia. - - len. rtoProp; solve_all. rewrite forallb_map; solve_all. + - len. rtoProp; solve_all. solve_all. now eapply isLambda_optimize. solve_all. - - len. rtoProp; solve_all. rewrite forallb_map; solve_all. + - len. rtoProp; repeat solve_all. + - rewrite test_prim_map. solve_all. Qed. Lemma optimize_csubst {a k b} n : @@ -787,10 +788,9 @@ Proof. rewrite hrel IHt //= andb_true_r. have hargs' := wellformed_projection_args wfΣ hl'. apply Nat.ltb_lt. len. - - cbn. unfold wf_fix; rtoProp; intuition auto; solve_all. now eapply isLambda_optimize. now len. - unfold test_def in *. len. eauto. - - cbn. unfold wf_fix; rtoProp; intuition auto; solve_all. now len. - unfold test_def in *. len. eauto. + - cbn. unfold wf_fix; rtoProp; intuition auto; solve_all. now eapply isLambda_optimize. + - cbn. unfold wf_fix; rtoProp; intuition auto; solve_all. + - cbn. rtoProp; intuition eauto; solve_all_k 6. Qed. Import EWellformed. diff --git a/erasure/theories/ELiftSubst.v b/erasure/theories/ELiftSubst.v index 83ad52ef7..a03c221eb 100644 --- a/erasure/theories/ELiftSubst.v +++ b/erasure/theories/ELiftSubst.v @@ -213,6 +213,7 @@ Ltac change_Sk := Ltac primProp := repeat match goal with + | [ H : (forall x, InPrim x _ -> _) |- _ ] => eapply InPrim_primProp in H | [ H : is_true (test_prim _ _) |- _ ] => eapply test_prim_impl_primProp in H | [ H : test_prim _ _ = true |- _ ] => eapply test_prim_impl_primProp in H | [ |- is_true (test_prim _ _) ] => eapply primProp_impl_test_prim @@ -223,7 +224,7 @@ Ltac primProp := Tactic Notation "solve_all_k" int(k) := unfold tFixProp in *; - primProp; + primProp; autorewrite with map; repeat toAll; try All_map; try close_Forall; change_Sk; auto with all; intuition eauto k with all. diff --git a/erasure/theories/EOptimizePropDiscr.v b/erasure/theories/EOptimizePropDiscr.v index 38cc0bbf5..c713427bd 100644 --- a/erasure/theories/EOptimizePropDiscr.v +++ b/erasure/theories/EOptimizePropDiscr.v @@ -7,7 +7,7 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICTyping PCUICInversion PCUICSafeLemmata. (* for welltyped *) From MetaCoq.SafeChecker Require Import PCUICWfEnvImpl. -From MetaCoq.Erasure Require Import EAst EAstUtils EDeps EExtends +From MetaCoq.Erasure Require Import EPrimitive EAst EAstUtils EDeps EExtends ELiftSubst ECSubst EGlobalEnv EWellformed EWcbvEval Extract Prelim EEnvMap EArities EProgram. @@ -68,7 +68,7 @@ Section remove_match_on_box. | tVar _ => t | tConst _ => t | tConstruct ind i args => tConstruct ind i (map remove_match_on_box args) - | tPrim _ => t + | tPrim p => tPrim (map_prim remove_match_on_box p) end. Lemma remove_match_on_box_mkApps f l : remove_match_on_box (mkApps f l) = mkApps (remove_match_on_box f) (map remove_match_on_box l). @@ -103,7 +103,6 @@ Section remove_match_on_box. - move/andP => []. intros. f_equal; solve_all; eauto. - move/andP => []. intros. f_equal; solve_all; eauto. - move/andP => []. intros. f_equal; solve_all; eauto. - destruct x0; cbn in *. f_equal; auto. Qed. Lemma closed_remove_match_on_box t k : closedn k t -> closedn k (remove_match_on_box t). @@ -128,6 +127,7 @@ Section remove_match_on_box. rtoProp; solve_all. solve_all. rtoProp; solve_all. solve_all. - destruct GlobalContextMap.inductive_isprop_and_pars as [[[|] _]|]; cbn; auto. + - solve_all_k 6. Qed. Lemma subst_csubst_comm l t k b : @@ -818,10 +818,9 @@ Proof. - cbn -[GlobalContextMap.inductive_isprop_and_pars lookup_inductive]. move/andP => [] /andP[]hasc hs ht. destruct GlobalContextMap.inductive_isprop_and_pars as [[[|] _]|] => /= //. all:rewrite hasc hs /=; eauto. - - cbn. unfold wf_fix; rtoProp; intuition auto; solve_all. now eapply isLambda_remove_match_on_box. now len. - unfold test_def in *. len. eauto. - - cbn. unfold wf_fix; rtoProp; intuition auto; solve_all. now len. - unfold test_def in *. len. eauto. + - cbn. unfold wf_fix; rtoProp; intuition auto; solve_all. now eapply isLambda_remove_match_on_box. + - cbn. unfold wf_fix; rtoProp; intuition auto; solve_all. + - cbn. rtoProp; intuition auto; solve_all_k 6. Qed. Import EWellformed. diff --git a/erasure/theories/EPrimitive.v b/erasure/theories/EPrimitive.v index 866b3c414..736744df1 100644 --- a/erasure/theories/EPrimitive.v +++ b/erasure/theories/EPrimitive.v @@ -32,7 +32,7 @@ Arguments primIntModel {term}. Arguments primFloatModel {term}. Arguments primArrayModel {term}. -Derive Signature NoConfusion for prim_model. +Derive Signature NoConfusion NoConfusionHom for prim_model. Definition prim_model_of (term : Set) (p : prim_tag) : Set := match p with @@ -165,6 +165,13 @@ Inductive primProp {term} P : prim_val term -> Type := primProp P (primArray; primArrayModel a). Derive Signature NoConfusion for primProp. +Definition fold_prim {A term} (f : A -> term -> A) (p : prim_val term) (acc : A) : A := + match p.π2 return A with + | primIntModel f => acc + | primFloatModel f => acc + | primArrayModel a => fold_left f a.(array_value) (f acc a.(array_default)) + end. + Section primtheory. Context {term term' term''} (g : term' -> term'') (f : term -> term') (p : prim_val term). @@ -175,7 +182,7 @@ Section primtheory. do 2 f_equal. destruct a; rewrite /map_array_model //= map_map_compose //. Qed. End primtheory. -#[global] Hint Rewrite @map_prim_compose : all. +#[global] Hint Rewrite @map_prim_compose : map all. Lemma reflectT_forallb {A} {P : A -> Type} {p l} : (forall x, reflectT (P x) (p x)) -> @@ -312,3 +319,58 @@ End primtheory. #[global] Hint Resolve map_prim_eq map_prim_eq_prop : all. #[global] Hint Resolve primProp_impl primProp_map : all. #[global] Hint Resolve test_prim_eq_prop : all. + +Set Equations Transparent. + +Section PrimIn. + Context {term : Set}. + + Equations InPrim (x : term) (p : prim_val term) : Prop := + | x | (primInt; primIntModel i) := False + | x | (primFloat; primFloatModel _) := False + | x | (primArray; primArrayModel a) := + x = a.(array_default) \/ In x a.(array_value). + + Lemma InPrim_primProp p P : (forall x, InPrim x p -> P x) -> primProp P p. + Proof. + intros hin. + destruct p as [? []]; constructor. + split. eapply hin; eauto. now left. + cbn in hin. + induction (array_value a); constructor. + eapply hin. now right; left. eapply IHl. + intros. eapply hin. intuition eauto. now right; right. + Qed. + + Equations test_primIn (p : prim_val term) (f : forall x : term, InPrim x p -> bool) : bool := + | (primInt; primIntModel i) | _ := true + | (primFloat; primFloatModel _) | _ := true + | (primArray; primArrayModel a) | f := + f a.(array_default) (or_introl eq_refl) && forallb_InP a.(array_value) (fun x H => f x (or_intror H)). + + Lemma test_primIn_spec p (f : term -> bool) : + test_primIn p (fun x H => f x) = test_prim f p. + Proof. + funelim (test_primIn p (fun x H => f x)); cbn => //. + rewrite forallb_InP_spec //. + Qed. + + Equations map_primIn (p : prim_val term) (f : forall x : term, InPrim x p -> term) : prim_val term := + | (primInt; primIntModel i) | _ := (primInt; primIntModel i) + | (primFloat; primFloatModel f) | _ := (primFloat; primFloatModel f) + | (primArray; primArrayModel a) | f := + (primArray; primArrayModel + {| array_default := f a.(array_default) (or_introl eq_refl); + array_value := map_InP a.(array_value) (fun x H => f x (or_intror H)) |}). + + Lemma map_primIn_spec p f : + map_primIn p (fun x _ => f x) = map_prim f p. + Proof. + funelim (map_primIn p _); cbn => //. + do 2 f_equal. unfold map_array_model; destruct a => //. + rewrite map_InP_spec //. + Qed. + +End PrimIn. + +#[global] Hint Rewrite @map_primIn_spec : map. \ No newline at end of file diff --git a/erasure/theories/ERemoveParams.v b/erasure/theories/ERemoveParams.v index bf4d33e52..42e43f62b 100644 --- a/erasure/theories/ERemoveParams.v +++ b/erasure/theories/ERemoveParams.v @@ -1,8 +1,8 @@ (* Distributed under the terms of the MIT license. *) From Coq Require Import Utf8 Program. From MetaCoq.Utils Require Import utils. -From MetaCoq.Common Require Import config Kernames BasicAst EnvMap. -From MetaCoq.Erasure Require Import EAst EAstUtils EInduction EArities +From MetaCoq.Common Require Import config Kernames Primitive BasicAst EnvMap. +From MetaCoq.Erasure Require Import EPrimitive EAst EAstUtils EInduction EArities ELiftSubst ESpineView EGlobalEnv EWellformed EEnvMap EWcbvEval EEtaExpanded ECSubst EWcbvEvalEtaInd EProgram. @@ -59,7 +59,7 @@ Section strip. | tVar n => EAst.tVar n | tConst n => EAst.tConst n | tConstruct ind i block_args => EAst.tConstruct ind i block_args - | tPrim p => EAst.tPrim p }. + | tPrim p => EAst.tPrim (map_primIn p (fun x H => strip x)) }. Proof. all:try lia. all:try apply (In_size); tea. @@ -72,10 +72,14 @@ Section strip. - pose proof (size_mkApps_l napp nnil). eapply (In_size id size) in H. change (fun x => size (id x)) with size in H. unfold id in H. lia. - eapply (In_size snd size) in H. cbn in H; lia. + - destruct p as [? []]; cbn in H; eauto. + intuition auto; subst; cbn; try lia. + eapply (In_size id size) in H0. unfold id in H0. + change (fun x => size x) with size in H0. lia. Qed. End Def. - Hint Rewrite @map_InP_spec : strip. + Hint Rewrite @map_primIn_spec @map_InP_spec : strip. Lemma map_repeat {A B} (f : A -> B) x n : map f (repeat x n) = repeat (f x) n. Proof using Type. @@ -107,6 +111,7 @@ Section strip. unfold test_def in *; simpl closed in *; try solve [simpl; subst; simpl closed; f_equal; auto; rtoProp; solve_all; solve_all]; try easy. + - solve_all_k 6. - rewrite !closedn_mkApps in H1 *. rtoProp; intuition auto. solve_all. @@ -557,6 +562,10 @@ Module Fast. | app, tConstruct kn c block_args with GlobalContextMap.lookup_inductive_pars Σ (inductive_mind kn) := { | Some npars => mkApps (EAst.tConstruct kn c block_args) (List.skipn npars app) | None => mkApps (EAst.tConstruct kn c block_args) app } + | app, tPrim (primInt; primIntModel i) => mkApps (tPrim (primInt; primIntModel i)) app + | app, tPrim (primFloat; primFloatModel f) => mkApps (tPrim (primFloat; primFloatModel f)) app + | app, tPrim (primArray; primArrayModel a) => + mkApps (tPrim (primArray; primArrayModel {| array_default := strip [] a.(array_default); array_value := strip_args a.(array_value) |})) app | app, x => mkApps x app } where strip_args (t : list term) : list term := @@ -597,6 +606,11 @@ Module Fast. specialize (IHv [] eq_refl). simpl in IHv. intros args ->. specialize (IHu (v :: args)). forward IHu. now rewrite -IHv. exact IHu. + - intros Hl hargs args ->. + rewrite strip_mkApps //=. simp_strip. + rewrite map_primIn_spec. f_equal. f_equal. cbn. + do 2 f_equal. rewrite /map_array_model. + specialize (Hl [] eq_refl). f_equal; eauto. - intros Hl args ->. rewrite strip_mkApps // /=. rewrite GlobalContextMap.lookup_inductive_pars_spec in Hl. now rewrite Hl. @@ -1072,7 +1086,7 @@ Proof. all: eapply H; auto. - unfold wf_fix_gen in *. rewrite map_length. rtoProp; intuition auto. toAll; solve_all. now rewrite -strip_isLambda. toAll; solve_all. - - unfold wf_fix in *. rewrite map_length; rtoProp; intuition auto. toAll; solve_all. + - rewrite map_primIn_spec. primProp. rtoProp; intuition eauto; solve_all_k 6. - move:H1; rewrite !wellformed_mkApps //. rtoProp; intuition auto. toAll; solve_all. toAll; solve_all. - move:H0; rewrite !wellformed_mkApps //. rtoProp; intuition auto. diff --git a/erasure/theories/ESubstitution.v b/erasure/theories/ESubstitution.v index af58c8a04..fb17df33d 100644 --- a/erasure/theories/ESubstitution.v +++ b/erasure/theories/ESubstitution.v @@ -113,10 +113,10 @@ Proof. split; eauto. - econstructor. induction H2; constructor. - induction H; constructor; depelim X2; eauto. + induction X3; constructor; depelim X2; eauto. depelim X1. - eapply All2_All_mix_left in X3; eauto. - eapply All2_impl. exact X3. + eapply All2_All_mix_left in a0; eauto. + eapply All2_impl. exact a0. cbn. intros ? ? [? ?]. eauto. Qed. @@ -315,6 +315,12 @@ Proof. unfold app_context in IH. rewrite <- !app_assoc in IH. rewrite (All2_length X3) in IH |- *. apply IH. apply IH'. + + - econstructor. depelim H3. + depelim X4; repeat constructor. + depelim X2; cbn. now eapply hdef. + depelim X2. cbn. eapply All2_map. + ELiftSubst.solve_all. Qed. Lemma erases_weakening (Σ : global_env_ext) (Γ Γ' : context) (t T : PCUICAst.term) t' : @@ -581,10 +587,10 @@ Proof. eapply is_type_subst; eauto. - cbn. depelim H1. * cbn; constructor. - depelim H1. depelim H1; repeat constructor. - depelim X2. - + eapply (@substitution _ Σ _ Γ Γ' s Δ) in H1. + depelim H1. depelim X5; depelim X2; repeat constructor; cbn; eauto. + ELiftSubst.solve_all. * constructor. eapply is_type_subst in X3; tea. + now cbn in X3. - eapply H; eauto. Qed. diff --git a/erasure/theories/EWcbvEvalCstrsAsBlocksFixLambdaInd.v b/erasure/theories/EWcbvEvalCstrsAsBlocksFixLambdaInd.v index 1ac8603df..d9a4ac6d7 100644 --- a/erasure/theories/EWcbvEvalCstrsAsBlocksFixLambdaInd.v +++ b/erasure/theories/EWcbvEvalCstrsAsBlocksFixLambdaInd.v @@ -2,7 +2,7 @@ From Coq Require Import Utf8 Program ssreflect ssrbool. From MetaCoq.Utils Require Import utils. From MetaCoq.Common Require Import config Kernames BasicAst EnvMap. -From MetaCoq.Erasure Require Import EAst EAstUtils EInduction ELiftSubst EWcbvEval EGlobalEnv +From MetaCoq.Erasure Require Import EPrimitive EAst EAstUtils EInduction ELiftSubst EWcbvEval EGlobalEnv EWellformed ECSubst EInduction EWcbvEvalInd EEtaExpanded. Set Asymmetric Patterns. @@ -42,7 +42,8 @@ Section OnSubterm. All (fun br => Q (#|br.1| + n) br.2) brs -> on_subterms Q n (tCase ci discr brs) | on_proj p c : has_tProj -> Q n c -> on_subterms Q n (tProj p c) | on_fix mfix idx : has_tFix -> All (fun d => Q (#|mfix| + n) d.(dbody)) mfix -> on_subterms Q n (tFix mfix idx) - | on_cofix mfix idx : has_tCoFix -> All (fun d => Q (#|mfix| + n) d.(dbody)) mfix -> on_subterms Q n (tCoFix mfix idx). + | on_cofix mfix idx : has_tCoFix -> All (fun d => Q (#|mfix| + n) d.(dbody)) mfix -> on_subterms Q n (tCoFix mfix idx) + | on_prim p : has_tPrim -> primProp (Q n) p -> on_subterms Q n (tPrim p). Derive Signature for on_subterms. End OnSubterm. @@ -551,7 +552,7 @@ Proof. - eapply Qatom; cbn; auto. - eapply Qatom; cbn; auto. - eapply Qatom; cbn; auto. - Unshelve. all: repeat econstructor. + - eapply Qatom; cbn; eauto. Qed. Lemma Qpreserves_wellformed (efl : EEnvFlags) Σ : wf_glob Σ -> Qpreserves (fun n x => wellformed Σ n x) Σ. @@ -576,7 +577,8 @@ Proof. eapply on_fix => //. move/andP: H0 => [] _ ha. solve_all. rtoProp; intuition auto. eapply on_cofix => //. move/andP: H0 => [] _ ha. solve_all. - - red. intros kn decl. + move/andP: H => [] hp ha. eapply on_prim => //. solve_all. + - red. intros kn decl. move/(lookup_env_wellformed clΣ). unfold wf_global_decl. destruct cst_body => //. - red. move=> hasapp n t args. rewrite wellformed_mkApps //. diff --git a/erasure/theories/EWcbvEvalCstrsAsBlocksInd.v b/erasure/theories/EWcbvEvalCstrsAsBlocksInd.v index f7b7d4325..4aa2b1de7 100644 --- a/erasure/theories/EWcbvEvalCstrsAsBlocksInd.v +++ b/erasure/theories/EWcbvEvalCstrsAsBlocksInd.v @@ -2,7 +2,7 @@ From Coq Require Import Utf8 Program ssreflect ssrbool. From MetaCoq.Utils Require Import utils. From MetaCoq.Common Require Import config Kernames BasicAst EnvMap. -From MetaCoq.Erasure Require Import EAst EAstUtils EInduction ELiftSubst EWcbvEval EGlobalEnv +From MetaCoq.Erasure Require Import EPrimitive EAst EAstUtils EInduction ELiftSubst EWcbvEval EGlobalEnv EWellformed ECSubst EInduction EWcbvEvalInd EEtaExpanded. Set Asymmetric Patterns. @@ -42,7 +42,8 @@ Section OnSubterm. All (fun br => Q (#|br.1| + n) br.2) brs -> on_subterms Q n (tCase ci discr brs) | on_proj p c : has_tProj -> Q n c -> on_subterms Q n (tProj p c) | on_fix mfix idx : has_tFix -> All (fun d => Q (#|mfix| + n) d.(dbody)) mfix -> on_subterms Q n (tFix mfix idx) - | on_cofix mfix idx : has_tCoFix -> All (fun d => Q (#|mfix| + n) d.(dbody)) mfix -> on_subterms Q n (tCoFix mfix idx). + | on_cofix mfix idx : has_tCoFix -> All (fun d => Q (#|mfix| + n) d.(dbody)) mfix -> on_subterms Q n (tCoFix mfix idx) + | on_prim p : has_tPrim -> primProp (Q n) p -> on_subterms Q n (tPrim p). Derive Signature for on_subterms. End OnSubterm. @@ -470,6 +471,7 @@ Proof. eapply on_fix => //. move/andP: H0 => [] _ ha. solve_all. rtoProp; intuition auto. eapply on_cofix => //. move/andP: H0 => [] _ ha. solve_all. + move/andP: H => [] hasp ht. eapply on_prim => //. solve_all. - red. intros kn decl. move/(lookup_env_wellformed clΣ). unfold wf_global_decl. destruct cst_body => //. diff --git a/erasure/theories/EWcbvEvalEtaInd.v b/erasure/theories/EWcbvEvalEtaInd.v index e30f2baf9..3dbd3814f 100644 --- a/erasure/theories/EWcbvEvalEtaInd.v +++ b/erasure/theories/EWcbvEvalEtaInd.v @@ -2,7 +2,7 @@ From Coq Require Import Utf8 Program ssreflect ssrbool. From MetaCoq.Utils Require Import utils. From MetaCoq.Common Require Import config Kernames BasicAst EnvMap. -From MetaCoq.Erasure Require Import EAst EAstUtils EInduction EWcbvEval EGlobalEnv +From MetaCoq.Erasure Require Import EPrimitive EAst EAstUtils EInduction EWcbvEval EGlobalEnv EWellformed ECSubst EInduction EWcbvEvalInd EEtaExpanded. Set Asymmetric Patterns. @@ -42,7 +42,8 @@ Section OnSubterm. All (fun br => Q (#|br.1| + n) br.2) brs -> on_subterms Q n (tCase ci discr brs) | on_proj p c : has_tProj -> Q n c -> on_subterms Q n (tProj p c) | on_fix mfix idx : has_tFix -> All (fun d => Q (#|mfix| + n) d.(dbody)) mfix -> on_subterms Q n (tFix mfix idx) - | on_cofix mfix idx : has_tCoFix -> All (fun d => Q (#|mfix| + n) d.(dbody)) mfix -> on_subterms Q n (tCoFix mfix idx). + | on_cofix mfix idx : has_tCoFix -> All (fun d => Q (#|mfix| + n) d.(dbody)) mfix -> on_subterms Q n (tCoFix mfix idx) + | on_prim p : has_tPrim -> primProp (Q n) p -> on_subterms Q n (tPrim p). Derive Signature for on_subterms. End OnSubterm. @@ -727,6 +728,7 @@ Proof. eapply on_proj; auto. eapply on_fix; eauto. move/andP: H0 => [] _ wf. solve_all. eapply on_cofix; eauto. move/andP: H0 => [] _ wf. solve_all. + eapply on_prim; eauto. solve_all. - red. intros kn decl. move/(lookup_env_wellformed clΣ). unfold wf_global_decl. destruct cst_body => //. diff --git a/erasure/theories/EWellformed.v b/erasure/theories/EWellformed.v index 034bd6334..6383c28ae 100644 --- a/erasure/theories/EWellformed.v +++ b/erasure/theories/EWellformed.v @@ -221,11 +221,6 @@ Section EEnvFlags. apply Nat.ltb_lt in H1. lia. - destruct cstr_as_blocks; eauto. destruct lookup_constructor_pars_args as [ [] | ]; rtoProp; repeat solve_all. destruct args; firstorder. - - solve_all. rewrite Nat.add_assoc. eauto. - - len. move/andP: H1 => [] -> ha. cbn. solve_all. - rewrite Nat.add_assoc; eauto. - - len. move/andP: H1 => [] -> ha. cbn. solve_all. - rewrite Nat.add_assoc; eauto. Qed. Lemma wellformed_subst_eq {s k k' t} : @@ -265,10 +260,8 @@ Section EEnvFlags. rewrite !Nat.add_assoc. eapply a => //. now rewrite !Nat.add_assoc in b. - destruct (dbody x) => //. - - intros. now len. - specialize (a (#|m| + k')). len. now rewrite !Nat.add_assoc !(Nat.add_comm k) in a, b0 |- *. - - intros. now len. - specialize (a (#|m| + k')); len. now rewrite !Nat.add_assoc !(Nat.add_comm k) in a, b |- *. Qed. diff --git a/erasure/theories/ErasureCorrectness.v b/erasure/theories/ErasureCorrectness.v index 8a3c835ba..df3a996f9 100644 --- a/erasure/theories/ErasureCorrectness.v +++ b/erasure/theories/ErasureCorrectness.v @@ -1379,4 +1379,11 @@ Proof. + constructor => //. eapply erases_deps_mkApps_inv in etaΣ as []. solve_all. + - intros Γ0 v etaΣ er. + depelim er; eauto. depelim H1. + depelim H0. 1-2:depelim X; repeat constructor. + depelim X0. eapply expanded_tPrim. constructor; split => //; cbn. + eapply H0; tea. now depelim etaΣ; cbn in *. + eapply Forall_All. depelim etaΣ. cbn in *. + solve_all. Qed. diff --git a/erasure/theories/ErasureFunction.v b/erasure/theories/ErasureFunction.v index 75ca06fa8..e579925c5 100644 --- a/erasure/theories/ErasureFunction.v +++ b/erasure/theories/ErasureFunction.v @@ -11,7 +11,7 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICPrimitive PCUICValidity PCUICPrincipality PCUICElimination PCUICOnFreeVars PCUICWellScopedCumulativity PCUICSN PCUICEtaExpand. From MetaCoq.SafeChecker Require Import PCUICErrors PCUICWfEnv PCUICSafeReduce PCUICSafeRetyping PCUICRetypingEnvIrrelevance. -From MetaCoq.Erasure Require Import EAstUtils EArities Extract Prelim EDeps ErasureProperties ErasureCorrectness. +From MetaCoq.Erasure Require Import EPrimitive EAstUtils ELiftSubst EArities Extract Prelim EDeps ErasureProperties ErasureCorrectness. Local Open Scope string_scope. Set Asymmetric Patterns. @@ -112,6 +112,18 @@ Proof. eapply All_In in X as [X]; tea. specialize (X _ _ H0). intuition auto; try knset. cbn in H0. apply KernameSet.union_spec in X as []; [knset|]. right; right. now exists cs. + - intros kn. + funelim (prim_global_deps term_global_deps p); cbn; simp prim_global_deps; try knset. + rewrite !knset_in_fold_left !KernameSet.union_spec; cbn. + dependent elimination X as [EPrimitive.primPropArray (pair s al)]. + rewrite knset_in_fold_left. + intuition eauto. + * eapply s in H0. rewrite KernameSet.union_spec in H0. intuition eauto. + * destruct H0 as [a [ha hkn]]. + eapply in_map_iff in ha as [x [heq hx]]. subst a. + eapply All_In in al as [al]; tea. + eapply al in hkn. + rewrite KernameSet.union_spec in hkn. destruct hkn; eauto. Qed. @@ -178,6 +190,10 @@ Proof. unfold EWellformed.wf_fix_gen in H1. move/andP: H1 => [_ hm]. solve_all. eapply All_In in hm as [hm]; tea. intuition auto. eapply (a0 (#|m| + k)) => //. + - primProp. + move: H0; funelim (prim_global_deps term_global_deps p); try knset. + rewrite knset_in_fold_left. depelim H2. intuition eauto. + destruct H1 as [ar []]. eapply All_In in b as [[]]; tea. eauto. Qed. Lemma knset_mem_spec kn s : reflect (KernameSet.In kn s) (KernameSet.mem kn s). @@ -998,8 +1014,15 @@ Section Erase. let Γ' := (fix_context mfix ++ Γ)%list in let mfix' := erase_cofix Γ' mfix _ in E.tCoFix mfix' n - | tPrim p := E.tPrim (erase_prim_val p) } + | tPrim (primInt; PCUICPrimitive.primIntModel i) := E.tPrim (primInt; EPrimitive.primIntModel i) ; + | tPrim (primFloat; PCUICPrimitive.primFloatModel f) := E.tPrim (primFloat; EPrimitive.primFloatModel f) ; + | tPrim (primArray; PCUICPrimitive.primArrayModel a) := + E.tPrim (primArray; EPrimitive.primArrayModel + {| EPrimitive.array_default := erase Γ a.(PCUICPrimitive.array_default) _; + EPrimitive.array_value := erase_terms Γ a.(PCUICPrimitive.array_value) _ |}) + } } } + where erase_terms (Γ : context) (l : list term) (Hl : forall Σ : global_env_ext, abstract_env_ext_rel X Σ -> ∥ All (welltyped Σ Γ) l ∥) : list E.term := { erase_terms Γ [] _ := []; erase_terms Γ (t :: ts) _ := @@ -1084,6 +1107,10 @@ Section Erase. solve_all. now eexists. - eapply inversion_CoFix in Ht as (? & ? & ? & ? & ? & ?); auto. sq. eapply All_impl; tea. cbn. intros d Ht; now eexists. + - eapply inversion_Prim in Ht as [prim_ty [decl []]]; eauto. + depelim p0. now eexists. + - eapply inversion_Prim in Ht as [prim_ty [decl []]]; eauto. + depelim p0. sq. solve_all. now eexists. - sq. now depelim Hl. - sq. now depelim Hl. - sq. now depelim Hbrs. @@ -1137,9 +1164,9 @@ Proof. Qed. Lemma erase_terms_eq X_type X {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} Γ ts wt : - erase_terms X_type X Γ ts wt = map_All (erase X_type X Γ) ts wt. + erase_terms X_type X Γ ts wt = All_Forall.map_All (erase X_type X Γ) ts wt. Proof. - funelim (map_All (erase X_type X Γ) ts wt); cbn; auto. + funelim (All_Forall.map_All (erase X_type X Γ) ts wt); cbn; auto. f_equal => //. apply erase_irrel. rewrite -H. eapply erase_irrel. Qed. @@ -1238,6 +1265,9 @@ Proof. intros isp. eapply isErasable_Proof in isp. eapply H'; intros. now rewrite (abstract_env_ext_irr _ H0 H). + - repeat constructor. + - repeat constructor. + - repeat constructor; eauto. - cbn. pose proof (abstract_env_ext_wf _ H) as [wf]. pose proof Hmfix as Hmfix'. specialize (Hmfix' _ H). destruct Hmfix'. @@ -1380,23 +1410,6 @@ Ltac iserasableb_irrel_env := generalize dependent H; rewrite (@is_erasableb_irrel_global_env _ _ _ _ _ _ _ _ wt wt') //; intros; rewrite Heq end. - (* - (forall Σ Σ' : global_env_ext, abstract_env_ext_rel X Σ -> - abstract_env_ext_rel X' Σ' -> ∥ extends_decls Σ' Σ ∥ /\ (Σ.2 = Σ'.2)) -> - - assert (hl : Hlookup X_type X X_type' X'). - { red. intros Σ H Σ' H0. specialize (ext _ _ H H0) as [[X0] H1]. - split. intros kn decl decl' H2 H3. - rewrite -(abstract_env_lookup_correct' _ _ H). - rewrite -(abstract_env_lookup_correct' _ _ H0). - rewrite H2 H3. - red in X0. specialize (proj1 X0 kn decl' H3). congruence. - destruct X0. intros tag. - rewrite (abstract_primitive_constant_correct _ _ _ H). - rewrite (abstract_primitive_constant_correct _ _ _ H0). - now symmetry. } -*) - Lemma erase_irrel_global_env {X_type X} {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} {X_type' X'} {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} {Γ t wt wt'} : forall (hl : Hlookup X_type X X_type' X'), erase X_type X Γ t wt = erase X_type' X' Γ t wt'. diff --git a/erasure/theories/ErasureFunctionProperties.v b/erasure/theories/ErasureFunctionProperties.v index 533e5104a..e617fc6fd 100644 --- a/erasure/theories/ErasureFunctionProperties.v +++ b/erasure/theories/ErasureFunctionProperties.v @@ -12,7 +12,7 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICPrimitive PCUICFirstorder. From MetaCoq.SafeChecker Require Import PCUICErrors PCUICWfEnv PCUICSafeReduce PCUICSafeRetyping PCUICRetypingEnvIrrelevance. -From MetaCoq.Erasure Require Import EAstUtils EArities Extract Prelim EDeps ErasureProperties ErasureCorrectness ErasureFunction. +From MetaCoq.Erasure Require Import EPrimitive EAstUtils ELiftSubst EArities Extract Prelim EDeps ErasureProperties ErasureCorrectness ErasureFunction. Local Open Scope string_scope. Set Asymmetric Patterns. @@ -152,6 +152,16 @@ Proof. eapply All2_In_right in H; eauto. destruct H as [[def [Hty Hdef]]]. eapply Hdef; eauto. + + - eapply inversion_Prim in wt as [prim_ty [decl []]]; eauto. + move: hin. + funelim (prim_global_deps _ _); try knset. + rewrite knset_in_fold_left. depelim H0; cbn in *. + depelim X. depelim p1. depelim H. depelim X. + intuition eauto. + destruct H as [a [hin inkn]]. cbn in hin. + eapply All2_All_mix_left in a2; tea. + eapply All2_In_right in a2 as [[x []]]; tea. eapply d1; tea. Qed. Global Remove Hints erases_deps_eval : core. @@ -289,7 +299,18 @@ Proof. eapply In_Forall in Σer. eapply Forall_All in Σer. eapply Forall2_All2 in H. - ELiftSubst.solve_all. Unshelve. + ELiftSubst.solve_all. + - eapply inversion_Prim in wt as [prim_ty [decl []]]; eauto. + depelim H; depelim H0. depelim X; depelim X0; depelim p1; constructor; + noconf H; cbn; simp prim_global_deps in Σer; simpl in *. + eapply e0; eauto. + red in Σer. intros kn; specialize (Σer kn). + rewrite knset_in_fold_left in Σer. intuition eauto. cbn. + eapply All_Forall. + eapply includes_deps_fold in Σer as [? Σer]. + eapply In_Forall in Σer. + eapply Forall_All in Σer. + solve_all. Qed. Lemma erases_deps_weaken kn d (Σ : global_env) (Σ' : EAst.global_declarations) t : @@ -343,6 +364,8 @@ Proof. red in H |- *. rewrite -H. simpl. unfold lookup_env; simpl; destruct (eqb_spec (inductive_mind p.(proj_ind)) kn); try congruence. eapply lookup_env_Some_fresh in H. subst kn. destruct X1. contradiction. + - solve_all. depelim H; econstructor; intuition eauto. + solve_all. Qed. Lemma lookup_env_ext {Σ kn kn' d d'} : @@ -1216,15 +1239,15 @@ From MetaCoq.Erasure Require Import EEtaExpandedFix. Lemma erase_brs_eq X_type X {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} Γ p ts wt : erase_brs X_type X Γ p ts wt = - map_All (fun br wt => (erase_context (bcontext br), erase X_type X _ (bbody br) wt)) ts wt. + All_Forall.map_All (fun br wt => (erase_context (bcontext br), erase X_type X _ (bbody br) wt)) ts wt. Proof. - funelim (map_All _ ts wt); cbn; auto. + funelim (All_Forall.map_All _ ts wt); cbn; auto. f_equal => //. f_equal => //. apply erase_irrel. rewrite -H. eapply erase_irrel. Qed. Lemma erase_fix_eq X_type X {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} Γ ts wt : - erase_fix X_type X Γ ts wt = map_All (fun d wt => + erase_fix X_type X Γ ts wt = All_Forall.map_All (fun d wt => let dbody' := erase X_type X _ (dbody d) (fun Σ abs => proj2 (wt Σ abs)) in let dbody' := if isBox dbody' then match d.(dbody) with @@ -1234,7 +1257,7 @@ Lemma erase_fix_eq X_type X {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_e in {| E.dname := d.(dname).(binder_name); E.rarg := d.(rarg); E.dbody := dbody' |}) ts wt. Proof. - funelim (map_All _ ts wt); cbn; auto. + funelim (All_Forall.map_All _ ts wt); cbn; auto. f_equal => //. f_equal => //. rewrite (fst (erase_irrel _ _) _ _ _ _ (fun (Σ : global_env_ext) (abs : abstract_env_ext_rel X Σ) => (map_All_obligation_1 x xs h Σ abs).p2)). @@ -1243,11 +1266,11 @@ Proof. Qed. Lemma erase_cofix_eq X_type X {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} Γ ts wt : - erase_cofix X_type X Γ ts wt = map_All (fun d wt => + erase_cofix X_type X Γ ts wt = All_Forall.map_All (fun d wt => let dbody' := erase X_type X _ (dbody d) wt in {| E.dname := d.(dname).(binder_name); E.rarg := d.(rarg); E.dbody := dbody' |}) ts wt. Proof. - funelim (map_All _ ts wt); cbn; auto. + funelim (All_Forall.map_All _ ts wt); cbn; auto. f_equal => //. f_equal => //. apply erase_irrel. rewrite -H. eapply erase_irrel. @@ -1407,7 +1430,7 @@ Section wffix. | tConstruct ind c _ => true | tVar _ => true | tBox => true - | tPrim _ => true + | tPrim p => test_prim wf_fixpoints p end. End wffix. @@ -1420,7 +1443,7 @@ Lemma erases_deps_wellformed (cf := config.extraction_checker_flags) (efl := all Proof. intros ed. induction ed using erases_deps_forall_ind; intros => //; - try solve [cbn in *; unfold wf_fix in *; rtoProp; intuition eauto; solve_all]. + try solve [cbn in *; unfold wf_fix in *; rtoProp; intuition eauto; PCUICAstUtils.solve_all]. - cbn. red in H0. rewrite H0 //. - cbn -[lookup_constructor]. cbn. now destruct H0 as [[-> ->] ->]. @@ -1428,21 +1451,24 @@ Proof. cbn. apply/andP; split. apply/andP; split. * now destruct H0 as [-> ->]. * now move/andP: H6. - * move/andP: H6; solve_all. + * move/andP: H6; PCUICAstUtils.solve_all. - cbn -[lookup_projection] in *. apply/andP; split; eauto. now rewrite (declared_projection_lookup H0). + - cbn in H, H0 |- *. solve_all_k 7. Qed. Lemma erases_wf_fixpoints Σ Γ t t' : Σ;;; Γ |- t ⇝ℇ t' -> ErasureProperties.wellformed Σ t -> wf_fixpoints t'. Proof. - induction 1 using erases_forall_list_ind; cbn; auto; try solve [rtoProp; repeat solve_all]. + induction 1 using erases_forall_list_ind; cbn; auto; try solve [rtoProp; repeat PCUICAstUtils.solve_all]. - move/andP => []. rewrite (All2_length X) => -> /=. unfold test_def in *. eapply Forall2_All2 in H. - eapply All2_All2_mix in X; tea. solve_all. + eapply All2_All2_mix in X; tea. PCUICAstUtils.solve_all. destruct b0; eapply erases_isLambda in H1; tea. - move/andP => []. rewrite (All2_length X) => -> /=. - unfold test_def in *. solve_all. + unfold test_def in *. PCUICAstUtils.solve_all. + - depelim H; depelim H0; depelim X; depelim X0; solve_all. noconf H. + solve_all_k 7. cbn in H0. constructor; cbn; rtoProp; intuition eauto. solve_all. Qed. Lemma erase_wf_fixpoints (efl := all_env_flags) {X_type X} univs wfΣ {Γ t} wt @@ -2117,7 +2143,7 @@ Lemma In_map_All {A B C : Type} {Q : C -> Type} {P : C -> A -> Prop} (fn : forall x : A, (forall y : C, Q y -> P y x) -> B) (l : list A) (Hl : forall y : C, Q y -> ∥ All (P y) l ∥) : forall x, In x l -> exists D : forall y : C, Q y -> P y x, - In (fn x D) (map_All fn l Hl). + In (fn x D) (All_Forall.map_All fn l Hl). Proof. induction l; cbn => //. intros x []. @@ -2849,8 +2875,7 @@ Proof. apply: (PCUICFirstorder.firstorder_value_inds Σ Γ). intros i n ui u args pandi hty hfo ih isp. assert (Forall2 (erases Σ Γ) args (map (flip compile_value_erase []) args)). - { solve_all. eapply All_All2; tea. - cbn. intros x [fo hx]. exact hx. } + { PCUICAstUtils.solve_all. } unshelve epose proof (erases_mkApps Σ Γ (tConstruct i n ui) (EAst.tConstruct i n []) args _ _ H). now constructor. now rewrite compile_value_erase_mkApps app_nil_r. diff --git a/erasure/theories/ErasureProperties.v b/erasure/theories/ErasureProperties.v index 12df078fd..a2af130af 100644 --- a/erasure/theories/ErasureProperties.v +++ b/erasure/theories/ErasureProperties.v @@ -2,7 +2,7 @@ From Coq Require Import Program ssreflect ssrbool. From MetaCoq.Utils Require Import utils. From MetaCoq.Common Require Import config. -From MetaCoq.Erasure Require Import ELiftSubst EGlobalEnv EWcbvEval Extract Prelim +From MetaCoq.Erasure Require Import EPrimitive ELiftSubst EGlobalEnv EWcbvEval Extract Prelim ESubstitution EArities EDeps. From MetaCoq.PCUIC Require Import PCUICTyping PCUICGlobalEnv PCUICAst PCUICAstUtils PCUICConversion PCUICSigmaCalculus @@ -252,6 +252,8 @@ Proof. + subst types. eapply conv_context_app_same; auto. + subst types. eapply conv_context_wf_local_app; eauto. + assumption. + - econstructor. depelim H2; constructor. depelim X5; depelim X2; constructor; eauto. + solve_all. Qed. @@ -408,6 +410,10 @@ Proof. eapply erases_ctx_ext. eassumption. unfold app_context. f_equal. eapply fix_context_subst_instance. all: eauto. + + - cbn; constructor. depelim H4. + depelim X2; depelim X5; repeat constructor; cbn; eauto. + solve_all. Qed. Lemma erases_subst_instance : @@ -484,6 +490,9 @@ Proof. unfold EAst.test_def; simpl; eauto. rewrite <-H. rewrite fix_context_length in b0. eapply b0. eauto. now rewrite app_length fix_context_length. + - depelim H. depelim X0; solve_all. + depelim X; solve_all. eapply primProp_impl_test_prim. + constructor; intuition eauto. solve_all. Qed. (** Auxilliary notion of wellformedness on PCUIC to prove that erased terms are wellformed *) @@ -505,7 +514,7 @@ Section wellscoped. Fixpoint wellformed (t : term) : bool := match t with | tRel i => true - | tPrim p => true + | tPrim p => test_prim wellformed p | tEvar ev args => List.forallb (wellformed) args | tLambda _ N M => wellformed N && wellformed M | tApp u v => wellformed u && wellformed v @@ -557,6 +566,7 @@ Section wellscoped. - now eapply nth_error_Some_length, Nat.ltb_lt in H0. - solve_all. destruct a as [s []], b. unfold test_def. len in i0. now rewrite i i0. + - depelim X1; solve_all; constructor; eauto. Qed. Lemma welltyped_wellformed {Σ : global_env_ext} {wfΣ : wf Σ} {Γ a} : welltyped Σ Γ a -> wellformed Σ a. @@ -692,6 +702,8 @@ Proof. unfold EAst.test_def; simpl; eauto. rewrite fix_context_length in b. eapply b. now move: b0 => /andP[]. eauto. now rewrite app_length fix_context_length. tea. + - depelim H. solve_all. primProp. + depelim X0; depelim X1; repeat constructor; cbn; intuition eauto. solve_all. Qed. diff --git a/erasure/theories/Extract.v b/erasure/theories/Extract.v index 67b4c5ff7..bf09291c7 100644 --- a/erasure/theories/Extract.v +++ b/erasure/theories/Extract.v @@ -68,7 +68,7 @@ Definition erase_context (Γ : context) : list name := map (fun d => d.(decl_name).(binder_name)) Γ. Inductive erase_prim_model (erase : term -> E.term -> Prop) : forall {t : prim_tag}, - @PCUICPrimitive.prim_model term t -> @prim_model E.term t -> Prop := + @PCUICPrimitive.prim_model term t -> @prim_model E.term t -> Type := | erase_primInt i : @erase_prim_model erase primInt (PCUICPrimitive.primIntModel i) (primIntModel i) | erase_primFloat f : @erase_prim_model erase primFloat (PCUICPrimitive.primFloatModel f) (primFloatModel f) | erase_primArray a ed ev : @@ -76,11 +76,13 @@ Inductive erase_prim_model (erase : term -> E.term -> Prop) : forall {t : prim_t All2 erase a.(PCUICPrimitive.array_value) ev -> @erase_prim_model erase primArray (PCUICPrimitive.primArrayModel a) (primArrayModel {| array_default := ed; array_value := ev |}). +Derive Signature NoConfusion for erase_prim_model. Inductive erase_prim_val (erase : term -> E.term -> Prop) : PCUICPrimitive.prim_val term -> prim_val E.term -> Prop := | erase_prim t m m' : @erase_prim_model erase t m m' -> erase_prim_val erase (t; m) (t; m'). +Derive Signature for erase_prim_val. Inductive erases (Σ : global_env_ext) (Γ : context) : term -> E.term -> Prop := erases_tRel : forall i : nat, Σ;;; Γ |- tRel i ⇝ℇ E.tRel i @@ -250,10 +252,10 @@ Proof. constructor; [now apply f|now apply f']. - eapply Hprim; tea. induction H. constructor. - induction H; constructor. + induction X; constructor. * now eapply f. * destruct a; cbn in *. - revert array_value ev X. + revert array_value ev a0. fix f' 3; intros mfix mfix' []; [now constructor|]. constructor; [now apply f|now apply f']. Defined. diff --git a/erasure/theories/Typed/Optimize.v b/erasure/theories/Typed/Optimize.v index dbd61b4ea..7d52b6110 100644 --- a/erasure/theories/Typed/Optimize.v +++ b/erasure/theories/Typed/Optimize.v @@ -4,7 +4,7 @@ From MetaCoq.Erasure.Typed Require Import ExAst. From MetaCoq.Erasure.Typed Require Import Transform. From MetaCoq.Erasure.Typed Require Import ResultMonad. From MetaCoq.Erasure.Typed Require Import Utils. -From MetaCoq.Erasure Require Import ELiftSubst. +From MetaCoq.Erasure Require Import EPrimitive ELiftSubst. From MetaCoq.Utils Require Import utils. Import Kernames. @@ -20,6 +20,7 @@ Definition map_subterms (f : term -> term) (t : term) : term := | tProj p t => tProj p (f t) | tFix def i => tFix (map (map_def f) def) i | tCoFix def i => tCoFix (map (map_def f) def) i + | tPrim p => tPrim (map_prim f p) | t => t end. @@ -309,6 +310,7 @@ Fixpoint is_dead (rel : nat) (t : term) : bool := | tFix defs _ | tCoFix defs _ => forallb (is_dead (#|defs| + rel) ∘ EAst.dbody) defs | tConstruct _ _ args => forallb (is_dead rel) args + | tPrim p => test_prim (is_dead rel) p | _ => true end. @@ -366,6 +368,7 @@ Fixpoint valid_cases (t : term) : bool := | tFix defs _ | tCoFix defs _ => forallb (valid_cases ∘ EAst.dbody) defs | tConstruct _ _ (_ :: _) => false (* check whether constructors are not blocks*) + | tPrim p => test_prim valid_cases p | _ => true end. @@ -407,7 +410,7 @@ Fixpoint is_expanded_aux (nargs : nat) (t : term) : bool := | tProj _ t => is_expanded_aux 0 t | tFix defs _ | tCoFix defs _ => forallb (is_expanded_aux 0 ∘ EAst.dbody) defs - | tPrim _ => true + | tPrim p => test_prim (is_expanded_aux 0) p end. (** Check if all applications are applied enough to be deboxed without eta expansion *) @@ -644,7 +647,7 @@ Fixpoint analyze (state : analyze_state) (t : term) {struct t} : analyze_state : let state := new_vars state #|defs| in let state := fold_left (fun state d => analyze state (dbody d)) defs state in remove_vars state #|defs| - | tPrim _ => state + | tPrim p => fold_prim analyze p state end. Fixpoint decompose_TArr (bt : box_type) : list box_type × box_type := diff --git a/erasure/theories/Typed/OptimizeCorrectness.v b/erasure/theories/Typed/OptimizeCorrectness.v index 8c08f5441..741357b87 100644 --- a/erasure/theories/Typed/OptimizeCorrectness.v +++ b/erasure/theories/Typed/OptimizeCorrectness.v @@ -9,7 +9,7 @@ From Coq Require Import List. From Coq Require Import ssrbool. From Coq Require Import PeanoNat. From Equations Require Import Equations. -From MetaCoq.Erasure Require Import EAstUtils. +From MetaCoq.Erasure Require Import EPrimitive EAstUtils. From MetaCoq.Erasure Require Import ECSubst. From MetaCoq.Erasure Require Import EInduction. From MetaCoq.Erasure Require Import ELiftSubst. @@ -220,6 +220,7 @@ Proof. split; [easy|]. rewrite <- Nat.add_succ_r in *. now eapply IHX. + - solve_all. Qed. Lemma is_dead_csubst k t u k' : @@ -285,6 +286,7 @@ Proof. + rewrite <- !Nat.add_succ_r in *. apply IHX; [easy|easy|]. now eapply closed_upwards. + - solve_all_k 6. Qed. Lemma valid_dearg_mask_nil t : valid_dearg_mask [] t. @@ -453,7 +455,7 @@ Proof. now f_equal. + rewrite <- !Nat.add_succ_r in *. now apply IHX. - - reflexivity. + - f_equal; solve_all. Qed. Lemma masked_nil {X} mask : @@ -802,7 +804,7 @@ Proof. cbn. f_equal. now rewrite p. - - now rewrite lift_mkApps. + - rewrite lift_mkApps. f_equal. simpl lift. f_equal. solve_all. Qed. Lemma lift_dearg n k t : @@ -853,7 +855,7 @@ Proof. f_equal. rewrite <- !Nat.add_succ_r. now apply IHX. - - reflexivity. + - solve_all. Qed. Lemma is_dead_lift_all k k' n t : @@ -894,6 +896,7 @@ Proof. cbn. rewrite <- !Nat.add_succ_r. now apply IHX. + - solve_all. Qed. Lemma is_dead_subst_other k k' s t : @@ -944,6 +947,7 @@ Proof. f_equal. rewrite <- !Nat.add_succ_r. now apply IHX. + - solve_all. Qed. Lemma valid_dearg_mask_lift mask n k t : @@ -1128,6 +1132,7 @@ Proof. cbn. rewrite <- Nat.add_succ_r. now rewrite p, IHX. + - solve_all. Qed. Lemma is_expanded_lift n k t : @@ -1190,6 +1195,7 @@ Proof. rewrite <- !Nat.add_succ_r. rewrite IHX by easy. now replace (S (k - n)) with (S k - n) by lia. + - solve_all. Qed. Lemma is_dead_dearg_single k mask t args : @@ -1288,6 +1294,19 @@ Proof. rewrite <- !Nat.add_succ_r. rewrite IHX. bia. + - solve_all. rtoProp; intuition eauto. + depelim X; cbn; eauto. + destruct a; unfold test_array_model; cbn. + destruct p. cbn in *. rewrite e; eauto. + rewrite <- !andb_assoc. f_equal. + rewrite andb_comm. + induction a; cbn. rewrite andb_true_r; auto. + rewrite <- !andb_assoc. + rewrite IHa. rewrite p;eauto. + rewrite <- !andb_assoc. f_equal. + rewrite andb_comm. + rewrite <- !andb_assoc. f_equal. + bia. now rewrite andb_true_r. Qed. Lemma is_dead_dearg_lambdas k mask t : @@ -1400,7 +1419,7 @@ Proof. rewrite p by easy. split; [easy|]. now apply IHX. - - now apply forallb_Forall. + - solve_all. rtoProp; intuition solve_all. Qed. Lemma valid_dearg_mask_dearg mask t : @@ -1557,7 +1576,9 @@ Proof. unfold map_def; cbn. f_equal. now apply (p _ _ []). - - now rewrite subst_mkApps, map_map. + - rewrite subst_mkApps, map_map; cbn; f_equal. f_equal. + solve_all. eapply map_prim_eq_prop; tea; cbn; intuition eauto. + specialize (a s k []). eauto. Qed. Lemma dearg_subst s k t : @@ -1673,7 +1694,7 @@ Proof. propify. rewrite <- !Nat.add_succ_r. now rewrite p, IHX. - - assumption. + - solve_all_k 6. Qed. Lemma is_expanded_aux_subst s n t k : @@ -1717,7 +1738,7 @@ Proof. propify. rewrite <- !Nat.add_succ_r. now rewrite p, IHX. - - assumption. + - solve_all_k 6. Qed. Lemma is_expanded_substl s n t : @@ -2001,6 +2022,7 @@ Proof. cbn in *. propify. now rewrite <- !Nat.add_succ_r. + - solve_all_k 6. Qed. Lemma valid_cases_subst s k t : @@ -2036,6 +2058,7 @@ Proof. - induction X in X, k, valid_t |- *; [easy|]. cbn in *; propify. now rewrite <- !Nat.add_succ_r. + - solve_all_k 6. Qed. Lemma closedn_dearg_single k t args mask : @@ -2113,6 +2136,7 @@ Proof. - rtoProp; solve_all. rewrite -> !Nat.add_assoc in *. specialize (a (#|m| + k')). unfold is_true. rewrite <- a. f_equal. lia. unfold is_true. rewrite <- b. f_equal. lia. + - solve_all_k 6. Qed. Lemma closedn_dearg_case_branch_body_rec i k mask t : @@ -2212,6 +2236,7 @@ Proof. split; [easy|]. rewrite <- !Nat.add_succ_r in *. now apply IHX. + - rewrite closedn_mkApps; cbn; rtoProp; intuition solve_all. solve_all_k 6. Qed. Hint Resolve diff --git a/erasure/theories/Typed/TypeAnnotations.v b/erasure/theories/Typed/TypeAnnotations.v index f37c98f9e..a59888cda 100644 --- a/erasure/theories/Typed/TypeAnnotations.v +++ b/erasure/theories/Typed/TypeAnnotations.v @@ -583,7 +583,7 @@ Proof. apply bigprod_map; [|exact ta.2]. intros. exact (f _ All_nil _ X). - - exact (annot_mkApps ta argsa). + - refine (annot_mkApps _ argsa). cbn. cbn in ta. exact ta. Defined. Definition annot_dearg im cm {t : term} (ta : annots box_type t) : annots box_type (dearg im cm t) := diff --git a/pcuic/theories/utils/PCUICPrimitive.v b/pcuic/theories/utils/PCUICPrimitive.v index ac205a86e..10d9a633b 100644 --- a/pcuic/theories/utils/PCUICPrimitive.v +++ b/pcuic/theories/utils/PCUICPrimitive.v @@ -31,7 +31,7 @@ Arguments primIntModel {term}. Arguments primFloatModel {term}. Arguments primArrayModel {term}. -Derive Signature NoConfusion for prim_model. +Derive Signature NoConfusion NoConfusionHom for prim_model. Definition prim_model_of (term : Type) (p : prim_tag) : Type := match p with diff --git a/utils/theories/All_Forall.v b/utils/theories/All_Forall.v index ad2966a7e..6c82bc714 100644 --- a/utils/theories/All_Forall.v +++ b/utils/theories/All_Forall.v @@ -3760,6 +3760,22 @@ Section map_All. Qed. End map_All. +Section make_All. + Context {A} {B : A -> Type} (f : forall x : A, B x). + + Equations make_All (l : list A) : All B l := + | [] := All_nil + | hd :: tl := All_cons (f hd) (make_All tl). +End make_All. + +Section make_All_All. + Context {A} {B : A -> Type} {C : A -> Type} (f : forall x : A, B x -> C x). + + Equations make_All_All {l : list A} (a : All B l) : All C l := + | All_nil := All_nil + | All_cons bhd btl := All_cons (f _ bhd) (make_All_All btl). +End make_All_All. + Lemma All_map_All {A B C} {Q : C -> Type} {P : C -> A -> Prop} {Q' : B -> Type} {R : C -> A -> Prop} f args (ha: forall y : C, Q y -> ∥ All (R y) args ∥) : From 8b4d7a9d451f7158073a7c4d6d6b7dfe23cabb96 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Fri, 24 Nov 2023 16:55:42 +0100 Subject: [PATCH 13/17] Disable quotation module build --- Makefile | 24 ++++++++++++------------ 1 file changed, 12 insertions(+), 12 deletions(-) diff --git a/Makefile b/Makefile index a9889ce64..30eb78ba6 100644 --- a/Makefile +++ b/Makefile @@ -1,5 +1,5 @@ -all: printconf template-coq pcuic safechecker erasure erasure-plugin safechecker-plugin quotation +all: printconf template-coq pcuic safechecker erasure erasure-plugin safechecker-plugin # quotation -include Makefile.conf @@ -12,7 +12,7 @@ ifeq '$(METACOQ_CONFIG)' 'local' export OCAMLPATH endif -.PHONY: printconf all utils template-coq pcuic erasure install uninstall html clean mrproper .merlin test-suite translations quotation +.PHONY: printconf all utils template-coq pcuic erasure install uninstall html clean mrproper .merlin test-suite translations # quotation printconf: ifeq '$(METACOQ_CONFIG)' 'local' @@ -26,14 +26,14 @@ else endif endif -install: all +install: all $(MAKE) -C utils install $(MAKE) -C common install $(MAKE) -C template-coq install $(MAKE) -C pcuic install $(MAKE) -C safechecker install $(MAKE) -C template-pcuic install - $(MAKE) -C quotation install +# $(MAKE) -C quotation install $(MAKE) -C safechecker-plugin install $(MAKE) -C erasure install $(MAKE) -C erasure-plugin install @@ -45,7 +45,7 @@ uninstall: $(MAKE) -C pcuic uninstall $(MAKE) -C safechecker uninstall $(MAKE) -C template-pcuic uninstall - $(MAKE) -C quotation uninstall +# $(MAKE) -C quotation uninstall $(MAKE) -C safechecker-plugin uninstall $(MAKE) -C erasure uninstall $(MAKE) -C erasure-plugin uninstall @@ -74,7 +74,7 @@ clean: $(MAKE) -C pcuic clean $(MAKE) -C safechecker clean $(MAKE) -C template-pcuic clean - $(MAKE) -C quotation clean +# $(MAKE) -C quotation clean $(MAKE) -C erasure clean $(MAKE) -C erasure-plugin clean $(MAKE) -C examples clean @@ -88,7 +88,7 @@ vos: $(MAKE) -C pcuic vos $(MAKE) -C safechecker vos $(MAKE) -C template-pcuic vos - $(MAKE) -C quotation vos +# $(MAKE) -C quotation vos $(MAKE) -C erasure vos $(MAKE) -C erasure-plugin vos $(MAKE) -C translations vos @@ -100,7 +100,7 @@ quick: $(MAKE) -C pcuic quick $(MAKE) -C safechecker quick $(MAKE) -C template-pcuic quick - $(MAKE) -C quotation vos # quick # we cannot unset universe checking in 8.16 due to COQBUG(https://github.com/coq/coq/issues/17361), and quick does not buy much in quotation anyway, where almost everything is transparent +# $(MAKE) -C quotation vos # quick # we cannot unset universe checking in 8.16 due to COQBUG(https://github.com/coq/coq/issues/17361), and quick does not buy much in quotation anyway, where almost everything is transparent $(MAKE) -C erasure quick $(MAKE) -C erasure-plugin quick $(MAKE) -C translations quick @@ -112,7 +112,7 @@ mrproper: $(MAKE) -C pcuic mrproper $(MAKE) -C safechecker mrproper $(MAKE) -C template-pcuic mrproper - $(MAKE) -C quotation mrproper +# $(MAKE) -C quotation mrproper $(MAKE) -C erasure mrproper $(MAKE) -C erasure-plugin mrproper $(MAKE) -C examples mrproper @@ -126,7 +126,7 @@ mrproper: $(MAKE) -C pcuic .merlin $(MAKE) -C safechecker .merlin $(MAKE) -C template-pcuic .merlin - $(MAKE) -C quotation .merlin +# $(MAKE) -C quotation .merlin $(MAKE) -C erasure .merlin $(MAKE) -C erasure-plugin .merin @@ -148,8 +148,8 @@ safechecker: pcuic template-pcuic: template-coq pcuic $(MAKE) -C template-pcuic -quotation: template-coq pcuic template-pcuic - $(MAKE) -C quotation +#quotation: template-coq pcuic template-pcuic +# $(MAKE) -C quotation safechecker-plugin: safechecker template-pcuic $(MAKE) -C safechecker-plugin From 133b5ad7ff8f9bff4c54e68444b88ae640faadb5 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Fri, 24 Nov 2023 16:55:49 +0100 Subject: [PATCH 14/17] Fixes after rebase --- erasure/theories/EImplementBox.v | 11 +++++------ erasure/theories/EInduction.v | 2 +- erasure/theories/EReflect.v | 6 +++--- quotation/theories/CommonUtils.v | 1 + .../ToPCUIC/QuotationOf/Common/Environment/Sig.v | 2 +- quotation/theories/ToTemplate/Init.v | 2 ++ .../ToTemplate/QuotationOf/Common/Environment/Sig.v | 2 +- 7 files changed, 14 insertions(+), 12 deletions(-) diff --git a/erasure/theories/EImplementBox.v b/erasure/theories/EImplementBox.v index 132539142..c5a3ad9fb 100644 --- a/erasure/theories/EImplementBox.v +++ b/erasure/theories/EImplementBox.v @@ -94,11 +94,11 @@ Section implement_box. rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; unfold test_def in *; simpl closed in *; - (*try solve [simpl; subst; simpl closed; f_equal; auto; rtoProp; solve_all; solve_all]; try easy. + try solve [simpl; subst; simpl closed; f_equal; auto; rtoProp; solve_all; solve_all]; try easy. rtoProp. split. eauto. solve_all. replace (#|x.1| + S k) with (#|x.1| + k + 1) by lia. - eapply closedn_lift. eauto.*) + eapply closedn_lift. eauto. try solve [simpl; subst; simpl closed; f_equal; auto; rtoProp; solve_all; solve_all_k 6]; try easy. Qed. @@ -124,10 +124,9 @@ Section implement_box. f_equal. lia. - cbn. f_equal. rewrite !map_map. solve_all. eapply In_All. intros ? ?. unfold map_def. cbn. f_equal. erewrite H; eauto. - f_equal. now rewrite map_length. - cbn. f_equal. rewrite !map_map. solve_all. eapply In_All. intros ? ?. unfold map_def. cbn. f_equal. erewrite H; eauto. - f_equal. now rewrite map_length. + - solve_all_k 6. Qed. (* Lemma implement_box_subst a k b : *) @@ -185,9 +184,9 @@ Section implement_box. f_equal. eapply H; eauto. - cbn. f_equal. rewrite !map_map. solve_all. - eapply In_All. intros ? ?. unfold map_def. cbn. f_equal. rewrite H; eauto. + eapply In_All. intros ? ?. unfold map_def. cbn. f_equal. solve_all. - cbn. f_equal. rewrite !map_map. solve_all. - eapply In_All. intros ? ?. unfold map_def. cbn. f_equal. rewrite H; eauto. + eapply In_All. intros ? ?. unfold map_def. cbn. f_equal. solve_all. - cbn. solve_all. Qed. diff --git a/erasure/theories/EInduction.v b/erasure/theories/EInduction.v index 208332ffc..306b6d25c 100644 --- a/erasure/theories/EInduction.v +++ b/erasure/theories/EInduction.v @@ -67,7 +67,7 @@ Proof. destruct mfix; constructor; [|apply auxm]. apply auxt. - destruct p as [? []]; cbn => //; constructor. + destruct prim as [? []]; cbn => //; constructor. destruct a as [def v]; cbn. split. eapply auxt. revert v; fix auxv 1. diff --git a/erasure/theories/EReflect.v b/erasure/theories/EReflect.v index 9fbc550b8..71284edc2 100644 --- a/erasure/theories/EReflect.v +++ b/erasure/theories/EReflect.v @@ -112,13 +112,13 @@ Proof. subst. inversion e0. subst. destruct (eq_dec rarg rarg0) ; nodec. subst. left. reflexivity. - - destruct p as [? []]; destruct p0 as [? []]; cbn ; nodec. + - destruct prim as [? []]; destruct p as [? []]; cbn ; nodec. + destruct (eq_dec i i0); nodec; subst; eauto. + destruct (eq_dec f f0); nodec; subst; eauto. + destruct a as [? ?], a0 as [? ?]; cbn. depelim X. destruct p as [hd hv]. cbn in *. - destruct (hd array_default0); nodec; subst; eauto. - revert array_value0; induction hv; intros []; eauto; nodec. + destruct (hd array_default); nodec; subst; eauto. + revert array_value; induction hv; intros []; eauto; nodec. destruct (p t); subst; nodec. destruct (IHhv l0); nodec. noconf e; eauto. Defined. diff --git a/quotation/theories/CommonUtils.v b/quotation/theories/CommonUtils.v index 18943d44e..27f703c07 100644 --- a/quotation/theories/CommonUtils.v +++ b/quotation/theories/CommonUtils.v @@ -188,6 +188,7 @@ Module WithTemplate. | tVar _ | tInt _ | tFloat _ + | tArray _ _ _ _ | tConst _ _ | tInd _ _ | tConstruct _ _ _ diff --git a/quotation/theories/ToPCUIC/QuotationOf/Common/Environment/Sig.v b/quotation/theories/ToPCUIC/QuotationOf/Common/Environment/Sig.v index 31de3ebc0..3b5d13ea9 100644 --- a/quotation/theories/ToPCUIC/QuotationOf/Common/Environment/Sig.v +++ b/quotation/theories/ToPCUIC/QuotationOf/Common/Environment/Sig.v @@ -43,7 +43,7 @@ Module Type QuoteEnvironmentSig (T : Term) (Import E : EnvironmentSig T). #[export] Declare Instance quote_extends {Σ Σ'} : ground_quotable (@extends Σ Σ'). #[export] Declare Instance quote_extends_decls {Σ Σ'} : ground_quotable (@extends_decls Σ Σ'). - #[export] Declare Instance quote_primitive_invariants {cdecl} : ground_quotable (primitive_invariants cdecl). + #[export] Declare Instance quote_primitive_invariants {cdecl ty} : ground_quotable (primitive_invariants cdecl ty). #[export] Declare Instance quote_All_decls {P t t'} {qP : quotation_of P} {quoteP : forall t t', ground_quotable (P t t')} : ground_quotable (All_decls P t t'). #[export] Declare Instance quote_All_decls_alpha {P t t'} {qP : quotation_of P} {quoteP : forall t t', ground_quotable (P t t')} : ground_quotable (All_decls_alpha P t t'). diff --git a/quotation/theories/ToTemplate/Init.v b/quotation/theories/ToTemplate/Init.v index 4bf46f2f7..487ea21ba 100644 --- a/quotation/theories/ToTemplate/Init.v +++ b/quotation/theories/ToTemplate/Init.v @@ -108,6 +108,7 @@ Proof. | tSort _ | tInt _ | tFloat _ + | tArray _ _ _ _ | tConst _ _ => if head_term_is_bound cur_modpath qt then tmMaybeInferQuotation tt @@ -212,6 +213,7 @@ Proof. | tConstruct _ _ _ | tInt _ | tFloat _ + | tArray _ _ _ _ | tInd _ _ => ret qt | tCast t kind v diff --git a/quotation/theories/ToTemplate/QuotationOf/Common/Environment/Sig.v b/quotation/theories/ToTemplate/QuotationOf/Common/Environment/Sig.v index 7bfba4ce3..86e7a9309 100644 --- a/quotation/theories/ToTemplate/QuotationOf/Common/Environment/Sig.v +++ b/quotation/theories/ToTemplate/QuotationOf/Common/Environment/Sig.v @@ -43,7 +43,7 @@ Module Type QuoteEnvironmentSig (T : Term) (Import E : EnvironmentSig T). #[export] Declare Instance quote_extends {Σ Σ'} : ground_quotable (@extends Σ Σ'). #[export] Declare Instance quote_extends_decls {Σ Σ'} : ground_quotable (@extends_decls Σ Σ'). - #[export] Declare Instance quote_primitive_invariants {cdecl} : ground_quotable (primitive_invariants cdecl). + #[export] Declare Instance quote_primitive_invariants {cdecl ty} : ground_quotable (primitive_invariants cdecl ty). #[export] Declare Instance quote_All_decls {P t t'} {qP : quotation_of P} {quoteP : forall t t', ground_quotable (P t t')} : ground_quotable (All_decls P t t'). #[export] Declare Instance quote_All_decls_alpha {P t t'} {qP : quotation_of P} {quoteP : forall t t', ground_quotable (P t t')} : ground_quotable (All_decls_alpha P t t'). From 8b5049ff2b3727575eea47bafb981c897c25e58a Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Thu, 30 Nov 2023 12:07:40 +0100 Subject: [PATCH 15/17] Fix translations, and a bit of quotation, up-to an issue with nested inductive definitions. --- .vscode/metacoq.code-workspace | 5 ++++- quotation/theories/ToPCUIC/Common/Environment.v | 2 +- quotation/theories/ToPCUIC/PCUIC/PCUICAst.v | 6 +++++- quotation/theories/ToPCUIC/PCUIC/PCUICEquality.v | 7 ++++++- quotation/theories/ToPCUIC/PCUIC/utils/PCUICPrimitive.v | 6 +++--- quotation/theories/ToTemplate/Common/Environment.v | 2 +- translations/param_binary.v | 2 +- translations/param_original.v | 2 +- 8 files changed, 22 insertions(+), 10 deletions(-) diff --git a/.vscode/metacoq.code-workspace b/.vscode/metacoq.code-workspace index 0c15d9de8..4abf21a94 100644 --- a/.vscode/metacoq.code-workspace +++ b/.vscode/metacoq.code-workspace @@ -12,6 +12,7 @@ + "-R", "safechecker-plugin/theories","MetaCoq.SafeCheckerPlugin", "-R", "utils/theories","MetaCoq.Utils", @@ -32,7 +33,9 @@ "-R", "erasure-plugin/theories","MetaCoq.ErasurePlugin", "-R", "translations","MetaCoq.Translations", "-R", "test-suite/plugin-demo/theories","MetaCoq.ExtractedPluginDemo", - "-R", "test-suite","MetaCoq.TestSuite" + "-R", "test-suite","MetaCoq.TestSuite", + "-R", "quotation/theories", "MetaCoq.Quotation", + ], // When enabled, will trim trailing whitespace when saving a file. diff --git a/quotation/theories/ToPCUIC/Common/Environment.v b/quotation/theories/ToPCUIC/Common/Environment.v index 7375dc51b..52dda940c 100644 --- a/quotation/theories/ToPCUIC/Common/Environment.v +++ b/quotation/theories/ToPCUIC/Common/Environment.v @@ -111,7 +111,7 @@ Module QuoteEnvironment (T : Term) (Import E : EnvironmentSig T) (Import QEH : Q #[export] Instance quote_extends_strictly_on_decls {Σ Σ'} : ground_quotable (@extends_strictly_on_decls Σ Σ') := ltac:(cbv [extends_strictly_on_decls]; exact _). #[export] Instance quote_strictly_extends_decls {Σ Σ'} : ground_quotable (@strictly_extends_decls Σ Σ') := ltac:(cbv [strictly_extends_decls]; exact _). - #[export] Instance quote_primitive_invariants {cdecl} : ground_quotable (primitive_invariants cdecl) := ltac:(cbv [primitive_invariants]; exact _). + #[export] Instance quote_primitive_invariants {prim_ty cdecl} : ground_quotable (primitive_invariants prim_ty cdecl) := ltac:(cbv [primitive_invariants]; exact _). #[export] Instance quote_All_decls {P t t'} {qP : quotation_of P} {quoteP : forall t t', ground_quotable (P t t')} : ground_quotable (All_decls P t t') := ltac:(induction 1; exact _). #[export] Instance quote_All_decls_alpha {P t t'} {qP : quotation_of P} {quoteP : forall t t', ground_quotable (P t t')} : ground_quotable (All_decls_alpha P t t') := ltac:(induction 1; exact _). diff --git a/quotation/theories/ToPCUIC/PCUIC/PCUICAst.v b/quotation/theories/ToPCUIC/PCUIC/PCUICAst.v index cacfba567..a983580f0 100644 --- a/quotation/theories/ToPCUIC/PCUIC/PCUICAst.v +++ b/quotation/theories/ToPCUIC/PCUIC/PCUICAst.v @@ -11,7 +11,11 @@ From MetaCoq.Quotation.ToPCUIC.QuotationOf.PCUIC Require Import PCUICAst.Instanc #[export] Instance quote_branch {term} {qterm : quotation_of term} {quote_term : ground_quotable term} : ground_quotable (branch term) := ltac:(destruct 1; exact _). #[local] Hint Extern 1 => assumption : typeclass_instances. -#[export] Instance quote_term : ground_quotable term := ltac:(induction 1 using term_forall_list_ind; exact _). +#[export] Instance quote_term : ground_quotable term. + induction 1 using term_forall_list_ind; try exact _. + destruct p as [? []]; cbn in X. exact _. exact _. destruct X as [? []]. + exact _. +Defined. Module QuotePCUICTerm <: QuoteTerm PCUICTerm. #[export] Instance quote_term : ground_quotable PCUICTerm.term := ltac:(cbv [PCUICTerm.term]; exact _). diff --git a/quotation/theories/ToPCUIC/PCUIC/PCUICEquality.v b/quotation/theories/ToPCUIC/PCUIC/PCUICEquality.v index 904baa88f..3c9b88ab2 100644 --- a/quotation/theories/ToPCUIC/PCUIC/PCUICEquality.v +++ b/quotation/theories/ToPCUIC/PCUIC/PCUICEquality.v @@ -29,6 +29,11 @@ End with_R. #[export] Instance quote_compare_decls {eq_term leq_term u u'} {qeq_term : quotation_of eq_term} {qleq_term : quotation_of leq_term} {quote_eq_term : forall x y, ground_quotable (eq_term x y)} {quote_leq_term : forall x y, ground_quotable (leq_term x y)} : ground_quotable (@compare_decls eq_term leq_term u u') := ltac:(destruct 1; exact _). +#[export] Instance quote_onPrims {eq_term leq_term u u'} {qeq_term : quotation_of eq_term} {qleq_term : quotation_of leq_term} + {quote_eq_term : forall x y, ground_quotable (eq_term x y)} + {quote_leq_term : forall x y, ground_quotable (leq_term x y)} : ground_quotable (@onPrims term eq_term leq_term u u') := + ltac:(destruct 1; exact _). + #[export] Hint Unfold eq_predicate : quotation. @@ -49,7 +54,7 @@ Proof. | forall (x1 : ?X1) (x2 : ?X2) (x3 : ?X3) (x4 : ?X4) (x5 : ?X5) (x6 : ?X6) (x7 : ?X7) (x8 : ?X8) (x9 : ?X9) (x10 : ?X10) (t : ?X11), quotation_of t => change (forall (x1 : X1) (x2 : X2) (x3 : X3) (x4 : X4) (x5 : X5) (x6 : X6) (x7 : X7) (x8 : X8) (x9 : X9) (x10 : X10), ground_quotable X11) in quote_eq_term_upto_univ_napp end. - destruct t; replace_quotation_of_goal (). + destruct t; try replace_quotation_of_goal (). Defined. #[export] Instance quote_compare_term {cf pb Σ ϕ x y} : ground_quotable (@compare_term cf pb Σ ϕ x y) := ltac:(cbv [compare_term]; exact _). diff --git a/quotation/theories/ToPCUIC/PCUIC/utils/PCUICPrimitive.v b/quotation/theories/ToPCUIC/PCUIC/utils/PCUICPrimitive.v index 95f4940c2..c74b76fe4 100644 --- a/quotation/theories/ToPCUIC/PCUIC/utils/PCUICPrimitive.v +++ b/quotation/theories/ToPCUIC/PCUIC/utils/PCUICPrimitive.v @@ -5,8 +5,8 @@ From MetaCoq.Quotation.ToPCUIC.Common Require Import (hints) Universes Primitive #[export] Instance quote_array_model {term} {qterm : quotation_of term} {quote_term : ground_quotable term} : ground_quotable (array_model term) := ltac:(destruct 1; exact _). -#[export] Instance quote_prim_model {term tag} {qterm : quotation_of term} : ground_quotable (prim_model term tag) := ltac:(destruct 1; exact _). +#[export] Instance quote_prim_model {term tag} {qterm : quotation_of term} {quote_term : ground_quotable term} : ground_quotable (prim_model term tag) := ltac:(destruct 1; eauto). -#[export] Instance quote_prim_model_of {term tag} : ground_quotable (prim_model_of term tag) := ltac:(cbv [prim_model_of]; exact _). +#[export] Instance quote_prim_model_of {term tag} {qterm : quotation_of term} {quote_term : ground_quotable term} : ground_quotable (prim_model_of term tag) := ltac:(cbv [prim_model_of]; destruct tag; exact _). -#[export] Instance quote_prim_val {term} {qterm : quotation_of term} : ground_quotable (prim_val term) := ltac:(cbv [prim_val]; exact _). +#[export] Instance quote_prim_val {term} {qterm : quotation_of term} {quote_term : ground_quotable term} : ground_quotable (prim_val term) := ltac:(cbv [prim_val]; exact _). diff --git a/quotation/theories/ToTemplate/Common/Environment.v b/quotation/theories/ToTemplate/Common/Environment.v index a2a3abdce..5b7544a12 100644 --- a/quotation/theories/ToTemplate/Common/Environment.v +++ b/quotation/theories/ToTemplate/Common/Environment.v @@ -111,7 +111,7 @@ Module QuoteEnvironment (T : Term) (Import E : EnvironmentSig T) (Import QEH : Q #[export] Instance quote_extends_strictly_on_decls {Σ Σ'} : ground_quotable (@extends_strictly_on_decls Σ Σ') := ltac:(cbv [extends_strictly_on_decls]; exact _). #[export] Instance quote_strictly_extends_decls {Σ Σ'} : ground_quotable (@strictly_extends_decls Σ Σ') := ltac:(cbv [strictly_extends_decls]; exact _). - #[export] Instance quote_primitive_invariants {cdecl} : ground_quotable (primitive_invariants cdecl) := ltac:(cbv [primitive_invariants]; exact _). + #[export] Instance quote_primitive_invariants {prim_ty cdecl} : ground_quotable (primitive_invariants prim_ty cdecl) := ltac:(cbv [primitive_invariants]; exact _). #[export] Instance quote_All_decls {P t t'} {qP : quotation_of P} {quoteP : forall t t', ground_quotable (P t t')} : ground_quotable (All_decls P t t') := ltac:(induction 1; exact _). #[export] Instance quote_All_decls_alpha {P t t'} {qP : quotation_of P} {quoteP : forall t t', ground_quotable (P t t')} : ground_quotable (All_decls_alpha P t t') := ltac:(induction 1; exact _). diff --git a/translations/param_binary.v b/translations/param_binary.v index 3de6343b0..ecf3c6be1 100644 --- a/translations/param_binary.v +++ b/translations/param_binary.v @@ -145,7 +145,7 @@ Fixpoint tsl_rec1_app (app : list term) (E : tsl_table) (t : term) : term := | tFix _ _ | tCoFix _ _ => todo "tsl" | tVar _ | tEvar _ _ => todo "tsl" | tLambda _ _ _ => tVar "impossible" - | tInt _ | tFloat _ => todo "tsl" + | tInt _ | tFloat _ | tArray _ _ _ _ => todo "tsl" end in apply app t1 end. diff --git a/translations/param_original.v b/translations/param_original.v index b817b9782..cbdcc3313 100644 --- a/translations/param_original.v +++ b/translations/param_original.v @@ -106,7 +106,7 @@ Fixpoint tsl_rec1_app (app : option term) (E : tsl_table) (t : term) : term := | tFix _ _ | tCoFix _ _ => todo "tsl" | tVar _ | tEvar _ _ => todo "tsl" | tLambda _ _ _ => tVar "impossible" - | tInt _ | tFloat _ => todo "tsl" + | tInt _ | tFloat _ | tArray _ _ _ _ => todo "tsl" end in match app with Some t' => mkApp t1 (t' {3 := tRel 1} {2 := tRel 0}) | None => t1 end From 1dd89940da55a7398ef1114d79396acbdcaac47d Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Thu, 30 Nov 2023 14:04:52 +0100 Subject: [PATCH 16/17] Disable quotation opam module dependency --- coq-metacoq.opam | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/coq-metacoq.opam b/coq-metacoq.opam index 0d05677db..b497f3e21 100644 --- a/coq-metacoq.opam +++ b/coq-metacoq.opam @@ -24,7 +24,7 @@ depends: [ "coq-metacoq-safechecker-plugin" {= version} "coq-metacoq-erasure-plugin" {= version} "coq-metacoq-translations" {= version} - "coq-metacoq-quotation" {= version} +# "coq-metacoq-quotation" {= version} ] build: [ ["bash" "./configure.sh" ] {with-test} From 7dd6de84f34317547573b7a90c95f58208f20706 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Thu, 30 Nov 2023 14:51:05 +0100 Subject: [PATCH 17/17] Disable the opam quotation package for now --- coq-metacoq-quotation.opam => coq-metacoq-quotation.opam.disabled | 0 1 file changed, 0 insertions(+), 0 deletions(-) rename coq-metacoq-quotation.opam => coq-metacoq-quotation.opam.disabled (100%) diff --git a/coq-metacoq-quotation.opam b/coq-metacoq-quotation.opam.disabled similarity index 100% rename from coq-metacoq-quotation.opam rename to coq-metacoq-quotation.opam.disabled