(* 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 ≝
definition l3_t1: prototerm ≝
(𝛌.@(⧣𝟏).@(𝛌.@(⧣𝟐).⧣𝟏).𝛌.(𝛌.@(⧣↑𝟐).⧣𝟏)).
+definition l3_t2: prototerm ≝
+ (𝛌.@(⧣𝟏).@(𝛌.@(⧣𝟐).⧣𝟏).𝛌.(𝛌.@(⧣↑𝟐).⧣↑𝟐)).
+
lemma l3_t01:
l3_t0 ➡𝐢𝐛𝐟[𝗟◗𝗔◗𝗔◗𝗟◗𝐞] l3_t1.
@ibfr_abst_hd
@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.
q ϵ 𝐂❨Ⓕ,n❩ → q◖𝗱↑n ϵ t →
@v.𝛌.t ➡𝐢𝐛𝐟[𝗔◗𝗟◗q] @v.𝛌.(t[⋔q←🠡[𝐮❨↑n❩]v]).
#v #t #q #n #Hn #Ht
-@(ex6_5_intro … (𝐞) (𝐞) q (𝟎) … Hn) -Hn //
-[ /3 width=1 by in_comp_appl_hd, in_comp_abst_hd/
+@(ex6_5_intro … (𝐞) (𝐞) q (𝟎) … Hn) -Hn
+[ //
+| //
+| //
+| /3 width=1 by in_comp_appl_hd, in_comp_abst_hd/
| @(subset_eq_canc_sn … (fsubst_appl_hd …)) @appl_eq_repl [ // ]
@(subset_eq_canc_sn … (fsubst_abst_hd …)) @abst_eq_repl
@fsubst_eq_repl // <nplus_zero_sn @lift_term_eq_repl_dx
@(subset_eq_canc_sn … (grafted_empty_dx … )) //
]
qed.
+
+lemma ibfr_beta_1 (v) (v1) (t) (q) (n):
+ q ϵ 𝐂❨Ⓕ,n❩ → q◖𝗱↑n ϵ t →
+ @v.@v1.𝛌.𝛌.t ➡𝐢𝐛𝐟[𝗔◗𝗔◗𝗟◗𝗟◗q] @v.@v1.𝛌.𝛌.(t[⋔q←🠡[𝐮❨↑↑n❩]v]).
+#v #v1 #t #q #n #Hn #Ht
+@(ex6_5_intro … (𝐞) (𝗔◗𝗟◗𝐞) q (𝟏) … Hn) -Hn
+[ //
+| /2 width=1 by pbc_empty, pbc_redex/
+| /3 width=1 by pcc_A_sn, pcc_L_sn, pcc_empty/
+| /5 width=1 by in_comp_appl_hd, in_comp_abst_hd/
+| @(subset_eq_canc_sn … (fsubst_appl_hd …)) @appl_eq_repl [ // ]
+ @(subset_eq_canc_sn … (fsubst_appl_hd …)) @appl_eq_repl [ // ]
+ @(subset_eq_canc_sn … (fsubst_abst_hd …)) @abst_eq_repl
+ @(subset_eq_canc_sn … (fsubst_abst_hd …)) @abst_eq_repl
+ @fsubst_eq_repl // <nplus_unit_sn @lift_term_eq_repl_dx
+ >list_cons_comm @(subset_eq_canc_sn … (grafted_appl_sd … ))
+ @(subset_eq_canc_sn … (grafted_empty_dx … )) //
+]
+qed.