]> matita.cs.unibo.it Git - helm.git/blob - matita/dama/ordered_fields_ch0.ma
More work on groups, real numbers and integration algebras.
[helm.git] / matita / dama / ordered_fields_ch0.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 set "baseuri" "cic:/matita/ordered_fields_ch0/".
16
17 include "fields.ma".
18
19 record is_total_order_relation (C:Type) (le:C→C→Prop) : Type \def
20  { to_cotransitive: ∀x,y,z:C. le x y → le x z ∨ le z y;
21    to_antisimmetry: ∀x,y:C. le x y → le y x → x=y
22  }.
23
24 theorem to_transitive:
25  ∀C.∀le:C→C→Prop. is_total_order_relation ? le → transitive ? le.
26  intros;
27  unfold transitive;
28  intros;
29  elim (to_cotransitive ? ? i ? ? z H);
30   [ assumption
31   | rewrite < (to_antisimmetry ? ? i ? ? H1 t);
32     assumption
33   ].
34 qed.
35
36 record is_ordered_field_ch0 (F:field) (le:F→F→Prop) : Type \def
37  { of_total_order_relation:> is_total_order_relation ? le;
38    of_mult_compat: ∀a,b. le 0 a → le 0 b → le 0 (a*b);
39    of_plus_compat: ∀a,b,c. le a b → le (a+c) (b+c);
40    of_weak_tricotomy : ∀a,b. a≠b → le a b ∨ le b a;
41    (* 0 characteristics *)
42    of_char0: ∀n. n > O → sum ? (plus F) 0 1 n ≠ 0
43  }.
44  
45 record ordered_field_ch0 : Type \def
46  { of_field:> field;
47    of_le: of_field → of_field → Prop;
48    of_ordered_field_properties:> is_ordered_field_ch0 ? of_le
49  }.
50
51 interpretation "Ordered field le" 'leq a b =
52  (cic:/matita/ordered_fields_ch0/of_le.con _ a b).
53  
54 definition lt \def λF:ordered_field_ch0.λa,b:F.a ≤ b ∧ a ≠ b.
55
56 interpretation "Ordered field lt" 'lt a b =
57  (cic:/matita/ordered_fields_ch0/lt.con _ a b).
58
59 lemma le_zero_x_to_le_opp_x_zero: ∀F:ordered_field_ch0.∀x:F. 0 ≤ x → -x ≤ 0.
60 intros;
61  generalize in match (of_plus_compat ? ? F ? ? (-x) H); intro;
62  rewrite > zero_neutral in H1;
63  rewrite > plus_comm in H1;
64  rewrite > opp_inverse in H1;
65  assumption.
66 qed.
67
68 lemma le_x_zero_to_le_zero_opp_x: ∀F:ordered_field_ch0.∀x:F. x ≤ 0 → 0 ≤ -x.
69  intros;
70  generalize in match (of_plus_compat ? ? F ? ? (-x) H); intro;
71  rewrite > zero_neutral in H1;
72  rewrite > plus_comm in H1;
73  rewrite > opp_inverse in H1;
74  assumption.
75 qed.
76
77 (*
78 lemma eq_opp_x_times_opp_one_x: ∀F:ordered_field_ch0.∀x:F.-x = -1*x.
79  intros;
80  
81 lemma not_eq_x_zero_to_lt_zero_mult_x_x:
82  ∀F:ordered_field_ch0.∀x:F. x ≠ 0 → 0 < x * x.
83  intros;
84  elim (of_weak_tricotomy ? ? ? ? ? ? ? ? F ? ? H);
85   [ generalize in match (le_x_zero_to_le_zero_opp_x F ? H1); intro;
86     generalize in match (of_mult_compat ? ? ? ? ? ? ? ?  F ? ? H2 H2); intro;
87 *)  
88
89 axiom lt_zero_to_lt_inv_zero:
90  ∀F:ordered_field_ch0.∀x:F.∀p:x≠0. 0 < x → 0 < inv ? x p.
91  
92 (* The ordering is not necessary. *)
93 axiom not_eq_sum_field_zero: ∀F:ordered_field_ch0.∀n. O<n → sum_field F n ≠ 0.
94 axiom le_zero_sum_field: ∀F:ordered_field_ch0.∀n. O<n → 0 < sum_field F n.
95
96 axiom lt_zero_to_le_inv_zero:
97  ∀F:ordered_field_ch0.∀n:nat.∀p:sum_field F n ≠ 0. 0 ≤ inv ? (sum_field ? n) p.
98
99 definition tends_to : ∀F:ordered_field_ch0.∀f:nat→F.∀l:F.Prop.
100  apply
101   (λF:ordered_field_ch0.λf:nat → F.λl:F.
102     ∀n:nat.∃m:nat.∀j:nat. m≤j →
103      l - (inv F (sum_field ? (S n)) ?) ≤ f j ∧
104      f j ≤ l + (inv F (sum_field ? (S n)) ?));
105  apply not_eq_sum_field_zero;
106  unfold;
107  auto new.
108 qed.
109
110 (*
111 definition is_cauchy_seq ≝
112  λF:ordered_field_ch0.λf:nat→F.
113   ∀eps:F. 0 < eps →
114    ∃n:nat.∀M. n ≤ M →
115     -eps ≤ f M - f n ∧ f M - f n ≤ eps.
116 *)
117
118
119
120 definition is_cauchy_seq : ∀F:ordered_field_ch0.∀f:nat→F.Prop.
121  apply
122   (λF:ordered_field_ch0.λf:nat→F.
123     ∀m:nat.
124      ∃n:nat.∀N. n ≤ N →
125       -(inv ? (sum_field F (S m)) ?) ≤ f N - f n ∧
126       f N - f n ≤ inv ? (sum_field F (S m)) ?);
127  apply not_eq_sum_field_zero;
128  unfold;
129  auto.
130 qed.
131
132 definition is_complete ≝
133  λF:ordered_field_ch0.
134   ∀f:nat→F. is_cauchy_seq ? f →
135    ex F (λl:F. tends_to ? f l).