@@ -1020,42 +1020,59 @@ Qed.
1020
1020
1021
1021
From MetaCoq.Erasure Require Import EReorderCstrs.
1022
1022
1023
- Axiom trust_reorder_cstrs_wf :
1024
- forall efl : EEnvFlags,
1025
- WcbvFlags ->
1026
- forall (m : inductives_mapping) (input : Transform.program E.global_context term),
1027
- wf_eprogram efl input -> wf_eprogram efl (reorder_program m input).
1028
- Axiom trust_reorder_cstrs_pres :
1029
- forall (efl : EEnvFlags) (wfl : WcbvFlags) (m : inductives_mapping) (p : Transform.program E.global_context term)
1030
- (v : term),
1031
- wf_eprogram efl p ->
1032
- eval_eprogram wfl p v -> exists v' : term, eval_eprogram wfl (reorder_program m p) v' /\ v' = reorder m v.
1023
+ Definition eval_eprogram_mapping (wfl : WcbvFlags) (p : inductives_mapping * eprogram) t :=
1024
+ eval_eprogram wfl p.2 t.
1033
1025
1034
- Program Definition reorder_cstrs_transformation (efl : EEnvFlags) (wfl : WcbvFlags) (m : inductives_mapping) :
1035
- Transform.t _ _ EAst.term EAst.term _ _
1036
- (eval_eprogram wfl) (eval_eprogram wfl) :=
1037
- {| name := "reoder inductive constructors ";
1038
- transform p _ := EReorderCstrs.reorder_program m p ;
1039
- pre p := wf_eprogram efl p ;
1040
- post p := wf_eprogram efl p ;
1041
- obseq p hp p' v v' := v' = EReorderCstrs.reorder m v |}.
1026
+ Definition eval_eprogram_env_mapping (wfl : WcbvFlags) (p : inductives_mapping * eprogram_env) t :=
1027
+ eval_eprogram_env wfl p.2 t.
1028
+
1029
+ Definition to_program (e : eprogram_env) : eprogram :=
1030
+ (e.1, e.2).
1031
+
1032
+ Program Definition reorder_cstrs_transformation {efl : EEnvFlags} {wca : cstr_as_blocks = false} {has_app : has_tApp}
1033
+ (wfl : WcbvFlags) {wcon : with_constructor_as_block = false} :
1034
+ Transform.t _ _ _ EAst.term _ _
1035
+ (eval_eprogram_env_mapping wfl) (eval_eprogram wfl) :=
1036
+ {| name := "reorder inductive constructors ";
1037
+ transform p _ := EReorderCstrs.reorder_program p.1 (to_program p.2) ;
1038
+ pre p := [/\ wf_eprogram_env efl p.2, EEtaExpandedFix.expanded_eprogram_env p.2 & wf_inductives_mapping p.2.1 p.1] ;
1039
+ post p := wf_eprogram efl p /\ EEtaExpandedFix.expanded_eprogram p;
1040
+ obseq p hp p' v v' := v' = EReorderCstrs.reorder p.1 v |}.
1042
1041
1043
1042
Next Obligation .
1044
- move=> efl wfl m. cbn. now apply trust_reorder_cstrs_wf.
1043
+ move=> efl wca hasapp wfl wcb [m p] [wfp exp wfm]. split => //.
1044
+ now unshelve eapply reorder_program_wf. cbn.
1045
+ now eapply reorder_program_expanded_fix.
1045
1046
Qed .
1046
1047
Next Obligation .
1047
- red. eapply trust_reorder_cstrs_pres.
1048
+ red. intros efl wca hasapp wfl wcb [m p] v [wfp wfm] evp.
1049
+ destruct evp as [ev].
1050
+ unshelve eapply EReorderCstrs.optimize_correct in ev; trea.
1051
+ 2,3:apply wfp.
1052
+ exists (reorder m v). split => //.
1048
1053
Qed .
1049
1054
1050
1055
#[global]
1051
- Axiom trust_reorder_cstrs_transformation_ext : forall (efl : EEnvFlags) (wfl : WcbvFlags) (m : inductives_mapping),
1052
- TransformExt.t (reorder_cstrs_transformation efl wfl m)
1053
- (fun p p' => extends p.1 p'.1) (fun p p' => extends p.1 p'.1).
1056
+ Instance reorder_cstrs_transformation_ext {efl : EEnvFlags} (wca : cstr_as_blocks = false) (has_app : has_tApp) (wfl : WcbvFlags)
1057
+ {wcon : with_constructor_as_block = false} :
1058
+ TransformExt.t (reorder_cstrs_transformation (wca := wca) (has_app := has_app) wfl (wcon:=wcon))
1059
+ (fun p p' => p.1 = p'.1 /\ extends p.2.1 p'.2.1) (fun p p' => extends p.1 p'.1).
1060
+ Proof .
1061
+ red. intros p p' pr pr' [eq ext].
1062
+ red; cbn. rewrite -eq. eapply EReorderCstrs.optimize_extends_env; eauto.
1063
+ move: pr'; cbn. now intros []. apply pr. apply pr'.
1064
+ Qed .
1054
1065
1055
1066
#[global]
1056
- Axiom trust_reorder_cstrs_transformation_ext' : forall (efl : EEnvFlags) (wfl : WcbvFlags) (m : inductives_mapping),
1057
- TransformExt.t (reorder_cstrs_transformation efl wfl m)
1058
- extends_eprogram extends_eprogram.
1067
+ Instance reorder_cstrs_transformation_ext' {efl : EEnvFlags} (wca : cstr_as_blocks = false) (has_app : has_tApp) (wfl : WcbvFlags)
1068
+ {wcon : with_constructor_as_block = false} :
1069
+ TransformExt.t (reorder_cstrs_transformation (wca := wca) (has_app := has_app) wfl (wcon:=wcon))
1070
+ (fun p p' => p.1 = p'.1 /\ extends_eprogram_env p.2 p'.2) extends_eprogram.
1071
+ Proof .
1072
+ red. intros p p' pr pr' [eq [ext eq']]. cbn.
1073
+ red. cbn. rewrite -eq -eq'. split => //. eapply EReorderCstrs.optimize_extends_env; eauto.
1074
+ move: pr'; cbn. now intros []. apply pr. apply pr'.
1075
+ Qed .
1059
1076
1060
1077
From MetaCoq.Erasure Require Import EUnboxing.
1061
1078
0 commit comments