]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/delayed_updating/reduction/dbfr_unprotected.ma
update in delayed_updating
[helm.git] / matita / matita / contribs / lambdadelta / delayed_updating / reduction / dbfr_unprotected.ma
index a0e5a6944858a6f549aba27fdf9452c82a9964b3..7c92760210cf1efaf2e1bb7edf61ba51e3870519 100644 (file)
@@ -18,6 +18,12 @@ include "ground/arith/pnat_two.ma".
 
 (* DELAYED BALANCED FOCUSED REDUCTION ***************************************)
 
+(* Prerequisites ************************************************************)
+
+lemma list_rcons_prop_1 (A) (a1) (a2):
+      ⓔ ⨭ a1 ⨭{A} a2 = a1 ⨮ (ⓔ ⨭ a2).
+// qed.
+
 (* Example of unprotected balanced β-reduction ******************************)
 
 definition l3_t0: prototerm ≝
@@ -26,6 +32,9 @@ definition l3_t0: prototerm ≝
 definition l3_t1: prototerm ≝
            (𝛌.@(⧣𝟏).@(𝛌.@(⧣𝟐).⧣𝟏).𝛌.(𝛌.@(⧣↑𝟐).⧣𝟏)).
 
+definition l3_t2: prototerm ≝
+           (𝛌.@(⧣𝟏).@(𝛌.@(⧣𝟐).⧣𝟏).𝛌.(𝛌.@(⧣↑𝟐).⧣↑𝟐)).
+
 lemma l3_t01:
       l3_t0 ➡𝐢𝐛𝐟[𝗟◗𝗔◗𝗔◗𝗟◗𝐞] l3_t1.
 @ibfr_abst_hd
@@ -40,3 +49,20 @@ lemma l3_t01:
 @appl_eq_repl
 @(subset_eq_canc_sn … (lift_term_oref_pap … )) //
 qed.
+
+lemma l3_t12:
+      l3_t1 ➡𝐢𝐛𝐟[𝗟◗𝗔◗𝗔◗𝗟◗𝗟◗𝗔◗𝐞] l3_t2.
+@ibfr_abst_hd
+@ibfr_eq_trans
+[| @(ibfr_beta_1 … (𝟎)) [| <list_rcons_prop_1 ]
+   /2 width=3 by pcc_A_sn, in_comp_appl_hd/
+]
+@appl_eq_repl [ // ]
+@appl_eq_repl [ // ]
+@abst_eq_repl
+@abst_eq_repl
+@(subset_eq_canc_sn … (fsubst_appl_hd …))
+@appl_eq_repl [ // ]
+@(subset_eq_canc_sn … (fsubst_empty_rc …))
+@(subset_eq_canc_sn … (lift_term_oref_pap … )) //
+qed.