From: matitaweb Date: Fri, 2 Mar 2012 17:42:35 +0000 (+0000) Subject: commit by user andrea X-Git-Tag: make_still_working~1912 X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=commitdiff_plain;h=000dc5a8de79b2ab63a49cf0f9db2b540cc05bcf commit by user andrea --- diff --git a/weblib/tutorial/chapter8.ma b/weblib/tutorial/chapter8.ma index 838a5f9f6..31acc919e 100644 --- a/weblib/tutorial/chapter8.ma +++ b/weblib/tutorial/chapter8.ma @@ -103,209 +103,206 @@ interpretation "eclose" 'eclose x = (eclose ? x). (* Here are a few simple properties of ▹ and •(-) *) -lemma pcl_true : - ∀S.∀i1,i2:pitem S. - 〈i1,true〉 ▸ i2 = i1 ◂ (•i2). +lemma pcl_true : ∀S.∀i1,i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. + a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1,a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a〉 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a i1 a title="pre_concat_r" href="cic:/fakeuri.def(1)"◃/a (a title="eclose" href="cic:/fakeuri.def(1)"•/ai2). // qed. -lemma pcl_true_bis : - ∀S.∀i1,i2:pitem S. - 〈i1,true〉 ▸ i2 = 〈i1 · \fst (•i2), \snd (•i2)〉. -#S #i1 #i2 normalize cases (•i2) // qed. +lemma pcl_true_bis : ∀S.∀i1,i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. + a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1,a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a〉 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1 a title="pitem cat" href="cic:/fakeuri.def(1)"·/a a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a title="eclose" href="cic:/fakeuri.def(1)"•/ai2), a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a (a title="eclose" href="cic:/fakeuri.def(1)"•/ai2)〉. +#S #i1 #i2 normalize cases (a title="eclose" href="cic:/fakeuri.def(1)"•/ai2) // qed. -lemma pcl_false: - ∀S.∀i1,i2:pitem S. - 〈i1,false〉 ▸ i2 = 〈i1 · i2, false〉. +lemma pcl_false: ∀S.∀i1,i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. + a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1,a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a〉 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1 a title="pitem cat" href="cic:/fakeuri.def(1)"·/a i2, a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a〉. // qed. -lemma eclose_plus: ∀S:DeqSet.∀i1,i2:pitem S. - •(i1 + i2) = •i1 ⊕ •i2. +lemma eclose_plus: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀i1,i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. + a title="eclose" href="cic:/fakeuri.def(1)"•/a(i1 a title="pitem or" href="cic:/fakeuri.def(1)"+/a i2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="eclose" href="cic:/fakeuri.def(1)"•/ai1 a title="oplus" href="cic:/fakeuri.def(1)"⊕/a a title="eclose" href="cic:/fakeuri.def(1)"•/ai2. // qed. -lemma eclose_dot: ∀S:DeqSet.∀i1,i2:pitem S. - •(i1 · i2) = •i1 ▹ i2. +lemma eclose_dot: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀i1,i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. + a title="eclose" href="cic:/fakeuri.def(1)"•/a(i1 a title="pitem cat" href="cic:/fakeuri.def(1)"·/a i2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="eclose" href="cic:/fakeuri.def(1)"•/ai1 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i2. // qed. -lemma eclose_star: ∀S:DeqSet.∀i:pitem S. - •i^* = 〈(\fst(•i))^*,true〉. +lemma eclose_star: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. + a title="eclose" href="cic:/fakeuri.def(1)"•/aia title="pitem star" href="cic:/fakeuri.def(1)"^/a* a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a(a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a(a title="eclose" href="cic:/fakeuri.def(1)"•/ai))a title="pitem star" href="cic:/fakeuri.def(1)"^/a*,a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a〉. // qed. (* The definition of •(-) (eclose) can then be lifted from items to pres in the obvious way. *) -definition lift ≝ λS.λf:pitem S →pre S.λe:pre S. +definition lift ≝ λS.λf:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S →a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S.λe:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. match e with - [ mk_Prod i b ⇒ 〈\fst (f i), \snd (f i) ∨ b〉]. + [ mk_Prod i b ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (f i), a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a (f i) a title="boolean or" href="cic:/fakeuri.def(1)"∨/a b〉]. -definition preclose ≝ λS. lift S (eclose S). +definition preclose ≝ λS. a href="cic:/matita/tutorial/chapter8/lift.def(2)"lift/a S (a href="cic:/matita/tutorial/chapter8/eclose.fix(0,1,4)"eclose/a S). interpretation "preclose" 'eclose x = (preclose ? x). (* Obviously, broadcasting does not change the carrier of the item, as it is easily proved by structural induction. *) -lemma erase_bull : ∀S.∀i:pitem S. |\fst (•i)| = |i|. +lemma erase_bull : ∀S.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a title="eclose" href="cic:/fakeuri.def(1)"•/ai)| a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="forget" href="cic:/fakeuri.def(1)"|/ai|. #S #i elim i // - [ #i1 #i2 #IH1 #IH2 >erase_dot eclose_dot - cases (•i1) #i11 #b1 cases b1 // odot_true_bis // - | #i1 #i2 #IH1 #IH2 >eclose_plus >(erase_plus … i1) eclose_star >(erase_star … i) a href="cic:/matita/tutorial/chapter7/erase_dot.def(4)"erase_dot/a a href="cic:/matita/tutorial/chapter8/eclose_dot.def(5)"eclose_dot/a + cases (a title="eclose" href="cic:/fakeuri.def(1)"•/ai1) #i11 #b1 cases b1 // a href="cic:/matita/tutorial/chapter8/pcl_true_bis.def(5)"pcl_true_bis/a // + | #i1 #i2 #IH1 #IH2 >a href="cic:/matita/tutorial/chapter8/eclose_plus.def(5)"eclose_plus/a >(a href="cic:/matita/tutorial/chapter7/erase_plus.def(4)"erase_plus/a … i1) a href="cic:/matita/tutorial/chapter8/eclose_star.def(5)"eclose_star/a >(a href="cic:/matita/tutorial/chapter7/erase_star.def(4)"erase_star/a … i) odot_false >sem_pre_false >sem_pre_false >sem_cat /2/ - |#H >odot_true >sem_pre_true @(eqP_trans … (sem_pre_concat_r …)) - >erase_bull @eqP_trans [|@(eqP_union_l … H)] - @eqP_trans [|@eqP_union_l[|@union_comm ]] - @eqP_trans [|@eqP_sym @union_assoc ] /3/ + [2:#th >a href="cic:/matita/tutorial/chapter8/pcl_false.def(5)"pcl_false/a >a href="cic:/matita/tutorial/chapter7/sem_pre_false.def(9)"sem_pre_false/a >a href="cic:/matita/tutorial/chapter7/sem_pre_false.def(9)"sem_pre_false/a >a href="cic:/matita/tutorial/chapter7/sem_cat.def(8)"sem_cat/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/tutorial/chapter8/eq_to_ex_eq.def(4)"eq_to_ex_eq/a/span/span/ + |#H >a href="cic:/matita/tutorial/chapter8/pcl_true.def(5)"pcl_true/a >a href="cic:/matita/tutorial/chapter7/sem_pre_true.def(9)"sem_pre_true/a @(a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a … (a href="cic:/matita/tutorial/chapter8/sem_pre_concat_r.def(10)"sem_pre_concat_r/a …)) + >a href="cic:/matita/tutorial/chapter8/erase_bull.def(6)"erase_bull/a @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@(a href="cic:/matita/tutorial/chapter4/eqP_union_l.def(3)"eqP_union_l/a … H)] + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@a href="cic:/matita/tutorial/chapter4/eqP_union_l.def(3)"eqP_union_l/a[|@a href="cic:/matita/tutorial/chapter4/union_comm.def(3)"union_comm/a ]] + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@a href="cic:/matita/tutorial/chapter4/eqP_sym.def(3)"eqP_sym/a @a href="cic:/matita/tutorial/chapter4/union_assoc.def(3)"union_assoc/a ] /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/tutorial/chapter4/eqP_union_r.def(3)"eqP_union_r/a, a href="cic:/matita/tutorial/chapter4/eqP_sym.def(3)"eqP_sym/a/span/span/ ] qed. -lemma minus_eps_pre_aux: ∀S.∀e:pre S.∀i:pitem S.∀A. - \sem{e} =1 \sem{i} ∪ A → \sem{\fst e} =1 \sem{i} ∪ (A - {[ ]}). +lemma minus_eps_pre_aux: ∀S.∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S.∀A. + a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e} a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{i} a title="union" href="cic:/fakeuri.def(1)"∪/a A → a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e} a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{i} a title="union" href="cic:/fakeuri.def(1)"∪/a (A a title="substraction" href="cic:/fakeuri.def(1)"-/a a title="singleton" href="cic:/fakeuri.def(1)"{/aa title="nil" href="cic:/fakeuri.def(1)"[/a ]}). #S #e #i #A #seme -@eqP_trans [|@minus_eps_pre] -@eqP_trans [||@eqP_union_r [|@eqP_sym @minus_eps_item]] -@eqP_trans [||@distribute_substract] -@eqP_substract_r // +@a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@a href="cic:/matita/tutorial/chapter7/minus_eps_pre.def(10)"minus_eps_pre/a] +@a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [||@a href="cic:/matita/tutorial/chapter4/eqP_union_r.def(3)"eqP_union_r/a [|@a href="cic:/matita/tutorial/chapter4/eqP_sym.def(3)"eqP_sym/a @a href="cic:/matita/tutorial/chapter7/minus_eps_item.def(9)"minus_eps_item/a]] +@a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [||@a href="cic:/matita/tutorial/chapter4/distribute_substract.def(3)"distribute_substract/a] +@a href="cic:/matita/tutorial/chapter4/eqP_substract_r.def(3)"eqP_substract_r/a // qed. -(* theorem 16: 1 *) -theorem sem_bull: ∀S:DeqSet. ∀i:pitem S. \sem{•i} =1 \sem{i} ∪ \sem{|i|}. +theorem sem_bull: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a. ∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{a title="eclose" href="cic:/fakeuri.def(1)"•/ai} a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{i} a title="union" href="cic:/fakeuri.def(1)"∪/a a title="in_l" href="cic:/fakeuri.def(1)"\sem/a{a title="forget" href="cic:/fakeuri.def(1)"|/ai|}. #S #e elim e - [#w normalize % [/2/ | * //] - |/2/ - |#x normalize #w % [ /2/ | * [@False_ind | //]] - |#x normalize #w % [ /2/ | * // ] - |#i1 #i2 #IH1 #IH2 >eclose_dot - @eqP_trans [|@odot_dot_aux //] >sem_cat - @eqP_trans - [|@eqP_union_r - [|@eqP_trans [|@(cat_ext_l … IH1)] @distr_cat_r]] - @eqP_trans [|@union_assoc] - @eqP_trans [||@eqP_sym @union_assoc] - @eqP_union_l // - |#i1 #i2 #IH1 #IH2 >eclose_plus - @eqP_trans [|@sem_oplus] >sem_plus >erase_plus - @eqP_trans [|@(eqP_union_l … IH2)] - @eqP_trans [|@eqP_sym @union_assoc] - @eqP_trans [||@union_assoc] @eqP_union_r - @eqP_trans [||@eqP_sym @union_assoc] - @eqP_trans [||@eqP_union_l [|@union_comm]] - @eqP_trans [||@union_assoc] /2/ - |#i #H >sem_pre_true >sem_star >erase_bull >sem_star - @eqP_trans [|@eqP_union_r [|@cat_ext_l [|@minus_eps_pre_aux //]]] - @eqP_trans [|@eqP_union_r [|@distr_cat_r]] - @eqP_trans [|@union_assoc] @eqP_union_l >erase_star - @eqP_sym @star_fix_eps + [#w normalize % [/span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a/span/span/ | * //] + |/span class="autotactic"2span class="autotrace" trace a href="cic:/matita/tutorial/chapter8/eq_to_ex_eq.def(4)"eq_to_ex_eq/a/span/span/ + |#x normalize #w % [ /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a/span/span/ | * [@a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a | //]] + |#x normalize #w % [ /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a/span/span/ | * // ] + |#i1 #i2 #IH1 #IH2 >a href="cic:/matita/tutorial/chapter8/eclose_dot.def(5)"eclose_dot/a + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@a href="cic:/matita/tutorial/chapter8/sem_pcl_aux.def(11)"sem_pcl_aux/a //] >a href="cic:/matita/tutorial/chapter7/sem_cat.def(8)"sem_cat/a + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a + [|@a href="cic:/matita/tutorial/chapter4/eqP_union_r.def(3)"eqP_union_r/a + [|@a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@(a href="cic:/matita/tutorial/chapter6/cat_ext_l.def(5)"cat_ext_l/a … IH1)] @a href="cic:/matita/tutorial/chapter6/distr_cat_r.def(5)"distr_cat_r/a]] + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@a href="cic:/matita/tutorial/chapter4/union_assoc.def(3)"union_assoc/a] + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [||@a href="cic:/matita/tutorial/chapter4/eqP_sym.def(3)"eqP_sym/a @a href="cic:/matita/tutorial/chapter4/union_assoc.def(3)"union_assoc/a] + @a href="cic:/matita/tutorial/chapter4/eqP_union_l.def(3)"eqP_union_l/a // + |#i1 #i2 #IH1 #IH2 >a href="cic:/matita/tutorial/chapter8/eclose_plus.def(5)"eclose_plus/a + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@a href="cic:/matita/tutorial/chapter8/sem_oplus.def(9)"sem_oplus/a] >a href="cic:/matita/tutorial/chapter7/sem_plus.def(8)"sem_plus/a >a href="cic:/matita/tutorial/chapter7/erase_plus.def(4)"erase_plus/a + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@(a href="cic:/matita/tutorial/chapter4/eqP_union_l.def(3)"eqP_union_l/a … IH2)] + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@a href="cic:/matita/tutorial/chapter4/eqP_sym.def(3)"eqP_sym/a @a href="cic:/matita/tutorial/chapter4/union_assoc.def(3)"union_assoc/a] + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [||@a href="cic:/matita/tutorial/chapter4/union_assoc.def(3)"union_assoc/a] @a href="cic:/matita/tutorial/chapter4/eqP_union_r.def(3)"eqP_union_r/a + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [||@a href="cic:/matita/tutorial/chapter4/eqP_sym.def(3)"eqP_sym/a @a href="cic:/matita/tutorial/chapter4/union_assoc.def(3)"union_assoc/a] + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [||@a href="cic:/matita/tutorial/chapter4/eqP_union_l.def(3)"eqP_union_l/a [|@a href="cic:/matita/tutorial/chapter4/union_comm.def(3)"union_comm/a]] + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [||@a href="cic:/matita/tutorial/chapter4/union_assoc.def(3)"union_assoc/a] /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/tutorial/chapter4/eqP_union_r.def(3)"eqP_union_r/a/span/span/ + |#i #H >a href="cic:/matita/tutorial/chapter7/sem_pre_true.def(9)"sem_pre_true/a >a href="cic:/matita/tutorial/chapter7/sem_star.def(8)"sem_star/a >a href="cic:/matita/tutorial/chapter8/erase_bull.def(6)"erase_bull/a >a href="cic:/matita/tutorial/chapter7/sem_star.def(8)"sem_star/a + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@a href="cic:/matita/tutorial/chapter4/eqP_union_r.def(3)"eqP_union_r/a [|@a href="cic:/matita/tutorial/chapter6/cat_ext_l.def(5)"cat_ext_l/a [|@a href="cic:/matita/tutorial/chapter8/minus_eps_pre_aux.def(11)"minus_eps_pre_aux/a //]]] + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@a href="cic:/matita/tutorial/chapter4/eqP_union_r.def(3)"eqP_union_r/a [|@a href="cic:/matita/tutorial/chapter6/distr_cat_r.def(5)"distr_cat_r/a]] + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@a href="cic:/matita/tutorial/chapter4/union_assoc.def(3)"union_assoc/a] @a href="cic:/matita/tutorial/chapter4/eqP_union_l.def(3)"eqP_union_l/a >a href="cic:/matita/tutorial/chapter7/erase_star.def(4)"erase_star/a + @a href="cic:/matita/tutorial/chapter4/eqP_sym.def(3)"eqP_sym/a @a href="cic:/matita/tutorial/chapter6/star_fix_eps.def(7)"star_fix_eps/a ] qed. -(* blank item *) -let rec blank (S: DeqSet) (i: re S) on i :pitem S ≝ +(* +h2Blank item/h2 +*) + +let rec blank (S: a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a) (i: a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"re/a S) on i :a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S ≝ match i with - [ z ⇒ `∅ - | e ⇒ ϵ - | s y ⇒ `y - | o e1 e2 ⇒ (blank S e1) + (blank S e2) - | c e1 e2 ⇒ (blank S e1) · (blank S e2) - | k e ⇒ (blank S e)^* ]. + [ z ⇒ a href="cic:/matita/tutorial/chapter7/pitem.con(0,1,1)"pz/a ? + | e ⇒ a title="pitem epsilon" href="cic:/fakeuri.def(1)"ϵ/a + | s y ⇒ a title="pitem ps" href="cic:/fakeuri.def(1)"`/ay + | o e1 e2 ⇒ (blank S e1) a title="pitem or" href="cic:/fakeuri.def(1)"+/a (blank S e2) + | c e1 e2 ⇒ (blank S e1) a title="pitem cat" href="cic:/fakeuri.def(1)"·/a (blank S e2) + | k e ⇒ (blank S e)a title="pitem star" href="cic:/fakeuri.def(1)"^/a* ]. -lemma forget_blank: ∀S.∀e:re S.|blank S e| = e. +lemma forget_blank: ∀S.∀e:a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"re/a S.a title="forget" href="cic:/fakeuri.def(1)"|/aa href="cic:/matita/tutorial/chapter8/blank.fix(0,1,3)"blank/a S e| a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a e. #S #e elim e normalize // qed. -lemma sem_blank: ∀S.∀e:re S.\sem{blank S e} =1 ∅. +lemma sem_blank: ∀S.∀e:a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"re/a S.a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{a href="cic:/matita/tutorial/chapter8/blank.fix(0,1,3)"blank/a S e} a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="empty set" href="cic:/fakeuri.def(1)"∅/a. #S #e elim e - [1,2:@eq_to_ex_eq // - |#s @eq_to_ex_eq // - |#e1 #e2 #Hind1 #Hind2 >sem_cat - @eqP_trans [||@(union_empty_r … ∅)] - @eqP_trans [|@eqP_union_l[|@Hind2]] @eqP_union_r - @eqP_trans [||@(cat_empty_l … ?)] @cat_ext_l @Hind1 - |#e1 #e2 #Hind1 #Hind2 >sem_plus - @eqP_trans [||@(union_empty_r … ∅)] - @eqP_trans [|@eqP_union_l[|@Hind2]] @eqP_union_r @Hind1 - |#e #Hind >sem_star - @eqP_trans [||@(cat_empty_l … ?)] @cat_ext_l @Hind + [1,2:@a href="cic:/matita/tutorial/chapter8/eq_to_ex_eq.def(4)"eq_to_ex_eq/a // + |#s @a href="cic:/matita/tutorial/chapter8/eq_to_ex_eq.def(4)"eq_to_ex_eq/a // + |#e1 #e2 #Hind1 #Hind2 >a href="cic:/matita/tutorial/chapter7/sem_cat.def(8)"sem_cat/a + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [||@(a href="cic:/matita/tutorial/chapter4/union_empty_r.def(3)"union_empty_r/a … a title="empty set" href="cic:/fakeuri.def(1)"∅/a)] + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@a href="cic:/matita/tutorial/chapter4/eqP_union_l.def(3)"eqP_union_l/a[|@Hind2]] @a href="cic:/matita/tutorial/chapter4/eqP_union_r.def(3)"eqP_union_r/a + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [||@(a href="cic:/matita/tutorial/chapter6/cat_empty_l.def(5)"cat_empty_l/a … ?)] @a href="cic:/matita/tutorial/chapter6/cat_ext_l.def(5)"cat_ext_l/a @Hind1 + |#e1 #e2 #Hind1 #Hind2 >a href="cic:/matita/tutorial/chapter7/sem_plus.def(8)"sem_plus/a + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [||@(a href="cic:/matita/tutorial/chapter4/union_empty_r.def(3)"union_empty_r/a … a title="empty set" href="cic:/fakeuri.def(1)"∅/a)] + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@a href="cic:/matita/tutorial/chapter4/eqP_union_l.def(3)"eqP_union_l/a[|@Hind2]] @a href="cic:/matita/tutorial/chapter4/eqP_union_r.def(3)"eqP_union_r/a @Hind1 + |#e #Hind >a href="cic:/matita/tutorial/chapter7/sem_star.def(8)"sem_star/a + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [||@(a href="cic:/matita/tutorial/chapter6/cat_empty_l.def(5)"cat_empty_l/a … ?)] @a href="cic:/matita/tutorial/chapter6/cat_ext_l.def(5)"cat_ext_l/a @Hind ] qed. -theorem re_embedding: ∀S.∀e:re S. - \sem{•(blank S e)} =1 \sem{e}. -#S #e @eqP_trans [|@sem_bull] >forget_blank -@eqP_trans [|@eqP_union_r [|@sem_blank]] -@eqP_trans [|@union_comm] @union_empty_r. +theorem re_embedding: ∀S.∀e:a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"re/a S. + a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{a title="eclose" href="cic:/fakeuri.def(1)"•/a(a href="cic:/matita/tutorial/chapter8/blank.fix(0,1,3)"blank/a S e)} a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_l" href="cic:/fakeuri.def(1)"\sem/a{e}. +#S #e @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@a href="cic:/matita/tutorial/chapter8/sem_bull.def(12)"sem_bull/a] >a href="cic:/matita/tutorial/chapter8/forget_blank.def(4)"forget_blank/a +@a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@a href="cic:/matita/tutorial/chapter4/eqP_union_r.def(3)"eqP_union_r/a [|@a href="cic:/matita/tutorial/chapter8/sem_blank.def(9)"sem_blank/a]] +@a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@a href="cic:/matita/tutorial/chapter4/union_comm.def(3)"union_comm/a] @a href="cic:/matita/tutorial/chapter4/union_empty_r.def(3)"union_empty_r/a. qed. -(* lefted operations *) -definition lifted_cat ≝ λS:DeqSet.λe:pre S. - lift S (pre_concat_l S eclose e). +(* +h2Lifted Operators/h2 +*) + +definition lifted_cat ≝ λS:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.λe:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. + a href="cic:/matita/tutorial/chapter8/lift.def(2)"lift/a S (a href="cic:/matita/tutorial/chapter8/pre_concat_l.def(3)"pre_concat_l/a S a href="cic:/matita/tutorial/chapter8/eclose.fix(0,1,4)"eclose/a e). notation "e1 ⊙ e2" left associative with precedence 70 for @{'odot $e1 $e2}. interpretation "lifted cat" 'odot e1 e2 = (lifted_cat ? e1 e2). -lemma odot_true_b : ∀S.∀i1,i2:pitem S.∀b. - 〈i1,true〉 ⊙ 〈i2,b〉 = 〈i1 · (\fst (•i2)),\snd (•i2) ∨ b〉. -#S #i1 #i2 #b normalize in ⊢ (??%?); cases (•i2) // +lemma odot_true_b : ∀S.∀i1,i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S.∀b. + a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1,a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a〉 a title="lifted cat" href="cic:/fakeuri.def(1)"⊙/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai2,b〉 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1 a title="pitem cat" href="cic:/fakeuri.def(1)"·/a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a title="eclose" href="cic:/fakeuri.def(1)"•/ai2)),a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a (a title="eclose" href="cic:/fakeuri.def(1)"•/ai2) a title="boolean or" href="cic:/fakeuri.def(1)"∨/a b〉. +#S #i1 #i2 #b normalize in ⊢ (??%?); cases (a title="eclose" href="cic:/fakeuri.def(1)"•/ai2) // qed. -lemma odot_false_b : ∀S.∀i1,i2:pitem S.∀b. - 〈i1,false〉 ⊙ 〈i2,b〉 = 〈i1 · i2 ,b〉. +lemma odot_false_b : ∀S.∀i1,i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S.∀b. + a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1,a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a〉 a title="lifted cat" href="cic:/fakeuri.def(1)"⊙/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai2,b〉 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1 a title="pitem cat" href="cic:/fakeuri.def(1)"·/a i2 ,b〉. // qed. -lemma erase_odot:∀S.∀e1,e2:pre S. - |\fst (e1 ⊙ e2)| = |\fst e1| · (|\fst e2|). -#S * #i1 * * #i2 #b2 // >odot_true_b >erase_dot // +lemma erase_odot:∀S.∀e1,e2:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. + a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (e1 a title="lifted cat" href="cic:/fakeuri.def(1)"⊙/a e2)| a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e1| a title="re cat" href="cic:/fakeuri.def(1)"·/a (a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e2|). +#S * #i1 * * #i2 #b2 // >a href="cic:/matita/tutorial/chapter8/odot_true_b.def(6)"odot_true_b/a >a href="cic:/matita/tutorial/chapter7/erase_dot.def(4)"erase_dot/a // qed. -definition lk ≝ λS:DeqSet.λe:pre S. +definition lk ≝ λS:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.λe:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. match e with [ mk_Prod i1 b1 ⇒ match b1 with - [true ⇒ 〈(\fst (eclose ? i1))^*, true〉 - |false ⇒ 〈i1^*,false〉 + [true ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a(a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a href="cic:/matita/tutorial/chapter8/eclose.fix(0,1,4)"eclose/a ? i1))a title="pitem star" href="cic:/fakeuri.def(1)"^/a*, a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a〉 + |false ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1a title="pitem star" href="cic:/fakeuri.def(1)"^/a*,a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a〉 ] ]. @@ -314,56 +311,55 @@ interpretation "lk" 'lk a = (lk ? a). notation "a^⊛" non associative with precedence 90 for @{'lk $a}. -lemma ostar_true: ∀S.∀i:pitem S. - 〈i,true〉^⊛ = 〈(\fst (•i))^*, true〉. +lemma ostar_true: ∀S.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. + a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai,a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a〉a title="lk" href="cic:/fakeuri.def(1)"^/a⊛ a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a(a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a title="eclose" href="cic:/fakeuri.def(1)"•/ai))a title="pitem star" href="cic:/fakeuri.def(1)"^/a*, a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a〉. // qed. -lemma ostar_false: ∀S.∀i:pitem S. - 〈i,false〉^⊛ = 〈i^*, false〉. +lemma ostar_false: ∀S.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. + a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai,a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a〉a title="lk" href="cic:/fakeuri.def(1)"^/a⊛ a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aia title="pitem star" href="cic:/fakeuri.def(1)"^/a*, a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a〉. // qed. -lemma erase_ostar: ∀S.∀e:pre S. - |\fst (e^⊛)| = |\fst e|^*. +lemma erase_ostar: ∀S.∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. + a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (ea title="lk" href="cic:/fakeuri.def(1)"^/a⊛)| a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e|a title="re star" href="cic:/fakeuri.def(1)"^/a*. #S * #i * // qed. -lemma sem_odot_true: ∀S:DeqSet.∀e1:pre S.∀i. - \sem{e1 ⊙ 〈i,true〉} =1 \sem{e1 ▹ i} ∪ { [ ] }. +lemma sem_odot_true: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀e1:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S.∀i. + a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1 a title="lifted cat" href="cic:/fakeuri.def(1)"⊙/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai,a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a〉} a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i} a title="union" href="cic:/fakeuri.def(1)"∪/a a title="singleton" href="cic:/fakeuri.def(1)"{/a a title="nil" href="cic:/fakeuri.def(1)"[/a ] }. #S #e1 #i -cut (e1 ⊙ 〈i,true〉 = 〈\fst (e1 ▹ i), \snd(e1 ▹ i) ∨ true〉) [//] -#H >H cases (e1 ▹ i) #i1 #b1 cases b1 - [>sem_pre_true @eqP_trans [||@eqP_sym @union_assoc] - @eqP_union_l /2/ - |/2/ +cut (e1 a title="lifted cat" href="cic:/fakeuri.def(1)"⊙/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai,a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a〉 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (e1 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i), a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a(e1 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i) a title="boolean or" href="cic:/fakeuri.def(1)"∨/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a〉) [//] +#H >H cases (e1 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i) #i1 #b1 cases b1 + [>a href="cic:/matita/tutorial/chapter7/sem_pre_true.def(9)"sem_pre_true/a @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [||@a href="cic:/matita/tutorial/chapter4/eqP_sym.def(3)"eqP_sym/a @a href="cic:/matita/tutorial/chapter4/union_assoc.def(3)"union_assoc/a] + @a href="cic:/matita/tutorial/chapter4/eqP_union_l.def(3)"eqP_union_l/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/tutorial/chapter4/eqP_sym.def(3)"eqP_sym/a/span/span/ + |/span class="autotactic"2span class="autotrace" trace a href="cic:/matita/tutorial/chapter8/eq_to_ex_eq.def(4)"eq_to_ex_eq/a/span/span/ ] qed. -lemma eq_odot_false: ∀S:DeqSet.∀e1:pre S.∀i. - e1 ⊙ 〈i,false〉 = e1 ▹ i. +lemma eq_odot_false: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀e1:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S.∀i. + e1 a title="lifted cat" href="cic:/fakeuri.def(1)"⊙/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai,a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a〉 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a e1 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i. #S #e1 #i -cut (e1 ⊙ 〈i,false〉 = 〈\fst (e1 ▹ i), \snd(e1 ▹ i) ∨ false〉) [//] -cases (e1 ▹ i) #i1 #b1 cases b1 #H @H +cut (e1 a title="lifted cat" href="cic:/fakeuri.def(1)"⊙/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai,a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a〉 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (e1 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i), a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a(e1 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i) a title="boolean or" href="cic:/fakeuri.def(1)"∨/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a〉) [//] +cases (e1 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i) #i1 #b1 cases b1 #H @H qed. lemma sem_odot: - ∀S.∀e1,e2: pre S. \sem{e1 ⊙ e2} =1 \sem{e1}· \sem{|\fst e2|} ∪ \sem{e2}. + ∀S.∀e1,e2: a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1 a title="lifted cat" href="cic:/fakeuri.def(1)"⊙/a e2} a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1}a title="cat lang" href="cic:/fakeuri.def(1)"·/a a title="in_l" href="cic:/fakeuri.def(1)"\sem/a{a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e2|} a title="union" href="cic:/fakeuri.def(1)"∪/a a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e2}. #S #e1 * #i2 * - [>sem_pre_true - @eqP_trans [|@sem_odot_true] - @eqP_trans [||@union_assoc] @eqP_union_r @odot_dot_aux // - |>sem_pre_false >eq_odot_false @odot_dot_aux // + [>a href="cic:/matita/tutorial/chapter7/sem_pre_true.def(9)"sem_pre_true/a + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@a href="cic:/matita/tutorial/chapter8/sem_odot_true.def(10)"sem_odot_true/a] + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [||@a href="cic:/matita/tutorial/chapter4/union_assoc.def(3)"union_assoc/a] @a href="cic:/matita/tutorial/chapter4/eqP_union_r.def(3)"eqP_union_r/a @a href="cic:/matita/tutorial/chapter8/sem_pcl_aux.def(11)"sem_pcl_aux/a // + |>a href="cic:/matita/tutorial/chapter7/sem_pre_false.def(9)"sem_pre_false/a >a href="cic:/matita/tutorial/chapter8/eq_odot_false.def(6)"eq_odot_false/a @a href="cic:/matita/tutorial/chapter8/sem_pcl_aux.def(11)"sem_pcl_aux/a // ] qed. -(* theorem 16: 4 *) -theorem sem_ostar: ∀S.∀e:pre S. - \sem{e^⊛} =1 \sem{e} · \sem{|\fst e|}^*. +theorem sem_ostar: ∀S.∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. + a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{ea title="lk" href="cic:/fakeuri.def(1)"^/a⊛} a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e} a title="cat lang" href="cic:/fakeuri.def(1)"·/a a title="in_l" href="cic:/fakeuri.def(1)"\sem/a{a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e|}a title="star lang" href="cic:/fakeuri.def(1)"^/a*. #S * #i #b cases b - [>sem_pre_true >sem_pre_true >sem_star >erase_bull - @eqP_trans [|@eqP_union_r[|@cat_ext_l [|@minus_eps_pre_aux //]]] - @eqP_trans [|@eqP_union_r [|@distr_cat_r]] - @eqP_trans [||@eqP_sym @distr_cat_r] - @eqP_trans [|@union_assoc] @eqP_union_l - @eqP_trans [||@eqP_sym @epsilon_cat_l] @eqP_sym @star_fix_eps - |>sem_pre_false >sem_pre_false >sem_star /2/ + [>a href="cic:/matita/tutorial/chapter7/sem_pre_true.def(9)"sem_pre_true/a >a href="cic:/matita/tutorial/chapter7/sem_pre_true.def(9)"sem_pre_true/a >a href="cic:/matita/tutorial/chapter7/sem_star.def(8)"sem_star/a >a href="cic:/matita/tutorial/chapter8/erase_bull.def(6)"erase_bull/a + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@a href="cic:/matita/tutorial/chapter4/eqP_union_r.def(3)"eqP_union_r/a[|@a href="cic:/matita/tutorial/chapter6/cat_ext_l.def(5)"cat_ext_l/a [|@a href="cic:/matita/tutorial/chapter8/minus_eps_pre_aux.def(11)"minus_eps_pre_aux/a //]]] + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@a href="cic:/matita/tutorial/chapter4/eqP_union_r.def(3)"eqP_union_r/a [|@a href="cic:/matita/tutorial/chapter6/distr_cat_r.def(5)"distr_cat_r/a]] + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [||@a href="cic:/matita/tutorial/chapter4/eqP_sym.def(3)"eqP_sym/a @a href="cic:/matita/tutorial/chapter6/distr_cat_r.def(5)"distr_cat_r/a] + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@a href="cic:/matita/tutorial/chapter4/union_assoc.def(3)"union_assoc/a] @a href="cic:/matita/tutorial/chapter4/eqP_union_l.def(3)"eqP_union_l/a + @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [||@a href="cic:/matita/tutorial/chapter4/eqP_sym.def(3)"eqP_sym/a @a href="cic:/matita/tutorial/chapter6/epsilon_cat_l.def(5)"epsilon_cat_l/a] @a href="cic:/matita/tutorial/chapter4/eqP_sym.def(3)"eqP_sym/a @a href="cic:/matita/tutorial/chapter6/star_fix_eps.def(7)"star_fix_eps/a + |>a href="cic:/matita/tutorial/chapter7/sem_pre_false.def(9)"sem_pre_false/a >a href="cic:/matita/tutorial/chapter7/sem_pre_false.def(9)"sem_pre_false/a >a href="cic:/matita/tutorial/chapter7/sem_star.def(8)"sem_star/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/tutorial/chapter8/eq_to_ex_eq.def(4)"eq_to_ex_eq/a/span/span/ ] -qed. \ No newline at end of file +qed. \ No newline at end of file