author | wenzelm |
Mon, 29 Nov 1999 15:52:49 +0100 | |
changeset 8039 | a901bafe4578 |
parent 7978 | 1b99ee57d131 |
child 8203 | 2fcc6017cb72 |
permissions | -rw-r--r-- |
7566 | 1 |
(* Title: HOL/Real/HahnBanach/Aux.thy |
2 |
ID: $Id$ |
|
3 |
Author: Gertrud Bauer, TU Munich |
|
4 |
*) |
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
5 |
|
7808 | 6 |
header {* Auxiliary theorems *}; |
7 |
||
7656 | 8 |
theory Aux = Real + Zorn:; |
7566 | 9 |
|
7917 | 10 |
text {* Some existing theorems are declared as extra introduction |
11 |
or elimination rules, respectively. *}; |
|
12 |
||
13 |
lemmas [intro!!] = isLub_isUb; |
|
7656 | 14 |
lemmas [intro!!] = chainD; |
15 |
lemmas chainE2 = chainD2 [elimify]; |
|
7917 | 16 |
|
17 |
text_raw {* \medskip *}; |
|
7978 | 18 |
text{* Lemmas about sets. *}; |
7917 | 19 |
|
20 |
lemma Int_singletonD: "[| A Int B = {v}; x:A; x:B |] ==> x = v"; |
|
21 |
by (fast elim: equalityE); |
|
22 |
||
23 |
lemma set_less_imp_diff_not_empty: "H < E ==> EX x0:E. x0 ~: H"; |
|
24 |
by (force simp add: psubset_eq); |
|
25 |
||
26 |
text_raw {* \medskip *}; |
|
7978 | 27 |
text{* Some lemmas about orders. *}; |
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
28 |
|
7917 | 29 |
lemma lt_imp_not_eq: "x < (y::'a::order) ==> x ~= y"; |
30 |
by (rule order_less_le[RS iffD1, RS conjunct2]); |
|
31 |
||
32 |
lemma le_noteq_imp_less: |
|
33 |
"[| x <= (r::'a::order); x ~= r |] ==> x < r"; |
|
34 |
proof -; |
|
35 |
assume "x <= (r::'a::order)" and ne:"x ~= r"; |
|
36 |
hence "x < r | x = r"; by (simp add: order_le_less); |
|
37 |
with ne; show ?thesis; by simp; |
|
38 |
qed; |
|
39 |
||
40 |
text_raw {* \medskip *}; |
|
41 |
text {* Some lemmas about linear orders. *}; |
|
7808 | 42 |
|
43 |
theorem linorder_linear_split: |
|
44 |
"[| x < a ==> Q; x = a ==> Q; a < (x::'a::linorder) ==> Q |] ==> Q"; |
|
7917 | 45 |
by (rule linorder_less_linear [of x a, elimify]) force+; |
7566 | 46 |
|
47 |
lemma le_max1: "x <= max x (y::'a::linorder)"; |
|
48 |
by (simp add: le_max_iff_disj[of x x y]); |
|
49 |
||
50 |
lemma le_max2: "y <= max x (y::'a::linorder)"; |
|
51 |
by (simp add: le_max_iff_disj[of y x y]); |
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
52 |
|
7917 | 53 |
text_raw {* \medskip *}; |
54 |
text{* Some lemmas for the reals. *}; |
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
55 |
|
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
56 |
lemma real_add_minus_eq: "x - y = 0r ==> x = y"; |
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
57 |
proof -; |
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
58 |
assume "x - y = 0r"; |
7808 | 59 |
have "x + - y = 0r"; by (simp!); |
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
60 |
hence "x = - (- y)"; by (rule real_add_minus_eq_minus); |
7656 | 61 |
also; have "... = y"; by simp; |
7808 | 62 |
finally; show "?thesis"; .; |
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
63 |
qed; |
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
64 |
|
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
65 |
lemma rabs_minus_one: "rabs (- 1r) = 1r"; |
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
66 |
proof -; |
7808 | 67 |
have "-1r < 0r"; |
68 |
by (rule real_minus_zero_less_iff[RS iffD1], simp, |
|
69 |
rule real_zero_less_one); |
|
70 |
hence "rabs (- 1r) = - (- 1r)"; |
|
71 |
by (rule rabs_minus_eqI2); |
|
7656 | 72 |
also; have "... = 1r"; by simp; |
7808 | 73 |
finally; show ?thesis; .; |
7656 | 74 |
qed; |
75 |
||
7808 | 76 |
lemma real_mult_le_le_mono2: |
77 |
"[| 0r <= z; x <= y |] ==> x * z <= y * z"; |
|
7656 | 78 |
proof -; |
7808 | 79 |
assume "0r <= z" "x <= y"; |
7656 | 80 |
hence "x < y | x = y"; by (force simp add: order_le_less); |
81 |
thus ?thesis; |
|
82 |
proof (elim disjE); |
|
83 |
assume "x < y"; show ?thesis; by (rule real_mult_le_less_mono1); |
|
84 |
next; |
|
7808 | 85 |
assume "x = y"; thus ?thesis;; by simp; |
7656 | 86 |
qed; |
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
87 |
qed; |
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
88 |
|
7808 | 89 |
lemma real_mult_less_le_anti: |
90 |
"[| z < 0r; x <= y |] ==> z * y <= z * x"; |
|
7656 | 91 |
proof -; |
7808 | 92 |
assume "z < 0r" "x <= y"; |
7656 | 93 |
hence "0r < - z"; by simp; |
94 |
hence "0r <= - z"; by (rule real_less_imp_le); |
|
7808 | 95 |
hence "(- z) * x <= (- z) * y"; |
96 |
by (rule real_mult_le_le_mono1); |
|
97 |
hence "- (z * x) <= - (z * y)"; |
|
98 |
by (simp only: real_minus_mult_eq1); |
|
7656 | 99 |
thus ?thesis; by simp; |
100 |
qed; |
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
101 |
|
7808 | 102 |
lemma real_mult_less_le_mono: |
103 |
"[| 0r < z; x <= y |] ==> z * x <= z * y"; |
|
7656 | 104 |
proof -; |
7808 | 105 |
assume "0r < z" "x <= y"; |
106 |
have "0r <= z"; by (rule real_less_imp_le); |
|
7656 | 107 |
thus ?thesis; by (rule real_mult_le_le_mono1); |
108 |
qed; |
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
109 |
|
7808 | 110 |
lemma real_mult_diff_distrib: |
111 |
"a * (- x - (y::real)) = - a * x - a * y"; |
|
7656 | 112 |
proof -; |
7808 | 113 |
have "- x - y = - x + - y"; by simp; |
114 |
also; have "a * ... = a * - x + a * - y"; |
|
115 |
by (simp only: real_add_mult_distrib2); |
|
7656 | 116 |
also; have "... = - a * x - a * y"; |
7808 | 117 |
by (simp add: real_minus_mult_eq2 [RS sym] real_minus_mult_eq1); |
7656 | 118 |
finally; show ?thesis; .; |
119 |
qed; |
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
120 |
|
7656 | 121 |
lemma real_mult_diff_distrib2: "a * (x - (y::real)) = a * x - a * y"; |
122 |
proof -; |
|
7808 | 123 |
have "x - y = x + - y"; by simp; |
124 |
also; have "a * ... = a * x + a * - y"; |
|
125 |
by (simp only: real_add_mult_distrib2); |
|
7656 | 126 |
also; have "... = a * x - a * y"; |
7808 | 127 |
by (simp add: real_minus_mult_eq2 [RS sym] real_minus_mult_eq1); |
7656 | 128 |
finally; show ?thesis; .; |
129 |
qed; |
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
130 |
|
7808 | 131 |
lemma real_minus_le: "- (x::real) <= y ==> - y <= x"; |
132 |
by simp; |
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
133 |
|
7808 | 134 |
lemma real_diff_ineq_swap: |
135 |
"(d::real) - b <= c + a ==> - a - b <= c - d"; |
|
136 |
by simp; |
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
137 |
|
7808 | 138 |
end; |