| author | paulson |
| Thu, 29 Nov 2001 17:39:23 +0100 | |
| changeset 12329 | 8743e8305611 |
| parent 12018 | ec054019c910 |
| child 12481 | ea5d6da573c5 |
| 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 |
|
| 9035 | 6 |
header {* Auxiliary theorems *}
|
| 7808 | 7 |
|
| 9035 | 8 |
theory Aux = Real + Zorn: |
| 7566 | 9 |
|
| 7917 | 10 |
text {* Some existing theorems are declared as extra introduction
|
| 9035 | 11 |
or elimination rules, respectively. *} |
| 7917 | 12 |
|
| 9408 | 13 |
lemmas [intro?] = isLub_isUb |
| 10687 | 14 |
lemmas [intro?] = chainD |
|
9941
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents:
9906
diff
changeset
|
15 |
lemmas chainE2 = chainD2 [elim_format, standard] |
| 7917 | 16 |
|
17 |
||
| 10687 | 18 |
text {* \medskip Lemmas about sets. *}
|
19 |
||
|
10783
2781ac7a4619
fixed two proofs that were affected by the removal of obsolete rules
paulson
parents:
10752
diff
changeset
|
20 |
lemma Int_singletonD: "A \<inter> B = {v} \<Longrightarrow> x \<in> A \<Longrightarrow> x \<in> B \<Longrightarrow> x = v"
|
| 9035 | 21 |
by (fast elim: equalityE) |
| 7917 | 22 |
|
|
10783
2781ac7a4619
fixed two proofs that were affected by the removal of obsolete rules
paulson
parents:
10752
diff
changeset
|
23 |
lemma set_less_imp_diff_not_empty: "H < E \<Longrightarrow> \<exists>x0 \<in> E. x0 \<notin> H" |
| 10687 | 24 |
by (auto simp add: psubset_eq) |
25 |
||
| 7917 | 26 |
|
| 10687 | 27 |
text{* \medskip Some lemmas about orders. *}
|
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
28 |
|
|
10783
2781ac7a4619
fixed two proofs that were affected by the removal of obsolete rules
paulson
parents:
10752
diff
changeset
|
29 |
lemma lt_imp_not_eq: "x < (y::'a::order) \<Longrightarrow> x \<noteq> y" |
| 9379 | 30 |
by (simp add: order_less_le) |
| 7917 | 31 |
|
| 10687 | 32 |
lemma le_noteq_imp_less: |
|
10783
2781ac7a4619
fixed two proofs that were affected by the removal of obsolete rules
paulson
parents:
10752
diff
changeset
|
33 |
"x \<le> (r::'a::order) \<Longrightarrow> x \<noteq> r \<Longrightarrow> x < r" |
| 9035 | 34 |
proof - |
|
10783
2781ac7a4619
fixed two proofs that were affected by the removal of obsolete rules
paulson
parents:
10752
diff
changeset
|
35 |
assume "x \<le> r" and ne:"x \<noteq> r" |
|
2781ac7a4619
fixed two proofs that were affected by the removal of obsolete rules
paulson
parents:
10752
diff
changeset
|
36 |
hence "x < r \<or> x = r" by (simp add: order_le_less) |
| 9035 | 37 |
with ne show ?thesis by simp |
38 |
qed |
|
| 7917 | 39 |
|
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
40 |
|
| 10687 | 41 |
text {* \medskip Some lemmas for the reals. *}
|
42 |
||
|
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11701
diff
changeset
|
43 |
lemma real_add_minus_eq: "x - y = (0::real) \<Longrightarrow> x = y" |
| 9035 | 44 |
by simp |
|
9013
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents:
8838
diff
changeset
|
45 |
|
|
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11701
diff
changeset
|
46 |
lemma abs_minus_one: "abs (- (1::real)) = 1" |
| 9035 | 47 |
by simp |
|
9013
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents:
8838
diff
changeset
|
48 |
|
| 10687 | 49 |
lemma real_mult_le_le_mono1a: |
|
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11701
diff
changeset
|
50 |
"(0::real) \<le> z \<Longrightarrow> x \<le> y \<Longrightarrow> z * x \<le> z * y" |
|
10783
2781ac7a4619
fixed two proofs that were affected by the removal of obsolete rules
paulson
parents:
10752
diff
changeset
|
51 |
by (simp add: real_mult_le_mono2) |
| 7656 | 52 |
|
| 10687 | 53 |
lemma real_mult_le_le_mono2: |
|
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11701
diff
changeset
|
54 |
"(0::real) \<le> z \<Longrightarrow> x \<le> y \<Longrightarrow> x * z \<le> y * z" |
| 9035 | 55 |
proof - |
|
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11701
diff
changeset
|
56 |
assume "(0::real) \<le> z" "x \<le> y" |
|
10783
2781ac7a4619
fixed two proofs that were affected by the removal of obsolete rules
paulson
parents:
10752
diff
changeset
|
57 |
hence "x < y \<or> x = y" by (auto simp add: order_le_less) |
| 9035 | 58 |
thus ?thesis |
| 10687 | 59 |
proof |
60 |
assume "x < y" |
|
61 |
show ?thesis by (rule real_mult_le_less_mono1) (simp!) |
|
62 |
next |
|
63 |
assume "x = y" |
|
64 |
thus ?thesis by simp |
|
| 9035 | 65 |
qed |
66 |
qed |
|
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
67 |
|
| 10687 | 68 |
lemma real_mult_less_le_anti: |
|
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11701
diff
changeset
|
69 |
"z < (0::real) \<Longrightarrow> x \<le> y \<Longrightarrow> z * y \<le> z * x" |
| 9035 | 70 |
proof - |
|
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11701
diff
changeset
|
71 |
assume "z < 0" "x \<le> y" |
|
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11701
diff
changeset
|
72 |
hence "0 < - z" by simp |
|
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11701
diff
changeset
|
73 |
hence "0 \<le> - z" by (rule order_less_imp_le) |
|
10783
2781ac7a4619
fixed two proofs that were affected by the removal of obsolete rules
paulson
parents:
10752
diff
changeset
|
74 |
hence "x * (- z) \<le> y * (- z)" |
| 9035 | 75 |
by (rule real_mult_le_le_mono2) |
|
10783
2781ac7a4619
fixed two proofs that were affected by the removal of obsolete rules
paulson
parents:
10752
diff
changeset
|
76 |
hence "- (x * z) \<le> - (y * z)" |
| 9035 | 77 |
by (simp only: real_minus_mult_eq2) |
78 |
thus ?thesis by (simp only: real_mult_commute) |
|
79 |
qed |
|
|
9013
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents:
8838
diff
changeset
|
80 |
|
| 10687 | 81 |
lemma real_mult_less_le_mono: |
|
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11701
diff
changeset
|
82 |
"(0::real) < z \<Longrightarrow> x \<le> y \<Longrightarrow> z * x \<le> z * y" |
| 10687 | 83 |
proof - |
|
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11701
diff
changeset
|
84 |
assume "0 < z" "x \<le> y" |
|
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11701
diff
changeset
|
85 |
have "0 \<le> z" by (rule order_less_imp_le) |
|
10783
2781ac7a4619
fixed two proofs that were affected by the removal of obsolete rules
paulson
parents:
10752
diff
changeset
|
86 |
hence "x * z \<le> y * z" |
| 9035 | 87 |
by (rule real_mult_le_le_mono2) |
88 |
thus ?thesis by (simp only: real_mult_commute) |
|
89 |
qed |
|
|
9013
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents:
8838
diff
changeset
|
90 |
|
|
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11701
diff
changeset
|
91 |
lemma real_mult_inv_right1: "(x::real) \<noteq> 0 \<Longrightarrow> x * inverse x = 1" |
| 9394 | 92 |
by simp |
|
9013
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents:
8838
diff
changeset
|
93 |
|
|
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11701
diff
changeset
|
94 |
lemma real_mult_inv_left1: "(x::real) \<noteq> 0 \<Longrightarrow> inverse x * x = 1" |
| 9394 | 95 |
by simp |
|
9013
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents:
8838
diff
changeset
|
96 |
|
| 10687 | 97 |
lemma real_le_mult_order1a: |
|
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11701
diff
changeset
|
98 |
"(0::real) \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> 0 \<le> x * y" |
|
10783
2781ac7a4619
fixed two proofs that were affected by the removal of obsolete rules
paulson
parents:
10752
diff
changeset
|
99 |
by (simp add: real_0_le_mult_iff) |
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
100 |
|
| 10687 | 101 |
lemma real_mult_diff_distrib: |
| 9035 | 102 |
"a * (- x - (y::real)) = - a * x - a * y" |
103 |
proof - |
|
104 |
have "- x - y = - x + - y" by simp |
|
| 10687 | 105 |
also have "a * ... = a * - x + a * - y" |
| 9035 | 106 |
by (simp only: real_add_mult_distrib2) |
| 10687 | 107 |
also have "... = - a * x - a * y" |
| 9394 | 108 |
by simp |
| 9035 | 109 |
finally show ?thesis . |
110 |
qed |
|
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
111 |
|
| 9035 | 112 |
lemma real_mult_diff_distrib2: "a * (x - (y::real)) = a * x - a * y" |
| 10687 | 113 |
proof - |
| 9035 | 114 |
have "x - y = x + - y" by simp |
| 10687 | 115 |
also have "a * ... = a * x + a * - y" |
| 9035 | 116 |
by (simp only: real_add_mult_distrib2) |
| 10687 | 117 |
also have "... = a * x - a * y" |
| 9394 | 118 |
by simp |
| 9035 | 119 |
finally show ?thesis . |
120 |
qed |
|
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
121 |
|
|
10783
2781ac7a4619
fixed two proofs that were affected by the removal of obsolete rules
paulson
parents:
10752
diff
changeset
|
122 |
lemma real_minus_le: "- (x::real) \<le> y \<Longrightarrow> - y \<le> x" |
| 9035 | 123 |
by simp |
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
124 |
|
| 10687 | 125 |
lemma real_diff_ineq_swap: |
|
10783
2781ac7a4619
fixed two proofs that were affected by the removal of obsolete rules
paulson
parents:
10752
diff
changeset
|
126 |
"(d::real) - b \<le> c + a \<Longrightarrow> - a - b \<le> c - d" |
| 9035 | 127 |
by simp |
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
128 |
|
| 10687 | 129 |
end |