]> matita.cs.unibo.it Git - helm.git/commitdiff
First proof finished (some tactics still not working).
authorClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Mon, 13 Jul 2009 09:13:12 +0000 (09:13 +0000)
committerClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Mon, 13 Jul 2009 09:13:12 +0000 (09:13 +0000)
helm/software/matita/nlibrary/algebra/magmas.ma
helm/software/matita/nlibrary/logic/equality.ma

index 7283cbc58f95d733587a3b15cfe9b5b01a8de5e6..4fac05b51c6e2926fdaa85f21e9edc8db03561a0 100644 (file)
@@ -47,6 +47,10 @@ ndefinition mmcarr ≝
  λA,B.λf: pre_magma_morphism A B. match f with [ mk_pre_magma_morphism f _ ⇒ f ].
 ncoercion mmcarr: ∀A,B.∀M: pre_magma_morphism A B. A → B ≝ mmcarr
  on _M: pre_magma_morphism ? ? to ∀_.?.
+ndefinition mmprop ≝
+ λA,B,M.
+  match M return λM:pre_magma_morphism A B.∀x,y. M (op ? x y) = op ? (M x) (M y) with
+   [ mk_pre_magma_morphism _ p ⇒ p ].
 
 nrecord magma_morphism (A) (B) (Ma: magma A) (Mb: magma B) : Type ≝
  { mmmcarr: pre_magma_morphism A B;
@@ -81,5 +85,8 @@ ndefinition mm_image:
      [ napply (op ? x0 y0) 
      | napply (conj ????)
         [ napply (op_closed ??????); nassumption
-        | nelim daemon ]##]
+        | (* nrewrite < Hx1; DOES NOT WORK *)
+          napply (eq_rect ?? (λ_.?) ?? Hx1);
+          napply (eq_rect ?? (λ_.?) ?? Hy1);
+          napply (mmprop ?? f ??)]##]
 nqed.
\ No newline at end of file
index ff8cc5a1392c5393d997f2f85dd0da48e8397942..2418f2a7c504da2a7f8b68957604c30ad224a4be 100644 (file)
@@ -16,6 +16,11 @@ include "logic/connectives.ma".
 
 ninductive eq (A: Type) (a: A) : A → CProp ≝
  refl: eq A a a.
+
+nlet rec eq_rect (A: Type) (x: A) (P: ∀y:A. eq A x y → CProp) (q: P x (refl A x))
+ (y: A) (p: eq A x y) on p : P y p ≝
+ match p with
+  [ refl ⇒ q ].
  
 interpretation "leibnitz's equality" 'eq t x y = (eq t x y).