definition left_inverse \def λC,op.λe:C.λinv:C→C. ∀x:C. op (inv x) x = e.
+definition distributive_right ≝
+ λA:Type.λf:A→A→A.λg:A→A→A.
+ ∀x,y,z. f (g x y) z = g (f x z) (f y z).
+
record is_abelian_group (C:Type) (plus:C→C→C) (zero:C) (opp:C→C) : Prop \def
{ (* abelian additive semigroup properties *)
plus_assoc: associative ? plus;
opp_inverse: left_inverse ? plus zero opp
}.
-record is_field (C:Type) (plus:C→C→C) (mult:C→C→C) (zero,one:C) (opp:C→C)
- (inv:∀x:C.x ≠ zero →C) : Prop
+record is_ring (C:Type) (plus:C→C→C) (mult:C→C→C) (zero:C) (opp:C→C) : Prop
≝
- { (* abelian group properties *)
+ { (* abelian group properties *)
abelian_group:> is_abelian_group ? plus zero opp;
- (* abelian multiplicative semigroup properties *)
+ (* multiplicative semigroup properties *)
mult_assoc: associative ? mult;
+ (* ring properties *)
+ mult_plus_distr_left: distributive ? mult plus;
+ mult_plus_distr_right: distributive_right C mult plus
+ }.
+
+record ring : Type \def
+ { r_carrier:> Type;
+ r_plus: r_carrier → r_carrier → r_carrier;
+ r_mult: r_carrier → r_carrier → r_carrier;
+ r_zero: r_carrier;
+ r_opp: r_carrier → r_carrier;
+ r_ring_properties:> is_ring ? r_plus r_mult r_zero r_opp
+ }.
+
+notation "0" with precedence 89
+for @{ 'zero }.
+
+interpretation "Ring zero" 'zero =
+ (cic:/matita/integration_algebras/r_zero.con _).
+
+interpretation "Ring plus" 'plus a b =
+ (cic:/matita/integration_algebras/r_plus.con _ a b).
+
+interpretation "Ring mult" 'times a b =
+ (cic:/matita/integration_algebras/r_mult.con _ a b).
+
+interpretation "Ring opp" 'uminus a =
+ (cic:/matita/integration_algebras/r_opp.con _ a).
+
+lemma eq_mult_zero_x_zero: ∀R:ring.∀x:R.0*x=0.
+ intros;
+ generalize in match (zero_neutral ? ? ? ? R 0); intro;
+ generalize in match (eq_f ? ? (λy.y*x) ? ? H); intro; clear H;
+ rewrite > (mult_plus_distr_right ? ? ? ? ? R) in H1;
+ generalize in match (eq_f ? ? (λy.-(0*x)+y) ? ? H1); intro; clear H1;
+ rewrite < (plus_assoc ? ? ? ? R) in H;
+ rewrite > (opp_inverse ? ? ? ? R) in H;
+ rewrite > (zero_neutral ? ? ? ? R) in H;
+ assumption.
+qed.
+
+record is_field (C:Type) (plus:C→C→C) (mult:C→C→C) (zero,one:C) (opp:C→C)
+ (inv:∀x:C.x ≠ zero →C) : Prop
+≝
+ { (* ring properties *)
+ ring_properties: is_ring ? plus mult zero opp;
+ (* multiplicative abelian properties *)
mult_comm: symmetric ? mult;
(* multiplicative monoid properties *)
one_neutral: left_neutral ? mult one;
(* multiplicative group properties *)
inv_inverse: ∀x.∀p: x ≠ zero. mult (inv x p) x = one;
- (* ring properties *)
- mult_plus_distr: distributive ? mult plus;
(* integral domain *)
not_eq_zero_one: zero ≠ one
}.
].
record field : Type \def
- { carrier:> Type;
- plus: carrier → carrier → carrier;
- mult: carrier → carrier → carrier;
- zero: carrier;
- one: carrier;
- opp: carrier → carrier;
- inv: ∀x:carrier. x ≠ zero → carrier;
- field_properties: is_field ? plus mult zero one opp inv
+ { f_ring:> ring;
+ one: f_ring;
+ inv: ∀x:f_ring. x ≠ 0 → f_ring;
+ field_properties:
+ is_field ? (r_plus f_ring) (r_mult f_ring) (r_zero f_ring) one
+ (r_opp f_ring) inv
}.
definition sum_field ≝
- λF:field. sum ? (plus F) (zero F) (one F).
+ λF:field. sum ? (r_plus F) (r_zero F) (one F).
-notation "0" with precedence 89
-for @{ 'zero }.
-
-interpretation "Field zero" 'zero =
- (cic:/matita/integration_algebras/zero.con _).
-
notation "1" with precedence 89
for @{ 'one }.
interpretation "Field one" 'one =
(cic:/matita/integration_algebras/one.con _).
-interpretation "Field plus" 'plus a b =
- (cic:/matita/integration_algebras/plus.con _ a b).
-
-interpretation "Field mult" 'times a b =
- (cic:/matita/integration_algebras/mult.con _ a b).
-
-interpretation "Field opp" 'uminus a =
- (cic:/matita/integration_algebras/opp.con _ a).
-
record is_ordered_field_ch0 (C:Type) (plus,mult:C→C→C) (zero,one:C) (opp:C→C)
(inv:∀x:C.x ≠ zero → C) (le:C→C→Prop) : Prop \def
{ (* field properties *)
{ of_field:> field;
of_le: of_field → of_field → Prop;
of_ordered_field_properties:>
- is_ordered_field_ch0 ? (plus of_field) (mult of_field) (zero of_field)
- (one of_field) (opp of_field) (inv of_field) of_le
+ is_ordered_field_ch0 ? (r_plus of_field) (r_mult of_field) (r_zero of_field)
+ (one of_field) (r_opp of_field) (inv of_field) of_le
}.
interpretation "Ordered field le" 'leq a b =
interpretation "Ordered field lt" 'lt a b =
(cic:/matita/integration_algebras/lt.con _ a b).
-lemma le_zero_x_to_le_opp_x_zero: ∀F:ordered_field_ch0.∀x:F. 0 ≤ x → -x ≤ 0.
- intros;
+axiom le_zero_x_to_le_opp_x_zero: ∀F:ordered_field_ch0.∀x:F. 0 ≤ x → -x ≤ 0.
+(* intros;
generalize in match (of_plus_compat ? ? ? ? ? ? ? ? F ? ? (-x) H); intro;
rewrite > (zero_neutral ? ? ? ? F) in H1;
rewrite > (plus_comm ? ? ? ? F) in H1;
rewrite > (opp_inverse ? ? ? ? F) in H1;
assumption.
-qed.
+qed.*)
-lemma le_x_zero_to_le_zero_opp_x: ∀F:ordered_field_ch0.∀x:F. x ≤ 0 → 0 ≤ -x.
- intros;
+axiom le_x_zero_to_le_zero_opp_x: ∀F:ordered_field_ch0.∀x:F. x ≤ 0 → 0 ≤ -x.
+(* intros;
generalize in match (of_plus_compat ? ? ? ? ? ? ? ? F ? ? (-x) H); intro;
rewrite > (zero_neutral ? ? ? ? F) in H1;
rewrite > (plus_comm ? ? ? ? F) in H1;
rewrite > (opp_inverse ? ? ? ? F) in H1;
assumption.
-qed.
-
-(* To be proved for rings only *)
-lemma eq_mult_zero_x_zero: ∀F:ordered_field_ch0.∀x:F.0*x=0.
- intros;
- generalize in match (zero_neutral ? ? ? ? F 0); intro;
- generalize in match (eq_f ? ? (λy.x*y) ? ? H); intro; clear H;
- rewrite > (mult_plus_distr ? ? ? ? ? ? ? F) in H1;
- generalize in match (eq_f ? ? (λy.-(x*0)+y) ? ? H1); intro; clear H1;
- rewrite < (plus_assoc ? ? ? ? F) in H;
- rewrite > (opp_inverse ? ? ? ? F) in H;
- rewrite > (zero_neutral ? ? ? ? F) in H;
- rewrite > (mult_comm ? ? ? ? ? ? ? F) in H;
- assumption.
-qed.
+qed.*)
(*
lemma eq_opp_x_times_opp_one_x: ∀F:ordered_field_ch0.∀x:F.-x = -1*x.
axiom not_eq_sum_field_zero: ∀F,n. n > O → sum_field F n ≠ 0.
record is_vector_space (K: field) (C:Type) (plus:C→C→C) (zero:C) (opp:C→C)
- (mult:K→C→C) : Prop
+ (emult:K→C→C) : Prop
≝
{ (* abelian group properties *)
vs_abelian_group: is_abelian_group ? plus zero opp;
(* other properties *)
- vs_nilpotent: ∀v. mult 0 v = zero;
- vs_neutral: ∀v. mult 1 v = v;
- vs_distributive: ∀a,b,v. mult (a + b) v = plus (mult a v) (mult b v);
- vs_associative: ∀a,b,v. mult (a * b) v = mult a (mult b v)
+ vs_nilpotent: ∀v. emult 0 v = zero;
+ vs_neutral: ∀v. emult 1 v = v;
+ vs_distributive: ∀a,b,v. emult (a + b) v = plus (emult a v) (emult b v);
+ vs_associative: ∀a,b,v. emult (a * b) v = emult a (emult b v)
}.
record is_lattice (C:Type) (join,meet:C→C→C) : Prop \def
definition le \def λC.λmeet:C→C→C.λf,g. meet f g = f.
record is_riesz_space (K:ordered_field_ch0) (C:Type) (plus:C→C→C) (zero:C)
- (opp:C→C) (mult:K→C→C) (join,meet:C→C→C) : Prop \def
+ (opp:C→C) (emult:K→C→C) (join,meet:C→C→C) : Prop \def
{ (* vector space properties *)
- rs_vector_space: is_vector_space K C plus zero opp mult;
+ rs_vector_space: is_vector_space K C plus zero opp emult;
(* lattice properties *)
rs_lattice: is_lattice C join meet;
(* other properties *)
rs_compat_le_plus: ∀f,g,h. le ? meet f g →le ? meet (plus f h) (plus g h);
- rs_compat_le_times: ∀a,f. 0≤a → le ? meet zero f → le ? meet zero (mult a f)
+ rs_compat_le_times: ∀a,f. 0≤a → le ? meet zero f → le ? meet zero (emult a f)
}.
definition absolute_value \def λC:Type.λopp.λjoin:C→C→C.λf.join f (opp f).
record is_archimedean_riesz_space (K:ordered_field_ch0) (C:Type) (plus:C→C→C)
- (zero:C) (opp:C→C) (mult:Type_OF_ordered_field_ch0 K→C→C) (join,meet:C→C→C) : Prop \def
+ (zero:C) (opp:C→C) (mult:Type_OF_ordered_field_ch0 K→C→C) (join,meet:C→C→C)
+ :Prop \def
{ ars_riesz_space: is_riesz_space ? ? plus zero opp mult join meet;
ars_archimedean: ∃u.∀n,a.∀p:n > O.
le C meet (absolute_value ? opp join a)
(mult (inv K (sum_field K n) (not_eq_sum_field_zero K n p)) u) →
a = zero
- }.
\ No newline at end of file
+ }.
+
+record is_algebra (K: field) (C:Type) (plus:C→C→C) (zero:C) (opp:C→C)
+ (emult:K→C→C) (mult:C→C→C) : Prop
+≝
+ { (* vector space properties *)
+ a_vector_space_properties: is_vector_space ? ? plus zero opp emult;
+ (* ring properties *)
+ a_ring: is_ring ? plus mult zero opp;
+ (* algebra properties *)
+ a_associative_left: ∀a,f,g. emult a (mult f g) = mult (emult a f) g;
+ a_associative_right: ∀a,f,g. emult a (mult f g) = mult f (emult a g)
+ }.
\ No newline at end of file