]> matita.cs.unibo.it Git - helm.git/commitdiff
16.2
authorEnrico Tassi <enrico.tassi@inria.fr>
Fri, 1 Oct 2010 16:31:08 +0000 (16:31 +0000)
committerEnrico Tassi <enrico.tassi@inria.fr>
Fri, 1 Oct 2010 16:31:08 +0000 (16:31 +0000)
helm/software/matita/nlibrary/re/re-setoids.ma

index 70b39c7c49af91faf9859037bd3c8d45c54a7156..729e8a0ea3a7806d55b8a334d952b613daee4aef 100644 (file)
@@ -161,15 +161,25 @@ unification hint 0 ≔ A:Alpha, a,b:re (carr (acarr 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".
@@ -184,13 +194,7 @@ unification hint 0 ≔ S:Alpha, A,B:re (carr (acarr 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; 
-##[##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),
@@ -202,6 +206,26 @@ unification hint 0 ≔ S:Alpha, A,B:re (carr (acarr 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.
+
+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 *)
 
 nlet rec conjunct S (l : list (list S)) (L : lang S) on l: CProp[0] ≝
@@ -738,6 +762,7 @@ 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;
@@ -766,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 (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.    
-(* XXX the previous hint does not work *)
+
 
 (* theorem 16: 1 → 3 *)
 nlemma odot_dot_aux : ∀S:Alpha.∀e1,e2: pre S.
@@ -800,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…));
@@ -845,8 +861,7 @@ ntheorem bull_cup : ∀S:Alpha. ∀e:pitem S.  𝐋\p (•e) =  𝐋\p e ∪ 
       napply (.=_1 (cupA …));
       napply (.=_1 # ╪_1 ((erase_dot ???)^-1 ╪_1 (cup0 ??)));
       napply (.=_1 # ╪_1 (cupC…));
-      napply (.=_1 (cupA …)^-1);
-      //;
+      napply (.=_1 (cupA …)^-1); //;
   ##| #e1 e2 IH1 IH2;
       nchange in match (•(?+?)) with (•e1 ⊕ •e2);
       napply (.=_1 (oplus_cup …));
@@ -856,25 +871,33 @@ ntheorem bull_cup : ∀S:Alpha. ∀e:pitem S.  𝐋\p (•e) =  𝐋\p e ∪ 
       napply (.=_1 # ╪_1 (cupA ????)^-1);
       napply (.=_1 # ╪_1 (cupC…));
       napply (.=_1 (cupA ????)^-1);
-      napply (.=_1 # ╪_1 (erase_plus ???)^-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'^* ) ∪ {[ ]});
-STOP
-      nchange in match (𝐋\p (pk ? e')) 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;##]
+      (* 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: 
   ∀S.∀e1,e2: pre S.  𝐋\p (e1 ⊙ e2) =  𝐋\p e1 · 𝐋 .|\fst e2| ∪ 𝐋\p e2.