X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=matita%2Fmatita%2Flib%2Fre%2Freb.ma;h=2c5b129d228e7c87bb4d537c39e47a28cb22dab1;hb=d71e53021b0c17e1a00c2d623e7139c6d18069d5;hp=e1d85641f3736aad7286038c31cfedcd0996bbac;hpb=eaaea3c18083de3e442e939768ff450d3b093911;p=helm.git diff --git a/matita/matita/lib/re/reb.ma b/matita/matita/lib/re/reb.ma index e1d85641f..2c5b129d2 100644 --- a/matita/matita/lib/re/reb.ma +++ b/matita/matita/lib/re/reb.ma @@ -14,6 +14,7 @@ include "arithmetics/nat.ma". include "basics/lists/list.ma". +include "basics/core_notation/singl_1.ma". interpretation "iff" 'iff a b = (iff a b). @@ -86,6 +87,8 @@ definition cat : ∀S,l1,l2,w.Prop ≝ λS.λl1,l2.λw:word S.∃w1,w2.w1 @ w2 = w ∧ l1 w1 ∧ l2 w2. interpretation "cat lang" 'pc a b = (cat ? a b). +(* BEGIN HERE + definition star ≝ λS.λl.λw:word S.∃lw.flatten ? lw = w ∧ conjunct ? lw l. interpretation "star lang" 'pk l = (star ? l). @@ -253,11 +256,11 @@ definition reclose ≝ lift eclose. interpretation "reclose" 'eclose x = (reclose ? x). definition eq_f1 ≝ λS.λa,b:word S → Prop.∀w.a w ↔ b w. -notation "A =1 B" non associative with precedence 45 for @{'eq_f1 $A $B}. +notation "A ≐ B" non associative with precedence 45 for @{'eq_f1 $A $B}. interpretation "eq f1" 'eq_f1 a b = (eq_f1 ? a b). (* -lemma epsilon_or : ∀S:Alpha.∀b1,b2. ϵ(b1 || b2) =1 ϵ b1 ∪ ϵ b2. +lemma epsilon_or : ∀S:Alpha.∀b1,b2. ϵ(b1 || b2) ≐ ϵ b1 ∪ ϵ b2. ##[##2: napply S] #S b1 b2; ncases b1; ncases b2; napply extP; #w; nnormalize; @; /2/; *; //; *; nqed. @@ -269,7 +272,7 @@ nlemma cupC : ∀S. ∀a,b:word S → Prop.a ∪ b = b ∪ a. #S a b; napply extP; #w; @; *; nnormalize; /2/; nqed.*) (* theorem 16: 2 *) -lemma oplus_cup : ∀S:Alpha.∀e1,e2:pre S.\sem{e1 ⊕ e2} =1 \sem{e1} ∪ \sem{e2}. +lemma oplus_cup : ∀S:Alpha.∀e1,e2:pre S.\sem{e1 ⊕ e2} ≐ \sem{e1} ∪ \sem{e2}. #S * #i1 #b1 * #i2 #b2 >lo_def normalize in ⊢ (?%?); #w cases b1 cases b2 normalize % #w r1; ncases r1; #e1 b1 r2; ncases r2; #e2 b2; @@ -776,3 +779,4 @@ ntheorem der1: ∀S,a,e,e',w. der S a e e' → in_l S w e' → in_l S (a::w) e. | #x; #y; #z; #w; #a; #b; #c; #d; #K; ncases (?:False); /2/ ] ##| #r1; #r2; #r1'; #r2'; #H1; #H2; #H3; #H4; #H5; #H6; +*)