]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/library/nat/plus.ma
ocaml 3.09 transition
[helm.git] / helm / matita / library / nat / plus.ma
index 6067ebcdcd6b87ffde1fe7bce0059627e6723fb7..d595dad19113cb6626ec4e58cc1e8a2d389be2b1 100644 (file)
@@ -21,8 +21,8 @@ let rec plus n m \def
  [ O \Rightarrow m
  | (S p) \Rightarrow S (plus p m) ].
 
+(*CSC: the URI must disappear: there is a bug now *)
 interpretation "natural plus" 'plus x y = (cic:/matita/nat/plus/plus.con x y).
-alias symbol "plus" (instance 0) = "natural plus".
 
 theorem plus_n_O: \forall n:nat. n = n+O.
 intros.elim n.
@@ -36,11 +36,6 @@ simplify.reflexivity.
 simplify.apply eq_f.assumption.
 qed.
 
-(* some problem here: confusion between relations/symmetric 
-and functions/symmetric; functions symmetric is not in 
-functions.moo why?
-theorem symmetric_plus: symmetric nat plus. *)
-
 theorem sym_plus: \forall n,m:nat. n+m = m+n.
 intros.elim n.
 simplify.apply plus_n_O.
@@ -48,7 +43,7 @@ simplify.rewrite > H.apply plus_n_Sm.
 qed.
 
 theorem associative_plus : associative nat plus.
-simplify.intros.elim x.
+unfold associative.intros.elim x.
 simplify.reflexivity.
 simplify.apply eq_f.assumption.
 qed.
@@ -67,10 +62,9 @@ theorem inj_plus_r: \forall p,n,m:nat. p+n = p+m \to n=m
 
 theorem injective_plus_l: \forall m:nat.injective nat nat (\lambda n.n+m).
 intro.simplify.intros.
-(* qui vorrei applicare injective_plus_r *)
-apply inj_plus_r m.
+apply (injective_plus_r m).
 rewrite < sym_plus.
-rewrite < sym_plus y.
+rewrite < (sym_plus y).
 assumption.
 qed.