]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/nlibrary/Plogic/connectives.ma
update in groud_2 and models
[helm.git] / helm / software / matita / nlibrary / Plogic / connectives.ma
index d840143e3835ca29080b6b51eb0542690a84762f..bf17771fb9890a632f0f49f2b8e223837712f36a 100644 (file)
@@ -23,18 +23,25 @@ ninductive False: Prop ≝ .
 
 default "false" cic:/matita/basics/connectives/False.ind.
 
+(*
 ndefinition Not: Prop → Prop ≝
-λA. A → False.
+λA. A → False. *)
+
+ninductive Not (A:Prop): Prop ≝
+nmk: (A → False) → Not A.
 
 interpretation "logical not" 'not x = (Not x).
 
+ntheorem absurd : ∀ A:Prop. A → ¬A → False.
+#A; #H; #Hn; nelim Hn;/2/; nqed.
+
+(*
 ntheorem absurd : ∀ A,C:Prop. A → ¬A → C.
 #A; #C; #H; #Hn; nelim (Hn H).
-nqed.
+nqed. *)
 
 ntheorem not_to_not : ∀A,B:Prop. (A → B) → ¬B →¬A.
-/3/.
-nqed.
+/4/; nqed.
 
 ninductive And (A,B:Prop) : Prop ≝
     conj : A → B → And A B.
@@ -68,3 +75,5 @@ ninductive ex2 (A:Type[0]) (P,Q:A \to Prop) : Prop ≝
 
 ndefinition iff :=
  λ A,B. (A → B) ∧ (B → A).
+
+interpretation "iff" 'iff a b = (iff a b).