X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2Fre%2Fre-setoids.ma;h=729e8a0ea3a7806d55b8a334d952b613daee4aef;hb=7761bd5341f81e4b7d335c0b47bc1e2271f02227;hp=ced520b05abf2276999e04ad0409131b9742cc9b;hpb=27c34e93fc35402111253325b93089a6308dd4bb;p=helm.git diff --git a/helm/software/matita/nlibrary/re/re-setoids.ma b/helm/software/matita/nlibrary/re/re-setoids.ma index ced520b05..729e8a0ea 100644 --- a/helm/software/matita/nlibrary/re/re-setoids.ma +++ b/helm/software/matita/nlibrary/re/re-setoids.ma @@ -12,76 +12,28 @@ (* *) (**************************************************************************) -include "datatypes/pairs.ma". -include "datatypes/bool.ma". +include "datatypes/pairs-setoids.ma". +include "datatypes/bool-setoids.ma". +include "datatypes/list-setoids.ma". include "sets/sets.ma". +(* ninductive Admit : CProp[0] ≝ . naxiom admit : Admit. +*) -(* single = is for the abstract equality of setoids, == is for concrete - equalities (that may be lifted to the setoid level when needed *) -notation < "hvbox(a break mpadded width -50% (=)= b)" non associative with precedence 45 for @{ 'eq_low $a $b }. -notation > "a == b" non associative with precedence 45 for @{ 'eq_low $a $b }. - - -(* XXX move to lists.ma *) -ninductive list (A:Type[0]) : Type[0] ≝ - | nil: list A - | cons: A -> list A -> list A. - -nlet rec eq_list (A : setoid) (l1, l2 : list A) on l1 : CProp[0] ≝ -match l1 with -[ nil ⇒ match l2 return λ_.CProp[0] with [ nil ⇒ True | _ ⇒ False ] -| cons x xs ⇒ match l2 with [ nil ⇒ False | cons y ys ⇒ x = y ∧ eq_list ? xs ys]]. - -interpretation "eq_list" 'eq_low a b = (eq_list ? a b). - -ndefinition LIST : setoid → setoid. -#S; @(list S); @(eq_list S); -##[ #l; nelim l; //; #; @; //; -##| #l1; nelim l1; ##[ #y; ncases y; //] #x xs H y; ncases y; ##[*] #y ys; *; #; @; /2/; -##| #l1; nelim l1; ##[ #l2 l3; ncases l2; ncases l3; /3/; #z zs y ys; *] - #x xs H l2 l3; ncases l2; ncases l3; /2/; #z zs y yz; *; #H1 H2; *; #H3 H4; @; /3/;##] -nqed. - -alias symbol "hint_decl" (instance 1) = "hint_decl_Type1". -unification hint 0 ≔ S : setoid; - T ≟ carr S, - P1 ≟ refl ? (eq0 (LIST S)), - P2 ≟ sym ? (eq0 (LIST S)), - P3 ≟ trans ? (eq0 (LIST S)), - X ≟ mk_setoid (list S) (mk_equivalence_relation ? (eq_list T) P1 P2 P3) -(*-----------------------------------------------------------------------*) ⊢ - carr X ≡ list T. - -unification hint 0 ≔ S:setoid,a,b:list S; - R ≟ eq0 (LIST S), - L ≟ (list S) -(* -------------------------------------------- *) ⊢ - eq_list S a b ≡ eq_rel L R a b. - -notation "hvbox(hd break :: tl)" - right associative with precedence 47 - for @{'cons $hd $tl}. - -notation "[ list0 x sep ; ]" - non associative with precedence 90 - for ${fold right @'nil rec acc @{'cons $x $acc}}. +(* XXX move somewere else *) +ndefinition if': ∀A,B:CPROP. A = B → A → B. +#A B; *; /2/. nqed. -notation "hvbox(l1 break @ l2)" - right associative with precedence 47 - for @{'append $l1 $l2 }. +ncoercion if : ∀A,B:CPROP. ∀p:A = B. A → B ≝ if' on _p : eq_rel1 ? (eq1 CPROP) ?? to ∀_:?.?. -interpretation "nil" 'nil = (nil ?). -interpretation "cons" 'cons hd tl = (cons ? hd tl). +ndefinition ifs': ∀S.∀A,B:Ω^S. A = B → ∀x. x ∈ A → x ∈ B. +#S A B; *; /2/. nqed. -nlet rec append A (l1: list A) l2 on l1 ≝ - match l1 with - [ nil ⇒ l2 - | cons hd tl ⇒ hd :: append A tl l2 ]. +ncoercion ifs : ∀S.∀A,B:Ω^S. ∀p:A = B.∀x. x ∈ A → x ∈ B ≝ ifs' on _p : eq_rel1 ? (eq1 (powerclass_setoid ?))?? to ∀_:?.?. -interpretation "append" 'append l1 l2 = (append ? l1 l2). +(* XXX move to list-setoids-theory.ma *) ntheorem append_nil: ∀A:setoid.∀l:list A.l @ [] = l. #A;#l;nelim l;//; #a;#l1;#IH;nnormalize;/2/;nqed. @@ -91,57 +43,37 @@ ndefinition associative ≝ λA:setoid.λf:A → A → A.∀x,y,z.f x (f y z) = ntheorem associative_append: ∀A:setoid.associative (list A) (append A). #A;#x;#y;#z;nelim x[ napply (refl ???) |#a;#x1;#H;nnormalize;/2/]nqed. -nlet rec flatten S (l : list (list S)) on l : list S ≝ -match l with [ nil ⇒ [ ] | cons w tl ⇒ w @ flatten ? tl ]. - (* end move to list *) + +(* XXX to undestand what I want inside Alpha + the eqb part should be split away, but when available it should be + possible to obtain a leibnitz equality on lemmas proved on setoids +*) interpretation "iff" 'iff a b = (iff a b). -ninductive eq (A:Type[0]) (x:A) : A → CProp[0] ≝ refl: eq A x x. +ninductive eq (A:Type[0]) (x:A) : A → CProp[0] ≝ erefl: eq A x x. nlemma eq_rect_Type0_r': - ∀A.∀a,x.∀p:eq ? x a.∀P: ∀x:A. eq ? x a → Type[0]. P a (refl A a) → P x p. + ∀A.∀a,x.∀p:eq ? x a.∀P: ∀x:A. eq ? x a → Type[0]. P a (erefl A a) → P x p. #A; #a; #x; #p; ncases p; #P; #H; nassumption. nqed. nlemma eq_rect_Type0_r: - ∀A.∀a.∀P: ∀x:A. eq ? x a → Type[0]. P a (refl A a) → ∀x.∀p:eq ? x a.P x p. + ∀A.∀a.∀P: ∀x:A. eq ? x a → Type[0]. P a (erefl A a) → ∀x.∀p:eq ? x a.P x p. #A; #a; #P; #p; #x0; #p0; napply (eq_rect_Type0_r' ??? p0); nassumption. nqed. nlemma eq_rect_CProp0_r': - ∀A.∀a,x.∀p:eq ? x a.∀P: ∀x:A. eq ? x a → CProp[0]. P a (refl A a) → P x p. + ∀A.∀a,x.∀p:eq ? x a.∀P: ∀x:A. eq ? x a → CProp[0]. P a (erefl A a) → P x p. #A; #a; #x; #p; ncases p; #P; #H; nassumption. nqed. nlemma eq_rect_CProp0_r: - ∀A.∀a.∀P: ∀x:A. eq ? x a → CProp[0]. P a (refl A a) → ∀x.∀p:eq ? x a.P x p. + ∀A.∀a.∀P: ∀x:A. eq ? x a → CProp[0]. P a (erefl A a) → ∀x.∀p:eq ? x a.P x p. #A; #a; #P; #p; #x0; #p0; napply (eq_rect_CProp0_r' ??? p0); nassumption. nqed. -(* XXX move to bool *) -interpretation "bool eq" 'eq_low a b = (eq bool a b). - -ndefinition BOOL : setoid. -@bool; @(eq bool); ncases admit.nqed. - -alias symbol "hint_decl" (instance 1) = "hint_decl_Type1". -alias id "refl" = "cic:/matita/ng/properties/relations/refl.fix(0,1,3)". -unification hint 0 ≔ ; - P1 ≟ refl ? (eq0 BOOL), - P2 ≟ sym ? (eq0 BOOL), - P3 ≟ trans ? (eq0 BOOL), - X ≟ mk_setoid bool (mk_equivalence_relation ? (eq bool) P1 P2 P3) -(*-----------------------------------------------------------------------*) ⊢ - carr X ≡ bool. - -unification hint 0 ≔ a,b; - R ≟ eq0 BOOL, - L ≟ bool -(* -------------------------------------------- *) ⊢ - eq bool a b ≡ eq_rel L R a b. - nrecord Alpha : Type[1] ≝ { acarr :> setoid; eqb: acarr → acarr → bool; @@ -149,7 +81,9 @@ nrecord Alpha : Type[1] ≝ { }. interpretation "eqb" 'eq_low a b = (eqb ? a b). +(* end alpha *) +(* re *) ninductive re (S: Type[0]) : Type[0] ≝ z: re S | e: re S @@ -158,6 +92,32 @@ ninductive re (S: Type[0]) : Type[0] ≝ | o: re S → re S → re S | k: re S → re S. +notation < "a \sup ⋇" 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). + +notation "a · b" non associative with precedence 60 for @{ 'pc $a $b}. +interpretation "cat" 'pc a b = (c ? a b). + +(* to get rid of \middot *) +ncoercion c : ∀S.∀p:re S. re S → re S ≝ c on _p : re ? to ∀_:?.?. + +notation < "a" non associative with precedence 90 for @{ 'ps $a}. +notation > "` term 90 a" non associative with precedence 90 for @{ 'ps $a}. +interpretation "atom" 'ps a = (s ? a). + +notation "ϵ" non associative with precedence 90 for @{ 'epsilon }. +interpretation "epsilon" 'epsilon = (e ?). + +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) }. + +(* setoid support for re *) + nlet rec eq_re (S:Alpha) (a,b : re S) on a : CProp[0] ≝ match a with [ z ⇒ match b with [ z ⇒ True | _ ⇒ False] @@ -186,8 +146,6 @@ ndefinition RE : Alpha → setoid. #r2 r3; /3/; ##]##] nqed. -alias symbol "hint_decl" (instance 1) = "hint_decl_Type1". -alias id "carr" = "cic:/matita/ng/sets/setoids/carr.fix(0,0,1)". unification hint 0 ≔ A : Alpha; S ≟ acarr A, T ≟ carr S, @@ -198,44 +156,81 @@ 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. -notation < "a \sup ⋇" 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). - -notation "a · b" non associative with precedence 60 for @{ 'pc $a $b}. -interpretation "cat" 'pc a b = (c ? a b). - -(* to get rid of \middot *) -ncoercion c : ∀S.∀p:re S. re S → re S ≝ c on _p : re ? to ∀_:?.?. - -notation < "a" non associative with precedence 90 for @{ 'ps $a}. -notation > "` term 90 a" non associative with precedence 90 for @{ 'ps $a}. -interpretation "atom" 'ps a = (s ? a). - -notation "ϵ" non associative with precedence 90 for @{ 'epsilon }. -interpretation "epsilon" 'epsilon = (e ?). +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; /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 (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 ≡ 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; /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. + +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 MM A ≡ k SS A. + +nlemma s_is_morph : ∀A:Alpha.A ⇒_0 (re A). +#A; @(λs1:A. s ? s1 ); #x y E; //; nqed. -notation "0" non associative with precedence 90 for @{ 'empty_r }. -interpretation "empty" 'empty_r = (z ?). +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. -notation > "'lang' S" non associative with precedence 90 for @{ Ω^(list $S) }. -notation > "'Elang' S" non associative with precedence 90 for @{ 𝛀^(list $S) }. +(* end setoids support for re *) nlet rec conjunct S (l : list (list S)) (L : lang S) on l: CProp[0] ≝ match l with [ nil ⇒ True | cons w tl ⇒ w ∈ L ∧ conjunct ? tl L ]. -(* -ndefinition sing_lang : ∀A:setoid.∀x:A.Ω^A ≝ λS.λx.{ w | x = w }. -interpretation "sing lang" 'singl x = (sing_lang ? x). -*) - interpretation "subset construction with type" 'comprehension t \eta.x = (mk_powerclass t x). @@ -243,10 +238,120 @@ ndefinition cat : ∀A:setoid.∀l1,l2:lang A.lang A ≝ λS.λl1,l2.{ w ∈ list S | ∃w1,w2.w =_0 w1 @ w2 ∧ w1 ∈ l1 ∧ w2 ∈ l2}. interpretation "cat lang" 'pc a b = (cat ? a b). +(* hints for cat *) +nlemma cat_is_morph : ∀A:setoid. (lang A) ⇒_1 (lang A) ⇒_1 (lang A). +#X; napply (mk_binary_morphism1 … (λA,B:lang X.A · B)); +#A1 A2 B1 B2 EA EB; napply ext_set; #x; +ncut (∀y,x:list X.(x ∈ B1) =_1 (x ∈ B2)); ##[ + #_; #y; ncases EA; ncases EB; #h1 h2 h3 h4; @; ##[ napply h1 | napply h2] ##] #YY; +ncut (∀x,y:list X.(x ∈ A1) =_1 (x ∈ A2)); ##[ + #y; #y; ncases EA; ncases EB; #h1 h2 h3 h4; @; ##[ napply h3 | napply h4] ##] #XX; +napply (.=_1 (∑w1, w2. XX w1 w2/ E ; (# ╪_1 E) ╪_1 #)); +napply (.=_1 (∑w1, w2. YY w1 w2/ E ; # ╪_1 E)); //; +nqed. + +nlemma cat_is_ext: ∀A:setoid. (Elang A) → (Elang A) → (Elang A). + #S A B; @ (ext_carr … A · ext_carr … B); (* XXX coercion ext_carr che non funge *) +#x y Exy; +ncut (∀w1,w2.(x == w1@w2) = (y == w1@w2)); ##[ + #w1 w2; @; #H; ##[ napply (.= Exy^-1) | napply (.= Exy)] // ] +#E; @; #H; +##[ napply (. (∑w1,w2. (E w1 w2)^-1 / E ; (E ╪_1 #) ╪_1 #)); napply H; +##| napply (. (∑w1,w2. E w1 w2 / E ; (E ╪_1 #) ╪_1 #)); napply H ] +nqed. + +alias symbol "hint_decl" = "hint_decl_Type1". +unification hint 0 ≔ A : setoid, B,C : Elang A; + AA ≟ LIST A, + 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 (carr S); + T ≟ powerclass_setoid (list (carr S)), + MM ≟ mk_unary_morphism1 T (unary_morphism1_setoid1 T T) + (λA:lang (carr S). + mk_unary_morphism1 T T + (λ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. + +nlemma cat_is_ext_morph:∀A:setoid.(Elang A) ⇒_1 (Elang A) ⇒_1 (Elang A). +#A; napply (mk_binary_morphism1 … (cat_is_ext …)); +#x1 x2 y1 y2 Ex Ey; napply (prop11 … (cat_is_morph A)); nassumption. +nqed. + +unification hint 1 ≔ AA : setoid, B,C : Elang AA; + AAS ≟ LIST 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 +(*------------------------------------------------------*) ⊢ + ext_carr AAS (fun11 T T (fun11 T (unary_morphism1_setoid1 T T) R B) C) ≡ cat AA BB CC. + +(* end hints for cat *) + 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 @@ -259,9 +364,72 @@ match r with notation "𝐋 term 70 E" non associative with precedence 75 for @{'L_re $E}. interpretation "in_l" 'L_re E = (L_re ? E). -notation "a || b" left associative with precedence 30 for @{'orb $a $b}. -ndefinition orb ≝ λa,b:bool. match a with [ true ⇒ true | _ ⇒ b ]. -interpretation "orb" 'orb a b = (orb a b). +(* support for 𝐋 as an extensional set *) +ndefinition L_re_is_ext : ∀S:Alpha.∀r:re S.Elang S. +#S r; @(𝐋 r); #w1 w2 E; nelim r; +##[ ##1,2: /2/; @; #defw1; napply (.=_0 (defw1 : [ ] = ?)); //; napply (?^-1); //; +##| #x; @; #defw1; napply (.=_0 (defw1 : [x] = ?)); //; napply (?^-1); //; +##| #e1 e2 H1 H2; (* not shure I shoud Inline *) + @; *; #s1; *; #s2; *; *; #defw1 s1L1 s2L2; + ##[ nlapply (trans … E^-1 defw1); #defw2; + ##| nlapply (trans … E defw1); #defw2; ##] @s1; @s2; /3/; +##| #e1 e2 H1 H2; napply (H1‡H2); (* good! *) +##| #e H; @; *; #l; *; #defw1 Pl; @l; @; //; napply (.=_1 defw1); /2/; ##] +nqed. + +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. + +nlemma L_re_is_morph:∀A:Alpha.(setoid1_of_setoid (re A)) ⇒_1 Ω^(list A). +#A; @; ##[ napply (λr:re A.𝐋 r); ##] #r1; nelim r1; +##[##1,2: #r2; ncases r2; //; ##[##1,6: *|##2,7,5,12,10: #a; *|##3,4,8,9: #a1 a2; *] +##|#x r2; ncases r2; ##[##1,2: *|##4,5: #a1 a2; *|##6: #a; *] #y E; @; #z defz; + ncases z in defz; ##[##1,3: *] #zh ztl; ncases ztl; ##[##2,4: #d dl; *; #_; *] + *; #defx; #_; @; //; napply (?^-1); napply (.= defx^-1); //; napply (?^-1); //; +##|#e1 e2 IH1 IH2 r2; ncases r2; ##[##1,2: *|##5: #a1 a2; *|##3,6: #a1; *] + #f1 f2; *; #E1 E2; nlapply (IH2 … E2); nlapply (IH1 … E1); #H1 H2; + nchange in match (𝐋 (e1 · e2)) with (?·?); + napply (.=_1 (H1 ╪_1 H2)); //; +##|#e1 e2 IH1 IH2 r2; ncases r2; ##[##1,2: *|##4: #a1 a2; *|##3,6: #a1; *] + #f1 f2; *; #E1 E2; nlapply (IH2 … E2); nlapply (IH1 … E1); #H1 H2; + napply (.=_1 H1╪_1H2); //; +##|#r IH r2; ncases r2; ##[##1,2: *|##4,5: #a1 a2; *|##3: #a1; *] + #e; #defe; nlapply (IH e defe); #H; + @; #x; *; #wl; *; #defx Px; @wl; @; //; nelim wl in Px; //; #l ls IH; *; #lr Pr; + ##[ nlapply (ifs' … H … lr) | nlapply (ifs' … H^-1 … lr) ] #le; + @; ##[##1,3: nassumption] /2/; ##] +nqed. + +unification hint 0 ≔ A:Alpha, a:re (carr (acarr A)); + T ≟ setoid1_of_setoid (RE A), + T2 ≟ powerclass_setoid (list (carr (acarr A))), + MM ≟ mk_unary_morphism1 ?? + (λa:carr1 (setoid1_of_setoid (RE A)).𝐋 a) (prop11 ?? (L_re_is_morph 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; +ncut (𝐋 b = 𝐋 a); ##[ napply (.=_1 (┼_1 E^-1)); // ] #EL; +@; #x H; nchange in H ⊢ % with (x ∈ 𝐋 ?); +##[ napply (. (# ╪_1 ?)); ##[##3: napply H |##2: ##skip ] napply EL; +##| napply (. (# ╪_1 ?)); ##[##3: napply H |##2: ##skip ] napply (EL^-1)] +nqed. + +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. + +(* end support for 𝐋 as an extensional set *) ninductive pitem (S: Type[0]) : Type[0] ≝ pz: pitem S @@ -272,6 +440,19 @@ ninductive pitem (S: Type[0]) : Type[0] ≝ | po: pitem S → pitem S → pitem S | pk: pitem S → pitem S. +interpretation "pstar" 'pk a = (pk ? a). +interpretation "por" 'plus a b = (po ? a b). +interpretation "pcat" 'pc a b = (pc ? a b). +notation < ".a" non associative with precedence 90 for @{ 'pp $a}. +notation > "`. term 90 a" non associative with precedence 90 for @{ 'pp $a}. +interpretation "ppatom" 'pp a = (pp ? a). +(* to get rid of \middot *) +ncoercion pc : ∀S.∀p:pitem S. pitem S → pitem S ≝ pc on _p : pitem ? to ∀_:?.?. +interpretation "patom" 'ps a = (ps ? a). +interpretation "pepsilon" 'epsilon = (pe ?). +interpretation "pempty" 'empty_r = (pz ?). + +(* setoids for pitem *) nlet rec eq_pitem (S : Alpha) (p1, p2 : pitem S) on p1 : CProp[0] ≝ match p1 with [ pz ⇒ match p2 with [ pz ⇒ True | _ ⇒ False] @@ -307,50 +488,18 @@ 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) (mk_equivalence_relation (pitem A) (eq_pitem SS) P1 P2 P3) -(*---------------------------*)⊢ + 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. -(* XXX move to pair.ma *) -nlet rec eq_pair (A, B : setoid) (a : A × B) (b : A × B) on a : CProp[0] ≝ - match a with [ mk_pair a1 a2 ⇒ - match b with [ mk_pair b1 b2 ⇒ a1 = b1 ∧ a2 = b2 ]]. - -interpretation "eq_pair" 'eq_low a b = (eq_pair ?? a b). - -nlemma PAIR : ∀A,B:setoid. setoid. -#A B; @(A × B); @(eq_pair …); -##[ #ab; ncases ab; #a b; @; napply #; -##| #ab cd; ncases ab; ncases cd; #a1 a2 b1 b2; *; #E1 E2; - @; napply (?^-1); //; -##| #a b c; ncases a; ncases b; ncases c; #c1 c2 b1 b2 a1 a2; - *; #E1 E2; *; #E3 E4; @; ##[ napply (.= E1); //] napply (.= E2); //.##] -nqed. - -alias symbol "hint_decl" (instance 1) = "hint_decl_Type1". -unification hint 0 ≔ AA, BB; - A ≟ carr AA, B ≟ carr BB, - P1 ≟ refl ? (eq0 (PAIR AA BB)), - P2 ≟ sym ? (eq0 (PAIR AA BB)), - P3 ≟ trans ? (eq0 (PAIR AA BB)), - R ≟ mk_setoid (A × B) (mk_equivalence_relation ? (eq_pair …) P1 P2 P3) -(*---------------------------------------------------------------------------*)⊢ - carr R ≡ A × B. - -unification hint 0 ≔ S1,S2,a,b; - R ≟ PAIR S1 S2, - L ≟ (pair S1 S2) -(* -------------------------------------------- *) ⊢ - eq_pair S1 S2 a b ≡ eq_rel L (eq0 R) a b. +(* end setoids for pitem *) -(* end move to pair *) - ndefinition pre ≝ λS.pitem S × bool. notation "\fst term 90 x" non associative with precedence 90 for @{'fst $x}. @@ -358,18 +507,6 @@ interpretation "fst" 'fst x = (fst ? ? x). notation > "\snd term 90 x" non associative with precedence 90 for @{'snd $x}. interpretation "snd" 'snd x = (snd ? ? x). -interpretation "pstar" 'pk a = (pk ? a). -interpretation "por" 'plus a b = (po ? a b). -interpretation "pcat" 'pc a b = (pc ? a b). -notation < ".a" non associative with precedence 90 for @{ 'pp $a}. -notation > "`. term 90 a" non associative with precedence 90 for @{ 'pp $a}. -interpretation "ppatom" 'pp a = (pp ? a). -(* to get rid of \middot *) -ncoercion pc : ∀S.∀p:pitem S. pitem S → pitem S ≝ pc on _p : pitem ? to ∀_:?.?. -interpretation "patom" 'ps a = (ps ? a). -interpretation "pepsilon" 'epsilon = (pe ?). -interpretation "pempty" 'empty_r = (pz ?). - notation > "|term 19 e|" non associative with precedence 70 for @{forget ? $e}. nlet rec forget (S: Alpha) (l : pitem S) on l: re S ≝ match l with @@ -380,6 +517,7 @@ nlet rec forget (S: Alpha) (l : pitem S) on l: re S ≝ | pc E1 E2 ⇒ (|E1| · |E2|) | po E1 E2 ⇒ (|E1| + |E2|) | pk E ⇒ |E|^* ]. + notation < "|term 19 e|" non associative with precedence 70 for @{'forget $e}. interpretation "forget" 'forget a = (forget ? a). @@ -397,92 +535,33 @@ notation > "𝐋\p term 70 E" non associative with precedence 75 for @{'L_pi $E} notation "𝐋\sub(\p) term 70 E" non associative with precedence 75 for @{'L_pi $E}. interpretation "in_pl" 'L_pi E = (L_pi ? E). -(* The caml, as some patches for it *) -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. - -(* XXX Ex setoid support *) -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. - -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 (.=_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 (.=_1 #╪_1(∑ E (λw,H.(H ╪_1 #) ╪_1 #))). -napply #. -nqed. - -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 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. -(* Ex setoid support end *) - +(* set support for 𝐋\p *) ndefinition L_pi_ext : ∀S:Alpha.∀r:pitem S.Elang S. #S r; @(𝐋\p r); #w1 w2 E; nelim r; ##[ ##1,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 (.=_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 #;##] + napply (.= (#‡H2)); + ncut (∀x1,x2. (w1 = (x1@x2)) = (w2 = (x1@x2)));##[ + #x1 x2; @; #X; ##[ napply ((.= E^-1) X) | napply ((.= E) X) ] ##] #X; + napply ((∑w1,w2. X w1 w2 / H ; (H╪_1#)╪_1#) ╪_1 #); +##| #e1 e2 H1 H2; napply (H1‡H2); +##| #e H; + ncut (∀x1,x2.(w1 = (x1@x2)) = (w2 = (x1@x2)));##[ + #x1 x2; @; #X; ##[ napply ((.= E^-1) X) | napply ((.= E) X) ] ##] #X; + napply (∑w1,w2. X w1 w2 / H ; (H╪_1#)╪_1#); +##] 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. + +(* end set support for 𝐋\p *) ndefinition epsilon ≝ λS:Alpha.λb.match b return λ_.lang S with [ true ⇒ { [ ] } | _ ⇒ ∅ ]. @@ -491,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). @@ -498,10 +622,10 @@ interpretation "L_pr" 'L_pi E = (L_pr ? E). nlemma append_eq_nil : ∀S:setoid.∀w1,w2:list S. [ ] = w1 @ w2 → w1 = [ ]. #S w1; ncases w1; //. nqed. -(* lemma 12 *) +(* lemma 12 *) (* XXX: a case where Leibnitz equality could be exploited for H *) 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;//; -*; ##[##2:*] nelim e; +#S r; ncases r; #e b; @; ##[##2: #H; ncases b in H; ##[##2:*] #; @2; /2/; ##] +ncases b; //; *; ##[##2:*] nelim e; ##[ ##1,2: *; ##| #c; *; ##| #c; *| ##7: #p H; ##| #r1 r2 H G; *; ##[##2: nassumption; ##] ##| #r1 r2 H1 H2; *; /2/ by {}] @@ -580,18 +704,9 @@ nchange in match (ϵ false) with ∅; ##| napply (.= (cup0 ? {[]})^-1); napply cupC; ##] nqed. -(* XXX move somewere else *) -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 ∀_:?.?. - (* 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)); napply (.=_1 #╪_1 (epsilon_or ???)); napply (.=_1 (cupA…)^-1); napply (.=_1 (cupA…)╪_1#); @@ -605,10 +720,12 @@ nqed. (* XXX problem: auto does not find # (refl) when it has a concrete == *) nlemma odotEt : ∀S:Alpha.∀e1,e2:pitem S.∀b2:bool. 〈e1,true〉 ⊙ 〈e2,b2〉 = 〈e1 · \fst (•e2),b2 || \snd (•e2)〉. -#S e1 e2 b2; ncases b2; nnormalize; @; //; @; napply refl; nqed. +#S e1 e2 b2; ncases b2; @; /3/ by refl, conj, I; nqed. +(* nlemma LcatE : ∀S:Alpha.∀e1,e2:pitem S. 𝐋\p (e1 · e2) = 𝐋\p e1 · 𝐋 |e2| ∪ 𝐋\p e2. //; nqed. +*) nlemma cup_dotD : ∀S:Alpha.∀p,q,r:lang S.(p ∪ q) · r = (p · r) ∪ (q · r). #S p q r; napply ext_set; #w; nnormalize; @; @@ -629,9 +746,6 @@ nlemma erase_star : ∀S:Alpha.∀e1:pitem S.𝐋 |e1|^* = 𝐋 |e1^*|. //; nqed nlemma mem_single : ∀S:setoid.∀a,b:S. a ∈ {(b)} → a = b. #S a b; nnormalize; /2/; nqed. -notation < "[\setoid\emsp\of\emsp term 19 x]" non associative with precedence 90 for @{'mk_setoid $x}. -interpretation "mk_setoid" 'mk_setoid x = (mk_setoid x ?). - nlemma cup_sub: ∀S.∀A,B:𝛀^S.∀x. ¬ (x ∈ A) → A ∪ (B - {(x)}) = (A ∪ B) - {(x)}. #S A B x H; napply ext_set; #w; @; ##[ *; ##[ #wa; @; ##[@;//] #H2; napply H; napply (. (mem_single ??? H2)^-1╪_1#); //] @@ -648,87 +762,86 @@ 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; nchange in ⊢ (???%) with (|e1| · |e2|); - -finqui: manca · morfismo, oppure un lemma che dice che == è come Leibnitz. - - nrewrite < IH1; nrewrite < IH2; - nchange in ⊢ (??(??%)?) with (\fst (•e1 ⊙ 〈e2,false〉)); - ncases (•e1); #e3 b; ncases b; nnormalize; - ##[ ncases (•e2); //; ##| nrewrite > IH2; //] -##| #e1 e2 IH1 IH2; nchange in ⊢ (???%) with (.|e1| + .|e2|); - nrewrite < IH2; nrewrite < IH1; - nchange in ⊢ (??(??%)?) with (\fst (•e1 ⊕ •e2)); +##[ #e1 e2 IH1 IH2; + 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] + napply (.=_0 #╪_0 (IH2)); //; +##| #e1 e2 IH1 IH2; napply (?^-1); + napply (.=_0 (IH1^-1)╪_0(IH2^-1)); + nchange in match (•(e1+?)) with (?⊕?); ncases (•e1); ncases (•e2); //] -##| #e IH; nchange in ⊢ (???%) with (.|e|^* ); nrewrite < IH; - nchange in ⊢ (??(??%)?) with (\fst (•e))^*; //; ##] nqed. -nlemma eta_lp : ∀S.∀p:pre S.𝐋\p p = 𝐋\p 〈\fst p, \snd p〉. +(* +nlemma eta_lp : ∀S:Alpha.∀p:pre S. 𝐋\p p = 𝐋\p 〈\fst p, \snd p〉. #S p; ncases p; //; nqed. +*) + +(* XXX coercion ext_carr non applica *) +nlemma epsilon_dot: ∀S:Alpha.∀p:Elang S. {[]} · (ext_carr ? p) = p. +#S e; napply ext_set; #w; @; ##[##2: #Hw; @[]; @w; @; //; @; //; napply #; (* XXX auto *) ##] +*; #w1; *; #w2; *; *; #defw defw1 Hw2; +napply (. defw╪_1#); +napply (. ((defw1 : [ ] = ?)^-1 ╪_0 #)╪_1#); +napply Hw2; +nqed. + -nlemma epsilon_dot: ∀S.∀p:word S → Prop. {[]} · p = p. -#S e; napply extP; #w; nnormalize; @; ##[##2: #Hw; @[]; @w; /3/; ##] -*; #w1; *; #w2; *; *; #defw defw1 Hw2; nrewrite < defw; nrewrite < defw1; -napply Hw2; nqed. (* theorem 16: 1 → 3 *) -nlemma odot_dot_aux : ∀S.∀e1,e2: pre S. - 𝐋\p (•(\fst e2)) = 𝐋\p (\fst e2) ∪ 𝐋 .|\fst e2| → - 𝐋\p (e1 ⊙ e2) = 𝐋\p e1 · 𝐋 .|\fst e2| ∪ 𝐋\p e2. +nlemma odot_dot_aux : ∀S:Alpha.∀e1,e2: pre S. + 𝐋\p (•(\fst e2)) = 𝐋\p (\fst e2) ∪ 𝐋 |\fst e2| → + 𝐋\p (e1 ⊙ e2) = 𝐋\p e1 · 𝐋 |\fst e2| ∪ 𝐋\p e2. #S e1 e2 th1; ncases e1; #e1' b1'; ncases b1'; -##[ nwhd in ⊢ (??(??%)?); nletin e2' ≝ (\fst e2); nletin b2' ≝ (\snd e2); +##[ nchange in match (〈?,true〉⊙?) with 〈?,?〉; + nletin e2' ≝ (\fst e2); nletin b2' ≝ (\snd e2); nletin e2'' ≝ (\fst (•(\fst e2))); nletin b2'' ≝ (\snd (•(\fst e2))); - nchange in ⊢ (??%?) with (?∪?); - nchange in ⊢ (??(??%?)?) with (?∪?); - nchange in match (𝐋\p 〈?,?〉) with (?∪?); - nrewrite > (epsilon_or …); nrewrite > (cupC ? (ϵ ?)…); - nrewrite > (cupA …);nrewrite < (cupA ?? (ϵ?)…); - nrewrite > (?: 𝐋\p e2'' ∪ ϵ b2'' = 𝐋\p e2' ∪ 𝐋 .|e2'|); ##[##2: - nchange with (𝐋\p 〈e2'',b2''〉 = 𝐋\p e2' ∪ 𝐋 .|e2'|); - ngeneralize in match th1; - nrewrite > (eta_lp…); #th1; nrewrite > th1; //;##] - nrewrite > (eta_lp ? e2); - nchange in match (𝐋\p 〈\fst e2,?〉) with (𝐋\p e2'∪ ϵ b2'); - nrewrite > (cup_dotD …); nrewrite > (epsilon_dot…); - nrewrite > (cupC ? (𝐋\p e2')…); nrewrite > (cupA…);nrewrite > (cupA…); - nrewrite < (erase_bull S e2') in ⊢ (???(??%?)); //; -##| ncases e2; #e2' b2'; nchange in match (〈e1',false〉⊙?) with 〈?,?〉; - nchange in match (𝐋\p ?) with (?∪?); - nchange in match (𝐋\p (e1'·?)) with (?∪?); - nchange in match (𝐋\p 〈e1',?〉) with (?∪?); - nrewrite > (cup0…); - nrewrite > (cupA…); //;##] + napply (.=_1 (# ╪_1 (epsilon_or …))); (* XXX … is too slow if combined with .= *) + nchange in match b2'' with b2''; (* XXX some unfoldings happened *) + nchange in match b2' with b2'; + napply (.=_1 (# ╪_1 (cupC …))); napply (.=_1 (cupA …)); + napply (.=_1 (# ╪_1 (cupA …)^-1)); (* XXX slow, but not because of disamb! *) + ncut (𝐋\p e2'' ∪ ϵ b2'' = 𝐋\p e2' ∪ 𝐋 |e2'|); ##[ + napply (?^-1); napply (.=_1 th1^-1); //;##] #E; + napply (.=_1 (# ╪_1 (E ╪_1 #))); + napply (?^-1); + napply (.=_1 (cup_dotD …) ╪_1 #); + napply (.=_1 (# ╪_1 (epsilon_dot …)) ╪_1 #); + napply (?^-1); + napply (.=_1 # ╪_1 ((cupC …) ╪_1 #)); + napply (.=_1 (cupA …)^-1); + napply (.=_1 (cupA …)^-1 ╪_1 #); + napply (.=_1 (cupA …)); + napply (.=_1 (((# ╪_1 (┼_1 (erase_bull S e2')) )╪_1 #)╪_1 #)); + //; +##| ncases e2; #e2' b2'; nchange in match (𝐋\p ?) with (?∪?∪?); + napply (.=_1 (cupA…)); + napply (?^-1); nchange in match (𝐋\p 〈?,false〉) with (?∪?); + napply (.=_1 ((cup0…)╪_1#)╪_1#); + //] nqed. + + 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 *) @@ -736,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: