X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2Fre%2Fre-setoids.ma;h=dcf1d85fe99fa8488b3e566c0b16d0dad7ed9e05;hb=d05dded8c907533b3aba2fcc75c82fa56478af0e;hp=561ff9d6a62b762f4e7119a5c7686b369bb3bd78;hpb=0ce50c5cdfd8433a29f05774a4977cd050ad474d;p=helm.git diff --git a/helm/software/matita/nlibrary/re/re-setoids.ma b/helm/software/matita/nlibrary/re/re-setoids.ma index 561ff9d6a..dcf1d85fe 100644 --- a/helm/software/matita/nlibrary/re/re-setoids.ma +++ b/helm/software/matita/nlibrary/re/re-setoids.ma @@ -12,95 +12,74 @@ (* *) (**************************************************************************) -include "datatypes/pairs.ma". -include "datatypes/bool.ma". -include "logic/cprop.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. +*) -ninductive list (A:setoid) : Type[0] ≝ - | nil: list A - | cons: A -> list A -> list A. - -nlet rec eq_list A (l1, l2 : list A) on l1 : CProp[0] ≝ -match l1 with -[ nil ⇒ match l2 with [ nil ⇒ ? | _ ⇒ ? ] -| cons x xs ⇒ match l2 with [ nil ⇒ ? | cons y ys ⇒ x = y ∧ eq_list ? xs ys]]. -##[ napply True|napply False|napply False]nqed. - -ndefinition LIST : setoid → setoid. -#S; @(list S); @(eq_list S); ncases admit; nqed. - -unification hint 0 ≔ S; - P1 ≟ refl ? (eq (LIST S)), - P2 ≟ sym ? (eq (LIST S)), - P3 ≟ trans ? (eq (LIST S)), - X ≟ mk_setoid (list S) (mk_equivalence_relation ? (eq_list S) P1 P2 P3) -(*-----------------------------------------------------------------------*) ⊢ - carr X ≡ list S. - -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}}. - -notation "hvbox(l1 break @ l2)" - right associative with precedence 47 - for @{'append $l1 $l2 }. - -interpretation "nil" 'nil = (nil ?). -interpretation "cons" 'cons hd tl = (cons ? hd tl). +(* XXX move somewere else *) +ndefinition if': ∀A,B:CPROP. A = B → A → B. +#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 if : ∀A,B:CPROP. ∀p:A = B. A → B ≝ if' on _p : eq_rel1 ???? to ∀_:?.?. -interpretation "append" 'append l1 l2 = (append ? l1 l2). +(* XXX move to list-setoids-theory.ma *) -ntheorem append_nil: ∀A.∀l:list A.l @ [] = l. +ntheorem append_nil: ∀A:setoid.∀l:list A.l @ [] = l. #A;#l;nelim l;//; #a;#l1;#IH;nnormalize;/2/;nqed. ndefinition associative ≝ λA:setoid.λf:A → A → A.∀x,y,z.f x (f y z) = f (f x y) z. -ninductive one : Type[0] ≝ unit : one. +ntheorem associative_append: ∀A:setoid.associative (list A) (append A). +#A;#x;#y;#z;nelim x[ napply (refl ???) |#a;#x1;#H;nnormalize;/2/]nqed. + +(* end move to list *) -ndefinition force ≝ - λS:Type[1].λs:S.λT:Type[1].λt:T.λlock:one. - match lock return λ_.Type[1] with [ unit ⇒ T]. -nlet rec lift (S:Type[1]) (s:S) (T:Type[1]) (t:T) (lock:one) on lock : force S s T t lock ≝ - match lock return λlock.force S s T t lock with [ unit ⇒ t ]. +(* 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). -ncoercion lift : ∀S:Type[1].∀s:S.∀T:Type[1].∀t:T.∀lock:one. force S s T t lock ≝ lift - on s : ? to force ?????. +ninductive eq (A:Type[0]) (x:A) : A → CProp[0] ≝ erefl: eq A x x. -unification hint 0 ≔ R : setoid; - TR ≟ setoid, MR ≟ (carr R), lock ≟ unit -(* ---------------------------------------- *) ⊢ - setoid ≡ force ?(*Type[0]*) MR TR R lock. +nlemma eq_rect_Type0_r': + ∀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. -ntheorem associative_append: ∀A.associative (list A) (append A). -#A;#x;#y;#z;nelim x[//|#a;#x1;#H;nnormalize;/2/]nqed. +nlemma eq_rect_Type0_r: + ∀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. -interpretation "iff" 'iff a b = (iff a b). +nlemma eq_rect_CProp0_r': + ∀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 (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. nrecord Alpha : Type[1] ≝ { - carr :> setoid; - eqb: carr → carr → bool; (* - eqb_true: ∀x,y. (eqb x y = true) = (x = y) *) + acarr :> setoid; + eqb: acarr → acarr → bool; + eqb_true: ∀x,y. (eqb x y = true) = (x = y) }. -notation "a == b" non associative with precedence 45 for @{ 'eqb $a $b }. -interpretation "eqb" 'eqb a b = (eqb ? a b). - -ndefinition word ≝ λS:Alpha.list S. +interpretation "eqb" 'eq_low a b = (eqb ? a b). +(* end alpha *) -ninductive re (S: Alpha) : Type[0] ≝ +(* re *) +ninductive re (S: Type[0]) : Type[0] ≝ z: re S | e: re S | s: S → re S @@ -109,7 +88,7 @@ ninductive re (S: Alpha) : Type[0] ≝ | k: re S → re S. 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). @@ -117,7 +96,7 @@ 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:Alpha.∀p:re S. re S → re S ≝ c on _p : re ? to ∀_:?.?. +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}. @@ -126,55 +105,123 @@ interpretation "atom" 'ps a = (s ? a). notation "ϵ" non associative with precedence 90 for @{ 'epsilon }. interpretation "epsilon" 'epsilon = (e ?). -notation "∅" non associative with precedence 90 for @{ 'empty }. -interpretation "empty" 'empty = (z ?). +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] + | e ⇒ match b with [ e ⇒ True | _ ⇒ False] + | s x ⇒ match b with [ s y ⇒ x = y | _ ⇒ False] + | c r1 r2 ⇒ match b with [ c s1 s2 ⇒ eq_re ? r1 s1 ∧ eq_re ? r2 s2 | _ ⇒ False] + | o r1 r2 ⇒ match b with [ o s1 s2 ⇒ eq_re ? r1 s1 ∧ eq_re ? r2 s2 | _ ⇒ False] + | k r1 ⇒ match b with [ k r2 ⇒ eq_re ? r1 r2 | _ ⇒ False]]. + +interpretation "eq_re" 'eq_low a b = (eq_re ? a b). + +ndefinition RE : Alpha → setoid. +#A; @(re A); @(eq_re A); +##[ #p; nelim p; /2/; +##| #p1; nelim p1; ##[##1,2: #p2; ncases p2; /2/; + ##|##2,3: #x p2; ncases p2; /2/; + ##|##4,5: #e1 e2 H1 H2 p2; ncases p2; /3/; #e3 e4; *; #; @; /2/; + ##|#r H p2; ncases p2; /2/;##] +##| #p1; nelim p1; + ##[ ##1,2: #y z; ncases y; ncases z; //; nnormalize; #; ncases (?:False); //; + ##| ##3: #a; #y z; ncases y; ncases z; /2/; nnormalize; #; ncases (?:False); //; + ##| ##4,5: #r1 r2 H1 H2 y z; ncases y; ncases z; //; nnormalize; + ##[##1,3,4,5,6,8: #; ncases (?:False); //;##] + #r1 r2 r3 r4; nnormalize; *; #H1 H2; *; #H3 H4; /3/; + ##| #r H y z; ncases y; ncases z; //; nnormalize; ##[##1,2,3: #; ncases (?:False); //] + #r2 r3; /3/; ##]##] +nqed. -nlet rec flatten (S : Alpha) (l : list (word S)) on l : word S ≝ -match l with [ nil ⇒ [ ] | cons w tl ⇒ w @ flatten ? tl ]. +unification hint 0 ≔ A : Alpha; + S ≟ acarr A, + T ≟ carr S, + P1 ≟ refl ? (eq0 (RE A)), + P2 ≟ sym ? (eq0 (RE A)), + P3 ≟ trans ? (eq0 (RE A)), + X ≟ mk_setoid (re T) (mk_equivalence_relation ? (eq_re A) P1 P2 P3) +(*-----------------------------------------------------------------------*) ⊢ + carr X ≡ re T. + +unification hint 0 ≔ A:Alpha,a,b:re A; + R ≟ eq0 (RE A), + L ≟ re A +(* -------------------------------------------- *) ⊢ + eq_re A a b ≡ eq_rel L R a b. + +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. -nlet rec conjunct (S : Alpha) (l : list (word S)) (r : word S → Prop) on l: Prop ≝ -match l with [ nil ⇒ ? | cons w tl ⇒ r w ∧ conjunct ? tl r ]. napply True. 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)), + T ≟ RE S +(*--------------------------------------------------------------------------*) ⊢ + fun1 T T (fun1 T (unary_morph_setoid T T) MM A) B ≡ 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. -ndefinition empty_lang ≝ λS.λw:word S.False. -notation "{}" non associative with precedence 90 for @{'empty_lang}. -interpretation "empty lang" 'empty_lang = (empty_lang ?). +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)), + T ≟ RE S +(*--------------------------------------------------------------------------*) ⊢ + fun1 T T (fun1 T (unary_morph_setoid T T) MM A) B ≡ A + B. -ndefinition sing_lang ≝ λS.λx,w:word S.x=w. -notation "{x}" non associative with precedence 90 for @{'sing_lang $x}. -interpretation "sing lang" 'sing_lang x = (sing_lang ? x). +(* end setoids support for re *) -ndefinition union : ∀S,l1,l2,w.Prop ≝ λS.λl1,l2.λw:word S.l1 w ∨ l2 w. -interpretation "union lang" 'union a b = (union ? a b). +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 cat : ∀S,l1,l2,w.Prop ≝ - λS.λl1,l2.λw:word S.∃w1,w2.w1 @ w2 = w ∧ l1 w1 ∧ l2 w2. +interpretation "subset construction with type" 'comprehension t \eta.x = + (mk_powerclass t x). + +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). -ndefinition star ≝ λS.λl.λw:word S.∃lw.flatten ? lw = w ∧ conjunct ? lw l. +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). -notation > "𝐋 term 90 E" non associative with precedence 90 for @{in_l ? $E}. -nlet rec in_l (S : Alpha) (r : re S) on r : word S → Prop ≝ +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 -[ z ⇒ {} +[ z ⇒ ∅ | e ⇒ { [ ] } | s x ⇒ { [x] } | c r1 r2 ⇒ 𝐋 r1 · 𝐋 r2 | o r1 r2 ⇒ 𝐋 r1 ∪ 𝐋 r2 | k r1 ⇒ (𝐋 r1) ^*]. -notation "𝐋 term 90 E" non associative with precedence 90 for @{'in_l $E}. -interpretation "in_l" 'in_l E = (in_l ? E). -interpretation "in_l mem" 'mem w l = (in_l ? l w). - -notation "a || b" left associative with precedence 30 for @{'orb $a $b}. -interpretation "orb" 'orb a b = (orb a b). - -ndefinition if_then_else ≝ λT:Type[0].λe,t,f.match e return λ_.T with [ true ⇒ t | false ⇒ f]. -notation > "'if' term 19 e 'then' term 19 t 'else' term 19 f" non associative with precedence 19 for @{ 'if_then_else $e $t $f }. -notation < "'if' \nbsp term 19 e \nbsp 'then' \nbsp term 19 t \nbsp 'else' \nbsp term 90 f \nbsp" non associative with precedence 19 for @{ 'if_then_else $e $t $f }. -interpretation "if_then_else" 'if_then_else e t f = (if_then_else ? e t f). +notation "𝐋 term 70 E" non associative with precedence 75 for @{'L_re $E}. +interpretation "in_l" 'L_re E = (L_re ? E). -ninductive pitem (S: Alpha) : Type[0] ≝ +ninductive pitem (S: Type[0]) : Type[0] ≝ pz: pitem S | pe: pitem S | ps: S → pitem S @@ -183,8 +230,6 @@ ninductive pitem (S: Alpha) : Type[0] ≝ | po: pitem S → pitem S → pitem S | pk: pitem S → pitem S. -ndefinition pre ≝ λS.pitem S × bool. - interpretation "pstar" 'pk a = (pk ? a). interpretation "por" 'plus a b = (po ? a b). interpretation "pcat" 'pc a b = (pc ? a b). @@ -195,73 +240,164 @@ interpretation "ppatom" 'pp a = (pp ? a). 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 = (pz ?). +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] + | pe ⇒ match p2 with [ pe ⇒ True | _ ⇒ False] + | ps x ⇒ match p2 with [ ps y ⇒ x = y | _ ⇒ False] + | pp x ⇒ match p2 with [ pp y ⇒ x = y | _ ⇒ False] + | pc a1 a2 ⇒ match p2 with [ pc b1 b2 ⇒ eq_pitem ? a1 b1 ∧ eq_pitem ? a2 b2| _ ⇒ False] + | po a1 a2 ⇒ match p2 with [ po b1 b2 ⇒ eq_pitem ? a1 b1 ∧ eq_pitem ? a2 b2| _ ⇒ False] + | pk a ⇒ match p2 with [ pk b ⇒ eq_pitem ? a b | _ ⇒ False]]. + +interpretation "eq_pitem" 'eq_low a b = (eq_pitem ? a b). + +nlemma PITEM : ∀S:Alpha.setoid. +#S; @(pitem S); @(eq_pitem …); +##[ #p; nelim p; //; nnormalize; #; @; //; +##| #p; nelim p; ##[##1,2: #y; ncases y; //; ##|##3,4: #x y; ncases y; //; #; napply (?^-1); nassumption; + ##|##5,6: #r1 r2 H1 H2 p2; ncases p2; //; #s1 s2; nnormalize; *; #; @; /2/; + ##| #r H y; ncases y; //; nnormalize; /2/;##] +##| #x; nelim x; + ##[ ##1,2: #y z; ncases y; ncases z; //; nnormalize; #; ncases (?:False); //; + ##| ##3,4: #a; #y z; ncases y; ncases z; /2/; nnormalize; #; ncases (?:False); //; + ##| ##5,6: #r1 r2 H1 H2 y z; ncases y; ncases z; //; nnormalize; + ##[##1,2,5,6,7,8,4,10: #; ncases (?:False); //;##] + #r1 r2 r3 r4; nnormalize; *; #H1 H2; *; #H3 H4; /3/; + ##| #r H y z; ncases y; ncases z; //; nnormalize; ##[##1,2,3,4: #; ncases (?:False); //] + #r2 r3; /3/; ##]##] +nqed. + +alias symbol "hint_decl" (instance 1) = "hint_decl_Type1". +unification hint 0 ≔ SS:Alpha; + S ≟ acarr SS, + A ≟ carr S, + 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) +(*---------------------------*)⊢ + carr R ≡ pitem A. + +unification hint 0 ≔ S:Alpha,a,b:pitem S; + R ≟ PITEM S, + L ≟ (pitem S) +(* -------------------------------------------- *) ⊢ + eq_pitem S a b ≡ eq_rel L (eq0 R) a b. + +(* end setoids for pitem *) + +ndefinition pre ≝ λS.pitem S × bool. + +notation "\fst term 90 x" non associative with precedence 90 for @{'fst $x}. +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). 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 - [ pz ⇒ ∅ + [ pz ⇒ 0 | pe ⇒ ϵ | ps x ⇒ `x | pp x ⇒ `x | 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}. + +notation < "|term 19 e|" non associative with precedence 70 for @{'forget $e}. interpretation "forget" 'forget a = (forget ? a). -notation "\fst term 90 x" non associative with precedence 90 for @{'fst $x}. -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). - -notation > "𝐋\p\ term 90 E" non associative with precedence 90 for @{in_pl ? $E}. -nlet rec in_pl (S : Alpha) (r : pitem S) on r : word S → Prop ≝ +notation > "𝐋\p\ term 70 E" non associative with precedence 75 for @{L_pi ? $E}. +nlet rec L_pi (S : Alpha) (r : pitem S) on r : lang S ≝ match r with -[ pz ⇒ {} -| pe ⇒ {} -| ps _ ⇒ {} +[ pz ⇒ ∅ +| pe ⇒ ∅ +| ps _ ⇒ ∅ | pp x ⇒ { [x] } -| pc r1 r2 ⇒ 𝐋\p\ r1 · 𝐋 .|r2| ∪ 𝐋\p\ r2 +| pc r1 r2 ⇒ 𝐋\p\ r1 · 𝐋 |r2| ∪ 𝐋\p\ r2 | po r1 r2 ⇒ 𝐋\p\ r1 ∪ 𝐋\p\ r2 -| pk r1 ⇒ 𝐋\p\ r1 · 𝐋 (.|r1|^* ) ]. -notation > "𝐋\p term 90 E" non associative with precedence 90 for @{'in_pl $E}. -notation "𝐋\sub(\p) term 90 E" non associative with precedence 90 for @{'in_pl $E}. -interpretation "in_pl" 'in_pl E = (in_pl ? E). -interpretation "in_pl mem" 'mem w l = (in_pl ? l w). +| pk r1 ⇒ 𝐋\p\ r1 · 𝐋 (|r1|^* ) ]. +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). + +(* 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 ((∃_.?)∨?); good! *) + 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); (* good! *) +##| #e H; + ncut (∀x1,x2.(w1 = (x1@x2)) = (w2 = (x1@x2)));##[ + #x1 x2; @; #X; ##[ napply ((.= E^-1) X) | napply ((.= E) X) ] ##] #X; + (* nnormalize in ⊢ (???%%); good! (but a bit too hard) *) + napply (∑w1,w2. X w1 w2 / H ; (H╪_1#)╪_1#); +##] +nqed. -ndefinition epsilon ≝ λS,b.if b then { ([ ] : word S) } else {}. +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. + +(* end set support for 𝐋\p *) + +ndefinition epsilon ≝ + λS:Alpha.λb.match b return λ_.lang S with [ true ⇒ { [ ] } | _ ⇒ ∅ ]. interpretation "epsilon" 'epsilon = (epsilon ?). notation < "ϵ b" non associative with precedence 90 for @{'app_epsilon $b}. interpretation "epsilon lang" 'app_epsilon b = (epsilon ? b). -ndefinition in_prl ≝ λS : Alpha.λp:pre S. 𝐋\p\ (\fst p) ∪ ϵ (\snd p). +ndefinition L_pr ≝ λS : Alpha.λp:pre S. 𝐋\p\ (\fst p) ∪ ϵ (\snd p). -interpretation "in_prl mem" 'mem w l = (in_prl ? l w). -interpretation "in_prl" 'in_pl E = (in_prl ? E). +interpretation "L_pr" 'L_pi E = (L_pr ? E). -nlemma append_eq_nil : ∀S.∀w1,w2:word S. w1 @ w2 = [ ] → w1 = [ ]. -#S w1; nelim w1; //. #x tl IH w2; nnormalize; #abs; ndestruct; nqed. - -(* lemma 12 *) -nlemma epsilon_in_true : ∀S.∀e:pre S. [ ] ∈ e ↔ \snd e = true. -#S r; ncases r; #e b; @; ##[##2: #H; nrewrite > H; @2; /2/; ##] ncases b;//; -nnormalize; *; ##[##2:*] nelim e; -##[ ##1,2: *; ##| #c; *; ##| #c; nnormalize; #; ndestruct; ##| ##7: #p H; -##| #r1 r2 H G; *; ##[##2: /3/ by or_intror] -##| #r1 r2 H1 H2; *; /3/ by or_intror, or_introl; ##] -*; #w1; *; #w2; *; *; #defw1; nrewrite > (append_eq_nil … w1 w2 …); /3/ by {};//; +nlemma append_eq_nil : ∀S:setoid.∀w1,w2:list S. [ ] = w1 @ w2 → w1 = [ ]. +#S w1; ncases w1; //. nqed. + +(* 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; 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 {}] +*; #w1; *; #w2; *; *; +##[ #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〉. @@ -286,12 +422,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 〉 @@ -305,112 +441,125 @@ 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). - -naxiom extP : ∀S.∀p,q:word S → Prop.(p =1 q) → p = q. - 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 ? {[]})^-1); +##| napply (.= (cup0 ? {[]})^-1); napply cupC; ##] 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 cupC : ∀S. ∀a,b:word S → Prop.a ∪ b = b ∪ a. -#S a b; napply extP; #w; @; *; nnormalize; /2/; nqed. - (* 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 …); //; +#S r1; ncases r1; #e1 b1 r2; ncases r2; #e2 b2; (* oh my! +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#); +napply (.=_1 (#╪_1(cupC…))╪_1#); +napply (.=_1 (cupA…)^-1╪_1#); +napply (.=_1 (cupA…)); +//; nqed. -nlemma odotEt : - ∀S.∀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. +(* 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; @; /3/ by refl, conj, I; nqed. -nlemma cup_dotD : ∀S.∀p,q,r:word S → Prop.(p ∪ q) · r = (p · r) ∪ (q · r). -#S p q r; napply extP; #w; nnormalize; @; +(* +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; @; ##[ *; #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; +nlemma erase_dot : ∀S:Alpha.∀e1,e2:pitem S.𝐋 |e1 · e2| = 𝐋 |e1| · 𝐋 |e2|. +#S e1 e2; napply ext_set; nnormalize; #w; @; *; #w1; *; #w2; *; *; /7/ by ex_intro, conj; nqed. -nlemma erase_plus : ∀S.∀e1,e2:pitem S.𝐋 .|e1 + e2| = 𝐋 .|e1| ∪ 𝐋 .|e2|. -#S e1 e2; napply extP; nnormalize; #w; @; *; /4/ by or_introl, or_intror; nqed. +nlemma erase_plus : ∀S:Alpha.∀e1,e2:pitem S.𝐋 |e1 + e2| = 𝐋 |e1| ∪ 𝐋 |e2|. +#S e1 e2; napply ext_set; nnormalize; #w; @; *; /4/ by or_introl, or_intror; nqed. -nlemma erase_star : ∀S.∀e1:pitem S.𝐋 .|e1|^* = 𝐋 .|e1^*|. //; nqed. +nlemma erase_star : ∀S:Alpha.∀e1:pitem S.𝐋 |e1|^* = 𝐋 |e1^*|. //; nqed. -ndefinition substract := λS.λp,q:word S → Prop.λw.p w ∧ ¬ q w. -interpretation "substract" 'minus a b = (substract ? a b). +nlemma mem_single : ∀S:setoid.∀a,b:S. a ∈ {(b)} → a = b. +#S a b; nnormalize; /2/; nqed. -nlemma cup_sub: ∀S.∀a,b:word S → Prop. ¬ (a []) → a ∪ (b - {[]}) = (a ∪ b) - {[]}. -#S a b c; napply extP; #w; nnormalize; @; *; /4/; *; /4/; nqed. +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#); //] + *; #wb nwn; @; ##[@2;//] //; +##| *; *; ##[ #wa nwn; @; //] #wb nwn; @2; @; //;##] +nqed. -nlemma sub0 : ∀S.∀a:word S → Prop. a - {} = a. -#S a; napply extP; #w; nnormalize; @; /3/; *; //; nqed. +nlemma sub0 : ∀S.∀a:Ω^S. a - ∅ = a. +#S a; napply ext_set; #w; nnormalize; @; /3/; *; //; nqed. -nlemma subK : ∀S.∀a:word S → Prop. a - a = {}. -#S a; napply extP; #w; nnormalize; @; *; /2/; nqed. +nlemma subK : ∀S.∀a:Ω^S. a - a = ∅. +#S a; napply ext_set; #w; nnormalize; @; *; /2/; nqed. -nlemma subW : ∀S.∀a,b:word S → Prop.∀w.(a - b) w → a w. +nlemma subW : ∀S.∀a,b:Ω^S.∀w.w ∈ (a - b) → w ∈ a. #S a b w; nnormalize; *; //; nqed. -nlemma erase_bull : ∀S.∀a:pitem S. .|\fst (•a)| = .|a|. +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|); - 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)); - 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〉. +##[ #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); //] +nqed. + +(* +nlemma eta_lp : ∀S:Alpha.∀p:pre S. 𝐋\p p = 𝐋\p 〈\fst p, \snd p〉. #S p; ncases p; //; 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. +(* 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. (* 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; //;##] + 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'|); ##[ + nchange with (𝐋\p 〈e2'',b2''〉 = 𝐋\p e2' ∪ 𝐋 |e2'|); + napply (?^-1); napply (.=_1 th1^-1); //;##] #E; + napply (.=_1 (# ╪_1 (E ╪_1 #))); + STOP + nrewrite > (eta_lp ? e2); nchange in match (𝐋\p 〈\fst e2,?〉) with (𝐋\p e2'∪ ϵ b2'); nrewrite > (cup_dotD …); nrewrite > (epsilon_dot…);