]> matita.cs.unibo.it Git - helm.git/commitdiff
Coercions via unification hints?
authorClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Wed, 6 Jan 2010 22:37:10 +0000 (22:37 +0000)
committerClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Wed, 6 Jan 2010 22:37:10 +0000 (22:37 +0000)
helm/software/matita/nlibrary/overlap/o-algebra.ma

index 227ef04517d4d65f92121d066d3240606928e7d6..89024fe9fad2c8b928c4ea8d5c269ac7d3608327 100644 (file)
@@ -237,11 +237,15 @@ unification hint 0 ≔ P, Q, r;
  R ≟ (mk_unary_morphism1 … (or_f_minus …) (prop11 … (or_f_minus_morphism1 …)))
 (* ------------------------ *) ⊢
   fun11 … R r ≡ or_f_minus P Q r.
+  
+naxiom daemon : False.
 
 nlemma ORelation_eq_respects_leq_or_f_star_:
  ∀P,Q:OAlgebra.∀r,r':ORelation P Q.
   r=r' → ∀x. r* x ≤ r'* x.
  #P; #Q; #a; #a'; #e; #x; (*CSC: una schifezza *)
+ ncases daemon.
+ (*
  ngeneralize in match (. (or_prop1 P Q a' (a* x) x)^-1) in ⊢ %; #H; napply H;
  nchange with (or_f P Q a' (a* x) ≤ x);
  napply (. ?‡#)
@@ -250,7 +254,7 @@ nlemma ORelation_eq_respects_leq_or_f_star_:
     nchange with (or_f P Q a' = or_f P Q a);
     napply (.= †e^-1); napply #]
  napply (. (or_prop1 …));
- napply oa_leq_refl.
+ napply oa_leq_refl.*)
 nqed.
 
 nlemma ORelation_eq3:
@@ -277,6 +281,7 @@ nlemma ORelation_eq_respects_leq_or_f_minus_star_:
  ∀P,Q:OAlgebra.∀r,r':ORelation P Q.
   r=r' → ∀x. r⎻* x ≤ r'⎻* x.
  #P; #Q; #a; #a'; #e; #x; (*CSC: una schifezza *)
+ ncases daemon. (*
  ngeneralize in match (. (or_prop2 P Q a' (a⎻* x) x)^-1) in ⊢ %; #H; napply H;
  nchange with (or_f_minus P Q a' (a⎻* x) ≤ x);
  napply (. ?‡#)
@@ -285,7 +290,7 @@ nlemma ORelation_eq_respects_leq_or_f_minus_star_:
     nchange with (a'⎻ = a⎻);
     napply (.= †e^-1); napply #]
  napply (. (or_prop2 …));
- napply oa_leq_refl.
+ napply oa_leq_refl.*)
 nqed.
 
 nlemma ORelation_eq4:
@@ -303,17 +308,105 @@ ndefinition or_f_minus_star_morphism1:
   | napply ORelation_eq4]
 nqed.
 
+
 unification hint 0 ≔ P, Q, r;
  R ≟ (mk_unary_morphism1 … (or_f_minus_star …) (prop11 … (or_f_minus_star_morphism1 …)))
 (* ------------------------ *) ⊢
   fun11 … R r ≡ or_f_minus_star P Q r.
+  
+ninductive one : Type[0] ≝ unit : one.
+
+(*   
+ndefinition force : ∀S:Type[2].∀T1,T2:Type[1].(T1 → T2) → ∀R:S.∀lock:one.Type[2] ≝   
+ λS,T1,T2,T,R,lock. match lock with [ unit => S ].
+*)
+
+ndefinition force : ∀S:Type[2]. S → ∀T:Type[2]. T → one → Type[2] ≝   
+ λS,s,T,t,lock. match lock with [ unit => S ].
+
+(*
+ndefinition enrich_as : 
+ ∀S:Type[2].∀T1,T2:Type[1]. ∀T:T1→T2. ∀R:S.∀lock:one.force S T1 T2 T R lock ≝ 
+ λS,T1,T2,T,R,lock. match lock return λlock.match lock with [ unit ⇒ S ] 
+                    with [ unit ⇒ R ].
+*)
 
+ndefinition enrich_as : 
+ ∀S:Type[2].∀s:S.∀T:Type[2].∀t:T.∀lock:one.force S s T t lock ≝ 
+ λS,s,T,t,lock. match lock return λlock.match lock with [ unit ⇒ S ] 
+                    with [ unit ⇒ s ].
+
+(*
+ncoercion enrich_as : ∀S:Type[2].∀T1,T2:Type[1].∀T:T1→T2.∀R:S.∀lock:one.force S T1 T2 T R lock
+ ≝ enrich_as on T : ? → ? to force ? ? ? ? ? ?.
+*)
+
+ncoercion enrich_as : ∀S:Type[2].∀s:S.∀T:Type[2].∀t:T.∀lock:one. force S s T t lock
+ ≝ enrich_as on t: ? to force ? ? ? ? ?.
+
+
+(* does not work here 
+nlemma foo : ∀A,B,C:setoid1.∀f:B ⇒ C.∀g:A ⇒ B. unary_morphism1 A C.
+#A; #B; #C; #f; #g; napply(f \circ g).
+nqed.*)
+
+(*
+unification hint 0 ≔ A,B,C : setoid1, f:B ⇒ C, g: A ⇒ B;
+   lock ≟ unit
+(* --------------------------------------------------------------- *) ⊢
+  (unary_morphism1 A C)
+ ≡
+  (force (unary_morphism1 A C) (carr1 A) (carr1 C)
+    (composition1 A B C f g) (comp1_unary_morphisms A B C f g) lock)
+  .
+*)
+
+unification hint 0 ≔ A,B,C : setoid1, f:B ⇒ C, g: A ⇒ B;
+   lock ≟ unit
+(* --------------------------------------------------------------- *) ⊢
+  (unary_morphism1 A C)
+ ≡
+  (force (unary_morphism1 A C) (comp1_unary_morphisms A B C f g)
+   (carr1 A → carr1 C) (composition1 A B C f g)  lock)
+  .
+
+(*
+unification hint 0 ≔ A,B,C : setoid1, f:B ⇒ C, g: A ⇒ B, X;
+   lock ≟ unit
+(* --------------------------------------------------------------- *) ⊢
+  (unary_morphism1 A C)
+ ≡
+  (force (unary_morphism1 A C) (carr1 A) (carr1 C)
+    (fun11 … X) (X) lock)
+  .
+*)
+
+nlemma foo : ∀A,B,C:setoid1.∀f:B ⇒ C.∀g:A ⇒ B. unary_morphism1 A C.
+#A; #B; #C; #f; #g; napply(f ∘ g).
+nqed.
+
+(*
+
+ndefinition uffa: ∀A,B. ∀U: unary_morphism1 A B. (A → B) → CProp[0].
+ #A;#B;#_;#_; napply True.
+nqed.
+ndefinition mk_uffa: ∀A,B.∀U: unary_morphism1 A B. ∀f: (A → B). uffa A B U f.
+ #A; #B; #U; #f; napply I.
+nqed.
+
+ndefinition coerc_to_unary_morphism1:
+ ∀A,B. ∀U: unary_morphism1 A B. uffa A B U (fun11 … U) → unary_morphism1 A B.
+ #A; #B; #U; #_; nassumption.
+nqed.
+
+ncheck (λA,B,C,f,g.coerc_to_unary_morphism1 ??? (mk_uffa ??? (composition1 A B C f g))). 
+*)
 ndefinition ORelation_composition : ∀P,Q,R. 
   binary_morphism1 (ORelation_setoid P Q) (ORelation_setoid Q R) (ORelation_setoid P R).
 #P; #Q; #R; @
 [ #F; #G; @
-  [ napply (comp1_unary_morphisms … G F) (*CSC: was (G ∘ F);*)
-  | napply (comp1_unary_morphisms … G⎻* F⎻* ) (*CSC: was (G⎻* ∘ F⎻* );*)
+  [ napply (G ∘ F) (* napply (comp1_unary_morphisms … G F) (*CSC: was (G ∘ F);*) *)
+  | napply (G⎻* ∘ F⎻* ) (* napply (comp1_unary_morphisms … G⎻* F⎻* ) (*CSC: was (G⎻* ∘ F⎻* );*)*)
   | napply (comp1_unary_morphisms … F* G* ) (*CSC: was (F* ∘ G* );*)
   | napply (comp1_unary_morphisms … F⎻ G⎻) (*CSC: was (F⎻ ∘ G⎻);*)
   | #p; #q; nnormalize;