Skip to content

Commit 0986967

Browse files
committed
Show that inhabitants of product types evaluate to functions through the erasure pipeline
1 parent 5488b36 commit 0986967

File tree

1 file changed

+142
-2
lines changed

1 file changed

+142
-2
lines changed

erasure-plugin/theories/ErasureCorrectness.v

+142-2
Original file line numberDiff line numberDiff line change
@@ -837,7 +837,7 @@ Section ErasureFunction.
837837

838838
Lemma pcuic_function_value (wfl := default_wcbv_flags)
839839
{guard_impl : abstract_guard_impl}
840-
(cf:=config.extraction_checker_flags) (Σ:global_env_ext) f na A B
840+
(cf:=config.extraction_checker_flags) {Σ : global_env_ext} {f na A B}
841841
(wfΣ : wf_ext Σ) (axfree : axiom_free Σ) (wf : Σ ;;; [] |- f : PCUICAst.tProd na A B) : { v & PCUICWcbvEval.eval Σ f v }.
842842
Proof.
843843
eapply (PCUICNormalization.wcbv_normalization wfΣ axfree wf). Unshelve.
@@ -856,7 +856,7 @@ Section ErasureFunction.
856856
Proof.
857857
intros axfree nise.
858858
destruct pr as [[]]. destruct wf.
859-
epose proof (pcuic_function_value Σ f na A B w.1 axfree X) as [v hv].
859+
epose proof (pcuic_function_value w.1 axfree X) as [v hv].
860860
eapply erase_function; tea. now sq.
861861
Qed.
862862

@@ -1444,6 +1444,33 @@ Proof.
14441444
tc.
14451445
Qed.
14461446

1447+
1448+
Lemma expand_lets_function (wfl := default_wcbv_flags)
1449+
{guard_impl : abstract_guard_impl}
1450+
(cf:=config.extraction_checker_flags) (Σ:global_env_ext_map) K f na A B
1451+
(wf : ∥ Σ ;;; [] |- f : PCUICAst.tProd na A B ∥) pr :
1452+
let (Σ', f') := transform (pcuic_expand_lets_transform K) (Σ, f) pr in
1453+
∥ Σ' ;;; [] |- f' : PCUICExpandLets.trans (PCUICAst.tProd na A B) ∥.
1454+
Proof.
1455+
unfold transform, pcuic_expand_lets_transform. cbn.
1456+
destruct pr as [[[]] ?].
1457+
sq.
1458+
eapply PCUICExpandLetsCorrectness.pcuic_expand_lets in wf; eauto.
1459+
now eapply PCUICExpandLetsCorrectness.trans_wf.
1460+
Qed.
1461+
1462+
Lemma transform_lambda_box_firstorder (Σer : EEnvMap.GlobalContextMap.t) p pre :
1463+
firstorder_evalue Σer p ->
1464+
(transform verified_lambdabox_pipeline (Σer, p) pre).2 = (compile_evalue_box (ERemoveParams.strip Σer p) []).
1465+
Proof.
1466+
intros fo.
1467+
destruct lambdabox_pres_fo as [fn [tr hfn]].
1468+
rewrite (ETransformPresFO.transform_fo _ _ _ _ (t:=tr)).
1469+
now rewrite hfn.
1470+
Qed.
1471+
1472+
1473+
14471474
Lemma transform_lambda_box_firstorder (Σer : EEnvMap.GlobalContextMap.t) p pre :
14481475
firstorder_evalue Σer p ->
14491476
(transform verified_lambdabox_pipeline (Σer, p) pre).2 = (compile_evalue_box (ERemoveParams.strip Σer p) []).
@@ -1810,6 +1837,119 @@ Section PCUICErase.
18101837
now eapply (PCUICExpandLetsCorrectness.expand_lets_sound (cf := extraction_checker_flags)) in Hsort.
18111838
Qed.
18121839

1840+
#[local] Existing Instance PCUICSN.extraction_normalizing.
1841+
1842+
(* Beware, this internally uses preservation of observations and determinism of evaluation
1843+
from the canonical evaluation of [f] in the source to the evaluation in the target.
1844+
*)
1845+
Lemma transform_erasure_pipeline_function
1846+
(wfl := default_wcbv_flags)
1847+
{guard_impl : abstract_guard_impl}
1848+
(cf:=config.extraction_checker_flags) (Σ:global_env_ext_map)
1849+
{f v' na A B}
1850+
(wf : ∥ Σ ;;; [] |- f : PCUICAst.tProd na A B ∥) pr :
1851+
axiom_free Σ ->
1852+
∥ nisErasable Σ [] f ∥ ->
1853+
let tr := transform verified_erasure_pipeline (Σ, f) pr in
1854+
eval (wfl := extraction_wcbv_flags) tr.1 tr.2 v' -> isFunction v' = true.
1855+
Proof.
1856+
intros axfree nise.
1857+
unfold verified_erasure_pipeline.
1858+
rewrite -!transform_compose_assoc.
1859+
pose proof (expand_lets_function Σ (fun p : global_env_ext_map =>
1860+
(wf_ext p -> PCUICSN.NormalizationIn p) /\
1861+
(wf_ext p -> PCUICWeakeningEnvSN.normalizationInAdjustUniversesIn p)) f na A B wf pr).
1862+
destruct_compose. intros pre.
1863+
set (trexp := transform (pcuic_expand_lets_transform _) _ _) in *.
1864+
eapply (PCUICExpandLetsCorrectness.trans_axiom_free Σ) in axfree.
1865+
have nise' : ∥ nisErasable trexp.1 [] trexp.2 ∥.
1866+
destruct pr as [[[]] ?], nise. sq; now eapply nisErasable_lets.
1867+
change (trans_global_env _) with (global_env_ext_map_global_env_ext trexp.1).1 in axfree.
1868+
clearbody trexp. clear nise pr wf Σ f. destruct trexp as [Σ f].
1869+
pose proof pre as pre'; destruct pre' as [[[wf _]] _].
1870+
pose proof (map_squash (pcuic_function_value wf axfree) H) as [[v ev]].
1871+
epose proof (Transform.preservation erase_transform).
1872+
specialize (H0 _ v pre (sq ev)).
1873+
revert H0.
1874+
destruct_compose. intros pre' htr.
1875+
destruct htr as [v'' [ev' _]].
1876+
epose proof (erase_function_to_function _ f v'' _ _ _ H pre axfree nise').
1877+
set (tre := transform erase_transform _ _) in *. clearbody tre.
1878+
cbn -[transform obseq].
1879+
intros ev2. red in ev'. destruct ev'.
1880+
epose proof (Transform.preservation verified_lambdabox_pipeline).
1881+
destruct tre as [Σ' f'].
1882+
specialize (H2 _ v'' pre' (sq H1)) as [finalv [[evfinal] obseq]].
1883+
pose proof (eval_deterministic evfinal ev2). subst v'.
1884+
have prev : Transform.pre verified_lambdabox_pipeline (Σ', v'').
1885+
{ clear -wfl pre' H1. cbn in H1.
1886+
destruct pre' as [[] []]. split; split => //=.
1887+
eapply Prelim.Ee.eval_wellformed; eauto.
1888+
eapply EEtaExpandedFix.isEtaExp_expanded.
1889+
eapply (@EEtaExpandedFix.eval_etaexp wfl); eauto.
1890+
now eapply EEtaExpandedFix.expanded_global_env_isEtaExp_env.
1891+
now eapply EEtaExpandedFix.expanded_isEtaExp. }
1892+
specialize (H0 H1).
1893+
eapply (obseq_lambdabox (Σ', f') (Σ', v'')) in obseq.
1894+
epose proof (ETransformPresAppLam.transform_lam _ _ _ (t := lambdabox_pres_app) (Σ', v'') prev H0).
1895+
rewrite -obseq. exact H2. cbn. red; tauto.
1896+
Qed.
1897+
1898+
(* This version provides the evaluation proof. *)
1899+
Lemma transform_erasure_pipeline_function'
1900+
(wfl := default_wcbv_flags)
1901+
{guard_impl : abstract_guard_impl}
1902+
(cf:=config.extraction_checker_flags) (Σ:global_env_ext_map)
1903+
{f na A B}
1904+
(wf : ∥ Σ ;;; [] |- f : PCUICAst.tProd na A B ∥) pr :
1905+
axiom_free Σ ->
1906+
∥ nisErasable Σ [] f ∥ ->
1907+
let tr := transform verified_erasure_pipeline (Σ, f) pr in
1908+
exists v, ∥ eval (wfl := extraction_wcbv_flags) tr.1 tr.2 v ∥ /\ isFunction v = true.
1909+
Proof.
1910+
intros axfree nise.
1911+
unfold verified_erasure_pipeline.
1912+
rewrite -!transform_compose_assoc.
1913+
pose proof (expand_lets_function Σ (fun p : global_env_ext_map =>
1914+
(wf_ext p -> PCUICSN.NormalizationIn p) /\
1915+
(wf_ext p -> PCUICWeakeningEnvSN.normalizationInAdjustUniversesIn p)) f na A B wf pr).
1916+
destruct_compose. intros pre.
1917+
set (trexp := transform (pcuic_expand_lets_transform _) _ _) in *.
1918+
eapply (PCUICExpandLetsCorrectness.trans_axiom_free Σ) in axfree.
1919+
have nise' : ∥ nisErasable trexp.1 [] trexp.2 ∥.
1920+
destruct pr as [[[]] ?], nise. sq; now eapply nisErasable_lets.
1921+
change (trans_global_env _) with (global_env_ext_map_global_env_ext trexp.1).1 in axfree.
1922+
clearbody trexp. clear nise pr wf Σ f. destruct trexp as [Σ f].
1923+
pose proof pre as pre'; destruct pre' as [[[wf _]] _].
1924+
pose proof (map_squash (pcuic_function_value wf axfree) H) as [[v ev]].
1925+
epose proof (Transform.preservation erase_transform).
1926+
specialize (H0 _ v pre (sq ev)).
1927+
revert H0.
1928+
destruct_compose. intros pre' htr.
1929+
destruct htr as [v'' [ev' _]].
1930+
epose proof (erase_function_to_function _ f v'' _ _ _ H pre axfree nise').
1931+
set (tre := transform erase_transform _ _) in *. clearbody tre.
1932+
cbn -[transform obseq].
1933+
red in ev'. destruct ev'.
1934+
epose proof (Transform.preservation verified_lambdabox_pipeline).
1935+
destruct tre as [Σ' f'].
1936+
specialize (H2 _ v'' pre' (sq H1)) as [finalv [[evfinal] obseq]].
1937+
exists finalv.
1938+
split. now sq.
1939+
have prev : Transform.pre verified_lambdabox_pipeline (Σ', v'').
1940+
{ clear -wfl pre' H1. cbn in H1.
1941+
destruct pre' as [[] []]. split; split => //=.
1942+
eapply Prelim.Ee.eval_wellformed; eauto.
1943+
eapply EEtaExpandedFix.isEtaExp_expanded.
1944+
eapply (@EEtaExpandedFix.eval_etaexp wfl); eauto.
1945+
now eapply EEtaExpandedFix.expanded_global_env_isEtaExp_env.
1946+
now eapply EEtaExpandedFix.expanded_isEtaExp. }
1947+
specialize (H0 H1).
1948+
eapply (obseq_lambdabox (Σ', f') (Σ', v'')) in obseq.
1949+
epose proof (ETransformPresAppLam.transform_lam _ _ _ (t := lambdabox_pres_app) (Σ', v'') prev H0).
1950+
rewrite -obseq. exact H2. cbn. red; tauto.
1951+
Qed.
1952+
18131953
Lemma expand_lets_transform_env K p p' pre pre' :
18141954
p.1 = p'.1 ->
18151955
(transform (pcuic_expand_lets_transform K) p pre).1 =

0 commit comments

Comments
 (0)