]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/static/lsubr_lsubr.ma
renaming in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / static / lsubr_lsubr.ma
index 787a1937799a4559422b4452f2c0b2cbe9567586..ac9691aa618c74b04152c412364aa98e026fbd3d 100644 (file)
@@ -16,38 +16,17 @@ include "basic_2/static/lsubr.ma".
 
 (* RESTRICTED REFINEMENT FOR LOCAL ENVIRONMENTS *****************************)
 
-(* Auxiliary inversion lemmas ***********************************************)
-
-fact lsubr_inv_pair1_aux: ∀L1,L2. L1 ⫃ L2 → ∀I,K1,X. L1 = K1.ⓑ{I}X →
-                          ∨∨ L2 = ⋆
-                           | ∃∃K2. K1 ⫃ K2 & L2 = K2.ⓑ{I}X
-                           | ∃∃K2,V,W. K1 ⫃ K2 & L2 = K2.ⓛW &
-                                       I = Abbr & X = ⓝW.V.
-#L1 #L2 * -L1 -L2
-[ #L #J #K1 #X #H destruct /2 width=1 by or3_intro0/
-| #I #L1 #L2 #V #HL12 #J #K1 #X #H destruct /3 width=3 by or3_intro1, ex2_intro/
-| #L1 #L2 #V #W #HL12 #J #K1 #X #H destruct /3 width=6 by or3_intro2, ex4_3_intro/
-]
-qed-.
-
-lemma lsubr_inv_pair1: ∀I,K1,L2,X. K1.ⓑ{I}X ⫃ L2 →
-                       ∨∨ L2 = ⋆
-                        | ∃∃K2. K1 ⫃ K2 & L2 = K2.ⓑ{I}X
-                        | ∃∃K2,V,W. K1 ⫃ K2 & L2 = K2.ⓛW &
-                                    I = Abbr & X = ⓝW.V.
-/2 width=3 by lsubr_inv_pair1_aux/ qed-.
-
 (* Main properties **********************************************************)
 
 theorem lsubr_trans: Transitive … lsubr.
-#L1 #L #H elim H -L1 -L
-[ #L1 #X #H
-  lapply (lsubr_inv_atom1 … H) -H //
-| #I #L1 #L #V #_ #IHL1 #X #H
-  elim (lsubr_inv_pair1 … H) -H // *
-  #L2 [2: #V2 #W2 ] #HL2 #H1 [ #H2 #H3 ] destruct /3 width=1 by lsubr_pair, lsubr_beta/
-| #L1 #L #V1 #W #_ #IHL1 #X #H
-  elim (lsubr_inv_abst1 … H) -H // *
-  #L2 #HL2 #H destruct /3 width=1 by lsubr_beta/
+#L1 #L #H elim H -L1 -L //
+[ #I #L1 #L #_ #IH #X #H elim (lsubr_inv_bind1 … H) -H *
+  [ #L2 #HL2 #H | #L2 #V #W #HL2 #H1 #H2 | #I1 #I2 #L2 #V #Hl2 #H1 #H2 ]
+  destruct /3 width=1 by lsubr_bind, lsubr_beta, lsubr_unit/
+| #L1 #L #V #W #_ #IH #X #H elim (lsubr_inv_abst1 … H) -H *
+  [ #L2 #HL2 #H | #I #L2 #HL2 #H ]
+  destruct /3 width=1 by lsubr_beta, lsubr_unit/
+| #I1 #I2 #L1 #L #V #_ #IH #X #H elim (lsubr_inv_unit1 … H) -H
+  /4 width=1 by lsubr_unit/
 ]
 qed-.