]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/reduction/fpbc.ma
A first example that uses a status monad where the status is a tree.
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / reduction / fpbc.ma
index cb391a6c932c96dc276ab03eacff43390a182698..45e671601a94511c930061b812f5c5be14089584 100644 (file)
@@ -13,7 +13,6 @@
 (**************************************************************************)
 
 include "basic_2/notation/relations/btpredproper_8.ma".
-include "basic_2/relocation/fquq_alt.ma".
 include "basic_2/reduction/fpb.ma".
 
 (* "BIG TREE" PROPER PARALLEL REDUCTION FOR CLOSURES ************************)
@@ -29,26 +28,14 @@ interpretation
 
 (* Basic properties *********************************************************)
 
-lemma fpbc_fpb: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, g] ⦃G2, L2, T2⦄ →
-                ⦃G1, L1, T1⦄ ≽[h, g] ⦃G2, L2, T2⦄.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2
-/3 width=1 by fpb_fquq, fpb_cpx, fqu_fquq/
-qed.
-
 lemma cpr_fpbc: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → (T1 = T2 → ⊥) →
                 ⦃G, L, T1⦄ ≻[h, g] ⦃G, L, T2⦄.
 /3 width=1 by fpbc_cpx, cpr_cpx/ qed.
 
-(* Inversion lemmas on "big tree" parallel reduction for closures ***********)
+(* Basic forward lemmas *****************************************************)
 
-lemma fpb_inv_fpbc: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≽[h, g] ⦃G2, L2, T2⦄ →
-                    ⦃G1, L1, T1⦄ ≻[h, g] ⦃G2, L2, T2⦄ ∨
-                    ∧∧ G1 = G2 & ⦃G1, L1⦄ ⊢ ➡[h, g] L2 & T1 = T2.
+lemma fpbc_fwd_fpb: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, g] ⦃G2, L2, T2⦄ →
+                    ⦃G1, L1, T1⦄ ≽[h, g] ⦃G2, L2, T2⦄.
 #h #g #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2
-/3 width=1 by and3_intro, or_intror/
-[ #G2 #L2 #T2 #H elim (fquq_inv_gen … H) -H [| * ]
-  /3 width=1 by fpbc_fqu, and3_intro, or_introl, or_intror/
-| #T2 #HT12 elim (term_eq_dec T1 T2) #H destruct
-  /4 width=1 by and3_intro, or_introl, or_intror, fpbc_cpx/
-]
+/3 width=1 by fpb_fquq, fpb_cpx, fqu_fquq/
 qed-.