]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/dama/fields.ma
removed dust
[helm.git] / matita / dama / fields.ma
index 0618f4923a92097fafcdd2199067bb4914679ccf..5ab17edae98c7e2eafd4a3a94e65125ab9e7c433 100644 (file)
@@ -21,10 +21,10 @@ record is_field (R:ring) (inv:∀x:R.x ≠ 0 → R) : Prop
  {  (* multiplicative abelian properties *)
     mult_comm_: symmetric ? (mult R);
     (* multiplicative group properties *)
-    inv_inverse_: ∀x.∀p: x ≠ 0. mult ? (inv x p) x = 1
+    inv_inverse_: ∀x.∀p: x ≠ 0. inv x p * x = 1
  }.
 
-lemma opp_opp: \forall R:ring. \forall x:R. (-(-x))=x.
+lemma opp_opp: ∀R:ring. ∀x:R. --x=x.
 intros; 
 apply (cancellationlaw ? (-x) ? ?); 
 rewrite > (opp_inverse R x); 
@@ -50,10 +50,11 @@ theorem mult_comm: ∀F:field.symmetric ? (mult F).
  apply (mult_comm_ ? ? (field_properties F)).
 qed.
 
-theorem inv_inverse: ∀F:field.∀x.∀p: x ≠ 0. mult ? (inv F x p) x = 1.
+theorem inv_inverse: ∀F:field.∀x:F.∀p: x ≠ 0. (inv ? x p)*x = 1.
  intro;
  apply (inv_inverse_ ? ? (field_properties F)).
 qed.
 
+(*CSC: qua funzionava anche mettendo ? al posto della prima F*)
 definition sum_field ≝
- λF:field. sum ? (plus F) (zero F) (one F).
+ λF:field. sum F (plus F) 0 1.