X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fstatic%2Faaa_fqus.ma;h=812bcc3656984c1fb38363dd1e635f01346f699b;hp=576b0814e3c99f9189441d05e8feabce7e0caef7;hb=222044da28742b24584549ba86b1805a87def070;hpb=52e675f555f559c047d5449db7fc89a51b977d35 diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_fqus.ma b/matita/matita/contribs/lambdadelta/basic_2/static/aaa_fqus.ma index 576b0814e..812bcc365 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_fqus.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/aaa_fqus.ma @@ -12,8 +12,8 @@ (* *) (**************************************************************************) -include "basic_2/multiple/fqus_alt.ma". -include "basic_2/static/aaa_lift.ma". +include "basic_2/s_computation/fqus_fqup.ma". +include "basic_2/static/aaa_drops.ma". (* ATONIC ARITY ASSIGNMENT ON TERMS *****************************************) @@ -22,30 +22,30 @@ include "basic_2/static/aaa_lift.ma". lemma aaa_fqu_conf: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → ∀A1. ⦃G1, L1⦄ ⊢ T1 ⁝ A1 → ∃A2. ⦃G2, L2⦄ ⊢ T2 ⁝ A2. #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 -[ #I #G #L #T #A #H elim (aaa_inv_lref … H) -H - #J #K #V #H #HA lapply (drop_inv_O2 … H) -H - #H destruct /2 width=2 by ex_intro/ -| * [ #a ] * #G #L #V #T #X #H +[ #I #G #L #T #A #H elim (aaa_inv_zero … H) -H + #J #K #V #H #HA destruct /2 width=2 by ex_intro/ +| * [ #p ] * #G #L #V #T #X #H [ elim (aaa_inv_abbr … H) | elim (aaa_inv_abst … H) | elim (aaa_inv_appl … H) | elim (aaa_inv_cast … H) ] -H /2 width=2 by ex_intro/ -| #a * #G #L #V #T #X #H +| #p * #G #L #V #T #X #H [ elim (aaa_inv_abbr … H) | elim (aaa_inv_abst … H) ] -H /2 width=2 by ex_intro/ +| #p #I #G #L #V #T #H destruct | * #G #L #V #T #X #H [ elim (aaa_inv_appl … H) | elim (aaa_inv_cast … H) ] -H /2 width=2 by ex_intro/ -| /3 width=9 by aaa_inv_lift, ex_intro/ +| /5 width=8 by aaa_inv_lifts, drops_refl, drops_drop, ex_intro/ ] qed-. lemma aaa_fquq_conf: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → ∀A1. ⦃G1, L1⦄ ⊢ T1 ⁝ A1 → ∃A2. ⦃G2, L2⦄ ⊢ T2 ⁝ A2. -#G1 #G2 #L1 #L2 #T1 #T2 #H elim(fquq_inv_gen … H) -H /2 width=6 by aaa_fqu_conf/ +#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -H /2 width=6 by aaa_fqu_conf/ * #H1 #H2 #H3 destruct /2 width=2 by ex_intro/ qed-. @@ -58,6 +58,6 @@ qed-. lemma aaa_fqus_conf: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ → ∀A1. ⦃G1, L1⦄ ⊢ T1 ⁝ A1 → ∃A2. ⦃G2, L2⦄ ⊢ T2 ⁝ A2. -#G1 #G2 #L1 #L2 #T1 #T2 #H elim(fqus_inv_gen … H) -H /2 width=6 by aaa_fqup_conf/ +#G1 #G2 #L1 #L2 #T1 #T2 #H elim(fqus_inv_fqup … H) -H /2 width=6 by aaa_fqup_conf/ * #H1 #H2 #H3 destruct /2 width=2 by ex_intro/ qed-.