author | paulson |
Mon, 11 Oct 1999 10:50:41 +0200 | |
changeset 7823 | 742715638a01 |
parent 7808 | fd019ac3485f |
child 7917 | 5e5b9813cce7 |
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 |
|
7656 | 10 |
lemmas [intro!!] = chainD; |
11 |
lemmas chainE2 = chainD2 [elimify]; |
|
12 |
lemmas [intro!!] = isLub_isUb; |
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
13 |
|
7808 | 14 |
theorem real_linear_split: |
15 |
"[| x < a ==> Q; x = a ==> Q; a < (x::real) ==> Q |] ==> Q"; |
|
16 |
by (rule real_linear [of x a, elimify], elim disjE, force+); |
|
17 |
||
18 |
theorem linorder_linear_split: |
|
19 |
"[| x < a ==> Q; x = a ==> Q; a < (x::'a::linorder) ==> Q |] ==> Q"; |
|
20 |
by (rule linorder_less_linear [of x a, elimify], elim disjE, force+); |
|
7566 | 21 |
|
22 |
lemma le_max1: "x <= max x (y::'a::linorder)"; |
|
23 |
by (simp add: le_max_iff_disj[of x x y]); |
|
24 |
||
25 |
lemma le_max2: "y <= max x (y::'a::linorder)"; |
|
26 |
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
|
27 |
|
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
28 |
lemma lt_imp_not_eq: "x < (y::'a::order) ==> x ~= y"; |
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
29 |
by (rule order_less_le[RS iffD1, RS conjunct2]); |
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
30 |
|
7566 | 31 |
lemma Int_singletonD: "[| A Int B = {v}; x:A; x:B |] ==> x = v"; |
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
32 |
by (fast elim: equalityE); |
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
33 |
|
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
34 |
lemma real_add_minus_eq: "x - y = 0r ==> x = y"; |
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
35 |
proof -; |
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
36 |
assume "x - y = 0r"; |
7808 | 37 |
have "x + - y = 0r"; by (simp!); |
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
38 |
hence "x = - (- y)"; by (rule real_add_minus_eq_minus); |
7656 | 39 |
also; have "... = y"; by simp; |
7808 | 40 |
finally; show "?thesis"; .; |
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
41 |
qed; |
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
42 |
|
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
43 |
lemma rabs_minus_one: "rabs (- 1r) = 1r"; |
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
44 |
proof -; |
7808 | 45 |
have "-1r < 0r"; |
46 |
by (rule real_minus_zero_less_iff[RS iffD1], simp, |
|
47 |
rule real_zero_less_one); |
|
48 |
hence "rabs (- 1r) = - (- 1r)"; |
|
49 |
by (rule rabs_minus_eqI2); |
|
7656 | 50 |
also; have "... = 1r"; by simp; |
7808 | 51 |
finally; show ?thesis; .; |
7656 | 52 |
qed; |
53 |
||
7808 | 54 |
lemma real_mult_le_le_mono2: |
55 |
"[| 0r <= z; x <= y |] ==> x * z <= y * z"; |
|
7656 | 56 |
proof -; |
7808 | 57 |
assume "0r <= z" "x <= y"; |
7656 | 58 |
hence "x < y | x = y"; by (force simp add: order_le_less); |
59 |
thus ?thesis; |
|
60 |
proof (elim disjE); |
|
61 |
assume "x < y"; show ?thesis; by (rule real_mult_le_less_mono1); |
|
62 |
next; |
|
7808 | 63 |
assume "x = y"; thus ?thesis;; by simp; |
7656 | 64 |
qed; |
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
65 |
qed; |
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
66 |
|
7808 | 67 |
lemma real_mult_less_le_anti: |
68 |
"[| z < 0r; x <= y |] ==> z * y <= z * x"; |
|
7656 | 69 |
proof -; |
7808 | 70 |
assume "z < 0r" "x <= y"; |
7656 | 71 |
hence "0r < - z"; by simp; |
72 |
hence "0r <= - z"; by (rule real_less_imp_le); |
|
7808 | 73 |
hence "(- z) * x <= (- z) * y"; |
74 |
by (rule real_mult_le_le_mono1); |
|
75 |
hence "- (z * x) <= - (z * y)"; |
|
76 |
by (simp only: real_minus_mult_eq1); |
|
7656 | 77 |
thus ?thesis; by simp; |
78 |
qed; |
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
79 |
|
7808 | 80 |
lemma real_mult_less_le_mono: |
81 |
"[| 0r < z; x <= y |] ==> z * x <= z * y"; |
|
7656 | 82 |
proof -; |
7808 | 83 |
assume "0r < z" "x <= y"; |
84 |
have "0r <= z"; by (rule real_less_imp_le); |
|
7656 | 85 |
thus ?thesis; by (rule real_mult_le_le_mono1); |
86 |
qed; |
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
87 |
|
7808 | 88 |
lemma real_mult_diff_distrib: |
89 |
"a * (- x - (y::real)) = - a * x - a * y"; |
|
7656 | 90 |
proof -; |
7808 | 91 |
have "- x - y = - x + - y"; by simp; |
92 |
also; have "a * ... = a * - x + a * - y"; |
|
93 |
by (simp only: real_add_mult_distrib2); |
|
7656 | 94 |
also; have "... = - a * x - a * y"; |
7808 | 95 |
by (simp add: real_minus_mult_eq2 [RS sym] real_minus_mult_eq1); |
7656 | 96 |
finally; show ?thesis; .; |
97 |
qed; |
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
98 |
|
7656 | 99 |
lemma real_mult_diff_distrib2: "a * (x - (y::real)) = a * x - a * y"; |
100 |
proof -; |
|
7808 | 101 |
have "x - y = x + - y"; by simp; |
102 |
also; have "a * ... = a * x + a * - y"; |
|
103 |
by (simp only: real_add_mult_distrib2); |
|
7656 | 104 |
also; have "... = a * x - a * y"; |
7808 | 105 |
by (simp add: real_minus_mult_eq2 [RS sym] real_minus_mult_eq1); |
7656 | 106 |
finally; show ?thesis; .; |
107 |
qed; |
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
108 |
|
7808 | 109 |
lemma le_noteq_imp_less: |
110 |
"[| x <= (r::'a::order); x ~= r |] ==> x < r"; |
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
111 |
proof -; |
7656 | 112 |
assume "x <= (r::'a::order)" and ne:"x ~= r"; |
113 |
then; have "x < r | x = r"; by (simp add: order_le_less); |
|
114 |
with ne; show ?thesis; by simp; |
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
115 |
qed; |
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
116 |
|
7808 | 117 |
lemma real_minus_le: "- (x::real) <= y ==> - y <= x"; |
118 |
by simp; |
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
119 |
|
7808 | 120 |
lemma real_diff_ineq_swap: |
121 |
"(d::real) - b <= c + a ==> - a - b <= c - d"; |
|
122 |
by simp; |
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
123 |
|
7656 | 124 |
lemma set_less_imp_diff_not_empty: "H < E ==> EX x0:E. x0 ~: H"; |
125 |
by (force simp add: psubset_eq); |
|
126 |
||
7808 | 127 |
end; |