]> matita.cs.unibo.it Git - helm.git/commitdiff
Added Streicher's K axiom.
authorWilmer Ricciotti <ricciott@cs.unibo.it>
Tue, 2 Feb 2010 18:50:29 +0000 (18:50 +0000)
committerWilmer Ricciotti <ricciott@cs.unibo.it>
Tue, 2 Feb 2010 18:50:29 +0000 (18:50 +0000)
helm/software/matita/nlibrary/logic/destruct_bb.ma
helm/software/matita/nlibrary/logic/equality.ma

index 269cfc765c8b93b3f15f95f372efb33037eb681b..308f8bebeb91ed3b554cca708bb4e6fbd0eb1124 100644 (file)
@@ -19,6 +19,7 @@ ninductive unit: Type[0] ≝ k: unit.
 ninductive bool: unit → Type[0] ≝ true : bool k | false : bool k.
 
 nlemma foo: true = false → False. #H; ndestruct;
+nqed.
 
 (* nlemma prova : ∀T:Type[0].∀a,b:T.∀e:a = b.
                ∀P:∀x,y:T.x=y→Prop.P a a (refl T a) → P a b e.
index 6fdac61b9cf424b4c7ebf2457167d1b959636a9f..8ae1bcde4c0c8640136d3232f2af7ed391726143 100644 (file)
@@ -117,6 +117,8 @@ napply (R3 ????????? e0 ? e1 ? e2);
 napply a4;
 nqed.
 
+naxiom streicherK : ∀T:Type[0].∀t:T.∀P:t = t → Type[2].P (refl ? t) → ∀p.P p. 
+
 ndefinition EQ: ∀A:Type[0]. equivalence_relation A.
  #A; napply mk_equivalence_relation
   [ napply eq