simplify.exact I.
elim y. simplify.exact I.
simplify.
-(*CSC: qui uso le perche' altrimenti ci sono troppe scelte
- per via delle coercions! *)
cut match (nat_compare n1 n) with
[ LT \Rightarrow n1<n
| EQ \Rightarrow n1=n
| GT \Rightarrow n<n1] \to
match (nat_compare n1 n) with
-[ LT \Rightarrow (le (S n1) n)
+[ LT \Rightarrow (S n1) \leq n
| EQ \Rightarrow neg n = neg n1
-| GT \Rightarrow (le (S n) n1)].
+| GT \Rightarrow (S n) \leq n1].
apply Hcut. apply nat_compare_to_Prop.
elim (nat_compare n1 n).
simplify.exact H.
elim y.simplify.exact I.
simplify.exact I.
simplify.
-(*CSC: qui uso le perche' altrimenti ci sono troppe scelte
- per via delle coercions! *)
cut match (nat_compare n n1) with
[ LT \Rightarrow n<n1
| EQ \Rightarrow n=n1
| GT \Rightarrow n1<n] \to
match (nat_compare n n1) with
-[ LT \Rightarrow (le (S n) n1)
+[ LT \Rightarrow (S n) \leq n1
| EQ \Rightarrow pos n = pos n1
-| GT \Rightarrow (le (S n1) n)].
+| GT \Rightarrow (S n1) \leq n].
apply Hcut. apply nat_compare_to_Prop.
elim (nat_compare n n1).
simplify.exact H.
theorem irrefl_Zlt: irreflexive Z Zlt
\def irreflexive_Zlt.
-(*CSC: qui uso lt perche' ho due istanze diverse di < *)
theorem Zlt_neg_neg_to_lt:
-\forall n,m:nat. neg n < neg m \to lt m n.
+\forall n,m:nat. neg n < neg m \to m < n.
intros.apply H.
qed.
-(*CSC: qui uso lt perche' ho due istanze diverse di < *)
-theorem lt_to_Zlt_neg_neg: \forall n,m:nat.lt m n \to neg n < neg m.
+theorem lt_to_Zlt_neg_neg: \forall n,m:nat.m < n \to neg n < neg m.
intros.
simplify.apply H.
qed.
-(*CSC: qui uso lt perche' ho due istanze diverse di < *)
theorem Zlt_pos_pos_to_lt:
-\forall n,m:nat. pos n < pos m \to lt n m.
+\forall n,m:nat. pos n < pos m \to n < m.
intros.apply H.
qed.
-(*CSC: qui uso lt perche' ho due istanze diverse di < *)
-theorem lt_to_Zlt_pos_pos: \forall n,m:nat.lt n m \to pos n < pos m.
+theorem lt_to_Zlt_pos_pos: \forall n,m:nat.n < m \to pos n < pos m.
intros.
simplify.apply H.
qed.
(*CSC: the URI must disappear: there is a bug now *)
interpretation "logical not" 'not x = (cic:/matita/logic/connectives/Not.con x).
-(*CSC: this alias should disappear. It is now required because the notation for Coq is pre-loaded *)
-alias symbol "not" (instance 0) = "logical not".
theorem absurd : \forall A,C:Prop. A \to \lnot A \to C.
intros. elim (H1 H).
(*CSC: the URI must disappear: there is a bug now *)
interpretation "logical and" 'and x y = (cic:/matita/logic/connectives/And.ind#xpointer(1/1) x y).
-(*CSC: this alias should disappear. It is now required because the notation for Coq is pre-loaded *)
-alias symbol "and" (instance 0) = "logical and".
theorem proj1: \forall A,B:Prop. A \land B \to A.
intros. elim H. assumption.
(*CSC: the URI must disappear: there is a bug now *)
interpretation "logical or" 'or x y = (cic:/matita/logic/connectives/Or.ind#xpointer(1/1) x y).
-(*CSC: this alias should disappear. It is now required because the notation for Coq is pre-loaded *)
-alias symbol "or" (instance 0) = "logical or".
definition decidable : Prop \to Prop \def \lambda A:Prop. A \lor \not A.
(*CSC: the URI must disappear: there is a bug now *)
interpretation "leibnitz's equality"
'eq x y = (cic:/matita/logic/equality/eq.ind#xpointer(1/1) _ x y).
-(*CSC: this alias should disappear. It is now required because the notation for Coq is pre-loaded *)
-alias symbol "eq" (instance 0) = "leibnitz's equality".
-
theorem reflexive_eq : \forall A:Type. reflexive A (eq A).
simplify.intros.apply refl_eq.
qed.