]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/computation/csx_aaa.ma
reorganization of the "static" component:
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / csx_aaa.ma
index f1b70f0c73ad6a8f87b07cb81fbd45ace604cb31..0a42fb1f10e12612b5b2f29ee9badbcbf52228b8 100644 (file)
@@ -32,7 +32,7 @@ fact aaa_ind_csx_aux: ∀h,g,G,L,A. ∀R:predicate term.
                             (∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → (T1 = T2 → ⊥) → R T2) → R T1
                       ) →
                       ∀T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → ⦃G, L⦄ ⊢ T ⁝ A → R T.
-#h #g #G #L #A #R #IH #T #H @(csx_ind … H) -T /4 width=5 by aaa_cpx_conf/
+#h #g #G #L #A #R #IH #T #H @(csx_ind … H) -T /4 width=5 by cpx_aaa_conf/
 qed-.
 
 lemma aaa_ind_csx: ∀h,g,G,L,A. ∀R:predicate term.
@@ -47,7 +47,7 @@ fact aaa_ind_csx_alt_aux: ∀h,g,G,L,A. ∀R:predicate term.
                                 (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → (T1 = T2 → ⊥) → R T2) → R T1
                           ) →
                           ∀T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → ⦃G, L⦄ ⊢ T ⁝ A → R T.
-#h #g #G #L #A #R #IH #T #H @(csx_ind_alt … H) -T /4 width=5 by aaa_cpxs_conf/
+#h #g #G #L #A #R #IH #T #H @(csx_ind_alt … H) -T /4 width=5 by cpxs_aaa_conf/
 qed-.
 
 lemma aaa_ind_csx_alt: ∀h,g,G,L,A. ∀R:predicate term.