]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/tests/pullback.ma
hints work better now
[helm.git] / helm / software / matita / tests / pullback.ma
index d224e6b0e25b1442085ed5dadea99a8e2a4b637f..54b784caae96b84d1ea88e9933938f8b1240c450 100644 (file)
@@ -1,16 +1,27 @@
-(**************************************************************************)
-(*       ___                                                              *)
-(*      ||M||                                                             *)
-(*      ||A||       A project by Andrea Asperti                           *)
-(*      ||T||                                                             *)
-(*      ||I||       Developers:                                           *)
-(*      ||T||         The HELM team.                                      *)
-(*      ||A||         http://helm.cs.unibo.it                             *)
-(*      \   /                                                             *)
-(*       \ /        This file is distributed under the terms of the       *)
-(*        v         GNU General Public License Version 2                  *)
-(*                                                                        *)
-(**************************************************************************)
+
+(*
+axiom A : Type.
+axiom B : Type.
+axiom B1 : Type.
+axiom C : Type.
+
+axiom f1 : A → B.
+axiom f2 : A → B1.
+axiom f3 : B → C.
+axiom f4 : B1 → C.
+
+coercion f1.
+coercion f2.
+coercion f3.
+coercion f4.
+
+
+
+
+axiom P : Prop.
+
+lemma x : P.
+*)
 
 include "logic/equality.ma".
 
@@ -25,7 +36,7 @@ record R : Type \def {
 }.
 
 record LR_ : Type \def {
-  lr_L : L ;
+  lr_L :> L ;
   lr_R_ : R ;
   lr_w : r_c lr_R_ = l_c lr_L
 }.
@@ -36,14 +47,40 @@ intros;constructor 1;
 |rewrite < (lr_w l);apply (r_op (lr_R_ l));]
 qed.
 
+(*
+axiom F : Prop → Prop.
+axiom C : Prop → Prop. 
+
+axiom daemon : ∀x.F x = C x.
+
+lemma xxx : ∀x.F x = C x. apply daemon; qed.
+
+axiom yyyy : ∀x.C (C x) = C (C x) → F (F x) = F (F x).
+
+coercion yyyy. 
+
+lemma x : ∀x. (λy:F (F x) = F (F x).True) (refl_eq ? (C (C x))).
+
+include "nat/factorial.ma".
+lemma xxx : 8! = 8 * 7!. intros; reflexivity; qed.
+
+lemma x : (λy:8!=8!.True) (refl_eq ? (8 * 7!)).
+apply (refl_eq ??);
+*)
+
+(*
 lemma xxx : \forall x. r_c (lr_R x) = l_c (lr_L x).
 intros; reflexivity;
 qed.
+*)
 
-coercion lr_L.
 
 
+definition Prop_OR_LR_1 : LR_ → Prop.
+apply (λx.r_c (lr_R x)).
+qed.
 
+coercion Prop_OR_LR_1.
 coercion lr_R.
 
 lemma foo : \forall x,y.l_op ? x (r_op ? x y) = x.