author | paulson <lp15@cam.ac.uk> |
Tue, 15 Mar 2016 14:08:25 +0000 | |
changeset 62623 | dbc62f86a1a9 |
parent 62378 | 85ed00c1fe7c |
child 62648 | ee48e0b4f669 |
permissions | -rw-r--r-- |
62375 | 1 |
(* Title: HOL/Library/Extended_Nonnegative_Real.thy |
2 |
Author: Johannes Hölzl |
|
3 |
*) |
|
4 |
||
5 |
subsection \<open>The type of non-negative extended real numbers\<close> |
|
6 |
||
7 |
theory Extended_Nonnegative_Real |
|
8 |
imports Extended_Real |
|
9 |
begin |
|
10 |
||
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
11 |
context linordered_nonzero_semiring |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
12 |
begin |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
13 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
14 |
lemma of_nat_nonneg [simp]: "0 \<le> of_nat n" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
15 |
by (induct n) simp_all |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
16 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
17 |
lemma of_nat_mono[simp]: "i \<le> j \<Longrightarrow> of_nat i \<le> of_nat j" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
18 |
by (auto simp add: le_iff_add intro!: add_increasing2) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
19 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
20 |
end |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
21 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
22 |
lemma of_nat_less[simp]: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
23 |
"i < j \<Longrightarrow> of_nat i < (of_nat j::'a::{linordered_nonzero_semiring, semiring_char_0})" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
24 |
by (auto simp: less_le) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
25 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
26 |
lemma of_nat_le_iff[simp]: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
27 |
"of_nat i \<le> (of_nat j::'a::{linordered_nonzero_semiring, semiring_char_0}) \<longleftrightarrow> i \<le> j" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
28 |
proof (safe intro!: of_nat_mono) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
29 |
assume "of_nat i \<le> (of_nat j::'a)" then show "i \<le> j" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
30 |
proof (intro leI notI) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
31 |
assume "j < i" from less_le_trans[OF of_nat_less[OF this] \<open>of_nat i \<le> of_nat j\<close>] show False |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
32 |
by blast |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
33 |
qed |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
34 |
qed |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
35 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
36 |
lemma (in complete_lattice) SUP_sup_const1: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
37 |
"I \<noteq> {} \<Longrightarrow> (SUP i:I. sup c (f i)) = sup c (SUP i:I. f i)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
38 |
using SUP_sup_distrib[of "\<lambda>_. c" I f] by simp |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
39 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
40 |
lemma (in complete_lattice) SUP_sup_const2: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
41 |
"I \<noteq> {} \<Longrightarrow> (SUP i:I. sup (f i) c) = sup (SUP i:I. f i) c" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
42 |
using SUP_sup_distrib[of f I "\<lambda>_. c"] by simp |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
43 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
44 |
lemma one_less_of_natD: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
45 |
"(1::'a::linordered_semidom) < of_nat n \<Longrightarrow> 1 < n" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
46 |
using zero_le_one[where 'a='a] |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
47 |
apply (cases n) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
48 |
apply simp |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
49 |
subgoal for n' |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
50 |
apply (cases n') |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
51 |
apply simp |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
52 |
apply simp |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
53 |
done |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
54 |
done |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
55 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
56 |
lemma setsum_le_suminf: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
57 |
fixes f :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
58 |
shows "summable f \<Longrightarrow> finite I \<Longrightarrow> \<forall>m\<in>- I. 0 \<le> f m \<Longrightarrow> setsum f I \<le> suminf f" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
59 |
by (rule sums_le[OF _ sums_If_finite_set summable_sums]) auto |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
60 |
|
62375 | 61 |
typedef ennreal = "{x :: ereal. 0 \<le> x}" |
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
62 |
morphisms enn2ereal e2ennreal' |
62375 | 63 |
by auto |
64 |
||
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
65 |
definition "e2ennreal x = e2ennreal' (max 0 x)" |
62375 | 66 |
|
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
67 |
lemma type_definition_ennreal': "type_definition enn2ereal e2ennreal {x. 0 \<le> x}" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
68 |
using type_definition_ennreal |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
69 |
by (auto simp: type_definition_def e2ennreal_def max_absorb2) |
62375 | 70 |
|
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
71 |
setup_lifting type_definition_ennreal' |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
72 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
73 |
lift_definition ennreal :: "real \<Rightarrow> ennreal" is "sup 0 \<circ> ereal" |
62375 | 74 |
by simp |
75 |
||
76 |
declare [[coercion ennreal]] |
|
77 |
declare [[coercion e2ennreal]] |
|
78 |
||
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
79 |
instantiation ennreal :: complete_linorder |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
80 |
begin |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
81 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
82 |
lift_definition top_ennreal :: ennreal is top by (rule top_greatest) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
83 |
lift_definition bot_ennreal :: ennreal is 0 by (rule order_refl) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
84 |
lift_definition sup_ennreal :: "ennreal \<Rightarrow> ennreal \<Rightarrow> ennreal" is sup by (rule le_supI1) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
85 |
lift_definition inf_ennreal :: "ennreal \<Rightarrow> ennreal \<Rightarrow> ennreal" is inf by (rule le_infI) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
86 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
87 |
lift_definition Inf_ennreal :: "ennreal set \<Rightarrow> ennreal" is "Inf" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
88 |
by (rule Inf_greatest) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
89 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
90 |
lift_definition Sup_ennreal :: "ennreal set \<Rightarrow> ennreal" is "sup 0 \<circ> Sup" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
91 |
by auto |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
92 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
93 |
lift_definition less_eq_ennreal :: "ennreal \<Rightarrow> ennreal \<Rightarrow> bool" is "op \<le>" . |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
94 |
lift_definition less_ennreal :: "ennreal \<Rightarrow> ennreal \<Rightarrow> bool" is "op <" . |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
95 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
96 |
instance |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
97 |
by standard |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
98 |
(transfer ; auto simp: Inf_lower Inf_greatest Sup_upper Sup_least le_max_iff_disj max.absorb1)+ |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
99 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
100 |
end |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
101 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
102 |
lemma ennreal_cases: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
103 |
fixes x :: ennreal |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
104 |
obtains (real) r :: real where "0 \<le> r" "x = ennreal r" | (top) "x = top" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
105 |
apply transfer |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
106 |
subgoal for x thesis |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
107 |
by (cases x) (auto simp: max.absorb2 top_ereal_def) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
108 |
done |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
109 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
110 |
instantiation ennreal :: infinity |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
111 |
begin |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
112 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
113 |
definition infinity_ennreal :: ennreal |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
114 |
where |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
115 |
[simp]: "\<infinity> = (top::ennreal)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
116 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
117 |
instance .. |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
118 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
119 |
end |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
120 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
121 |
instantiation ennreal :: "{semiring_1_no_zero_divisors, comm_semiring_1}" |
62375 | 122 |
begin |
123 |
||
124 |
lift_definition one_ennreal :: ennreal is 1 by simp |
|
125 |
lift_definition zero_ennreal :: ennreal is 0 by simp |
|
126 |
lift_definition plus_ennreal :: "ennreal \<Rightarrow> ennreal \<Rightarrow> ennreal" is "op +" by simp |
|
127 |
lift_definition times_ennreal :: "ennreal \<Rightarrow> ennreal \<Rightarrow> ennreal" is "op *" by simp |
|
128 |
||
129 |
instance |
|
130 |
by standard (transfer; auto simp: field_simps ereal_right_distrib)+ |
|
131 |
||
132 |
end |
|
133 |
||
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
134 |
instantiation ennreal :: minus |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
135 |
begin |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
136 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
137 |
lift_definition minus_ennreal :: "ennreal \<Rightarrow> ennreal \<Rightarrow> ennreal" is "\<lambda>a b. max 0 (a - b)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
138 |
by simp |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
139 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
140 |
instance .. |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
141 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
142 |
end |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
143 |
|
62375 | 144 |
instance ennreal :: numeral .. |
145 |
||
146 |
instantiation ennreal :: inverse |
|
147 |
begin |
|
148 |
||
149 |
lift_definition inverse_ennreal :: "ennreal \<Rightarrow> ennreal" is inverse |
|
150 |
by (rule inverse_ereal_ge0I) |
|
151 |
||
152 |
definition divide_ennreal :: "ennreal \<Rightarrow> ennreal \<Rightarrow> ennreal" |
|
153 |
where "x div y = x * inverse (y :: ennreal)" |
|
154 |
||
155 |
instance .. |
|
156 |
||
157 |
end |
|
158 |
||
159 |
lemma ennreal_zero_less_one: "0 < (1::ennreal)" |
|
160 |
by transfer auto |
|
161 |
||
62376
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents:
62375
diff
changeset
|
162 |
instance ennreal :: dioid |
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents:
62375
diff
changeset
|
163 |
proof (standard; transfer) |
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents:
62375
diff
changeset
|
164 |
fix a b :: ereal assume "0 \<le> a" "0 \<le> b" then show "(a \<le> b) = (\<exists>c\<in>Collect (op \<le> 0). b = a + c)" |
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents:
62375
diff
changeset
|
165 |
unfolding ereal_ex_split Bex_def |
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents:
62375
diff
changeset
|
166 |
by (cases a b rule: ereal2_cases) (auto intro!: exI[of _ "real_of_ereal (b - a)"]) |
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents:
62375
diff
changeset
|
167 |
qed |
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents:
62375
diff
changeset
|
168 |
|
62375 | 169 |
instance ennreal :: ordered_comm_semiring |
170 |
by standard |
|
171 |
(transfer ; auto intro: add_mono mult_mono mult_ac ereal_left_distrib ereal_mult_left_mono)+ |
|
172 |
||
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
173 |
instance ennreal :: linordered_nonzero_semiring |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
174 |
proof qed (transfer; simp) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
175 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
176 |
declare [[coercion "of_nat :: nat \<Rightarrow> ennreal"]] |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
177 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
178 |
lemma e2ennreal_neg: "x \<le> 0 \<Longrightarrow> e2ennreal x = 0" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
179 |
unfolding zero_ennreal_def e2ennreal_def by (simp add: max_absorb1) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
180 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
181 |
lemma e2ennreal_mono: "x \<le> y \<Longrightarrow> e2ennreal x \<le> e2ennreal y" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
182 |
by (cases "0 \<le> x" "0 \<le> y" rule: bool.exhaust[case_product bool.exhaust]) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
183 |
(auto simp: e2ennreal_neg less_eq_ennreal.abs_eq eq_onp_def) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
184 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
185 |
subsection \<open>Cancellation simprocs\<close> |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
186 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
187 |
lemma ennreal_add_left_cancel: "a + b = a + c \<longleftrightarrow> a = (\<infinity>::ennreal) \<or> b = c" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
188 |
unfolding infinity_ennreal_def by transfer (simp add: top_ereal_def ereal_add_cancel_left) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
189 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
190 |
lemma ennreal_add_left_cancel_le: "a + b \<le> a + c \<longleftrightarrow> a = (\<infinity>::ennreal) \<or> b \<le> c" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
191 |
unfolding infinity_ennreal_def by transfer (simp add: ereal_add_le_add_iff top_ereal_def disj_commute) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
192 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
193 |
lemma ereal_add_left_cancel_less: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
194 |
fixes a b c :: ereal |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
195 |
shows "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> a + b < a + c \<longleftrightarrow> a \<noteq> \<infinity> \<and> b < c" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
196 |
by (cases a b c rule: ereal3_cases) auto |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
197 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
198 |
lemma ennreal_add_left_cancel_less: "a + b < a + c \<longleftrightarrow> a \<noteq> (\<infinity>::ennreal) \<and> b < c" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
199 |
unfolding infinity_ennreal_def |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
200 |
by transfer (simp add: top_ereal_def ereal_add_left_cancel_less) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
201 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
202 |
ML \<open> |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
203 |
structure Cancel_Ennreal_Common = |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
204 |
struct |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
205 |
(* copied from src/HOL/Tools/nat_numeral_simprocs.ML *) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
206 |
fun find_first_t _ _ [] = raise TERM("find_first_t", []) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
207 |
| find_first_t past u (t::terms) = |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
208 |
if u aconv t then (rev past @ terms) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
209 |
else find_first_t (t::past) u terms |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
210 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
211 |
fun dest_summing (Const (@{const_name Groups.plus}, _) $ t $ u, ts) = |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
212 |
dest_summing (t, dest_summing (u, ts)) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
213 |
| dest_summing (t, ts) = t :: ts |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
214 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
215 |
val mk_sum = Arith_Data.long_mk_sum |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
216 |
fun dest_sum t = dest_summing (t, []) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
217 |
val find_first = find_first_t [] |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
218 |
val trans_tac = Numeral_Simprocs.trans_tac |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
219 |
val norm_ss = |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
220 |
simpset_of (put_simpset HOL_basic_ss @{context} |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
221 |
addsimps @{thms ac_simps add_0_left add_0_right}) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
222 |
fun norm_tac ctxt = ALLGOALS (simp_tac (put_simpset norm_ss ctxt)) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
223 |
fun simplify_meta_eq ctxt cancel_th th = |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
224 |
Arith_Data.simplify_meta_eq [] ctxt |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
225 |
([th, cancel_th] MRS trans) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
226 |
fun mk_eq (a, b) = HOLogic.mk_Trueprop (HOLogic.mk_eq (a, b)) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
227 |
end |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
228 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
229 |
structure Eq_Ennreal_Cancel = ExtractCommonTermFun |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
230 |
(open Cancel_Ennreal_Common |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
231 |
val mk_bal = HOLogic.mk_eq |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
232 |
val dest_bal = HOLogic.dest_bin @{const_name HOL.eq} @{typ ennreal} |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
233 |
fun simp_conv _ _ = SOME @{thm ennreal_add_left_cancel} |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
234 |
) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
235 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
236 |
structure Le_Ennreal_Cancel = ExtractCommonTermFun |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
237 |
(open Cancel_Ennreal_Common |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
238 |
val mk_bal = HOLogic.mk_binrel @{const_name Orderings.less_eq} |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
239 |
val dest_bal = HOLogic.dest_bin @{const_name Orderings.less_eq} @{typ ennreal} |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
240 |
fun simp_conv _ _ = SOME @{thm ennreal_add_left_cancel_le} |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
241 |
) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
242 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
243 |
structure Less_Ennreal_Cancel = ExtractCommonTermFun |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
244 |
(open Cancel_Ennreal_Common |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
245 |
val mk_bal = HOLogic.mk_binrel @{const_name Orderings.less} |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
246 |
val dest_bal = HOLogic.dest_bin @{const_name Orderings.less} @{typ ennreal} |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
247 |
fun simp_conv _ _ = SOME @{thm ennreal_add_left_cancel_less} |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
248 |
) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
249 |
\<close> |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
250 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
251 |
simproc_setup ennreal_eq_cancel |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
252 |
("(l::ennreal) + m = n" | "(l::ennreal) = m + n") = |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
253 |
\<open>fn phi => fn ctxt => fn ct => Eq_Ennreal_Cancel.proc ctxt (Thm.term_of ct)\<close> |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
254 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
255 |
simproc_setup ennreal_le_cancel |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
256 |
("(l::ennreal) + m \<le> n" | "(l::ennreal) \<le> m + n") = |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
257 |
\<open>fn phi => fn ctxt => fn ct => Le_Ennreal_Cancel.proc ctxt (Thm.term_of ct)\<close> |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
258 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
259 |
simproc_setup ennreal_less_cancel |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
260 |
("(l::ennreal) + m < n" | "(l::ennreal) < m + n") = |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
261 |
\<open>fn phi => fn ctxt => fn ct => Less_Ennreal_Cancel.proc ctxt (Thm.term_of ct)\<close> |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
262 |
|
62375 | 263 |
instantiation ennreal :: linear_continuum_topology |
264 |
begin |
|
265 |
||
266 |
definition open_ennreal :: "ennreal set \<Rightarrow> bool" |
|
267 |
where "(open :: ennreal set \<Rightarrow> bool) = generate_topology (range lessThan \<union> range greaterThan)" |
|
268 |
||
269 |
instance |
|
270 |
proof |
|
271 |
show "\<exists>a b::ennreal. a \<noteq> b" |
|
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
272 |
using zero_neq_one by (intro exI) |
62375 | 273 |
show "\<And>x y::ennreal. x < y \<Longrightarrow> \<exists>z>x. z < y" |
274 |
proof transfer |
|
275 |
fix x y :: ereal assume "0 \<le> x" "x < y" |
|
276 |
moreover from dense[OF this(2)] guess z .. |
|
277 |
ultimately show "\<exists>z\<in>Collect (op \<le> 0). x < z \<and> z < y" |
|
278 |
by (intro bexI[of _ z]) auto |
|
279 |
qed |
|
280 |
qed (rule open_ennreal_def) |
|
281 |
||
282 |
end |
|
283 |
||
284 |
lemma ennreal_rat_dense: |
|
285 |
fixes x y :: ennreal |
|
286 |
shows "x < y \<Longrightarrow> \<exists>r::rat. x < real_of_rat r \<and> real_of_rat r < y" |
|
287 |
proof transfer |
|
288 |
fix x y :: ereal assume xy: "0 \<le> x" "0 \<le> y" "x < y" |
|
289 |
moreover |
|
290 |
from ereal_dense3[OF \<open>x < y\<close>] |
|
291 |
obtain r where "x < ereal (real_of_rat r)" "ereal (real_of_rat r) < y" |
|
292 |
by auto |
|
293 |
moreover then have "0 \<le> r" |
|
294 |
using le_less_trans[OF \<open>0 \<le> x\<close> \<open>x < ereal (real_of_rat r)\<close>] by auto |
|
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
295 |
ultimately show "\<exists>r. x < (sup 0 \<circ> ereal) (real_of_rat r) \<and> (sup 0 \<circ> ereal) (real_of_rat r) < y" |
62375 | 296 |
by (intro exI[of _ r]) (auto simp: max_absorb2) |
297 |
qed |
|
298 |
||
299 |
lemma enn2ereal_range: "e2ennreal ` {0..} = UNIV" |
|
300 |
proof - |
|
301 |
have "\<exists>y\<ge>0. x = e2ennreal y" for x |
|
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
302 |
by (cases x) (auto simp: e2ennreal_def max_absorb2) |
62375 | 303 |
then show ?thesis |
304 |
by (auto simp: image_iff Bex_def) |
|
305 |
qed |
|
306 |
||
307 |
lemma enn2ereal_nonneg: "0 \<le> enn2ereal x" |
|
308 |
using ennreal.enn2ereal[of x] by simp |
|
309 |
||
310 |
lemma ereal_ennreal_cases: |
|
311 |
obtains b where "0 \<le> a" "a = enn2ereal b" | "a < 0" |
|
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
312 |
using e2ennreal'_inverse[of a, symmetric] by (cases "0 \<le> a") (auto intro: enn2ereal_nonneg) |
62375 | 313 |
|
314 |
lemma enn2ereal_Iio: "enn2ereal -` {..<a} = (if 0 \<le> a then {..< e2ennreal a} else {})" |
|
315 |
using enn2ereal_nonneg |
|
316 |
by (cases a rule: ereal_ennreal_cases) |
|
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
317 |
(auto simp add: vimage_def set_eq_iff ennreal.enn2ereal_inverse less_ennreal.rep_eq e2ennreal_def max_absorb2 |
62375 | 318 |
intro: le_less_trans less_imp_le) |
319 |
||
320 |
lemma enn2ereal_Ioi: "enn2ereal -` {a <..} = (if 0 \<le> a then {e2ennreal a <..} else UNIV)" |
|
321 |
using enn2ereal_nonneg |
|
322 |
by (cases a rule: ereal_ennreal_cases) |
|
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
323 |
(auto simp add: vimage_def set_eq_iff ennreal.enn2ereal_inverse less_ennreal.rep_eq e2ennreal_def max_absorb2 |
62375 | 324 |
intro: less_le_trans) |
325 |
||
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
326 |
lemma continuous_on_e2ennreal: "continuous_on A e2ennreal" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
327 |
proof (rule continuous_on_subset) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
328 |
show "continuous_on ({0..} \<union> {..0}) e2ennreal" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
329 |
proof (rule continuous_on_closed_Un) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
330 |
show "continuous_on {0 ..} e2ennreal" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
331 |
by (rule continuous_onI_mono) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
332 |
(auto simp add: less_eq_ennreal.abs_eq eq_onp_def enn2ereal_range) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
333 |
show "continuous_on {.. 0} e2ennreal" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
334 |
by (subst continuous_on_cong[OF refl, of _ _ "\<lambda>_. 0"]) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
335 |
(auto simp add: e2ennreal_neg continuous_on_const) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
336 |
qed auto |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
337 |
show "A \<subseteq> {0..} \<union> {..0::ereal}" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
338 |
by auto |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
339 |
qed |
62375 | 340 |
|
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
341 |
lemma continuous_at_e2ennreal: "continuous (at x within A) e2ennreal" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
342 |
by (rule continuous_on_imp_continuous_within[OF continuous_on_e2ennreal, of _ UNIV]) auto |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
343 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
344 |
lemma continuous_on_enn2ereal: "continuous_on UNIV enn2ereal" |
62375 | 345 |
by (rule continuous_on_generate_topology[OF open_generated_order]) |
346 |
(auto simp add: enn2ereal_Iio enn2ereal_Ioi) |
|
347 |
||
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
348 |
lemma continuous_at_enn2ereal: "continuous (at x within A) enn2ereal" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
349 |
by (rule continuous_on_imp_continuous_within[OF continuous_on_enn2ereal]) auto |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
350 |
|
62375 | 351 |
lemma transfer_enn2ereal_continuous_on [transfer_rule]: |
352 |
"rel_fun (op =) (rel_fun (rel_fun op = pcr_ennreal) op =) continuous_on continuous_on" |
|
353 |
proof - |
|
354 |
have "continuous_on A f" if "continuous_on A (\<lambda>x. enn2ereal (f x))" for A and f :: "'a \<Rightarrow> ennreal" |
|
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
355 |
using continuous_on_compose2[OF continuous_on_e2ennreal[of "{0..}"] that] |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
356 |
by (auto simp: ennreal.enn2ereal_inverse subset_eq enn2ereal_nonneg e2ennreal_def max_absorb2) |
62375 | 357 |
moreover |
358 |
have "continuous_on A (\<lambda>x. enn2ereal (f x))" if "continuous_on A f" for A and f :: "'a \<Rightarrow> ennreal" |
|
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
359 |
using continuous_on_compose2[OF continuous_on_enn2ereal that] by auto |
62375 | 360 |
ultimately |
361 |
show ?thesis |
|
362 |
by (auto simp add: rel_fun_def ennreal.pcr_cr_eq cr_ennreal_def) |
|
363 |
qed |
|
364 |
||
365 |
lemma continuous_on_add_ennreal[continuous_intros]: |
|
366 |
fixes f g :: "'a::topological_space \<Rightarrow> ennreal" |
|
367 |
shows "continuous_on A f \<Longrightarrow> continuous_on A g \<Longrightarrow> continuous_on A (\<lambda>x. f x + g x)" |
|
368 |
by (transfer fixing: A) (auto intro!: tendsto_add_ereal_nonneg simp: continuous_on_def) |
|
369 |
||
370 |
lemma continuous_on_inverse_ennreal[continuous_intros]: |
|
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
371 |
fixes f :: "'a::topological_space \<Rightarrow> ennreal" |
62375 | 372 |
shows "continuous_on A f \<Longrightarrow> continuous_on A (\<lambda>x. inverse (f x))" |
373 |
proof (transfer fixing: A) |
|
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
374 |
show "pred_fun (\<lambda>_. True) (op \<le> 0) f \<Longrightarrow> continuous_on A (\<lambda>x. inverse (f x))" if "continuous_on A f" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
375 |
for f :: "'a \<Rightarrow> ereal" |
62375 | 376 |
using continuous_on_compose2[OF continuous_on_inverse_ereal that] by (auto simp: subset_eq) |
377 |
qed |
|
378 |
||
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
379 |
instance ennreal :: topological_comm_monoid_add |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
380 |
proof |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
381 |
show "((\<lambda>x. fst x + snd x) \<longlongrightarrow> a + b) (nhds a \<times>\<^sub>F nhds b)" for a b :: ennreal |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
382 |
using continuous_on_add_ennreal[of UNIV fst snd] |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
383 |
using tendsto_at_iff_tendsto_nhds[symmetric, of "\<lambda>x::(ennreal \<times> ennreal). fst x + snd x"] |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
384 |
by (auto simp: continuous_on_eq_continuous_at) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
385 |
(simp add: isCont_def nhds_prod[symmetric]) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
386 |
qed |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
387 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
388 |
lemma ennreal_zero_less_top[simp]: "0 < (top::ennreal)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
389 |
by transfer (simp add: top_ereal_def) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
390 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
391 |
lemma ennreal_one_less_top[simp]: "1 < (top::ennreal)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
392 |
by transfer (simp add: top_ereal_def) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
393 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
394 |
lemma ennreal_zero_neq_top[simp]: "0 \<noteq> (top::ennreal)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
395 |
by transfer (simp add: top_ereal_def) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
396 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
397 |
lemma ennreal_top_neq_zero[simp]: "(top::ennreal) \<noteq> 0" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
398 |
by transfer (simp add: top_ereal_def) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
399 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
400 |
lemma ennreal_top_neq_one[simp]: "top \<noteq> (1::ennreal)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
401 |
by transfer (simp add: top_ereal_def one_ereal_def ereal_max[symmetric] del: ereal_max) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
402 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
403 |
lemma ennreal_one_neq_top[simp]: "1 \<noteq> (top::ennreal)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
404 |
by transfer (simp add: top_ereal_def one_ereal_def ereal_max[symmetric] del: ereal_max) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
405 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
406 |
lemma ennreal_add_less_top[simp]: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
407 |
fixes a b :: ennreal |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
408 |
shows "a + b < top \<longleftrightarrow> a < top \<and> b < top" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
409 |
by transfer (auto simp: top_ereal_def) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
410 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
411 |
lemma ennreal_add_eq_top[simp]: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
412 |
fixes a b :: ennreal |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
413 |
shows "a + b = top \<longleftrightarrow> a = top \<or> b = top" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
414 |
by transfer (auto simp: top_ereal_def) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
415 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
416 |
lemma ennreal_setsum_less_top[simp]: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
417 |
fixes f :: "'a \<Rightarrow> ennreal" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
418 |
shows "finite I \<Longrightarrow> (\<Sum>i\<in>I. f i) < top \<longleftrightarrow> (\<forall>i\<in>I. f i < top)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
419 |
by (induction I rule: finite_induct) auto |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
420 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
421 |
lemma ennreal_setsum_eq_top[simp]: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
422 |
fixes f :: "'a \<Rightarrow> ennreal" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
423 |
shows "finite I \<Longrightarrow> (\<Sum>i\<in>I. f i) = top \<longleftrightarrow> (\<exists>i\<in>I. f i = top)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
424 |
by (induction I rule: finite_induct) auto |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
425 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
426 |
lemma enn2ereal_eq_top_iff[simp]: "enn2ereal x = \<infinity> \<longleftrightarrow> x = top" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
427 |
by transfer (simp add: top_ereal_def) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
428 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
429 |
lemma ennreal_top_top: "top - top = (top::ennreal)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
430 |
by transfer (auto simp: top_ereal_def max_def) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
431 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
432 |
lemma ennreal_minus_zero[simp]: "a - (0::ennreal) = a" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
433 |
by transfer (auto simp: max_def) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
434 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
435 |
lemma ennreal_add_diff_cancel_right[simp]: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
436 |
fixes x y z :: ennreal shows "y \<noteq> top \<Longrightarrow> (x + y) - y = x" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
437 |
apply transfer |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
438 |
subgoal for x y |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
439 |
apply (cases x y rule: ereal2_cases) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
440 |
apply (auto split: split_max simp: top_ereal_def) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
441 |
done |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
442 |
done |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
443 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
444 |
lemma ennreal_add_diff_cancel_left[simp]: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
445 |
fixes x y z :: ennreal shows "y \<noteq> top \<Longrightarrow> (y + x) - y = x" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
446 |
by (simp add: add.commute) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
447 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
448 |
lemma |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
449 |
fixes a b :: ennreal |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
450 |
shows "a - b = 0 \<Longrightarrow> a \<le> b" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
451 |
apply transfer |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
452 |
subgoal for a b |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
453 |
apply (cases a b rule: ereal2_cases) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
454 |
apply (auto simp: not_le max_def split: if_splits) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
455 |
done |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
456 |
done |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
457 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
458 |
lemma ennreal_minus_cancel: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
459 |
fixes a b c :: ennreal |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
460 |
shows "c \<noteq> top \<Longrightarrow> a \<le> c \<Longrightarrow> b \<le> c \<Longrightarrow> c - a = c - b \<Longrightarrow> a = b" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
461 |
apply transfer |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
462 |
subgoal for a b c |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
463 |
by (cases a b c rule: ereal3_cases) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
464 |
(auto simp: top_ereal_def max_def split: if_splits) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
465 |
done |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
466 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
467 |
lemma enn2ereal_ennreal[simp]: "0 \<le> x \<Longrightarrow> enn2ereal (ennreal x) = x" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
468 |
by transfer (simp add: max_absorb2) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
469 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
470 |
lemma tendsto_ennrealD: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
471 |
assumes lim: "((\<lambda>x. ennreal (f x)) \<longlongrightarrow> ennreal x) F" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
472 |
assumes *: "\<forall>\<^sub>F x in F. 0 \<le> f x" and x: "0 \<le> x" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
473 |
shows "(f \<longlongrightarrow> x) F" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
474 |
using continuous_on_tendsto_compose[OF continuous_on_enn2ereal lim] |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
475 |
apply simp |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
476 |
apply (subst (asm) tendsto_cong) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
477 |
using * |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
478 |
apply eventually_elim |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
479 |
apply (auto simp: max_absorb2 \<open>0 \<le> x\<close>) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
480 |
done |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
481 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
482 |
lemma tendsto_ennreal_iff[simp]: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
483 |
"\<forall>\<^sub>F x in F. 0 \<le> f x \<Longrightarrow> 0 \<le> x \<Longrightarrow> ((\<lambda>x. ennreal (f x)) \<longlongrightarrow> ennreal x) F \<longleftrightarrow> (f \<longlongrightarrow> x) F" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
484 |
by (auto dest: tendsto_ennrealD) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
485 |
(auto simp: ennreal_def |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
486 |
intro!: continuous_on_tendsto_compose[OF continuous_on_e2ennreal[of UNIV]] tendsto_max) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
487 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
488 |
lemma ennreal_zero[simp]: "ennreal 0 = 0" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
489 |
by (simp add: ennreal_def max.absorb1 zero_ennreal.abs_eq) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
490 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
491 |
lemma ennreal_plus[simp]: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
492 |
"0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> ennreal (a + b) = ennreal a + ennreal b" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
493 |
by (transfer fixing: a b) (auto simp: max_absorb2) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
494 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
495 |
lemma ennreal_inj[simp]: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
496 |
"0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> ennreal a = ennreal b \<longleftrightarrow> a = b" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
497 |
by (transfer fixing: a b) (auto simp: max_absorb2) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
498 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
499 |
lemma ennreal_le_iff[simp]: "0 \<le> y \<Longrightarrow> ennreal x \<le> ennreal y \<longleftrightarrow> x \<le> y" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
500 |
by (auto simp: ennreal_def zero_ereal_def less_eq_ennreal.abs_eq eq_onp_def split: split_max) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
501 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
502 |
lemma setsum_ennreal[simp]: "(\<And>i. i \<in> I \<Longrightarrow> 0 \<le> f i) \<Longrightarrow> (\<Sum>i\<in>I. ennreal (f i)) = ennreal (setsum f I)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
503 |
by (induction I rule: infinite_finite_induct) (auto simp: setsum_nonneg) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
504 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
505 |
lemma sums_ennreal[simp]: "(\<And>i. 0 \<le> f i) \<Longrightarrow> 0 \<le> x \<Longrightarrow> (\<lambda>i. ennreal (f i)) sums ennreal x \<longleftrightarrow> f sums x" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
506 |
unfolding sums_def by (simp add: always_eventually setsum_nonneg) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
507 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
508 |
lemma summable_suminf_not_top: "(\<And>i. 0 \<le> f i) \<Longrightarrow> (\<Sum>i. ennreal (f i)) \<noteq> top \<Longrightarrow> summable f" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
509 |
using summable_sums[OF summableI, of "\<lambda>i. ennreal (f i)"] |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
510 |
by (cases "\<Sum>i. ennreal (f i)" rule: ennreal_cases) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
511 |
(auto simp: summable_def) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
512 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
513 |
lemma suminf_ennreal[simp]: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
514 |
"(\<And>i. 0 \<le> f i) \<Longrightarrow> (\<Sum>i. ennreal (f i)) \<noteq> top \<Longrightarrow> (\<Sum>i. ennreal (f i)) = ennreal (\<Sum>i. f i)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
515 |
by (rule sums_unique[symmetric]) (simp add: summable_suminf_not_top suminf_nonneg summable_sums) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
516 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
517 |
lemma suminf_less_top: "(\<Sum>i. f i :: ennreal) < top \<Longrightarrow> f i < top" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
518 |
using le_less_trans[OF setsum_le_suminf[OF summableI, of "{i}" f]] by simp |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
519 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
520 |
lemma add_top: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
521 |
fixes x :: "'a::{order_top, ordered_comm_monoid_add}" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
522 |
shows "0 \<le> x \<Longrightarrow> x + top = top" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
523 |
by (intro top_le add_increasing order_refl) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
524 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
525 |
lemma top_add: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
526 |
fixes x :: "'a::{order_top, ordered_comm_monoid_add}" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
527 |
shows "0 \<le> x \<Longrightarrow> top + x = top" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
528 |
by (intro top_le add_increasing2 order_refl) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
529 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
530 |
lemma enn2ereal_top: "enn2ereal top = \<infinity>" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
531 |
by transfer (simp add: top_ereal_def) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
532 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
533 |
lemma e2ennreal_infty: "e2ennreal \<infinity> = top" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
534 |
by (simp add: top_ennreal.abs_eq top_ereal_def) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
535 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
536 |
lemma sup_const_add_ennreal: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
537 |
fixes a b c :: "ennreal" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
538 |
shows "sup (c + a) (c + b) = c + sup a b" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
539 |
apply transfer |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
540 |
subgoal for a b c |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
541 |
apply (cases a b c rule: ereal3_cases) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
542 |
apply (auto simp: ereal_max[symmetric] simp del: ereal_max) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
543 |
apply (auto simp: top_ereal_def[symmetric] sup_ereal_def[symmetric] |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
544 |
simp del: sup_ereal_def) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
545 |
apply (auto simp add: top_ereal_def) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
546 |
done |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
547 |
done |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
548 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
549 |
lemma bot_ennreal: "bot = (0::ennreal)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
550 |
by transfer rule |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
551 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
552 |
lemma le_lfp: "mono f \<Longrightarrow> x \<le> lfp f \<Longrightarrow> f x \<le> lfp f" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
553 |
by (subst lfp_unfold) (auto dest: monoD) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
554 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
555 |
lemma lfp_transfer: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
556 |
assumes \<alpha>: "sup_continuous \<alpha>" and f: "sup_continuous f" and mg: "mono g" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
557 |
assumes bot: "\<alpha> bot \<le> lfp g" and eq: "\<And>x. x \<le> lfp f \<Longrightarrow> \<alpha> (f x) = g (\<alpha> x)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
558 |
shows "\<alpha> (lfp f) = lfp g" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
559 |
proof (rule antisym) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
560 |
note mf = sup_continuous_mono[OF f] |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
561 |
have f_le_lfp: "(f ^^ i) bot \<le> lfp f" for i |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
562 |
by (induction i) (auto intro: le_lfp mf) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
563 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
564 |
have "\<alpha> ((f ^^ i) bot) \<le> lfp g" for i |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
565 |
by (induction i) (auto simp: bot eq f_le_lfp intro!: le_lfp mg) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
566 |
then show "\<alpha> (lfp f) \<le> lfp g" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
567 |
unfolding sup_continuous_lfp[OF f] |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
568 |
by (subst \<alpha>[THEN sup_continuousD]) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
569 |
(auto intro!: mono_funpow sup_continuous_mono[OF f] SUP_least) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
570 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
571 |
show "lfp g \<le> \<alpha> (lfp f)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
572 |
by (rule lfp_lowerbound) (simp add: eq[symmetric] lfp_unfold[OF mf, symmetric]) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
573 |
qed |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
574 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
575 |
lemma sup_continuous_applyD: "sup_continuous f \<Longrightarrow> sup_continuous (\<lambda>x. f x h)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
576 |
using sup_continuous_apply[THEN sup_continuous_compose] . |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
577 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
578 |
lemma INF_ennreal_add_const: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
579 |
fixes f g :: "nat \<Rightarrow> ennreal" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
580 |
shows "(INF i. f i + c) = (INF i. f i) + c" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
581 |
using continuous_at_Inf_mono[of "\<lambda>x. x + c" "f`UNIV"] |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
582 |
using continuous_add[of "at_right (Inf (range f))", of "\<lambda>x. x" "\<lambda>x. c"] |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
583 |
by (auto simp: mono_def) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
584 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
585 |
lemma INF_ennreal_const_add: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
586 |
fixes f g :: "nat \<Rightarrow> ennreal" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
587 |
shows "(INF i. c + f i) = c + (INF i. f i)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
588 |
using INF_ennreal_add_const[of f c] by (simp add: ac_simps) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
589 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
590 |
lemma sup_continuous_e2ennreal[order_continuous_intros]: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
591 |
assumes f: "sup_continuous f" shows "sup_continuous (\<lambda>x. e2ennreal (f x))" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
592 |
apply (rule sup_continuous_compose[OF _ f]) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
593 |
apply (rule continuous_at_left_imp_sup_continuous) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
594 |
apply (auto simp: mono_def e2ennreal_mono continuous_at_e2ennreal) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
595 |
done |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
596 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
597 |
lemma sup_continuous_enn2ereal[order_continuous_intros]: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
598 |
assumes f: "sup_continuous f" shows "sup_continuous (\<lambda>x. enn2ereal (f x))" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
599 |
apply (rule sup_continuous_compose[OF _ f]) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
600 |
apply (rule continuous_at_left_imp_sup_continuous) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
601 |
apply (simp_all add: mono_def less_eq_ennreal.rep_eq continuous_at_enn2ereal) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
602 |
done |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
603 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
604 |
lemma ennreal_1[simp]: "ennreal 1 = 1" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
605 |
by transfer (simp add: max_absorb2) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
606 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
607 |
lemma ennreal_of_nat_eq_real_of_nat: "of_nat i = ennreal (of_nat i)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
608 |
by (induction i) simp_all |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
609 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
610 |
lemma ennreal_SUP_of_nat_eq_top: "(SUP x. of_nat x :: ennreal) = top" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
611 |
proof (intro antisym top_greatest le_SUP_iff[THEN iffD2] allI impI) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
612 |
fix y :: ennreal assume "y < top" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
613 |
then obtain r where "y = ennreal r" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
614 |
by (cases y rule: ennreal_cases) auto |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
615 |
then show "\<exists>i\<in>UNIV. y < of_nat i" |
62623
dbc62f86a1a9
rationalisation of theorem names esp about "real Archimedian" etc.
paulson <lp15@cam.ac.uk>
parents:
62378
diff
changeset
|
616 |
using reals_Archimedean2[of "max 1 r"] zero_less_one |
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
617 |
by (auto simp: ennreal_of_nat_eq_real_of_nat ennreal_def less_ennreal.abs_eq eq_onp_def max.absorb2 |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
618 |
dest!: one_less_of_natD intro: less_trans) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
619 |
qed |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
620 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
621 |
lemma ennreal_SUP_eq_top: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
622 |
fixes f :: "'a \<Rightarrow> ennreal" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
623 |
assumes "\<And>n. \<exists>i\<in>I. of_nat n \<le> f i" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
624 |
shows "(SUP i : I. f i) = top" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
625 |
proof - |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
626 |
have "(SUP x. of_nat x :: ennreal) \<le> (SUP i : I. f i)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
627 |
using assms by (auto intro!: SUP_least intro: SUP_upper2) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
628 |
then show ?thesis |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
629 |
by (auto simp: ennreal_SUP_of_nat_eq_top top_unique) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
630 |
qed |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
631 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
632 |
lemma sup_continuous_SUP[order_continuous_intros]: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
633 |
fixes M :: "_ \<Rightarrow> _ \<Rightarrow> 'a::complete_lattice" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
634 |
assumes M: "\<And>i. i \<in> I \<Longrightarrow> sup_continuous (M i)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
635 |
shows "sup_continuous (SUP i:I. M i)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
636 |
unfolding sup_continuous_def by (auto simp add: sup_continuousD[OF M] intro: SUP_commute) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
637 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
638 |
lemma sup_continuous_apply_SUP[order_continuous_intros]: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
639 |
fixes M :: "_ \<Rightarrow> _ \<Rightarrow> 'a::complete_lattice" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
640 |
shows "(\<And>i. i \<in> I \<Longrightarrow> sup_continuous (M i)) \<Longrightarrow> sup_continuous (\<lambda>x. SUP i:I. M i x)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
641 |
unfolding SUP_apply[symmetric] by (rule sup_continuous_SUP) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
642 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
643 |
lemma sup_continuous_lfp'[order_continuous_intros]: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
644 |
assumes 1: "sup_continuous f" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
645 |
assumes 2: "\<And>g. sup_continuous g \<Longrightarrow> sup_continuous (f g)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
646 |
shows "sup_continuous (lfp f)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
647 |
proof - |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
648 |
have "sup_continuous ((f ^^ i) bot)" for i |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
649 |
proof (induction i) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
650 |
case (Suc i) then show ?case |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
651 |
by (auto intro!: 2) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
652 |
qed (simp add: bot_fun_def sup_continuous_const) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
653 |
then show ?thesis |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
654 |
unfolding sup_continuous_lfp[OF 1] by (intro order_continuous_intros) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
655 |
qed |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
656 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
657 |
lemma sup_continuous_lfp''[order_continuous_intros]: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
658 |
assumes 1: "\<And>s. sup_continuous (f s)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
659 |
assumes 2: "\<And>g. sup_continuous g \<Longrightarrow> sup_continuous (\<lambda>s. f s (g s))" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
660 |
shows "sup_continuous (\<lambda>x. lfp (f x))" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
661 |
proof - |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
662 |
have "sup_continuous (\<lambda>x. (f x ^^ i) bot)" for i |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
663 |
proof (induction i) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
664 |
case (Suc i) then show ?case |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
665 |
by (auto intro!: 2) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
666 |
qed (simp add: bot_fun_def sup_continuous_const) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
667 |
then show ?thesis |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
668 |
unfolding sup_continuous_lfp[OF 1] by (intro order_continuous_intros) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
669 |
qed |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
670 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
671 |
lemma ennreal_INF_const_minus: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
672 |
fixes f :: "'a \<Rightarrow> ennreal" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
673 |
shows "I \<noteq> {} \<Longrightarrow> (SUP x:I. c - f x) = c - (INF x:I. f x)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
674 |
by (transfer fixing: I) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
675 |
(simp add: sup_max[symmetric] SUP_sup_const1 SUP_ereal_minus_right del: sup_ereal_def) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
676 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
677 |
lemma ennreal_diff_add_assoc: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
678 |
fixes a b c :: ennreal |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
679 |
shows "a \<le> b \<Longrightarrow> c + b - a = c + (b - a)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
680 |
apply transfer |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
681 |
subgoal for a b c |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
682 |
apply (cases a b c rule: ereal3_cases) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
683 |
apply (auto simp: field_simps max_absorb2) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
684 |
done |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
685 |
done |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
686 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
687 |
lemma ennreal_top_minus[simp]: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
688 |
fixes c :: ennreal |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
689 |
shows "top - c = top" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
690 |
by transfer (auto simp: top_ereal_def max_absorb2) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
691 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
692 |
lemma le_ennreal_iff: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
693 |
"0 \<le> r \<Longrightarrow> x \<le> ennreal r \<longleftrightarrow> (\<exists>q\<ge>0. x = ennreal q \<and> q \<le> r)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
694 |
apply (transfer fixing: r) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
695 |
subgoal for x |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
696 |
by (cases x) (auto simp: max_absorb2 cong: conj_cong) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
697 |
done |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
698 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
699 |
lemma ennreal_minus: "0 \<le> q \<Longrightarrow> q \<le> r \<Longrightarrow> ennreal r - ennreal q = ennreal (r - q)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
700 |
by transfer (simp add: max_absorb2) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
701 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
702 |
lemma ennreal_tendsto_const_minus: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
703 |
fixes g :: "'a \<Rightarrow> ennreal" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
704 |
assumes ae: "\<forall>\<^sub>F x in F. g x \<le> c" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
705 |
assumes g: "((\<lambda>x. c - g x) \<longlongrightarrow> 0) F" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
706 |
shows "(g \<longlongrightarrow> c) F" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
707 |
proof (cases c rule: ennreal_cases) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
708 |
case top with tendsto_unique[OF _ g, of "top"] show ?thesis |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
709 |
by (cases "F = bot") auto |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
710 |
next |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
711 |
case (real r) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
712 |
then have "\<forall>x. \<exists>q\<ge>0. g x \<le> c \<longrightarrow> (g x = ennreal q \<and> q \<le> r)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
713 |
by (auto simp: le_ennreal_iff) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
714 |
then obtain f where *: "\<And>x. g x \<le> c \<Longrightarrow> 0 \<le> f x" "\<And>x. g x \<le> c \<Longrightarrow> g x = ennreal (f x)" "\<And>x. g x \<le> c \<Longrightarrow> f x \<le> r" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
715 |
by metis |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
716 |
from ae have ae2: "\<forall>\<^sub>F x in F. c - g x = ennreal (r - f x) \<and> f x \<le> r \<and> g x = ennreal (f x) \<and> 0 \<le> f x" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
717 |
proof eventually_elim |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
718 |
fix x assume "g x \<le> c" with *[of x] \<open>0 \<le> r\<close> show "c - g x = ennreal (r - f x) \<and> f x \<le> r \<and> g x = ennreal (f x) \<and> 0 \<le> f x" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
719 |
by (auto simp: real ennreal_minus) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
720 |
qed |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
721 |
with g have "((\<lambda>x. ennreal (r - f x)) \<longlongrightarrow> ennreal 0) F" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
722 |
by (auto simp add: tendsto_cong eventually_conj_iff) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
723 |
with ae2 have "((\<lambda>x. r - f x) \<longlongrightarrow> 0) F" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
724 |
by (subst (asm) tendsto_ennreal_iff) (auto elim: eventually_mono) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
725 |
then have "(f \<longlongrightarrow> r) F" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
726 |
by (rule Lim_transform2[OF tendsto_const]) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
727 |
with ae2 have "((\<lambda>x. ennreal (f x)) \<longlongrightarrow> ennreal r) F" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
728 |
by (subst tendsto_ennreal_iff) (auto elim: eventually_mono simp: real) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
729 |
with ae2 show ?thesis |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
730 |
by (auto simp: real tendsto_cong eventually_conj_iff) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
731 |
qed |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
732 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
733 |
lemma ereal_add_diff_cancel: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
734 |
fixes a b :: ereal |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
735 |
shows "\<bar>b\<bar> \<noteq> \<infinity> \<Longrightarrow> (a + b) - b = a" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
736 |
by (cases a b rule: ereal2_cases) auto |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
737 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
738 |
lemma ennreal_add_diff_cancel: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
739 |
fixes a b :: ennreal |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
740 |
shows "b \<noteq> \<infinity> \<Longrightarrow> (a + b) - b = a" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
741 |
unfolding infinity_ennreal_def |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
742 |
by transfer (simp add: max_absorb2 top_ereal_def ereal_add_diff_cancel) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
743 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
744 |
lemma ennreal_mult_eq_top_iff: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
745 |
fixes a b :: ennreal |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
746 |
shows "a * b = top \<longleftrightarrow> (a = top \<and> b \<noteq> 0) \<or> (b = top \<and> a \<noteq> 0)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
747 |
by transfer (auto simp: top_ereal_def) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
748 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
749 |
lemma ennreal_top_eq_mult_iff: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
750 |
fixes a b :: ennreal |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
751 |
shows "top = a * b \<longleftrightarrow> (a = top \<and> b \<noteq> 0) \<or> (b = top \<and> a \<noteq> 0)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
752 |
using ennreal_mult_eq_top_iff[of a b] by auto |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
753 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
754 |
lemma ennreal_mult: "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> ennreal (a * b) = ennreal a * ennreal b" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
755 |
by transfer (simp add: max_absorb2) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
756 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
757 |
lemma setsum_enn2ereal[simp]: "(\<And>i. i \<in> I \<Longrightarrow> 0 \<le> f i) \<Longrightarrow> (\<Sum>i\<in>I. enn2ereal (f i)) = enn2ereal (setsum f I)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
758 |
by (induction I rule: infinite_finite_induct) (auto simp: setsum_nonneg zero_ennreal.rep_eq plus_ennreal.rep_eq) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
759 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
760 |
lemma e2ennreal_enn2ereal[simp]: "e2ennreal (enn2ereal x) = x" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
761 |
by (simp add: e2ennreal_def max_absorb2 enn2ereal_nonneg ennreal.enn2ereal_inverse) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
762 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
763 |
lemma tendsto_enn2ereal_iff[simp]: "((\<lambda>i. enn2ereal (f i)) \<longlongrightarrow> enn2ereal x) F \<longleftrightarrow> (f \<longlongrightarrow> x) F" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
764 |
using continuous_on_enn2ereal[THEN continuous_on_tendsto_compose, of f x F] |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
765 |
continuous_on_e2ennreal[THEN continuous_on_tendsto_compose, of "\<lambda>x. enn2ereal (f x)" "enn2ereal x" F UNIV] |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
766 |
by auto |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
767 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
768 |
lemma sums_enn2ereal[simp]: "(\<lambda>i. enn2ereal (f i)) sums enn2ereal x \<longleftrightarrow> f sums x" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
769 |
unfolding sums_def by (simp add: always_eventually setsum_nonneg setsum_enn2ereal) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
770 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
771 |
lemma suminf_enn2real[simp]: "(\<Sum>i. enn2ereal (f i)) = enn2ereal (suminf f)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
772 |
by (rule sums_unique[symmetric]) (simp add: summable_sums) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
773 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
774 |
lemma pcr_ennreal_enn2ereal[simp]: "pcr_ennreal (enn2ereal x) x" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
775 |
by (simp add: ennreal.pcr_cr_eq cr_ennreal_def) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
776 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
777 |
lemma rel_fun_eq_pcr_ennreal: "rel_fun op = pcr_ennreal f g \<longleftrightarrow> f = enn2ereal \<circ> g" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
778 |
by (auto simp: rel_fun_def ennreal.pcr_cr_eq cr_ennreal_def) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
779 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
780 |
lemma transfer_e2ennreal_suminf [transfer_rule]: "rel_fun (rel_fun op = pcr_ennreal) pcr_ennreal suminf suminf" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
781 |
by (auto simp: rel_funI rel_fun_eq_pcr_ennreal comp_def) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
782 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
783 |
lemma ennreal_suminf_cmult[simp]: "(\<Sum>i. r * f i) = r * (\<Sum>i. f i::ennreal)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
784 |
by transfer (auto intro!: suminf_cmult_ereal) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
785 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
786 |
lemma ennreal_suminf_SUP_eq_directed: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
787 |
fixes f :: "'a \<Rightarrow> nat \<Rightarrow> ennreal" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
788 |
assumes *: "\<And>N i j. i \<in> I \<Longrightarrow> j \<in> I \<Longrightarrow> finite N \<Longrightarrow> \<exists>k\<in>I. \<forall>n\<in>N. f i n \<le> f k n \<and> f j n \<le> f k n" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
789 |
shows "(\<Sum>n. SUP i:I. f i n) = (SUP i:I. \<Sum>n. f i n)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
790 |
proof cases |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
791 |
assume "I \<noteq> {}" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
792 |
then obtain i where "i \<in> I" by auto |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
793 |
from * show ?thesis |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
794 |
by (transfer fixing: I) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
795 |
(auto simp: max_absorb2 SUP_upper2[OF \<open>i \<in> I\<close>] suminf_nonneg summable_ereal_pos \<open>I \<noteq> {}\<close> |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
796 |
intro!: suminf_SUP_eq_directed) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
797 |
qed (simp add: bot_ennreal) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
798 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
799 |
lemma ennreal_eq_zero_iff[simp]: "0 \<le> x \<Longrightarrow> ennreal x = 0 \<longleftrightarrow> x = 0" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
800 |
by transfer (auto simp: max_absorb2) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
801 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
802 |
lemma ennreal_neq_top[simp]: "ennreal r \<noteq> top" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
803 |
by transfer (simp add: top_ereal_def zero_ereal_def ereal_max[symmetric] del: ereal_max) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
804 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
805 |
lemma ennreal_of_nat_neq_top[simp]: "of_nat i \<noteq> (top::ennreal)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
806 |
by (induction i) auto |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
807 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
808 |
lemma ennreal_suminf_neq_top: "summable f \<Longrightarrow> (\<And>i. 0 \<le> f i) \<Longrightarrow> (\<Sum>i. ennreal (f i)) \<noteq> top" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
809 |
using sums_ennreal[of f "suminf f"] |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
810 |
by (simp add: suminf_nonneg sums_unique[symmetric] summable_sums_iff[symmetric] del: sums_ennreal) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
811 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
812 |
instance ennreal :: semiring_char_0 |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
813 |
proof (standard, safe intro!: linorder_injI) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
814 |
have *: "1 + of_nat k \<noteq> (0::ennreal)" for k |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
815 |
using add_pos_nonneg[OF zero_less_one, of "of_nat k :: ennreal"] by auto |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
816 |
fix x y :: nat assume "x < y" "of_nat x = (of_nat y::ennreal)" then show False |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
817 |
by (auto simp add: less_iff_Suc_add *) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
818 |
qed |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
819 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
820 |
lemma ennreal_suminf_lessD: "(\<Sum>i. f i :: ennreal) < x \<Longrightarrow> f i < x" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
821 |
using le_less_trans[OF setsum_le_suminf[OF summableI, of "{i}" f]] by simp |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
822 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
823 |
lemma ennreal_less_top[simp]: "ennreal x < top" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
824 |
by transfer (simp add: top_ereal_def max_def) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
825 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
826 |
lemma ennreal_le_epsilon: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
827 |
"(\<And>e::real. y < top \<Longrightarrow> 0 < e \<Longrightarrow> x \<le> y + ennreal e) \<Longrightarrow> x \<le> y" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
828 |
apply (cases y rule: ennreal_cases) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
829 |
apply (cases x rule: ennreal_cases) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
830 |
apply (auto simp del: ennreal_plus simp add: top_unique ennreal_plus[symmetric] |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
831 |
intro: zero_less_one field_le_epsilon) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
832 |
done |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
833 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
834 |
lemma ennreal_less_zero_iff[simp]: "0 < ennreal x \<longleftrightarrow> 0 < x" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
835 |
by transfer (auto simp: max_def) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
836 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
837 |
lemma suminf_ennreal_eq: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
838 |
"(\<And>i. 0 \<le> f i) \<Longrightarrow> f sums x \<Longrightarrow> (\<Sum>i. ennreal (f i)) = ennreal x" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
839 |
using suminf_nonneg[of f] sums_unique[of f x] |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
840 |
by (intro sums_unique[symmetric]) (auto simp: summable_sums_iff) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
841 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
842 |
lemma transfer_e2ennreal_sumset [transfer_rule]: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
843 |
"rel_fun (rel_fun op = pcr_ennreal) (rel_fun op = pcr_ennreal) setsum setsum" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
844 |
by (auto intro!: rel_funI simp: rel_fun_eq_pcr_ennreal comp_def) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
845 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
846 |
lemma ennreal_suminf_bound_add: |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
847 |
fixes f :: "nat \<Rightarrow> ennreal" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
848 |
shows "(\<And>N. (\<Sum>n<N. f n) + y \<le> x) \<Longrightarrow> suminf f + y \<le> x" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
849 |
by transfer (auto intro!: suminf_bound_add) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
850 |
|
62375 | 851 |
end |