From: Ferruccio Guidi Date: Sat, 8 Dec 2012 23:26:45 +0000 (+0000) Subject: - list.ma: improved notation for constant lists (a "term 19" was missing) X-Git-Tag: make_still_working~1402 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=23e75ebc00553e178e090ca1373ac075ee650a60;hp=178820be7648a60af17837727e51fd1f3f2791db;p=helm.git - list.ma: improved notation for constant lists (a "term 19" was missing) - lambda: more properties on pointers, extra xoa quantifier removed, bugfix in notations and names --- diff --git a/matita/matita/contribs/lambda/labelled_hap_computation.ma b/matita/matita/contribs/lambda/labelled_hap_computation.ma index 70cc06b79..5cd7678d1 100644 --- a/matita/matita/contribs/lambda/labelled_hap_computation.ma +++ b/matita/matita/contribs/lambda/labelled_hap_computation.ma @@ -29,7 +29,7 @@ notation "hvbox( M break ⓗ⇀* [ term 46 p ] break term 46 N )" non associative with precedence 45 for @{ 'HApStar $M $p $N }. -lemma lhap1_lhap: ∀p,M1,M2. M1 ⓗ⇀[p] M2 → M1 ⓗ⇀*[p::◊] M2. +lemma lhap_step_rc: ∀p,M1,M2. M1 ⓗ⇀[p] M2 → M1 ⓗ⇀*[p::◊] M2. /2 width=1/ qed. @@ -42,7 +42,7 @@ lemma lhap_inv_cons: ∀s,M1,M2. M1 ⓗ⇀*[s] M2 → ∀q,r. q::r = s → /2 width=3 by lstar_inv_cons/ qed-. -lemma lhap_inv_lhap1: ∀p,M1,M2. M1 ⓗ⇀*[p::◊] M2 → M1 ⓗ⇀[p] M2. +lemma lhap_inv_step_rc: ∀p,M1,M2. M1 ⓗ⇀*[p::◊] M2 → M1 ⓗ⇀[p] M2. /2 width=1 by lstar_inv_step/ qed-. @@ -87,12 +87,3 @@ lemma lhap_inv_lsreds: ∀s,M1,M2. M1 ⓗ⇀*[s] M2 → M1 ⇀*[s] M2. #p #s #M1 #M #HM1 #_ #IHM2 lapply (lhap1_inv_lsred … HM1) -HM1 /2 width=3/ qed-. - -lemma lhap_fwd_le: ∀s,M1,M2. M1 ⓗ⇀*[s] M2 → is_le s. -#s #M1 #M2 #H @(lstar_ind_l ????????? H) -s -M1 /3 width=3/ -#a1 #s #M1 #M #HM1 #IHM1 -generalize in match HM1; -HM1 -cases IHM1 -s -M -M2 // -#a #M0 #M #HM0 #s #M2 #_ #HM10 #H -M2 -lapply (lhap1_fwd_le … HM10 … HM0) -M -M0 -M1 /2 width=1/ -qed-. diff --git a/matita/matita/contribs/lambda/labelled_hap_reduction.ma b/matita/matita/contribs/lambda/labelled_hap_reduction.ma index 575abd342..e1f279f67 100644 --- a/matita/matita/contribs/lambda/labelled_hap_reduction.ma +++ b/matita/matita/contribs/lambda/labelled_hap_reduction.ma @@ -48,30 +48,6 @@ lemma lhap1_inv_cons: ∀p,M,N. M ⓗ⇀[p] N → ∀c,q. c::q = p → ] qed-. -lemma lhap1_inv_abst_sn: ∀p,M,N. M ⓗ⇀[p] N → ∀A. 𝛌.A = M → ⊥. -#p #M #N * -p -M -N -[ #B #A #A0 #H destruct -| #p #B #A1 #A2 #_ #A0 #H destruct -] -qed-. - -lemma lhap1_inv_appl_sn: ∀p,M,N. M ⓗ⇀[p] N → ∀B,A. @B.A = M → - (∃∃C. ◊ = p & 𝛌.C = A & [⬐B]C = N) ∨ - ∃∃q,C. A ⓗ⇀[q] C & dx::q = p & @B.C = N. -#p #M #N * -p -M -N -[ #B #A #B0 #A0 #H destruct /3 width=3/ -| #p #B #A1 #A2 #HA12 #B0 #A0 #H destruct /3 width=5/ -] -qed-. - -lemma lhap1_inv_abst_dx: ∀p,M,N. M ⓗ⇀[p] N → ∀C. 𝛌.C = N → - ∃∃B,A. ◊ = p & @B.𝛌.A = M & 𝛌.C = [⬐B]A. -#p #M #N * -p -M -N -[ #B #A #C #H /2 width=4/ -| #p #B #A1 #A2 #_ #C #H destruct -] -qed-. - lemma lhap1_lift: ∀p. liftable (lhap1 p). #p #h #M1 #M2 #H elim H -p -M1 -M2 normalize /2 width=1/ #B #A #d xoa_notation basics/pts.ma 2 2 - 3 1 3 2 3 3 4 3 diff --git a/matita/matita/contribs/lambda/xoa.ma b/matita/matita/contribs/lambda/xoa.ma index 9d4ee880b..1e82edec7 100644 --- a/matita/matita/contribs/lambda/xoa.ma +++ b/matita/matita/contribs/lambda/xoa.ma @@ -24,14 +24,6 @@ inductive ex2_2 (A0,A1:Type[0]) (P0,P1:A0→A1→Prop) : Prop ≝ interpretation "multiple existental quantifier (2, 2)" 'Ex P0 P1 = (ex2_2 ? ? P0 P1). -(* multiple existental quantifier (3, 1) *) - -inductive ex3_1 (A0:Type[0]) (P0,P1,P2:A0→Prop) : Prop ≝ - | ex3_1_intro: ∀x0. P0 x0 → P1 x0 → P2 x0 → ex3_1 ? ? ? ? -. - -interpretation "multiple existental quantifier (3, 1)" 'Ex P0 P1 P2 = (ex3_1 ? P0 P1 P2). - (* multiple existental quantifier (3, 2) *) inductive ex3_2 (A0,A1:Type[0]) (P0,P1,P2:A0→A1→Prop) : Prop ≝ diff --git a/matita/matita/contribs/lambda/xoa_notation.ma b/matita/matita/contribs/lambda/xoa_notation.ma index 6151aa044..b14f46cec 100644 --- a/matita/matita/contribs/lambda/xoa_notation.ma +++ b/matita/matita/contribs/lambda/xoa_notation.ma @@ -24,16 +24,6 @@ notation < "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0 break & term 19 non associative with precedence 20 for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.$P1) }. -(* multiple existental quantifier (3, 1) *) - -notation > "hvbox(∃∃ ident x0 break . term 19 P0 break & term 19 P1 break & term 19 P2)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}.$P0) (λ${ident x0}.$P1) (λ${ident x0}.$P2) }. - -notation < "hvbox(∃∃ ident x0 break . term 19 P0 break & term 19 P1 break & term 19 P2)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}:$T0.$P0) (λ${ident x0}:$T0.$P1) (λ${ident x0}:$T0.$P2) }. - (* multiple existental quantifier (3, 2) *) notation > "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0 break & term 19 P1 break & term 19 P2)" diff --git a/matita/matita/lib/basics/lists/list.ma b/matita/matita/lib/basics/lists/list.ma index 7429625dc..c1f6ab738 100644 --- a/matita/matita/lib/basics/lists/list.ma +++ b/matita/matita/lib/basics/lists/list.ma @@ -20,7 +20,7 @@ notation "hvbox(hd break :: tl)" right associative with precedence 47 for @{'cons $hd $tl}. -notation "[ list0 x sep ; ]" +notation "[ list0 term 19 x sep ; ]" non associative with precedence 90 for ${fold right @'nil rec acc @{'cons $x $acc}}.