X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2Fre%2Fre-setoids.ma;h=ab2d13458928808becfd8616589c91bd2a5bce26;hb=b15a4df27469ee4b64d1b3b8fc996cd15e8a61f0;hp=06791066ab1d21f3d486bc26ff943bc46deeebb7;hpb=070b44c9c2344967ca8c4531909614a0d4da2fbe;p=helm.git diff --git a/helm/software/matita/nlibrary/re/re-setoids.ma b/helm/software/matita/nlibrary/re/re-setoids.ma index 06791066a..ab2d13458 100644 --- a/helm/software/matita/nlibrary/re/re-setoids.ma +++ b/helm/software/matita/nlibrary/re/re-setoids.ma @@ -31,6 +31,7 @@ match l1 with ndefinition LIST : setoid → setoid. #S; @(list S); @(eq_list S); ncases admit; nqed. +alias symbol "hint_decl" (instance 1) = "hint_decl_Type1". unification hint 0 ≔ S : setoid; P1 ≟ refl ? (eq0 (LIST S)), P2 ≟ sym ? (eq0 (LIST S)), @@ -168,7 +169,7 @@ unification hint 0 ≔ A : Alpha; carr X ≡ (re T). notation < "a \sup ⋇" non associative with precedence 90 for @{ 'pk $a}. -notation > "a ^ *" non associative with precedence 90 for @{ 'pk $a}. +notation > "a ^ *" non associative with precedence 75 for @{ 'pk $a}. interpretation "star" 'pk a = (k ? a). interpretation "or" 'plus a b = (o ? a b). @@ -293,9 +294,25 @@ notation "𝐋\sub(\p) term 70 E" non associative with precedence 75 for @{'L_pi interpretation "in_pl" 'L_pi E = (L_pi ? E). unification hint 0 ≔ S,a,b; - R ≟ LIST S + R ≟ LIST S, + L ≟ (list S) (* -------------------------------------------- *) ⊢ - eq_list S a b ≡ eq_rel (list S) (eq0 R) a b. + eq_list S a b ≡ eq_rel L (eq0 R) a b. + +ncoercion setoid1_of_setoid : ∀s:setoid. setoid1 ≝ setoid1_of_setoid on _s: setoid to setoid1. + +alias symbol "hint_decl" (instance 1) = "hint_decl_CProp2". +unification hint 0 ≔ S : setoid, x,y; + SS ≟ LIST S, + TT ≟ setoid1_of_setoid SS +(*-----------------------------------------*) ⊢ + eq_list S x y ≡ eq_rel1 ? (eq1 TT) x y. + +unification hint 0 ≔ SS : setoid; + S ≟ carr SS, + TT ≟ setoid1_of_setoid (LIST SS) +(*-----------------------------------------------------------------*) ⊢ + list S ≡ carr1 TT. notation > "B ⇒_0 C" right associative with precedence 72 for @{'umorph0 $B $C}. notation > "B ⇒_1 C" right associative with precedence 72 for @{'umorph1 $B $C}. @@ -305,223 +322,77 @@ notation "B ⇒\sub 1 C" right associative with precedence 72 for @{'umorph1 $B interpretation "unary morphism 0" 'umorph0 A B = (unary_morphism A B). interpretation "unary morphism 1" 'umorph1 A B = (unary_morphism1 A B). -ncoercion setoid1_of_setoid : ∀s:setoid. setoid1 ≝ setoid1_of_setoid on _s: setoid to setoid1. -nlemma exists_is_morph: (* BUG *) ∀S,T:setoid.∀P: S ⇒_1 (T ⇒_1 (CProp[0]:?)). - ∀y,z:S.y =_0 z → (Ex T (P y)) = (Ex T (P z)). -#S T P y z E; @; -##[ *; #x Px; @x; alias symbol "refl" (instance 4) = "refl". - alias symbol "prop2" (instance 2) = "prop21". - napply (. E^-1‡#); napply Px; -##| *; #x Px; @x; napply (. E‡#); napply Px;##] -nqed. - -ndefinition ex_morph : ∀S:setoid. (S ⇒_1 CPROP) ⇒_1 CPROP. -#S; @; ##[ #P; napply (Ex ? P); ##| #P1 P2 E; @; -*; #x; #H; @ x; nlapply (E x x ?); //; *; /2/; -nqed. - nlemma Sig: ∀S,T:setoid.∀P: S → (T → CPROP). ∀y,z:T.y = z → (∀x.y=z → P x y = P x z) → (Ex S (λx.P x y)) =_1 (Ex S (λx.P x z)). #S T P y z Q E; @; *; #x Px; @x; nlapply (E x Q); *; /2/; nqed. -(* desiderata : Σ(λx.H‡#) - ottenute : Σ H (λx,H.H‡#) -- quindi monoriscrittura. H toplevel permette inferenza di y e z in Sig -*) +notation "∑" non associative with precedence 90 for @{Sig ?????}. nlemma test : ∀S:setoid. ∀ee: S ⇒_1 S ⇒_1 CPROP. ∀x,y:setoid1_of_setoid S.x =_1 y → (Ex S (λw.ee x w ∧ True)) =_1 (Ex S (λw.ee y w ∧ True)). #S m x y E; - -napply (trans1 ????? (Sig ????? E (λw,H. - (prop11 ?(unary_morphism1_setoid1 ??)??? - (prop11 ?(unary_morphism1_setoid1 ??)??? - H - ?? (refl ???)) - ?? (refl1 ???))))). -napply refl1. +napply (.=_1 (∑ E (λw,H.(H ╪_1 #)╪_1 #))). +napply #. nqed. nlemma test2 : ∀S:setoid. ∀ee: S ⇒_1 S ⇒_1 CPROP. ∀x,y:setoid1_of_setoid S.x =_1 y → (True ∧ (Ex S (λw.ee x w ∧ True))) =_1 (True ∧ (Ex S (λw.ee y w ∧ True))). #S m x y E; -napply (trans1 ????? -(#‡ - (Sig ????? E (λw,H. - (prop11 ?(unary_morphism1_setoid1 ??)??? - (prop11 ?(unary_morphism1_setoid1 ??)??? - H - ?? (refl ???)) - ?? (refl1 ???)))))). -napply refl1. - -alias symbol "trans" (instance 1) = "trans1". -alias symbol "refl" (instance 5) = "refl". -alias symbol "prop2" (instance 3) = "prop21". -alias symbol "trans" (instance 1) = "trans1". -alias symbol "prop2" (instance 3) = "prop21". -(* super bug 1+1: - eseguire senza modificare per ottenrere degli alias, fare back di 1 passo - e ri-eseguire. Se riesegue senza aggiungere altri alias allora hai gli - alias giusti (ma se fate back di più passi, gli alias non vanno più bene...). - ora (m x w) e True possono essere sostituiti da ?, se invece - si toglie anche l'∧, allora un super bug si scatena (meta contesto locale - scazzato, con y al posto di x, unificata con se stessa ma col contesto - locale corretto...). lo stesso (o simile) bug salta fuori se esegui - senza gli alias giusti con ? al posto di True o (m x w). - - bug a parte, pare inferisca tutto lui... - la E astratta nella prova è solo per fargli inferire x e y, se Sig - lo si riformula in modo più naturale (senza y=z) allora bisogna passare - x e y esplicitamente. *) -alias symbol "trans" (instance 1) = "trans1". -alias symbol "prop2" (instance 3) = "prop21". -alias symbol "trans" (instance 1) = "trans1". -alias symbol "prop2" (instance 3) = "prop21". - -alias symbol "trans" (instance 1) = "trans1". -alias symbol "prop2" (instance 3) = "prop21". -(*napply (.= (Sig ? S (λw,x. m x w ∧ True) ?? E (λw,H.(H‡#)‡#))); *) - -(* OK *) -napply (.= (Sig ? ? (?) ?? E (λw,H. - (prop11 ?(unary_morphism1_setoid1 ??)??? - (prop11 ?(unary_morphism1_setoid1 ??)??? - H - ?? (refl ???)) - ?? (refl1 ???))))). -napply refl1. - - (prop11 (setoid1_of_setoid ?) (ext_powerclass_setoid ?) ? ? ???? - (prop11 (setoid1_of_setoid ?) (ext_powerclass_setoid ?) ? ? ???? - H - (refl ??)) - (refl ? True))))). - -prop11 ???????? (prop11 ???????? H (refl ??)) (refl ??)))); - -napply (.= (Sig ? S (λw,x. m x w ∧ True) ?? E (λw,H.(H‡#)‡#))); - - -napply (.= (Sig ? S (λw,x.?) ?? E (λw,H.(H‡#)‡#))); -(Ex S (λx.?P x y)) ==?== -//; nqed. -STOP -napply (λw:S.(.= ((E‡#)‡#)) #); ##[##2: napply w| napply m. #H; napply H; - - - - let form ≝ comp1_unary_morphisms ??? (ex_morph (list S)) ee in - form x = form y. - #S ee x y E; - nletin F ≝ (comp1_unary_morphisms ??? (ex_morph (list S)) ee); - - nnormalize; - nchange in E with (eq_rel1 ? (eq1 (setoid1_of_setoid (LIST S))) x y); - - ncheck (exists_is_morph (LIST S) (LIST S) ? ?? (E‡#)). - nletin xxx ≝ (exists_is_morph); (LIST S)); (LIST S) ee x y E); - - nchange with (F x = F y); - nchange in E with (eq_rel1 ? (eq1 (setoid1_of_setoid (LIST S))) x y); - napply (.= † E); - napply #. +napply (.=_1 #╪_1(∑ E (λw,H.(H ╪_1 #) ╪_1 #))). +napply #. nqed. - -E : w = w2 - - - Σ(λx.(#‡E)‡#) : ∃x.x = w ∧ m → ∃x.x = w2 ∧ m - λx.(#‡E)‡# : ∀x.x = w ∧ m → x = w2 ∧ m - - -w; -F ≟ ex_moprh ∘ G -g ≟ fun11 G ------------------------------- -ex (λx.g w x) ==?== fun11 F w - -x ⊢ fun11 ?h ≟ λw. ?g w x -?G ≟ morphism_leibniz (T → S) ∘ ?h ------------------------------- -(λw. (λx:T.?g w x)) ==?== fun11 ?G - -... -x ⊢ fun11 ?h ==?== λw. eq x w ∧ m [w] -(λw,x.eq x w ∧ m[w]) ==?== fun11 ?G -ex (λx.?g w x) ==?== ex (λx.x = w ∧ m[w]) - -ndefinition ex_morph : ∀S:setoid. (S ⇒_1 CPROP) ⇒_1 CPROP. -#S; @; ##[ #P; napply (Ex ? P); ##| ncases admit. ##] nqed. - -ndefinition ex_morph1 : ∀S:setoid. (S ⇒_1 CPROP) ⇒_1 CPROP. -#S; @; ##[ #P; napply (Ex ? (λx.P); ##| ncases admit. ##] nqed. - - -nlemma d : ∀S:Alpha. - ∀ee: (setoid1_of_setoid (list S)) ⇒_1 (setoid1_of_setoid (list S)) ⇒_1 CPROP. - ∀x,y:list S.x = y → - let form ≝ comp1_unary_morphisms ??? (ex_morph (list S)) ee in - form x = form y. - #S ee x y E; - nletin F ≝ (comp1_unary_morphisms ??? (ex_morph (list S)) ee); - - nnormalize; - - nchange with (F x = F y); - nchange in E with (eq_rel1 ? (eq1 (setoid1_of_setoid (LIST S))) x y); - napply (.= † E); - napply #. +nlemma ex_setoid : ∀S:setoid.(S ⇒_1 CPROP) → setoid. +#T P; @ (Ex T (λx:T.P x)); @; +##[ #H1 H2; napply True |##*: //; ##] nqed. +unification hint 0 ≔ T,P ; S ≟ (ex_setoid T P) ⊢ + Ex T (λx:T.P x) ≡ carr S. -nlemma d : ∀S:Alpha.∀ee: (setoid1_of_setoid (list S)) ⇒_1 (setoid1_of_setoid (list S)) ⇒_1 CPROP.∀x,y:list S.x = y → - let form ≝ comp1_unary_morphisms ??? (ex_morph (list S)) ee in - form x = form y. - #S ee x y E; - nletin F ≝ (comp1_unary_morphisms ??? (ex_morph (list S)) ee); - - nnormalize; - - nchange with (F x = F y); - nchange in E with (eq_rel1 ? (eq1 (setoid1_of_setoid (LIST S))) x y); - napply (.= † E); - napply #. +nlemma test3 : ∀S:setoid. ∀ee: S ⇒_1 S ⇒_1 CPROP. + ∀x,y:setoid1_of_setoid S.x =_1 y → + ((Ex S (λw.ee x w ∧ True) ∨ True)) =_1 ((Ex S (λw.ee y w ∧ True) ∨ True)). +#S m x y E; +napply (.=_1 (∑ E (λw,H.(H ╪_1 #) ╪_1 #)) ╪_1 #). +napply #. nqed. - - -nlemma d : ∀S:Alpha.(setoid1_of_setoid (list S)) ⇒_1 CPROP. - #S; napply (comp1_unary_morphisms ??? (ex_morph (list S)) ?); - napply (eq1). - - - -(* -ndefinition comp_ex : ∀X,S,T,K.∀P:X ⇒_1 (S ⇒_1 T).∀Pc : (S ⇒_1 T) ⇒_1 K. X ⇒_1 K. - *) ndefinition L_pi_ext : ∀S:Alpha.∀r:pitem S.Elang S. -#S r; @(𝐋\p r); #w1 w2 E; nelim r; /2/; -##[ #x; @; #H; ##[ nchange in H ⊢ % with ([?]=?); napply ((.= H) E)] - nchange in H ⊢ % with ([?]=?); napply ((.= H) E^-1); +#S r; @(𝐋\p r); #w1 w2 E; nelim r; +##[ /2/; +##| /2/; +##| #x; @; *; +##| #x; @; #H; nchange in H with ([?] =_0 ?); ##[ napply ((.=_0 H) E); ##] + napply ((.=_0 H) E^-1); ##| #e1 e2 H1 H2; nchange in match (w1 ∈ 𝐋\p (?·?)) with ((∃_.?)∨?); + nchange in match (w2 ∈ 𝐋\p (?·?)) with ((∃_.?)∨?); napply (.= (#‡H2)); - napply (.= (Eexists ?? ? w1 w2 E)‡#); - - - nchange in match (w2 ∈ 𝐋\p (?·?)) with (?∨?); - napply (.= - - - //; napply (trans ?? ??? H E); - napply (trans (list S) (eq0 (LIST S)) ? [?] ?(*w1 [x] w2*) H E); - nlapply (trans (list S) (eq0 (LIST S))). - -napply (.= H); nnormalize; nlapply (trans ? [x] w1 w2 E H); napply (.= ?) [napply (w1 = [x])] ##[##2: napply #; napply sym1; napply refl1 ] + napply (.=_1 (∑ E (λx1,H1.∑ E (λx2,H2.?)))╪_1 #); ##[ + ncut ((w1 = (x1@x2)) = (w2 = (x1@x2)));##[ + @; #X; ##[ napply ((.= H1^-1) X) | napply ((.= H2) X) ] ##] #X; + napply ( (X‡#)‡#); ##] + napply #; +##| #e1 e2 H1 H2; + nnormalize in ⊢ (???%%); + napply (H1‡H2); +##| #e H; nnormalize in ⊢ (???%%); + napply (.=_1 (∑ E (λx1,H1.∑ E (λx2,H2.?)))); ##[ + ncut ((w1 = (x1@x2)) = (w2 = (x1@x2)));##[ + @; #X; ##[ napply ((.= H1^-1) X) | napply ((.= H2) X) ] ##] #X; + napply ((X‡#)‡#); ##] + napply #;##] +nqed. +unification hint 0 ≔ S : Alpha,e : pitem S; + SS ≟ (list S), + X ≟ (mk_ext_powerclass SS (𝐋\p e) (ext_prop SS (L_pi_ext S e))) +(*-----------------------------------------------------------------*)⊢ + ext_carr SS X ≡ 𝐋\p e. + ndefinition epsilon ≝ λS:Alpha.λb.match b return λ_.lang S with [ true ⇒ { [ ] } | _ ⇒ ∅ ]. @@ -533,9 +404,9 @@ ndefinition L_pr ≝ λS : Alpha.λp:pre S. 𝐋\p\ (\fst p) ∪ ϵ (\snd p). interpretation "L_pr" 'L_pi E = (L_pr ? E). -nlemma append_eq_nil : ∀S:setoid.∀w1,w2:list S. w1 @ w2 = [ ] → w1 = [ ]. +nlemma append_eq_nil : ∀S:setoid.∀w1,w2:list S. [ ] = w1 @ w2 → w1 = [ ]. #S w1; ncases w1; //. nqed. - + (* lemma 12 *) nlemma epsilon_in_true : ∀S:Alpha.∀e:pre S. [ ] ∈ 𝐋\p e = (\snd e = true). #S r; ncases r; #e b; @; ##[##2: #H; nrewrite > H; @2; /2/; ##] ncases b;//; @@ -544,19 +415,26 @@ nlemma epsilon_in_true : ∀S:Alpha.∀e:pre S. [ ] ∈ 𝐋\p e = (\snd e = tru ##| #r1 r2 H G; *; ##[##2: nassumption; ##] ##| #r1 r2 H1 H2; *; /2/ by {}] *; #w1; *; #w2; *; *; -##[ #defw1 H1 foo; napply H; napply (. #‡#); (append_eq_nil … defw1)^-1‡#); - - nrewrite > (append_eq_nil ? … w1 w2 …); /3/ by {};//; +##[ #defw1 H1 foo; napply H; + napply (. (append_eq_nil ? ?? defw1)^-1╪_1#); + nassumption; +##| #defw1 H1 foo; napply H; + napply (. (append_eq_nil ? ?? defw1)^-1╪_1#); + nassumption; +##] nqed. -nlemma not_epsilon_lp : ∀S.∀e:pitem S. ¬ (𝐋\p e [ ]). -#S e; nelim e; nnormalize; /2/ by nmk; -##[ #; @; #; ndestruct; -##| #r1 r2 n1 n2; @; *; /2/; *; #w1; *; #w2; *; *; #H; - nrewrite > (append_eq_nil …H…); /2/; -##| #r1 r2 n1 n2; @; *; /2/; -##| #r n; @; *; #w1; *; #w2; *; *; #H; - nrewrite > (append_eq_nil …H…); /2/;##] +nlemma not_epsilon_lp : ∀S:Alpha.∀e:pitem S. ¬ ([ ] ∈ (𝐋\p e)). +#S e; nelim e; ##[##1,2,3,4: nnormalize;/2/] +##[ #p1 p2 np1 np2; *; ##[##2: napply np2] *; #w1; *; #w2; *; *; #abs; + nlapply (append_eq_nil ??? abs); # defw1; #; napply np1; + napply (. defw1^-1╪_1#); + nassumption; +##| #p1 p2 np1 np2; *; nchange with (¬?); //; +##| #r n; *; #w1; *; #w2; *; *; #abs; #; napply n; + nlapply (append_eq_nil ??? abs); # defw1; #; + napply (. defw1^-1╪_1#); + nassumption;##] nqed. ndefinition lo ≝ λS:Alpha.λa,b:pre S.〈\fst a + \fst b,\snd a || \snd b〉. @@ -581,12 +459,12 @@ ndefinition lk ≝ λS:Alpha.λbcast:∀S:Alpha.∀E:pitem S.pre S.λa:pre S. notation < "a \sup ⊛" non associative with precedence 90 for @{'lk $op $a}. interpretation "lk" 'lk op a = (lk ? op a). -notation > "a^⊛" non associative with precedence 90 for @{'lk eclose $a}. +notation > "a ^ ⊛" non associative with precedence 75 for @{'lk eclose $a}. notation > "•" non associative with precedence 60 for @{eclose ?}. nlet rec eclose (S: Alpha) (E: pitem S) on E : pre S ≝ match E with - [ pz ⇒ 〈 ∅, false 〉 + [ pz ⇒ 〈 0, false 〉 | pe ⇒ 〈 ϵ, true 〉 | ps x ⇒ 〈 `.x, false 〉 | pp x ⇒ 〈 `.x, false 〉 @@ -600,48 +478,123 @@ notation > "• x" non associative with precedence 60 for @{'eclose $x}. ndefinition reclose ≝ λS:Alpha.λp:pre S.let p' ≝ •\fst p in 〈\fst p',\snd p || \snd p'〉. interpretation "reclose" 'eclose x = (reclose ? x). -ndefinition 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 =\sub 1 B" non associative with precedence 45 for @{'eq_f1 $A $B}. -interpretation "eq f1" 'eq_f1 a b = (eq_f1 ? a b). +nlemma cup0 :∀S.∀p:lang S.p ∪ ∅ = p. +#S p; @; #w; ##[*; //| #; @1; //] *; nqed. + +nlemma cupC : ∀S. ∀a,b:lang S.a ∪ b = b ∪ a. +#S a b; @; #w; *; nnormalize; /2/; nqed. -naxiom extP : ∀S.∀p,q:word S → Prop.(p =1 q) → p = q. +nlemma cupID : ∀S. ∀a:lang S.a ∪ a = a. +#S a; @; #w; ##[*; //] /2/; nqed. nlemma 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/; *; //; *; +#S b1 b2; ncases b1; ncases b2; +nchange in match (true || true) with true; +nchange in match (true || false) with true; +nchange in match (ϵ true) with {[]}; +nchange in match (ϵ false) with ∅; +##[##1,4: napply ((cupID…)^-1); +##| napply ((cup0 S {[]})^-1); +##| napply (.= (cup0 S {[]})^-1); napply cupC; ##] +nqed. + +nlemma cupA : ∀S.∀a,b,c:lang S.a ∪ b ∪ c = a ∪ (b ∪ c). +#S a b c; @; #w; *; /3/; *; /3/; nqed. + +nlemma setP : ∀S.∀A,B:Ω^S.∀x:S. A = B → (x ∈ A) = (x ∈ B). +#S A B x; *; #H1 H2; @; ##[ napply H1 | napply H2] nqed. + +nlemma pset_ext : ∀S.∀A,B:Ω^S.(∀x:S. (x ∈ A) = (x ∈ B)) → A = B. +#S A B H; @; #x; ncases (H x); #H1 H2; ##[ napply H1 | napply H2] nqed. + +ndefinition if': ∀A,B:CPROP. A = B → A → B. +#A B; *; /2/. nqed. + +ncoercion if : ∀A,B:CPROP. ∀p:A = B. A → B ≝ if' on _p : eq_rel1 ???? to ∀_:?.?. + +(* move in sets.ma? *) +nlemma union_morphism : ∀A.Ω^A ⇒_1 (Ω^A ⇒_1 Ω^A). +#X; napply (mk_binary_morphism1 … (λA,B.A ∪ B)); +#A1 A2 B1 B2 EA EB; napply pset_ext; #x; +nchange in match (x ∈ (A1 ∪ B1)) with (?∨?); +napply (.= (setP ??? x EA)‡#); +napply (.= #‡(setP ??? x EB)); //; nqed. -nlemma cupA : ∀S.∀a,b,c:word S → Prop.a ∪ b ∪ c = a ∪ (b ∪ c). -#S a b c; napply extP; #w; nnormalize; @; *; /3/; *; /3/; nqed. +nlemma union_is_ext: ∀A. 𝛀^A → 𝛀^A → 𝛀^A. + #S A B; @ (A ∪ B); #x y Exy; @; *; #H1; +##[##1,3: @; ##|##*: @2 ] +##[##1,3: napply (. (Exy^-1)╪_1#) +##|##2,4: napply (. Exy╪_1#)] +nassumption; +nqed. -nlemma cupC : ∀S. ∀a,b:word S → Prop.a ∪ b = b ∪ a. -#S a b; napply extP; #w; @; *; nnormalize; /2/; nqed. +alias symbol "hint_decl" = "hint_decl_Type1". +unification hint 0 ≔ + A : setoid, B,C : 𝛀^A; + R ≟ (mk_ext_powerclass ? (B ∪ C) (ext_prop ? (union_is_ext ? B C))) + (* ----------------------------------------------------------------*) ⊢ + ext_carr A R ≡ union ? (ext_carr ? B) (ext_carr ? C). + +unification hint 0 ≔ S:Type[0], A,B:Ω^S; + MM ≟ mk_unary_morphism1 ?? + (λA.mk_unary_morphism1 ?? (λB.A ∪ B) (prop11 ?? (union_morphism S A))) + (prop11 ?? (union_morphism S)) + (*-----------------------------------------------------*) ⊢ + fun11 ?? (fun11 ?? MM A) B ≡ A ∪ B. + +nlemma union_is_ext_morph:∀A:setoid.𝛀^A ⇒_1 (𝛀^A ⇒_1 𝛀^A). +#A; napply (mk_binary_morphism1 … (union_is_ext …)); +#x1 x2 y1 y2 Ex Ey; napply (prop11 … (union_morphism A)); nassumption. +nqed. + +unification hint 1 ≔ + AA : setoid, B,C : 𝛀^AA; + A ≟ carr AA, + R ≟ (mk_unary_morphism1 ?? + (λS:(*ext_powerclass_setoid AA*)𝛀^AA. + mk_unary_morphism1 ?? + (λS':(*ext_powerclass_setoid AA*)𝛀^AA. + mk_ext_powerclass AA (S ∪ S') (ext_prop AA (union_is_ext ? S S'))) + (prop11 ?? (union_is_ext_morph AA S))) + (prop11 ?? (union_is_ext_morph AA))) , + BB ≟ (ext_carr ? B), + CC ≟ (ext_carr ? C) + (* ------------------------------------------------------*) ⊢ + ext_carr AA (R B C) ≡ union A BB CC. + (* theorem 16: 2 *) nlemma oplus_cup : ∀S:Alpha.∀e1,e2:pre S.𝐋\p (e1 ⊕ e2) = 𝐋\p e1 ∪ 𝐋\p e2. #S r1; ncases r1; #e1 b1 r2; ncases r2; #e2 b2; -nwhd in ⊢ (??(??%)?); -nchange in ⊢(??%?) with (𝐋\p (e1 + e2) ∪ ϵ (b1 || b2)); -nchange in ⊢(??(??%?)?) with (𝐋\p (e1) ∪ 𝐋\p (e2)); -nrewrite > (epsilon_or S …); nrewrite > (cupA S (𝐋\p e1) …); -nrewrite > (cupC ? (ϵ b1) …); nrewrite < (cupA S (𝐋\p e2) …); -nrewrite > (cupC ? ? (ϵ b1) …); nrewrite < (cupA …); //; +nwhd in ⊢ (???(??%)?); +nchange in ⊢(???%?) with (𝐋\p (e1 + e2) ∪ ϵ (b1 || b2)); +nchange in ⊢(???(??%?)?) with (𝐋\p (e1) ∪ 𝐋\p (e2)); +napply (.= #╪_1 (epsilon_or ???)); +napply (.= (cupA…)^-1); +napply (.= (cupA…)╪_1#); +napply (.= (#╪_1(cupC…))╪_1#); +napply (.= (cupA…)^-1╪_1#); +napply (.= (cupA…)); +//; nqed. +FINQUI + +manca setoide per pair (e pre) + nlemma odotEt : - ∀S.∀e1,e2:pitem S.∀b2. 〈e1,true〉 ⊙ 〈e2,b2〉 = 〈e1 · \fst (•e2),b2 || \snd (•e2)〉. + ∀S:Alpha.∀e1,e2:pitem S.∀b2. 〈e1,true〉 ⊙ 〈e2,b2〉 = ?.〈e1 · \fst (•e2),b2 || \snd (•e2)〉. #S e1 e2 b2; nnormalize; ncases (•e2); //; nqed. nlemma LcatE : ∀S.∀e1,e2:pitem S.𝐋\p (e1 · e2) = 𝐋\p e1 · 𝐋 .|e2| ∪ 𝐋\p e2. //; nqed. -nlemma cup_dotD : ∀S.∀p,q,r:word S → Prop.(p ∪ q) · r = (p · r) ∪ (q · r). +nlemma cup_dotD : ∀S.∀p,q,r:lang S.(p ∪ q) · r = (p · r) ∪ (q · r). #S p q r; napply extP; #w; nnormalize; @; ##[ *; #x; *; #y; *; *; #defw; *; /7/ by or_introl, or_intror, ex_intro, conj; ##| *; *; #x; *; #y; *; *; /7/ by or_introl, or_intror, ex_intro, conj; ##] nqed. -nlemma cup0 :∀S.∀p:word S → Prop.p ∪ {} = p. -#S p; napply extP; #w; nnormalize; @; /2/; *; //; *; nqed. nlemma erase_dot : ∀S.∀e1,e2:pitem S.𝐋 .|e1 · e2| = 𝐋 .|e1| · 𝐋 .|e2|. #S e1 e2; napply extP; nnormalize; #w; @; *; #w1; *; #w2; *; *; /7/ by ex_intro, conj;