]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/library/nat/compare.ma
ported to new syntactic requirement about terms being surrounded by parens
[helm.git] / helm / matita / library / nat / compare.ma
index 55e526c8fad0f80eade4775141c99326975efee8..6431f5f293e219fcc5af7ebe84c315c4da5d0ae4 100644 (file)
@@ -34,19 +34,19 @@ match (eqb n m) with
 [ true  \Rightarrow n = m 
 | false \Rightarrow n \neq m].
 intros.
-apply nat_elim2
+apply (nat_elim2
 (\lambda n,m:nat.match (eqb n m) with
 [ true  \Rightarrow n = m 
-| false \Rightarrow n \neq m]).
+| false \Rightarrow n \neq m])).
 intro.elim n1.
 simplify.reflexivity.
 simplify.apply not_eq_O_S.
 intro.
 simplify.
-intro. apply not_eq_O_S n1.apply sym_eq.assumption.
+intro. apply (not_eq_O_S n1).apply sym_eq.assumption.
 intros.simplify.
 generalize in match H.
-elim (eqb n1 m1).
+elim ((eqb n1 m1)).
 simplify.apply eq_f.apply H1.
 simplify.intro.apply H1.apply inj_S.assumption.
 qed.
@@ -55,13 +55,13 @@ theorem eqb_elim : \forall n,m:nat.\forall P:bool \to Prop.
 (n=m \to (P true)) \to (n \neq m \to (P false)) \to (P (eqb n m)). 
 intros.
 cut 
-match (eqb n m) with
+(match (eqb n m) with
 [ true  \Rightarrow n = m
-| false \Rightarrow n \neq m] \to (P (eqb n m)).
+| false \Rightarrow n \neq m] \to (P (eqb n m))).
 apply Hcut.apply eqb_to_Prop.
-elim eqb n m.
-apply (H H2).
-apply (H1 H2).
+elim (eqb n m).
+apply ((H H2)).
+apply ((H1 H2)).
 qed.
 
 theorem eqb_n_n: \forall n. eqb n n = true.
@@ -71,15 +71,15 @@ qed.
 
 theorem eq_to_eqb_true: \forall n,m:nat.
 n = m \to eqb n m = true.
-intros.apply eqb_elim n m.
+intros.apply (eqb_elim n m).
 intros. reflexivity.
-intros.apply False_ind.apply H1 H.
+intros.apply False_ind.apply (H1 H).
 qed.
 
 theorem not_eq_to_eqb_false: \forall n,m:nat.
 \lnot (n = m) \to eqb n m = false.
-intros.apply eqb_elim n m.
-intros. apply False_ind.apply H H1.
+intros.apply (eqb_elim n m).
+intros. apply False_ind.apply (H H1).
 intros.reflexivity.
 qed.
 
@@ -96,13 +96,13 @@ match (leb n m) with
 [ true  \Rightarrow n \leq m 
 | false \Rightarrow n \nleq m].
 intros.
-apply nat_elim2
+apply (nat_elim2
 (\lambda n,m:nat.match (leb n m) with
 [ true  \Rightarrow n \leq m 
-| false \Rightarrow n \nleq m]).
+| false \Rightarrow n \nleq m])).
 simplify.exact le_O_n.
 simplify.exact not_le_Sn_O.
-intros 2.simplify.elim (leb n1 m1).
+intros 2.simplify.elim ((leb n1 m1)).
 simplify.apply le_S_S.apply H.
 simplify.intros.apply H.apply le_S_S_to_le.assumption.
 qed.
@@ -112,13 +112,13 @@ theorem leb_elim: \forall n,m:nat. \forall P:bool \to Prop.
 P (leb n m).
 intros.
 cut 
-match (leb n m) with
+(match (leb n m) with
 [ true  \Rightarrow n \leq m
-| false \Rightarrow n \nleq m] \to (P (leb n m)).
+| false \Rightarrow n \nleq m] \to (P (leb n m))).
 apply Hcut.apply leb_to_Prop.
-elim leb n m.
-apply (H H2).
-apply (H1 H2).
+elim (leb n m).
+apply ((H H2)).
+apply ((H1 H2)).
 qed.
 
 let rec nat_compare n m: compare \def
@@ -144,7 +144,7 @@ intros.simplify.reflexivity.
 qed.
 
 theorem S_pred: \forall n:nat.lt O n \to eq nat n (S (pred n)).
-intro.elim n.apply False_ind.exact not_le_Sn_O O H.
+intro.elim n.apply False_ind.exact (not_le_Sn_O O H).
 apply eq_f.apply pred_Sn.
 qed.
 
@@ -152,8 +152,8 @@ theorem nat_compare_pred_pred:
 \forall n,m:nat.lt O n \to lt O m \to 
 eq compare (nat_compare n m) (nat_compare (pred n) (pred m)).
 intros.
-apply lt_O_n_elim n H.
-apply lt_O_n_elim m H1.
+apply (lt_O_n_elim n H).
+apply (lt_O_n_elim m H1).
 intros.
 simplify.reflexivity.
 qed.
@@ -164,14 +164,14 @@ match (nat_compare n m) with
   | EQ \Rightarrow n=m
   | GT \Rightarrow m < n ].
 intros.
-apply nat_elim2 (\lambda n,m.match (nat_compare n m) with
+apply (nat_elim2 (\lambda n,m.match (nat_compare n m) with
   [ LT \Rightarrow n < m
   | EQ \Rightarrow n=m
-  | GT \Rightarrow m < n ]).
+  | GT \Rightarrow m < n ])).
 intro.elim n1.simplify.reflexivity.
 simplify.apply le_S_S.apply le_O_n.
 intro.simplify.apply le_S_S. apply le_O_n.
-intros 2.simplify.elim (nat_compare n1 m1).
+intros 2.simplify.elim ((nat_compare n1 m1)).
 simplify. apply le_S_S.apply H.
 simplify. apply eq_f. apply H.
 simplify. apply le_S_S.apply H.
@@ -180,7 +180,7 @@ qed.
 theorem nat_compare_n_m_m_n: \forall n,m:nat. 
 nat_compare n m = compare_invert (nat_compare m n).
 intros. 
-apply nat_elim2 (\lambda n,m. nat_compare n m = compare_invert (nat_compare m n)).
+apply (nat_elim2 (\lambda n,m. nat_compare n m = compare_invert (nat_compare m n))).
 intros.elim n1.simplify.reflexivity.
 simplify.reflexivity.
 intro.elim n1.simplify.reflexivity.
@@ -192,14 +192,14 @@ theorem nat_compare_elim : \forall n,m:nat. \forall P:compare \to Prop.
 (n < m \to P LT) \to (n=m \to P EQ) \to (m < n \to P GT) \to 
 (P (nat_compare n m)).
 intros.
-cut match (nat_compare n m) with
+cut (match (nat_compare n m) with
 [ LT \Rightarrow n < m
 | EQ \Rightarrow n=m
 | GT \Rightarrow m < n] \to
-(P (nat_compare n m)).
+(P (nat_compare n m))).
 apply Hcut.apply nat_compare_to_Prop.
-elim (nat_compare n m).
-apply (H H3).
-apply (H1 H3).
-apply (H2 H3).
+elim ((nat_compare n m)).
+apply ((H H3)).
+apply ((H1 H3)).
+apply ((H2 H3)).
 qed.