]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/nlibrary/re/re-setoids.ma
update in groud_2 and models
[helm.git] / helm / software / matita / nlibrary / re / re-setoids.ma
index cfbd7cb31d8c0d5b203654d388d6e64958882ba7..729e8a0ea3a7806d55b8a334d952b613daee4aef 100644 (file)
@@ -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: