1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 set "baseuri" "cic:/matita/integration_algebras/".
17 include "vector_spaces.ma".
18 include "lattices.ma".
20 (**************** Riesz Spaces ********************)
22 record pre_riesz_space (K:ordered_field_ch0) : Type \def
23 { rs_vector_space:> vector_space K;
25 rs_ordered_abelian_group_: ordered_abelian_group;
27 og_abelian_group rs_ordered_abelian_group_ = vs_abelian_group ? rs_vector_space;
29 og_ordered_set rs_ordered_abelian_group_ = ordered_set_of_lattice rs_lattice_
32 lemma rs_lattice: ∀K.pre_riesz_space K → lattice.
34 cut (os_carrier (rs_lattice_ ? V) = V);
37 | apply (eq_rect ? ? (λC:Type.C→C→C) ? ? Hcut);
39 | apply (eq_rect ? ? (λC:Type.C→C→C) ? ? Hcut);
43 (λa:Type.λH:os_carrier (rs_lattice_ ? V)=a.
45 (eq_rect Type (rs_lattice_ K V) (λC:Type.C→C→C)
46 (l_join (rs_lattice_ K V)) a H)
47 (eq_rect Type (rs_lattice_ K V) (λC:Type.C→C→C)
48 (l_meet (rs_lattice_ K V)) a H))
51 apply l_lattice_properties
53 | transitivity (os_carrier (rs_ordered_abelian_group_ ? V));
54 [ apply (eq_f ? ? os_carrier);
57 | apply (eq_f ? ? carrier);
63 coercion cic:/matita/integration_algebras/rs_lattice.con.
65 lemma rs_ordered_abelian_group: ∀K.pre_riesz_space K → ordered_abelian_group.
67 apply mk_ordered_abelian_group;
68 [ apply mk_pre_ordered_abelian_group;
69 [ apply (vs_abelian_group ? (rs_vector_space ? V))
70 | apply (ordered_set_of_lattice (rs_lattice ? V))
75 (og_ordered_abelian_group_properties (rs_ordered_abelian_group_ ? V));
85 (plus (abelian_group_OF_pre_riesz_space K V) f h)
86 (plus (abelian_group_OF_pre_riesz_space K V) g h))
92 os_le (ordered_set_OF_pre_riesz_space K V) f g →
93 os_le (ordered_set_OF_pre_riesz_space K V)
94 (plus (abelian_group_OF_pre_riesz_space K V) f h)
95 (plus (abelian_group_OF_pre_riesz_space K V) g h))
98 apply og_ordered_abelian_group_properties*)
102 coercion cic:/matita/integration_algebras/rs_ordered_abelian_group.con.
104 record is_riesz_space (K:ordered_field_ch0) (V:pre_riesz_space K) : Prop ≝
105 { rs_compat_le_times: ∀a:K.∀f:V. 0≤a → 0≤f → 0≤a*f
108 record riesz_space (K:ordered_field_ch0) : Type \def
109 { rs_pre_riesz_space:> pre_riesz_space K;
110 rs_riesz_space_properties: is_riesz_space ? rs_pre_riesz_space
113 record is_positive_linear (K) (V:riesz_space K) (T:V→K) : Prop ≝
114 { positive: ∀u:V. 0≤u → 0≤T u;
115 linear1: ∀u,v:V. T (u+v) = T u + T v;
116 linear2: ∀u:V.∀k:K. T (k*u) = k*(T u)
119 record sequentially_order_continuous (K) (V:riesz_space K) (T:V→K) : Prop ≝
121 ∀a:nat→V.∀l:V.is_increasing ? a → is_sup V a l →
122 is_increasing K (λn.T (a n)) ∧ tends_to ? (λn.T (a n)) (T l)
125 definition absolute_value ≝ λK.λS:riesz_space K.λf:S.f ∨ -f.
127 (**************** Normed Riesz spaces ****************************)
129 definition is_riesz_norm ≝
130 λR:real.λV:riesz_space R.λnorm:norm R V.
131 ∀f,g:V. absolute_value ? V f ≤ absolute_value ? V g →
132 n_function R V norm f ≤ n_function R V norm g.
134 record riesz_norm (R:real) (V:riesz_space R) : Type ≝
135 { rn_norm:> norm R V;
136 rn_riesz_norm_property: is_riesz_norm ? ? rn_norm
139 (*CSC: non fa la chiusura delle coercion verso funclass *)
140 definition rn_function ≝
141 λR:real.λV:riesz_space R.λnorm:riesz_norm ? V.
142 n_function R V (rn_norm ? ? norm).
144 coercion cic:/matita/integration_algebras/rn_function.con 1.
146 (************************** L-SPACES *************************************)
148 record is_l_space (R:real) (V:riesz_space R) (norm:riesz_norm ? V) : Prop ≝
149 { ls_banach: is_complete ? V (induced_distance ? ? norm);
150 ls_linear: ∀f,g:V. le ? V 0 f → le ? V 0 g → norm (f+g) = norm f + norm g
153 (******************** ARCHIMEDEAN RIESZ SPACES ***************************)
155 record is_archimedean_riesz_space (K) (S:riesz_space K) : Prop
157 { ars_archimedean: ∃u:S.∀n.∀a.∀p:n > O.
158 absolute_value ? S a ≤
159 (inv K (sum_field K n) (not_eq_sum_field_zero K n p))* u →
163 record archimedean_riesz_space (K:ordered_field_ch0) : Type \def
164 { ars_riesz_space:> riesz_space K;
165 ars_archimedean_property: is_archimedean_riesz_space ? ars_riesz_space
168 definition is_weak_unit ≝
169 (* This definition is by Spitters. He cites Fremlin 353P, but:
170 1. that theorem holds only in f-algebras (as in Spitters, but we are
171 defining it on Riesz spaces)
172 2. Fremlin proves |x|/\u=0 \to u=0. How do we remove the absolute value?
173 λR:real.λV:archimedean_riesz_space R.λunit: V.
174 ∀x:V. meet x unit = 0 → u = 0.
175 3. Fremlin proves u > 0 implies x /\ u > 0 > 0 for Archimedean spaces
176 only. We pick this definition for now.
177 *) λR:real.λV:archimedean_riesz_space R.λe:V.
178 ∀v:V. 0<v → 0 < (v ∧ e).
180 (* Here we are avoiding a construction (the quotient space to define
181 f=g iff I(|f-g|)=0 *)
182 record integration_riesz_space (R:real) : Type \def
183 { irs_archimedean_riesz_space:> archimedean_riesz_space R;
184 irs_unit: irs_archimedean_riesz_space;
185 irs_weak_unit: is_weak_unit ? ? irs_unit;
186 integral: irs_archimedean_riesz_space → R;
187 irs_positive_linear: is_positive_linear ? ? integral;
189 ∀f:irs_archimedean_riesz_space.
191 (λn.integral (f ∧ ((sum_field R n)*irs_unit)))
194 ∀f:irs_archimedean_riesz_space.
198 ((inv ? (sum_field R (S n))
199 (not_eq_sum_field_zero R (S n) (le_S_S O n (le_O_n n)))
202 ∀f,g:irs_archimedean_riesz_space.
203 integral (absolute_value ? irs_archimedean_riesz_space (f - g)) = 0 → f=g
206 definition induced_norm_fun ≝
207 λR:real.λV:integration_riesz_space R.λf:V.
208 integral ? V (absolute_value ? ? f).
210 lemma induced_norm_is_norm:
211 ∀R:real.∀V:integration_riesz_space R.is_norm R V (induced_norm_fun ? V).
215 [ apply mk_is_semi_norm;
216 [ unfold induced_norm_fun;
219 [ apply (irs_positive_linear ? V)
224 unfold induced_norm_fun;
228 unfold induced_norm_fun;
233 unfold induced_norm_fun in H;
234 apply irs_quotient_space1;
237 rewrite < eq_zero_opp_zero;
238 rewrite > zero_neutral;
243 definition induced_norm ≝
244 λR:real.λV:integration_riesz_space R.
245 mk_norm ? ? (induced_norm_fun ? V) (induced_norm_is_norm ? V).
247 lemma is_riesz_norm_induced_norm:
248 ∀R:real.∀V:integration_riesz_space R.
249 is_riesz_norm ? ? (induced_norm ? V).
251 unfold is_riesz_norm;
255 unfold induced_norm_fun;
260 definition induced_riesz_norm ≝
261 λR:real.λV:integration_riesz_space R.
262 mk_riesz_norm ? ? (induced_norm ? V) (is_riesz_norm_induced_norm ? V).
264 definition distance_induced_by_integral ≝
265 λR:real.λV:integration_riesz_space R.
266 induced_distance ? ? (induced_norm R V).
268 definition is_complete_integration_riesz_space ≝
269 λR:real.λV:integration_riesz_space R.
270 is_complete ? ? (distance_induced_by_integral ? V).
272 record complete_integration_riesz_space (R:real) : Type ≝
273 { cirz_integration_riesz_space:> integration_riesz_space R;
274 cirz_complete_integration_riesz_space_property:
275 is_complete_integration_riesz_space ? cirz_integration_riesz_space
278 (* now we prove that any complete integration riesz space is an L-space *)
280 (*theorem is_l_space_l_space_induced_by_integral:
281 ∀R:real.∀V:complete_integration_riesz_space R.
282 is_l_space ? ? (induced_riesz_norm ? V).
285 [ apply cirz_complete_integration_riesz_space_property
287 unfold induced_riesz_norm;
291 unfold induced_norm_fun;
297 (**************************** f-ALGEBRAS ********************************)
299 record is_algebra (K: field) (V:vector_space K) (mult:V→V→V) (one:V) : Prop
301 { (* ring properties *)
302 a_ring: is_ring V mult one;
303 (* algebra properties *)
304 a_associative_left: ∀a,f,g. a * (mult f g) = mult (a * f) g;
305 a_associative_right: ∀a,f,g. a * (mult f g) = mult f (a * g)
308 record algebra (K: field) : Type \def
309 { a_vector_space:> vector_space K;
310 a_one: a_vector_space;
311 a_mult: a_vector_space → a_vector_space → a_vector_space;
312 a_algebra_properties: is_algebra ? ? a_mult a_one
315 interpretation "Algebra product" 'times a b =
316 (cic:/matita/integration_algebras/a_mult.con _ a b).
318 definition ring_of_algebra ≝
320 mk_ring A (a_mult ? A) (a_one ? A)
321 (a_ring ? ? ? ? (a_algebra_properties ? A)).
323 coercion cic:/matita/integration_algebras/ring_of_algebra.con.
325 record pre_f_algebra (K:ordered_field_ch0) : Type ≝
326 { fa_archimedean_riesz_space:> archimedean_riesz_space K;
327 fa_algebra_: algebra K;
328 fa_with: a_vector_space ? fa_algebra_ = rs_vector_space ? fa_archimedean_riesz_space
331 lemma fa_algebra: ∀K:ordered_field_ch0.pre_f_algebra K → algebra K.
334 [ apply (rs_vector_space ? A)
341 coercion cic:/matita/integration_algebras/fa_algebra.con.
343 record is_f_algebra (K) (A:pre_f_algebra K) : Prop ≝
344 { compat_mult_le: ∀f,g:A.0 ≤ f → 0 ≤ g → 0 ≤ f*g;
346 ∀f,g,h:A.(f ∧ g) = 0 → ((h*f) ∧ g) = 0
349 record f_algebra (K:ordered_field_ch0) : Type ≝
350 { fa_pre_f_algebra:> pre_f_algebra K;
351 fa_f_algebra_properties: is_f_algebra ? fa_pre_f_algebra
354 (* to be proved; see footnote 2 in the paper by Spitters *)
355 axiom symmetric_a_mult:
356 ∀K.∀A:f_algebra K. symmetric ? (a_mult ? A).
358 record integration_f_algebra (R:real) : Type \def
359 { ifa_integration_riesz_space:> integration_riesz_space R;
360 ifa_f_algebra_: f_algebra R;
362 fa_archimedean_riesz_space ? ifa_f_algebra_ =
363 irs_archimedean_riesz_space ? ifa_integration_riesz_space
366 axiom ifa_f_algebra: ∀R:real.integration_f_algebra R → f_algebra R.
368 coercion cic:/matita/integration_algebras/ifa_f_algebra.con.