X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2Fre%2Fre-setoids.ma;h=729e8a0ea3a7806d55b8a334d952b613daee4aef;hb=e880d6eab5e1700f4a625ddcd7d0fa8f0cce2dcc;hp=cfbd7cb31d8c0d5b203654d388d6e64958882ba7;hpb=6f2f5039ef719f60ebcf24d7ee17c83eac6cc635;p=helm.git diff --git a/helm/software/matita/nlibrary/re/re-setoids.ma b/helm/software/matita/nlibrary/re/re-setoids.ma index cfbd7cb31..729e8a0ea 100644 --- a/helm/software/matita/nlibrary/re/re-setoids.ma +++ b/helm/software/matita/nlibrary/re/re-setoids.ma @@ -114,7 +114,7 @@ notation "0" non associative with precedence 90 for @{ 'empty_r }. interpretation "empty" 'empty_r = (z ?). notation > "'lang' S" non associative with precedence 90 for @{ Ω^(list $S) }. -notation > "'Elang' S" non associative with precedence 90 for @{ 𝛀^(list $S) }. +notation > "'Elang' S" non associative with precedence 90 for @{ 𝛀^(LIST $S) }. (* setoid support for re *) @@ -156,47 +156,75 @@ unification hint 0 ≔ A : Alpha; (*-----------------------------------------------------------------------*) ⊢ carr X ≡ re T. -unification hint 0 ≔ A:Alpha,a,b:re A; +unification hint 0 ≔ A:Alpha, a,b:re (carr (acarr A)); R ≟ eq0 (RE A), - L ≟ re A + L ≟ re (carr (acarr A)) (* -------------------------------------------- *) ⊢ eq_re A a b ≡ eq_rel L R a b. + +(* XXX This seems to be a pattern for equations in setoid(0) *) +unification hint 0 ≔ AA; + A ≟ carr (acarr AA), + R ≟ setoid1_of_setoid (RE AA) +(*-----------------------------------------------*) ⊢ + re A ≡ carr1 R. + +alias symbol "hint_decl" (instance 1) = "hint_decl_CProp2". +unification hint 0 ≔ S : Alpha, x,y: re (carr (acarr S)); + SS ≟ RE S, + TT ≟ setoid1_of_setoid SS, + T ≟ carr1 TT +(*-----------------------------------------*) ⊢ + eq_re S x y ≡ eq_rel1 T (eq1 TT) x y. +(* contructors are morphisms *) nlemma c_is_morph : ∀A:Alpha.(re A) ⇒_0 (re A) ⇒_0 (re A). -#A; napply (mk_binary_morphism … (λs1,s2:re A. s1 · s2)); -#a; nelim a; -##[##1,2: #a' b b'; ncases a'; nnormalize; /2/ by conj; -##|#x a' b b'; ncases a'; /2/ by conj; -##|##4,5: #r1 r2 IH1 IH2 a'; ncases a'; nnormalize; /2/ by conj; -##|#r IH a'; ncases a'; nnormalize; /2/ by conj; ##] -nqed. +#A; napply (mk_binary_morphism … (λs1,s2:re A. s1 · s2)); #a; nelim a; /2/ by conj; nqed. (* XXX This is the good format for hints about morphisms, fix the others *) alias symbol "hint_decl" (instance 1) = "hint_decl_Type0". -unification hint 0 ≔ S:Alpha, A,B:re S; - MM ≟ mk_unary_morphism ?? - (λA:re S.mk_unary_morphism ?? (λB.A · B) (prop1 ?? (c_is_morph S A))) - (prop1 ?? (c_is_morph S)), +unification hint 0 ≔ S:Alpha, A,B:re (carr (acarr S)); + SS ≟ carr (acarr S), + MM ≟ mk_unary_morphism ?? (λA. + mk_unary_morphism ?? + (λB.A · B) (prop1 ?? (fun1 ?? (c_is_morph S) A))) + (prop1 ?? (c_is_morph S)), T ≟ RE S (*--------------------------------------------------------------------------*) ⊢ - fun1 T T (fun1 T (unary_morph_setoid T T) MM A) B ≡ A · B. + fun1 T T (fun1 T (unary_morph_setoid T T) MM A) B ≡ c SS A B. nlemma o_is_morph : ∀A:Alpha.(re A) ⇒_0 (re A) ⇒_0 (re A). -#A; napply (mk_binary_morphism … (λs1,s2:re A. s1 + s2)); -#a; nelim a; -##[##1,2: #a' b b'; ncases a'; nnormalize; /2/ by conj; -##|#x a' b b'; ncases a'; /2/ by conj; -##|##4,5: #r1 r2 IH1 IH2 a'; ncases a'; nnormalize; /2/ by conj; -##|#r IH a'; ncases a'; nnormalize; /2/ by conj; ##] -nqed. +#A; napply (mk_binary_morphism … (λs1,s2:re A. s1 + s2)); #a; nelim a; /2/ by conj; nqed. + +unification hint 0 ≔ S:Alpha, A,B:re (carr (acarr S)); + SS ≟ carr (acarr S), + MM ≟ mk_unary_morphism ?? (λA. + mk_unary_morphism ?? + (λB.A + B) (prop1 ?? (fun1 ?? (o_is_morph S) A))) + (prop1 ?? (o_is_morph S)), + T ≟ RE S +(*--------------------------------------------------------------------------*) ⊢ + fun1 T T (fun1 T (unary_morph_setoid T T) MM A) B ≡ o SS A B. -unification hint 0 ≔ S:Alpha, A,B:re S; - MM ≟ mk_unary_morphism ?? - (λA:re S.mk_unary_morphism ?? (λB.A + B) (prop1 ?? (o_is_morph S A))) - (prop1 ?? (o_is_morph S)), +nlemma k_is_morph : ∀A:Alpha.(re A) ⇒_0 (re A). +#A; @(λs1:re A. s1^* ); #a; nelim a; /2/ by conj; nqed. + +unification hint 0 ≔ S:Alpha, A:re (carr (acarr S)); + SS ≟ carr (acarr S), + MM ≟ mk_unary_morphism ?? (λB.B^* ) (prop1 ?? (k_is_morph S)), T ≟ RE S (*--------------------------------------------------------------------------*) ⊢ - fun1 T T (fun1 T (unary_morph_setoid T T) MM A) B ≡ A + B. + fun1 T T MM A ≡ k SS A. + +nlemma s_is_morph : ∀A:Alpha.A ⇒_0 (re A). +#A; @(λs1:A. s ? s1 ); #x y E; //; nqed. + +unification hint 0 ≔ S:Alpha, a: carr (acarr S); + SS ≟ carr (acarr S), + MM ≟ mk_unary_morphism ?? (λb.s ? b ) (prop1 ?? (s_is_morph S)), + T ≟ RE S, T1 ≟ acarr S +(*--------------------------------------------------------------------------*) ⊢ + fun1 T1 T MM a ≡ s SS a. (* end setoids support for re *) @@ -233,19 +261,23 @@ ncut (∀w1,w2.(x == w1@w2) = (y == w1@w2)); ##[ nqed. alias symbol "hint_decl" = "hint_decl_Type1". -unification hint 0 ≔ A : setoid, B,C : Elang A; +unification hint 0 ≔ A : setoid, B,C : Elang A; AA ≟ LIST A, - R ≟ mk_ext_powerclass ? - (ext_carr ? B · ext_carr ? C) (ext_prop ? (cat_is_ext ? B C)) -(*--------------------------------------------------------------------*) ⊢ - ext_carr AA R ≡ cat A (ext_carr AA B) (ext_carr AA C). + BB ≟ ext_carr AA B, + CC ≟ ext_carr AA C, + R ≟ mk_ext_powerclass AA + (cat A (ext_carr AA B) (ext_carr AA C)) + (ext_prop AA (cat_is_ext A B C)) +(*----------------------------------------------------------*) ⊢ + ext_carr AA R ≡ cat A BB CC. -unification hint 0 ≔ S:setoid, A,B:lang S; - T ≟ powerclass_setoid (list S), +unification hint 0 ≔ S:setoid, A,B:lang (carr S); + T ≟ powerclass_setoid (list (carr S)), MM ≟ mk_unary_morphism1 T (unary_morphism1_setoid1 T T) - (λA:lang S. + (λA:lang (carr S). mk_unary_morphism1 T T - (λB:lang S.cat S A B) (prop11 T T (cat_is_morph S A))) + (λB:lang (carr S).cat S A B) + (prop11 T T (fun11 ?? (cat_is_morph S) A))) (prop11 T (unary_morphism1_setoid1 T T) (cat_is_morph S)) (*--------------------------------------------------------------------------*) ⊢ fun11 T T (fun11 T (unary_morphism1_setoid1 T T) MM A) B ≡ cat S A B. @@ -257,15 +289,14 @@ nqed. unification hint 1 ≔ AA : setoid, B,C : Elang AA; AAS ≟ LIST AA, - T ≟ ext_powerclass_setoid (list AA), - R ≟ mk_unary_morphism1 T (unary_morphism1_setoid1 T T) - (λS:Elang AA. - mk_unary_morphism1 T T - (λS':Elang AA. - mk_ext_powerclass (list AA) (cat AA (ext_carr ? S) (ext_carr ? S')) - (ext_prop (list AA) (cat_is_ext AA S S'))) - (prop11 T T (cat_is_ext_morph AA S))) - (prop11 T (unary_morphism1_setoid1 T T) (cat_is_ext_morph AA)), + T ≟ ext_powerclass_setoid AAS, + R ≟ mk_unary_morphism1 T (unary_morphism1_setoid1 T T) (λX:Elang AA. + mk_unary_morphism1 T T (λY:Elang AA. + mk_ext_powerclass AAS + (cat AA (ext_carr ? X) (ext_carr ? Y)) + (ext_prop AAS (cat_is_ext AA X Y))) + (prop11 T T (fun11 ?? (cat_is_ext_morph AA) X))) + (prop11 T (unary_morphism1_setoid1 T T) (cat_is_ext_morph AA)), BB ≟ ext_carr ? B, CC ≟ ext_carr ? C (*------------------------------------------------------*) ⊢ @@ -277,6 +308,50 @@ ndefinition star : ∀A:setoid.∀l:lang A.lang A ≝ λS.λl.{ w ∈ list S | ∃lw.flatten ? lw = w ∧ conjunct ? lw l}. interpretation "star lang" 'pk l = (star ? l). +(* hints for star *) +nlemma star_is_morph : ∀A:setoid. (lang A) ⇒_1 (lang A). +#X; @(λA:lang X.A^* ); #a1 a2 E; @; #x; *; #wl; *; #defx Px; @wl; @; //; +nelim wl in Px; //; #s tl IH; *; #a1s a1tl; /4/; nqed. + +nlemma star_is_ext: ∀A:setoid. (Elang A) → (Elang A). + #S A; @ ((ext_carr … A) ^* ); #w1 w2 E; @; *; #wl; *; #defw1 Pwl; + @wl; @; //; napply (.=_0 defw1); /2/; nqed. + +alias symbol "hint_decl" = "hint_decl_Type1". +unification hint 0 ≔ A : setoid, B : Elang A; + AA ≟ LIST A, + BB ≟ ext_carr AA B, + R ≟ mk_ext_powerclass ? + ((ext_carr ? B)^* ) (ext_prop ? (star_is_ext ? B)) +(*--------------------------------------------------------------------*) ⊢ + ext_carr AA R ≡ star A BB. + +unification hint 0 ≔ S:setoid, A:lang (carr S); + T ≟ powerclass_setoid (list (carr S)), + MM ≟ mk_unary_morphism1 T T + (λB:lang (carr S).star S B) (prop11 T T (star_is_morph S)) +(*--------------------------------------------------------------------------*) ⊢ + fun11 T T MM A ≡ star S A. + +nlemma star_is_ext_morph:∀A:setoid.(Elang A) ⇒_1 (Elang A). +#A; @(star_is_ext …); +#x1 x2 Ex; napply (prop11 … (star_is_morph A)); nassumption. +nqed. + +unification hint 1 ≔ AA : setoid, B : Elang AA; + AAS ≟ LIST AA, + T ≟ ext_powerclass_setoid AAS, + R ≟ mk_unary_morphism1 T T + (λS:Elang AA. + mk_ext_powerclass AAS (star AA (ext_carr ? S)) + (ext_prop AAS (star_is_ext AA S))) + (prop11 T T (star_is_ext_morph AA)), + BB ≟ ext_carr ? B +(*------------------------------------------------------*) ⊢ + ext_carr AAS (fun11 T T R B) ≡ star AA BB. + +(* end hints for star *) + notation > "𝐋 term 70 E" non associative with precedence 75 for @{L_re ? $E}. nlet rec L_re (S : Alpha) (r : re S) on r : lang S ≝ match r with @@ -302,8 +377,8 @@ ndefinition L_re_is_ext : ∀S:Alpha.∀r:re S.Elang S. ##| #e H; @; *; #l; *; #defw1 Pl; @l; @; //; napply (.=_1 defw1); /2/; ##] nqed. -unification hint 0 ≔ S : Alpha,e : re S; - SS ≟ LIST S, +unification hint 0 ≔ S : Alpha,e : re (carr (acarr S)); + SS ≟ LIST (acarr S), X ≟ mk_ext_powerclass SS (𝐋 e) (ext_prop SS (L_re_is_ext S e)) (*-----------------------------------------------------------------*)⊢ ext_carr SS X ≡ L_re S e. @@ -328,14 +403,13 @@ nlemma L_re_is_morph:∀A:Alpha.(setoid1_of_setoid (re A)) ⇒_1 Ω^(list A). @; ##[##1,3: nassumption] /2/; ##] nqed. -unification hint 0 ≔ A:Alpha, a:re A; +unification hint 0 ≔ A:Alpha, a:re (carr (acarr A)); T ≟ setoid1_of_setoid (RE A), - T1 ≟ LIST A, - T2 ≟ powerclass_setoid T1, + T2 ≟ powerclass_setoid (list (carr (acarr A))), MM ≟ mk_unary_morphism1 ?? - (λa:setoid1_of_setoid (RE A).𝐋 a) (prop11 ?? (L_re_is_morph A)) + (λa:carr1 (setoid1_of_setoid (RE A)).𝐋 a) (prop11 ?? (L_re_is_morph A)) (*--------------------------------------------------------------------------*) ⊢ - fun11 T T2 MM a ≡ 𝐋 a. + fun11 T T2 MM a ≡ L_re A a. nlemma L_re_is_ext_morph:∀A:Alpha.(setoid1_of_setoid (re A)) ⇒_1 𝛀^(list A). #A; @; ##[ #a; napply (L_re_is_ext ? a); ##] #a b E; @@ -345,12 +419,13 @@ ncut (𝐋 b = 𝐋 a); ##[ napply (.=_1 (┼_1 E^-1)); // ] #EL; ##| napply (. (# ╪_1 ?)); ##[##3: napply H |##2: ##skip ] napply (EL^-1)] nqed. -unification hint 1 ≔ AA : Alpha, a: re AA; - T ≟ RE AA, T1 ≟ LIST AA, TT ≟ ext_powerclass_setoid T1, - R ≟ mk_unary_morphism1 ?? - (λa:setoid1_of_setoid T. - mk_ext_powerclass ? (𝐋 a) (ext_prop ? (L_re_is_ext AA a))) - (prop11 ?? (L_re_is_ext_morph AA)) +unification hint 1 ≔ AA : Alpha, a: re (carr (acarr AA)); + T ≟ RE AA, T1 ≟ LIST (acarr AA), T2 ≟ setoid1_of_setoid T, + TT ≟ ext_powerclass_setoid T1, + R ≟ mk_unary_morphism1 T2 TT + (λa:carr1 (setoid1_of_setoid T). + mk_ext_powerclass T1 (𝐋 a) (ext_prop T1 (L_re_is_ext AA a))) + (prop11 T2 TT (L_re_is_ext_morph AA)) (*------------------------------------------------------*) ⊢ ext_carr T1 (fun11 (setoid1_of_setoid T) TT R a) ≡ L_re AA a. @@ -413,13 +488,13 @@ unification hint 0 ≔ SS:Alpha; P1 ≟ refl ? (eq0 (PITEM SS)), P2 ≟ sym ? (eq0 (PITEM SS)), P3 ≟ trans ? (eq0 (PITEM SS)), - R ≟ mk_setoid (pitem S) + R ≟ mk_setoid (pitem (carr S)) (mk_equivalence_relation (pitem A) (eq_pitem SS) P1 P2 P3) (*-----------------------------------------------------------------*)⊢ carr R ≡ pitem A. -unification hint 0 ≔ S:Alpha,a,b:pitem S; - R ≟ PITEM S, L ≟ (pitem S) +unification hint 0 ≔ S:Alpha,a,b:pitem (carr (acarr S)); + R ≟ PITEM S, L ≟ pitem (carr (acarr S)) (* -------------------------------------------- *) ⊢ eq_pitem S a b ≡ eq_rel L (eq0 R) a b. @@ -480,9 +555,9 @@ ndefinition L_pi_ext : ∀S:Alpha.∀r:pitem S.Elang S. ##] 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))) +unification hint 0 ≔ S : Alpha,e : pitem (carr (acarr S)); + SS ≟ LIST (acarr S), + X ≟ mk_ext_powerclass SS (𝐋\p e) (ext_prop SS (L_pi_ext S e)) (*-----------------------------------------------------------------*)⊢ ext_carr SS X ≡ 𝐋\p e. @@ -495,6 +570,51 @@ interpretation "epsilon" 'epsilon = (epsilon ?). notation < "ϵ b" non associative with precedence 90 for @{'app_epsilon $b}. interpretation "epsilon lang" 'app_epsilon b = (epsilon ? b). +(* hints for epsilon *) +nlemma epsilon_is_morph : ∀A:Alpha. (setoid1_of_setoid bool) ⇒_1 (lang A). +#X; @; ##[#b; napply(ϵ b)] #a1 a2; ncases a1; ncases a2; //; *; nqed. + +nlemma epsilon_is_ext: ∀A:Alpha. (setoid1_of_setoid bool) → (Elang A). + #S b; @(ϵ b); #w1 w2 E; ncases b; @; ##[##3,4:*] +nchange in match (w1 ∈ ϵ true) with ([] =_0 w1); +nchange in match (w2 ∈ ϵ true) with ([] =_0 w2); #H; napply (.= H); /2/; +nqed. + +alias symbol "hint_decl" = "hint_decl_Type1". +unification hint 0 ≔ A : Alpha, B : bool; + AA ≟ LIST (acarr A), + R ≟ mk_ext_powerclass ? + (ϵ B) (ext_prop ? (epsilon_is_ext ? B)) +(*--------------------------------------------------------------------*) ⊢ + ext_carr AA R ≡ epsilon A B. + +unification hint 0 ≔ S:Alpha, A:bool; + B ≟ setoid1_of_setoid BOOL, + T ≟ powerclass_setoid (list (carr (acarr S))), + MM ≟ mk_unary_morphism1 B T + (λB.ϵ B) (prop11 B T (epsilon_is_morph S)) +(*--------------------------------------------------------------------------*) ⊢ + fun11 B T MM A ≡ epsilon S A. + +nlemma epsilon_is_ext_morph:∀A:Alpha. (setoid1_of_setoid bool) ⇒_1 (Elang A). +#A; @(epsilon_is_ext …); +#x1 x2 Ex; napply (prop11 … (epsilon_is_morph A)); nassumption. +nqed. + +unification hint 1 ≔ AA : Alpha, B : bool; + AAS ≟ LIST (acarr AA), + BB ≟ setoid1_of_setoid BOOL, + T ≟ ext_powerclass_setoid AAS, + R ≟ mk_unary_morphism1 BB T + (λS. + mk_ext_powerclass AAS (epsilon AA S) + (ext_prop AAS (epsilon_is_ext AA S))) + (prop11 BB T (epsilon_is_ext_morph AA)) +(*------------------------------------------------------*) ⊢ + ext_carr AAS (fun11 BB T R B) ≡ epsilon AA B. + +(* end hints for epsilon *) + ndefinition L_pr ≝ λS : Alpha.λp:pre S. 𝐋\p\ (\fst p) ∪ ϵ (\snd p). interpretation "L_pr" 'L_pi E = (L_pr ? E). @@ -642,10 +762,11 @@ nlemma subK : ∀S.∀a:Ω^S. a - a = ∅. nlemma subW : ∀S.∀a,b:Ω^S.∀w.w ∈ (a - b) → w ∈ a. #S a b w; nnormalize; *; //; nqed. +alias symbol "eclose" (instance 3) = "eclose". nlemma erase_bull : ∀S:Alpha.∀a:pitem S. |\fst (•a)| = |a|. #S a; nelim a; // by {}; ##[ #e1 e2 IH1 IH2; - napply (?^-1); + napply (?^-1); napply (.=_0 (IH1^-1)╪_0 (IH2^-1)); nchange in match (•(e1 · ?)) with (?⊙?); ncases (•e1); #e3 b; ncases b; ##[ nnormalize; ncases (•e2); /3/ by refl, conj] @@ -670,15 +791,7 @@ napply (. ((defw1 : [ ] = ?)^-1 ╪_0 #)╪_1#); napply Hw2; nqed. -(* XXX This seems to be a pattern for equations *) -alias symbol "hint_decl" (instance 1) = "hint_decl_CProp2". -unification hint 0 ≔ S : Alpha, x,y: re S; - SS ≟ RE S, - TT ≟ setoid1_of_setoid SS, - T ≟ carr1 TT -(*-----------------------------------------*) ⊢ - eq_re S x y ≡ eq_rel1 T (eq1 TT) x y. -(* XXX the previous hint does not work *) + (* theorem 16: 1 → 3 *) nlemma odot_dot_aux : ∀S:Alpha.∀e1,e2: pre S. @@ -704,8 +817,7 @@ nlemma odot_dot_aux : ∀S:Alpha.∀e1,e2: pre S. napply (.=_1 (cupA …)^-1); napply (.=_1 (cupA …)^-1 ╪_1 #); napply (.=_1 (cupA …)); - nlapply (erase_bull S e2'); #XX; - napply (.=_1 (((# ╪_1 (┼_1 ?) )╪_1 #)╪_1 #)); ##[##2: napply XX; ##| ##skip] + napply (.=_1 (((# ╪_1 (┼_1 (erase_bull S e2')) )╪_1 #)╪_1 #)); //; ##| ncases e2; #e2' b2'; nchange in match (𝐋\p ?) with (?∪?∪?); napply (.=_1 (cupA…)); @@ -714,34 +826,22 @@ nlemma odot_dot_aux : ∀S:Alpha.∀e1,e2: pre S. //] nqed. -STOP + nlemma sub_dot_star : - ∀S.∀X:word S → Prop.∀b. (X - ϵ b) · X^* ∪ {[]} = X^*. -#S X b; napply extP; #w; @; -##[ *; ##[##2: nnormalize; #defw; nrewrite < defw; @[]; @; //] + ∀S:Alpha.∀X:Elang S.∀b. (X - ϵ b) · (ext_carr … X)^* ∪ {[]} = (ext_carr … X)^*. +#S X b; napply ext_set; #w; @; +##[ *; ##[##2: #defw; @[]; @; //] *; #w1; *; #w2; *; *; #defw sube; *; #lw; *; #flx cj; - @ (w1 :: lw); nrewrite < defw; nrewrite < flx; @; //; + @ (w1 :: lw); @; ##[ napply (.=_0 # ╪_0 flx); napply (?^-1); //] @; //; napply (subW … sube); -##| *; #wl; *; #defw Pwl; nrewrite < defw; nelim wl in Pwl; ##[ #_; @2; //] - #w' wl' IH; *; #Pw' IHp; nlapply (IH IHp); *; - ##[ *; #w1; *; #w2; *; *; #defwl' H1 H2; - @; ncases b in H1; #H1; - ##[##2: nrewrite > (sub0…); @w'; @(w1@w2); - nrewrite > (associative_append ? w' w1 w2); - nrewrite > defwl'; @; ##[@;//] @(wl'); @; //; - ##| ncases w' in Pw'; - ##[ #ne; @w1; @w2; nrewrite > defwl'; @; //; @; //; - ##| #x xs Px; @(x::xs); @(w1@w2); - nrewrite > (defwl'); @; ##[@; //; @; //; @; nnormalize; #; ndestruct] - @wl'; @; //; ##] ##] - ##| #wlnil; nchange in match (flatten ? (w'::wl')) with (w' @ flatten ? wl'); - nrewrite < (wlnil); nrewrite > (append_nil…); ncases b; - ##[ ncases w' in Pw'; /2/; #x xs Pxs; @; @(x::xs); @([]); - nrewrite > (append_nil…); @; ##[ @; //;@; //; nnormalize; @; #; ndestruct] - @[]; @; //; - ##| @; @w'; @[]; nrewrite > (append_nil…); @; ##[##2: @[]; @; //] - @; //; @; //; @; *;##]##]##] +##| *; #wl; *; #defw Pwl; napply (. (defw^-1 ╪_1 #)); + nelim wl in Pwl; /2/; + #s tl IH; *; #Xs Ptl; ncases s in Xs; ##[ #; napply IH; //] #x xs Xxxs; + @; @(x :: xs); @(flatten ? tl); @; + ##[ @; ##[ napply #] @; ##[nassumption] ncases b; *; ##] + nelim tl in Ptl; ##[ #; @[]; /2/] #w ws IH; *; #Xw Pws; @(w :: ws); @; ##[ napply #] + @; //;##] nqed. (* theorem 16: 1 *) @@ -749,38 +849,54 @@ alias symbol "pc" (instance 13) = "cat lang". alias symbol "in_pl" (instance 23) = "in_pl". alias symbol "in_pl" (instance 5) = "in_pl". alias symbol "eclose" (instance 21) = "eclose". -ntheorem bull_cup : ∀S:Alpha. ∀e:pitem S. 𝐋\p (•e) = 𝐋\p e ∪ 𝐋 .|e|. +ntheorem bull_cup : ∀S:Alpha. ∀e:pitem S. 𝐋\p (•e) = 𝐋\p e ∪ 𝐋 |e|. #S e; nelim e; //; - ##[ #a; napply extP; #w; nnormalize; @; *; /3/ by or_introl, or_intror; - ##| #a; napply extP; #w; nnormalize; @; *; /3/ by or_introl; *; + ##[ #a; napply ext_set; #w; @; *; /3/ by or_introl, or_intror; + ##| #a; napply ext_set; #w; @; *; /3/ by or_introl; *; ##| #e1 e2 IH1 IH2; - nchange in ⊢ (??(??(%))?) with (•e1 ⊙ 〈e2,false〉); - nrewrite > (odot_dot_aux S (•e1) 〈e2,false〉 IH2); - nrewrite > (IH1 …); nrewrite > (cup_dotD …); - nrewrite > (cupA …); nrewrite > (cupC ?? (𝐋\p ?) …); - nchange in match (𝐋\p 〈?,?〉) with (𝐋\p e2 ∪ {}); nrewrite > (cup0 …); - nrewrite < (erase_dot …); nrewrite < (cupA …); //; + nchange in match (•(e1·e2)) with (•e1 ⊙ 〈e2,false〉); + napply (.=_1 (odot_dot_aux ?? 〈e2,false〉 IH2)); + napply (.=_1 (IH1 ╪_1 #) ╪_1 #); + napply (.=_1 (cup_dotD …) ╪_1 #); + napply (.=_1 (cupA …)); + napply (.=_1 # ╪_1 ((erase_dot ???)^-1 ╪_1 (cup0 ??))); + napply (.=_1 # ╪_1 (cupC…)); + napply (.=_1 (cupA …)^-1); //; ##| #e1 e2 IH1 IH2; - nchange in match (•(?+?)) with (•e1 ⊕ •e2); nrewrite > (oplus_cup …); - nrewrite > (IH1 …); nrewrite > (IH2 …); nrewrite > (cupA …); - nrewrite > (cupC ? (𝐋\p e2)…);nrewrite < (cupA ??? (𝐋\p e2)…); - nrewrite > (cupC ?? (𝐋\p e2)…); nrewrite < (cupA …); - nrewrite < (erase_plus …); //. + nchange in match (•(?+?)) with (•e1 ⊕ •e2); + napply (.=_1 (oplus_cup …)); + napply (.=_1 IH1 ╪_1 IH2); + napply (.=_1 (cupA …)); + napply (.=_1 # ╪_1 (# ╪_1 (cupC…))); + napply (.=_1 # ╪_1 (cupA ????)^-1); + napply (.=_1 # ╪_1 (cupC…)); + napply (.=_1 (cupA ????)^-1); + napply (.=_1 # ╪_1 (erase_plus ???)^-1); //; ##| #e; nletin e' ≝ (\fst (•e)); nletin b' ≝ (\snd (•e)); #IH; - nchange in match (𝐋\p ?) with (𝐋\p 〈e'^*,true〉); + nchange in match (𝐋\p ?) with (𝐋\p 〈e'^*,true〉); nchange in match (𝐋\p ?) with (𝐋\p (e'^* ) ∪ {[ ]}); - nchange in ⊢ (??(??%?)?) with (𝐋\p e' · 𝐋 .|e'|^* ); - nrewrite > (erase_bull…e); - nrewrite > (erase_star …); - nrewrite > (?: 𝐋\p e' = 𝐋\p e ∪ (𝐋 .|e| - ϵ b')); ##[##2: - nchange in IH : (??%?) with (𝐋\p e' ∪ ϵ b'); ncases b' in IH; - ##[ #IH; nrewrite > (cup_sub…); //; nrewrite < IH; - nrewrite < (cup_sub…); //; nrewrite > (subK…); nrewrite > (cup0…);//; - ##| nrewrite > (sub0 …); #IH; nrewrite < IH; nrewrite > (cup0 …);//; ##]##] - nrewrite > (cup_dotD…); nrewrite > (cupA…); - nrewrite > (?: ((?·?)∪{[]} = 𝐋 .|e^*|)); //; - nchange in match (𝐋 .|e^*|) with ((𝐋. |e|)^* ); napply sub_dot_star;##] - nqed. + (* nwhd in match (𝐋\p e'^* ); (* XXX bug uncertain *) *) + nchange in ⊢ (???(??%?)?) with (𝐋\p e' · ?); + napply (.=_1 (# ╪_1 (┼_1 (┼_0 (erase_bull S e)))) ╪_1 #); + napply (.=_1 (# ╪_1 (erase_star …)) ╪_1 #); + ncut ( 𝐋\p e' = 𝐋\p e ∪ (𝐋 |e| - ϵ b')); ##[ + nchange in IH : (???%?) with (𝐋\p e' ∪ ϵ b'); ncases b' in IH; + ##[ #IH; napply (?^-1); napply (.=_1 (cup_sub … (not_epsilon_lp…))); + napply (.=_1 (IH^-1 ╪_1 #)); + alias symbol "invert" = "setoid1 symmetry". + (* XXX too slow if ambiguous, since it tries with a ? (takes 12s) then + tries with sym0 and fails immediately, then with sym1 that is OK *) + napply (.=_1 (cup_sub …(not_epsilon_lp …))^-1); + napply (.=_1 # ╪_1 (subK…)); napply (.=_1 (cup0…)); //; + ##| #IH; napply (?^-1); napply (.=_1 # ╪_1 (sub0 …)); + napply (.=_1 IH^-1); napply (.=_1 (cup0 …)); //; ##]##] #EE; + napply (.=_1 (EE ╪_1 #) ╪_1 #); + napply (.=_1 (cup_dotD…) ╪_1 #); + napply (.=_1 (cupA…)); + napply (.=_1 # ╪_1 (sub_dot_star…)); //; ##] +nqed. + +STOP (* theorem 16: 3 *) nlemma odot_dot: