author | nipkow |
Fri, 16 Jul 2004 17:32:34 +0200 | |
changeset 15055 | aed573241bea |
parent 15013 | 34264f5e4691 |
child 15131 | c69542757a4d |
permissions | -rw-r--r-- |
7219 | 1 |
(* Title : PReal.thy |
2 |
ID : $Id$ |
|
5078 | 3 |
Author : Jacques D. Fleuriot |
4 |
Copyright : 1998 University of Cambridge |
|
5 |
Description : The positive reals as Dedekind sections of positive |
|
14335 | 6 |
rationals. Fundamentals of Abstract Analysis [Gleason- p. 121] |
5078 | 7 |
provides some of the definitions. |
8 |
*) |
|
9 |
||
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
10 |
theory PReal = Rational: |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
11 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
12 |
text{*Could be generalized and moved to @{text Ring_and_Field}*} |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
13 |
lemma add_eq_exists: "\<exists>x. a+x = (b::rat)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
14 |
by (rule_tac x="b-a" in exI, simp) |
5078 | 15 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
16 |
text{*As a special case, the sum of two positives is positive. One of the |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
17 |
premises could be weakened to the relation @{text "\<le>"}.*} |
14738 | 18 |
lemma pos_add_strict: "[|0<a; b<c|] ==> b < a + (c::'a::ordered_semidom)" |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
19 |
by (insert add_strict_mono [of 0 a b c], simp) |
14335 | 20 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
21 |
lemma interval_empty_iff: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
22 |
"({y::'a::ordered_field. x < y & y < z} = {}) = (~(x < z))" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
23 |
by (blast dest: dense intro: order_less_trans) |
14335 | 24 |
|
5078 | 25 |
|
26 |
constdefs |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
27 |
cut :: "rat set => bool" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
28 |
"cut A == {} \<subset> A & |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
29 |
A < {r. 0 < r} & |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
30 |
(\<forall>y \<in> A. ((\<forall>z. 0<z & z < y --> z \<in> A) & (\<exists>u \<in> A. y < u)))" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
31 |
|
5078 | 32 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
33 |
lemma cut_of_rat: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
34 |
assumes q: "0 < q" shows "cut {r::rat. 0 < r & r < q}" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
35 |
proof - |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
36 |
let ?A = "{r::rat. 0 < r & r < q}" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
37 |
from q have pos: "?A < {r. 0 < r}" by force |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
38 |
have nonempty: "{} \<subset> ?A" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
39 |
proof |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
40 |
show "{} \<subseteq> ?A" by simp |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
41 |
show "{} \<noteq> ?A" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
42 |
by (force simp only: q eq_commute [of "{}"] interval_empty_iff) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
43 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
44 |
show ?thesis |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
45 |
by (simp add: cut_def pos nonempty, |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
46 |
blast dest: dense intro: order_less_trans) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
47 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
48 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
49 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
50 |
typedef preal = "{A. cut A}" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
51 |
by (blast intro: cut_of_rat [OF zero_less_one]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
52 |
|
14691 | 53 |
instance preal :: "{ord, plus, minus, times, inverse}" .. |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
54 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
55 |
constdefs |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
56 |
preal_of_rat :: "rat => preal" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
57 |
"preal_of_rat q == Abs_preal({x::rat. 0 < x & x < q})" |
5078 | 58 |
|
14335 | 59 |
psup :: "preal set => preal" |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
60 |
"psup(P) == Abs_preal(\<Union>X \<in> P. Rep_preal(X))" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
61 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
62 |
add_set :: "[rat set,rat set] => rat set" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
63 |
"add_set A B == {w. \<exists>x \<in> A. \<exists>y \<in> B. w = x + y}" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
64 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
65 |
diff_set :: "[rat set,rat set] => rat set" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
66 |
"diff_set A B == {w. \<exists>x. 0 < w & 0 < x & x \<notin> B & x + w \<in> A}" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
67 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
68 |
mult_set :: "[rat set,rat set] => rat set" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
69 |
"mult_set A B == {w. \<exists>x \<in> A. \<exists>y \<in> B. w = x * y}" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
70 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
71 |
inverse_set :: "rat set => rat set" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
72 |
"inverse_set A == {x. \<exists>y. 0 < x & x < y & inverse y \<notin> A}" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
73 |
|
5078 | 74 |
|
14335 | 75 |
defs (overloaded) |
5078 | 76 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
77 |
preal_less_def: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
78 |
"R < (S::preal) == Rep_preal R < Rep_preal S" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
79 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
80 |
preal_le_def: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
81 |
"R \<le> (S::preal) == Rep_preal R \<subseteq> Rep_preal S" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
82 |
|
14335 | 83 |
preal_add_def: |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
84 |
"R + S == Abs_preal (add_set (Rep_preal R) (Rep_preal S))" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
85 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
86 |
preal_diff_def: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
87 |
"R - S == Abs_preal (diff_set (Rep_preal R) (Rep_preal S))" |
5078 | 88 |
|
14335 | 89 |
preal_mult_def: |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
90 |
"R * S == Abs_preal(mult_set (Rep_preal R) (Rep_preal S))" |
5078 | 91 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
92 |
preal_inverse_def: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
93 |
"inverse R == Abs_preal(inverse_set (Rep_preal R))" |
14335 | 94 |
|
95 |
||
96 |
lemma inj_on_Abs_preal: "inj_on Abs_preal preal" |
|
97 |
apply (rule inj_on_inverseI) |
|
98 |
apply (erule Abs_preal_inverse) |
|
99 |
done |
|
100 |
||
101 |
declare inj_on_Abs_preal [THEN inj_on_iff, simp] |
|
102 |
||
103 |
lemma inj_Rep_preal: "inj(Rep_preal)" |
|
104 |
apply (rule inj_on_inverseI) |
|
105 |
apply (rule Rep_preal_inverse) |
|
106 |
done |
|
107 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
108 |
lemma preal_nonempty: "A \<in> preal ==> \<exists>x\<in>A. 0 < x" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
109 |
by (unfold preal_def cut_def, blast) |
14335 | 110 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
111 |
lemma preal_imp_psubset_positives: "A \<in> preal ==> A < {r. 0 < r}" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
112 |
by (force simp add: preal_def cut_def) |
14335 | 113 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
114 |
lemma preal_exists_bound: "A \<in> preal ==> \<exists>x. 0 < x & x \<notin> A" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
115 |
by (drule preal_imp_psubset_positives, auto) |
14335 | 116 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
117 |
lemma preal_exists_greater: "[| A \<in> preal; y \<in> A |] ==> \<exists>u \<in> A. y < u" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
118 |
by (unfold preal_def cut_def, blast) |
14335 | 119 |
|
120 |
lemma mem_Rep_preal_Ex: "\<exists>x. x \<in> Rep_preal X" |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
121 |
apply (insert Rep_preal [of X]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
122 |
apply (unfold preal_def cut_def, blast) |
14335 | 123 |
done |
124 |
||
125 |
declare Abs_preal_inverse [simp] |
|
126 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
127 |
lemma preal_downwards_closed: "[| A \<in> preal; y \<in> A; 0 < z; z < y |] ==> z \<in> A" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
128 |
by (unfold preal_def cut_def, blast) |
14335 | 129 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
130 |
text{*Relaxing the final premise*} |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
131 |
lemma preal_downwards_closed': |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
132 |
"[| A \<in> preal; y \<in> A; 0 < z; z \<le> y |] ==> z \<in> A" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
133 |
apply (simp add: order_le_less) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
134 |
apply (blast intro: preal_downwards_closed) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
135 |
done |
14335 | 136 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
137 |
lemma Rep_preal_exists_bound: "\<exists>x. 0 < x & x \<notin> Rep_preal X" |
14335 | 138 |
apply (cut_tac x = X in Rep_preal) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
139 |
apply (drule preal_imp_psubset_positives) |
14335 | 140 |
apply (auto simp add: psubset_def) |
141 |
done |
|
142 |
||
143 |
||
144 |
subsection{*@{term preal_of_prat}: the Injection from prat to preal*} |
|
145 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
146 |
lemma rat_less_set_mem_preal: "0 < y ==> {u::rat. 0 < u & u < y} \<in> preal" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
147 |
apply (auto simp add: preal_def cut_def intro: order_less_trans) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
148 |
apply (force simp only: eq_commute [of "{}"] interval_empty_iff) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
149 |
apply (blast dest: dense intro: order_less_trans) |
14335 | 150 |
done |
151 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
152 |
lemma rat_subset_imp_le: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
153 |
"[|{u::rat. 0 < u & u < x} \<subseteq> {u. 0 < u & u < y}; 0<x|] ==> x \<le> y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
154 |
apply (simp add: linorder_not_less [symmetric]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
155 |
apply (blast dest: dense intro: order_less_trans) |
14335 | 156 |
done |
157 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
158 |
lemma rat_set_eq_imp_eq: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
159 |
"[|{u::rat. 0 < u & u < x} = {u. 0 < u & u < y}; |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
160 |
0 < x; 0 < y|] ==> x = y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
161 |
by (blast intro: rat_subset_imp_le order_antisym) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
162 |
|
14335 | 163 |
|
164 |
||
165 |
subsection{*Theorems for Ordering*} |
|
166 |
||
167 |
text{*A positive fraction not in a positive real is an upper bound. |
|
168 |
Gleason p. 122 - Remark (1)*} |
|
169 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
170 |
lemma not_in_preal_ub: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
171 |
assumes A: "A \<in> preal" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
172 |
and notx: "x \<notin> A" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
173 |
and y: "y \<in> A" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
174 |
and pos: "0 < x" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
175 |
shows "y < x" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
176 |
proof (cases rule: linorder_cases) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
177 |
assume "x<y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
178 |
with notx show ?thesis |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
179 |
by (simp add: preal_downwards_closed [OF A y] pos) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
180 |
next |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
181 |
assume "x=y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
182 |
with notx and y show ?thesis by simp |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
183 |
next |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
184 |
assume "y<x" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
185 |
thus ?thesis by assumption |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
186 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
187 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
188 |
lemmas not_in_Rep_preal_ub = not_in_preal_ub [OF Rep_preal] |
14335 | 189 |
|
190 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
191 |
subsection{*The @{text "\<le>"} Ordering*} |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
192 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
193 |
lemma preal_le_refl: "w \<le> (w::preal)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
194 |
by (simp add: preal_le_def) |
14335 | 195 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
196 |
lemma preal_le_trans: "[| i \<le> j; j \<le> k |] ==> i \<le> (k::preal)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
197 |
by (force simp add: preal_le_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
198 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
199 |
lemma preal_le_anti_sym: "[| z \<le> w; w \<le> z |] ==> z = (w::preal)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
200 |
apply (simp add: preal_le_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
201 |
apply (rule Rep_preal_inject [THEN iffD1], blast) |
14335 | 202 |
done |
203 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
204 |
(* Axiom 'order_less_le' of class 'order': *) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
205 |
lemma preal_less_le: "((w::preal) < z) = (w \<le> z & w \<noteq> z)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
206 |
by (simp add: preal_le_def preal_less_def Rep_preal_inject psubset_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
207 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
208 |
instance preal :: order |
14691 | 209 |
by intro_classes |
210 |
(assumption | |
|
211 |
rule preal_le_refl preal_le_trans preal_le_anti_sym preal_less_le)+ |
|
14335 | 212 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
213 |
lemma preal_imp_pos: "[|A \<in> preal; r \<in> A|] ==> 0 < r" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
214 |
by (insert preal_imp_psubset_positives, blast) |
14335 | 215 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
216 |
lemma preal_le_linear: "x <= y | y <= (x::preal)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
217 |
apply (auto simp add: preal_le_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
218 |
apply (rule ccontr) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
219 |
apply (blast dest: not_in_Rep_preal_ub intro: preal_imp_pos [OF Rep_preal] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
220 |
elim: order_less_asym) |
14335 | 221 |
done |
222 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
223 |
instance preal :: linorder |
14691 | 224 |
by intro_classes (rule preal_le_linear) |
14335 | 225 |
|
226 |
||
227 |
||
228 |
subsection{*Properties of Addition*} |
|
229 |
||
230 |
lemma preal_add_commute: "(x::preal) + y = y + x" |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
231 |
apply (unfold preal_add_def add_set_def) |
14335 | 232 |
apply (rule_tac f = Abs_preal in arg_cong) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
233 |
apply (force simp add: add_commute) |
14335 | 234 |
done |
235 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
236 |
text{*Lemmas for proving that addition of two positive reals gives |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
237 |
a positive real*} |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
238 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
239 |
lemma empty_psubset_nonempty: "a \<in> A ==> {} \<subset> A" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
240 |
by blast |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
241 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
242 |
text{*Part 1 of Dedekind sections definition*} |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
243 |
lemma add_set_not_empty: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
244 |
"[|A \<in> preal; B \<in> preal|] ==> {} \<subset> add_set A B" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
245 |
apply (insert preal_nonempty [of A] preal_nonempty [of B]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
246 |
apply (auto simp add: add_set_def) |
14335 | 247 |
done |
248 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
249 |
text{*Part 2 of Dedekind sections definition. A structured version of |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
250 |
this proof is @{text preal_not_mem_mult_set_Ex} below.*} |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
251 |
lemma preal_not_mem_add_set_Ex: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
252 |
"[|A \<in> preal; B \<in> preal|] ==> \<exists>q. 0 < q & q \<notin> add_set A B" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
253 |
apply (insert preal_exists_bound [of A] preal_exists_bound [of B], auto) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
254 |
apply (rule_tac x = "x+xa" in exI) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
255 |
apply (simp add: add_set_def, clarify) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
256 |
apply (drule not_in_preal_ub, assumption+)+ |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
257 |
apply (force dest: add_strict_mono) |
14335 | 258 |
done |
259 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
260 |
lemma add_set_not_rat_set: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
261 |
assumes A: "A \<in> preal" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
262 |
and B: "B \<in> preal" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
263 |
shows "add_set A B < {r. 0 < r}" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
264 |
proof |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
265 |
from preal_imp_pos [OF A] preal_imp_pos [OF B] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
266 |
show "add_set A B \<subseteq> {r. 0 < r}" by (force simp add: add_set_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
267 |
next |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
268 |
show "add_set A B \<noteq> {r. 0 < r}" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
269 |
by (insert preal_not_mem_add_set_Ex [OF A B], blast) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
270 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
271 |
|
14335 | 272 |
text{*Part 3 of Dedekind sections definition*} |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
273 |
lemma add_set_lemma3: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
274 |
"[|A \<in> preal; B \<in> preal; u \<in> add_set A B; 0 < z; z < u|] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
275 |
==> z \<in> add_set A B" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
276 |
proof (unfold add_set_def, clarify) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
277 |
fix x::rat and y::rat |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
278 |
assume A: "A \<in> preal" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
279 |
and B: "B \<in> preal" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
280 |
and [simp]: "0 < z" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
281 |
and zless: "z < x + y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
282 |
and x: "x \<in> A" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
283 |
and y: "y \<in> B" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
284 |
have xpos [simp]: "0<x" by (rule preal_imp_pos [OF A x]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
285 |
have ypos [simp]: "0<y" by (rule preal_imp_pos [OF B y]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
286 |
have xypos [simp]: "0 < x+y" by (simp add: pos_add_strict) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
287 |
let ?f = "z/(x+y)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
288 |
have fless: "?f < 1" by (simp add: zless pos_divide_less_eq) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
289 |
show "\<exists>x' \<in> A. \<exists>y'\<in>B. z = x' + y'" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
290 |
proof |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
291 |
show "\<exists>y' \<in> B. z = x*?f + y'" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
292 |
proof |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
293 |
show "z = x*?f + y*?f" |
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14387
diff
changeset
|
294 |
by (simp add: left_distrib [symmetric] divide_inverse mult_ac |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
295 |
order_less_imp_not_eq2) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
296 |
next |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
297 |
show "y * ?f \<in> B" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
298 |
proof (rule preal_downwards_closed [OF B y]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
299 |
show "0 < y * ?f" |
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14387
diff
changeset
|
300 |
by (simp add: divide_inverse zero_less_mult_iff) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
301 |
next |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
302 |
show "y * ?f < y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
303 |
by (insert mult_strict_left_mono [OF fless ypos], simp) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
304 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
305 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
306 |
next |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
307 |
show "x * ?f \<in> A" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
308 |
proof (rule preal_downwards_closed [OF A x]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
309 |
show "0 < x * ?f" |
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14387
diff
changeset
|
310 |
by (simp add: divide_inverse zero_less_mult_iff) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
311 |
next |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
312 |
show "x * ?f < x" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
313 |
by (insert mult_strict_left_mono [OF fless xpos], simp) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
314 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
315 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
316 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
317 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
318 |
text{*Part 4 of Dedekind sections definition*} |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
319 |
lemma add_set_lemma4: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
320 |
"[|A \<in> preal; B \<in> preal; y \<in> add_set A B|] ==> \<exists>u \<in> add_set A B. y < u" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
321 |
apply (auto simp add: add_set_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
322 |
apply (frule preal_exists_greater [of A], auto) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
323 |
apply (rule_tac x="u + y" in exI) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
324 |
apply (auto intro: add_strict_left_mono) |
14335 | 325 |
done |
326 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
327 |
lemma mem_add_set: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
328 |
"[|A \<in> preal; B \<in> preal|] ==> add_set A B \<in> preal" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
329 |
apply (simp (no_asm_simp) add: preal_def cut_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
330 |
apply (blast intro!: add_set_not_empty add_set_not_rat_set |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
331 |
add_set_lemma3 add_set_lemma4) |
14335 | 332 |
done |
333 |
||
334 |
lemma preal_add_assoc: "((x::preal) + y) + z = x + (y + z)" |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
335 |
apply (simp add: preal_add_def mem_add_set Rep_preal) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
336 |
apply (force simp add: add_set_def add_ac) |
14335 | 337 |
done |
338 |
||
339 |
lemma preal_add_left_commute: "x + (y + z) = y + ((x + z)::preal)" |
|
340 |
apply (rule mk_left_commute [of "op +"]) |
|
341 |
apply (rule preal_add_assoc) |
|
342 |
apply (rule preal_add_commute) |
|
343 |
done |
|
344 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
345 |
text{* Positive Real addition is an AC operator *} |
14335 | 346 |
lemmas preal_add_ac = preal_add_assoc preal_add_commute preal_add_left_commute |
347 |
||
348 |
||
349 |
subsection{*Properties of Multiplication*} |
|
350 |
||
351 |
text{*Proofs essentially same as for addition*} |
|
352 |
||
353 |
lemma preal_mult_commute: "(x::preal) * y = y * x" |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
354 |
apply (unfold preal_mult_def mult_set_def) |
14335 | 355 |
apply (rule_tac f = Abs_preal in arg_cong) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
356 |
apply (force simp add: mult_commute) |
14335 | 357 |
done |
358 |
||
15055 | 359 |
text{*Multiplication of two positive reals gives a positive real.*} |
14335 | 360 |
|
361 |
text{*Lemmas for proving positive reals multiplication set in @{typ preal}*} |
|
362 |
||
363 |
text{*Part 1 of Dedekind sections definition*} |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
364 |
lemma mult_set_not_empty: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
365 |
"[|A \<in> preal; B \<in> preal|] ==> {} \<subset> mult_set A B" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
366 |
apply (insert preal_nonempty [of A] preal_nonempty [of B]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
367 |
apply (auto simp add: mult_set_def) |
14335 | 368 |
done |
369 |
||
370 |
text{*Part 2 of Dedekind sections definition*} |
|
371 |
lemma preal_not_mem_mult_set_Ex: |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
372 |
assumes A: "A \<in> preal" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
373 |
and B: "B \<in> preal" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
374 |
shows "\<exists>q. 0 < q & q \<notin> mult_set A B" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
375 |
proof - |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
376 |
from preal_exists_bound [OF A] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
377 |
obtain x where [simp]: "0 < x" "x \<notin> A" by blast |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
378 |
from preal_exists_bound [OF B] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
379 |
obtain y where [simp]: "0 < y" "y \<notin> B" by blast |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
380 |
show ?thesis |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
381 |
proof (intro exI conjI) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
382 |
show "0 < x*y" by (simp add: mult_pos) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
383 |
show "x * y \<notin> mult_set A B" |
14377 | 384 |
proof - |
385 |
{ fix u::rat and v::rat |
|
14550 | 386 |
assume "u \<in> A" and "v \<in> B" and "x*y = u*v" |
387 |
moreover |
|
388 |
with prems have "u<x" and "v<y" by (blast dest: not_in_preal_ub)+ |
|
389 |
moreover |
|
390 |
with prems have "0\<le>v" |
|
391 |
by (blast intro: preal_imp_pos [OF B] order_less_imp_le prems) |
|
392 |
moreover |
|
393 |
from calculation |
|
394 |
have "u*v < x*y" by (blast intro: mult_strict_mono prems) |
|
395 |
ultimately have False by force } |
|
14377 | 396 |
thus ?thesis by (auto simp add: mult_set_def) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
397 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
398 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
399 |
qed |
14335 | 400 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
401 |
lemma mult_set_not_rat_set: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
402 |
assumes A: "A \<in> preal" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
403 |
and B: "B \<in> preal" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
404 |
shows "mult_set A B < {r. 0 < r}" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
405 |
proof |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
406 |
show "mult_set A B \<subseteq> {r. 0 < r}" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
407 |
by (force simp add: mult_set_def |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
408 |
intro: preal_imp_pos [OF A] preal_imp_pos [OF B] mult_pos) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
409 |
next |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
410 |
show "mult_set A B \<noteq> {r. 0 < r}" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
411 |
by (insert preal_not_mem_mult_set_Ex [OF A B], blast) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
412 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
413 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
414 |
|
14335 | 415 |
|
416 |
text{*Part 3 of Dedekind sections definition*} |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
417 |
lemma mult_set_lemma3: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
418 |
"[|A \<in> preal; B \<in> preal; u \<in> mult_set A B; 0 < z; z < u|] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
419 |
==> z \<in> mult_set A B" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
420 |
proof (unfold mult_set_def, clarify) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
421 |
fix x::rat and y::rat |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
422 |
assume A: "A \<in> preal" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
423 |
and B: "B \<in> preal" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
424 |
and [simp]: "0 < z" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
425 |
and zless: "z < x * y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
426 |
and x: "x \<in> A" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
427 |
and y: "y \<in> B" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
428 |
have [simp]: "0<y" by (rule preal_imp_pos [OF B y]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
429 |
show "\<exists>x' \<in> A. \<exists>y' \<in> B. z = x' * y'" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
430 |
proof |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
431 |
show "\<exists>y'\<in>B. z = (z/y) * y'" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
432 |
proof |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
433 |
show "z = (z/y)*y" |
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14387
diff
changeset
|
434 |
by (simp add: divide_inverse mult_commute [of y] mult_assoc |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
435 |
order_less_imp_not_eq2) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
436 |
show "y \<in> B" . |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
437 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
438 |
next |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
439 |
show "z/y \<in> A" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
440 |
proof (rule preal_downwards_closed [OF A x]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
441 |
show "0 < z/y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
442 |
by (simp add: zero_less_divide_iff) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
443 |
show "z/y < x" by (simp add: pos_divide_less_eq zless) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
444 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
445 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
446 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
447 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
448 |
text{*Part 4 of Dedekind sections definition*} |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
449 |
lemma mult_set_lemma4: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
450 |
"[|A \<in> preal; B \<in> preal; y \<in> mult_set A B|] ==> \<exists>u \<in> mult_set A B. y < u" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
451 |
apply (auto simp add: mult_set_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
452 |
apply (frule preal_exists_greater [of A], auto) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
453 |
apply (rule_tac x="u * y" in exI) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
454 |
apply (auto intro: preal_imp_pos [of A] preal_imp_pos [of B] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
455 |
mult_strict_right_mono) |
14335 | 456 |
done |
457 |
||
458 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
459 |
lemma mem_mult_set: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
460 |
"[|A \<in> preal; B \<in> preal|] ==> mult_set A B \<in> preal" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
461 |
apply (simp (no_asm_simp) add: preal_def cut_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
462 |
apply (blast intro!: mult_set_not_empty mult_set_not_rat_set |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
463 |
mult_set_lemma3 mult_set_lemma4) |
14335 | 464 |
done |
465 |
||
466 |
lemma preal_mult_assoc: "((x::preal) * y) * z = x * (y * z)" |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
467 |
apply (simp add: preal_mult_def mem_mult_set Rep_preal) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
468 |
apply (force simp add: mult_set_def mult_ac) |
14335 | 469 |
done |
470 |
||
471 |
lemma preal_mult_left_commute: "x * (y * z) = y * ((x * z)::preal)" |
|
472 |
apply (rule mk_left_commute [of "op *"]) |
|
473 |
apply (rule preal_mult_assoc) |
|
474 |
apply (rule preal_mult_commute) |
|
475 |
done |
|
476 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
477 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
478 |
text{* Positive Real multiplication is an AC operator *} |
14335 | 479 |
lemmas preal_mult_ac = |
480 |
preal_mult_assoc preal_mult_commute preal_mult_left_commute |
|
481 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
482 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
483 |
text{* Positive real 1 is the multiplicative identity element *} |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
484 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
485 |
lemma rat_mem_preal: "0 < q ==> {r::rat. 0 < r & r < q} \<in> preal" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
486 |
by (simp add: preal_def cut_of_rat) |
14335 | 487 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
488 |
lemma preal_mult_1: "(preal_of_rat 1) * z = z" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
489 |
proof (induct z) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
490 |
fix A :: "rat set" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
491 |
assume A: "A \<in> preal" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
492 |
have "{w. \<exists>u. 0 < u \<and> u < 1 & (\<exists>v \<in> A. w = u * v)} = A" (is "?lhs = A") |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
493 |
proof |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
494 |
show "?lhs \<subseteq> A" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
495 |
proof clarify |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
496 |
fix x::rat and u::rat and v::rat |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
497 |
assume upos: "0<u" and "u<1" and v: "v \<in> A" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
498 |
have vpos: "0<v" by (rule preal_imp_pos [OF A v]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
499 |
hence "u*v < 1*v" by (simp only: mult_strict_right_mono prems) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
500 |
thus "u * v \<in> A" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
501 |
by (force intro: preal_downwards_closed [OF A v] mult_pos upos vpos) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
502 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
503 |
next |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
504 |
show "A \<subseteq> ?lhs" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
505 |
proof clarify |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
506 |
fix x::rat |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
507 |
assume x: "x \<in> A" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
508 |
have xpos: "0<x" by (rule preal_imp_pos [OF A x]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
509 |
from preal_exists_greater [OF A x] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
510 |
obtain v where v: "v \<in> A" and xlessv: "x < v" .. |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
511 |
have vpos: "0<v" by (rule preal_imp_pos [OF A v]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
512 |
show "\<exists>u. 0 < u \<and> u < 1 \<and> (\<exists>v\<in>A. x = u * v)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
513 |
proof (intro exI conjI) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
514 |
show "0 < x/v" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
515 |
by (simp add: zero_less_divide_iff xpos vpos) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
516 |
show "x / v < 1" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
517 |
by (simp add: pos_divide_less_eq vpos xlessv) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
518 |
show "\<exists>v'\<in>A. x = (x / v) * v'" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
519 |
proof |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
520 |
show "x = (x/v)*v" |
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14387
diff
changeset
|
521 |
by (simp add: divide_inverse mult_assoc vpos |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
522 |
order_less_imp_not_eq2) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
523 |
show "v \<in> A" . |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
524 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
525 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
526 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
527 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
528 |
thus "preal_of_rat 1 * Abs_preal A = Abs_preal A" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
529 |
by (simp add: preal_of_rat_def preal_mult_def mult_set_def |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
530 |
rat_mem_preal A) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
531 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
532 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
533 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
534 |
lemma preal_mult_1_right: "z * (preal_of_rat 1) = z" |
14335 | 535 |
apply (rule preal_mult_commute [THEN subst]) |
536 |
apply (rule preal_mult_1) |
|
537 |
done |
|
538 |
||
539 |
||
540 |
subsection{*Distribution of Multiplication across Addition*} |
|
541 |
||
542 |
lemma mem_Rep_preal_add_iff: |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
543 |
"(z \<in> Rep_preal(R+S)) = (\<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. z = x + y)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
544 |
apply (simp add: preal_add_def mem_add_set Rep_preal) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
545 |
apply (simp add: add_set_def) |
14335 | 546 |
done |
547 |
||
548 |
lemma mem_Rep_preal_mult_iff: |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
549 |
"(z \<in> Rep_preal(R*S)) = (\<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. z = x * y)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
550 |
apply (simp add: preal_mult_def mem_mult_set Rep_preal) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
551 |
apply (simp add: mult_set_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
552 |
done |
14335 | 553 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
554 |
lemma distrib_subset1: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
555 |
"Rep_preal (w * (x + y)) \<subseteq> Rep_preal (w * x + w * y)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
556 |
apply (auto simp add: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_mult_iff) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
557 |
apply (force simp add: right_distrib) |
14335 | 558 |
done |
559 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
560 |
lemma linorder_le_cases [case_names le ge]: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
561 |
"((x::'a::linorder) <= y ==> P) ==> (y <= x ==> P) ==> P" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
562 |
apply (insert linorder_linear, blast) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
563 |
done |
14335 | 564 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
565 |
lemma preal_add_mult_distrib_mean: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
566 |
assumes a: "a \<in> Rep_preal w" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
567 |
and b: "b \<in> Rep_preal w" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
568 |
and d: "d \<in> Rep_preal x" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
569 |
and e: "e \<in> Rep_preal y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
570 |
shows "\<exists>c \<in> Rep_preal w. a * d + b * e = c * (d + e)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
571 |
proof |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
572 |
let ?c = "(a*d + b*e)/(d+e)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
573 |
have [simp]: "0<a" "0<b" "0<d" "0<e" "0<d+e" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
574 |
by (blast intro: preal_imp_pos [OF Rep_preal] a b d e pos_add_strict)+ |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
575 |
have cpos: "0 < ?c" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
576 |
by (simp add: zero_less_divide_iff zero_less_mult_iff pos_add_strict) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
577 |
show "a * d + b * e = ?c * (d + e)" |
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14387
diff
changeset
|
578 |
by (simp add: divide_inverse mult_assoc order_less_imp_not_eq2) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
579 |
show "?c \<in> Rep_preal w" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
580 |
proof (cases rule: linorder_le_cases) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
581 |
assume "a \<le> b" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
582 |
hence "?c \<le> b" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
583 |
by (simp add: pos_divide_le_eq right_distrib mult_right_mono |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
584 |
order_less_imp_le) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
585 |
thus ?thesis by (rule preal_downwards_closed' [OF Rep_preal b cpos]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
586 |
next |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
587 |
assume "b \<le> a" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
588 |
hence "?c \<le> a" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
589 |
by (simp add: pos_divide_le_eq right_distrib mult_right_mono |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
590 |
order_less_imp_le) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
591 |
thus ?thesis by (rule preal_downwards_closed' [OF Rep_preal a cpos]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
592 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
593 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
594 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
595 |
lemma distrib_subset2: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
596 |
"Rep_preal (w * x + w * y) \<subseteq> Rep_preal (w * (x + y))" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
597 |
apply (auto simp add: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_mult_iff) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
598 |
apply (drule_tac w=w and x=x and y=y in preal_add_mult_distrib_mean, auto) |
14335 | 599 |
done |
600 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
601 |
lemma preal_add_mult_distrib2: "(w * ((x::preal) + y)) = (w * x) + (w * y)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
602 |
apply (rule inj_Rep_preal [THEN injD]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
603 |
apply (rule equalityI [OF distrib_subset1 distrib_subset2]) |
14335 | 604 |
done |
605 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
606 |
lemma preal_add_mult_distrib: "(((x::preal) + y) * w) = (x * w) + (y * w)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
607 |
by (simp add: preal_mult_commute preal_add_mult_distrib2) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
608 |
|
14335 | 609 |
|
610 |
subsection{*Existence of Inverse, a Positive Real*} |
|
611 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
612 |
lemma mem_inv_set_ex: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
613 |
assumes A: "A \<in> preal" shows "\<exists>x y. 0 < x & x < y & inverse y \<notin> A" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
614 |
proof - |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
615 |
from preal_exists_bound [OF A] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
616 |
obtain x where [simp]: "0<x" "x \<notin> A" by blast |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
617 |
show ?thesis |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
618 |
proof (intro exI conjI) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
619 |
show "0 < inverse (x+1)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
620 |
by (simp add: order_less_trans [OF _ less_add_one]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
621 |
show "inverse(x+1) < inverse x" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
622 |
by (simp add: less_imp_inverse_less less_add_one) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
623 |
show "inverse (inverse x) \<notin> A" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
624 |
by (simp add: order_less_imp_not_eq2) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
625 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
626 |
qed |
14335 | 627 |
|
628 |
text{*Part 1 of Dedekind sections definition*} |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
629 |
lemma inverse_set_not_empty: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
630 |
"A \<in> preal ==> {} \<subset> inverse_set A" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
631 |
apply (insert mem_inv_set_ex [of A]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
632 |
apply (auto simp add: inverse_set_def) |
14335 | 633 |
done |
634 |
||
635 |
text{*Part 2 of Dedekind sections definition*} |
|
636 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
637 |
lemma preal_not_mem_inverse_set_Ex: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
638 |
assumes A: "A \<in> preal" shows "\<exists>q. 0 < q & q \<notin> inverse_set A" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
639 |
proof - |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
640 |
from preal_nonempty [OF A] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
641 |
obtain x where x: "x \<in> A" and xpos [simp]: "0<x" .. |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
642 |
show ?thesis |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
643 |
proof (intro exI conjI) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
644 |
show "0 < inverse x" by simp |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
645 |
show "inverse x \<notin> inverse_set A" |
14377 | 646 |
proof - |
647 |
{ fix y::rat |
|
648 |
assume ygt: "inverse x < y" |
|
649 |
have [simp]: "0 < y" by (simp add: order_less_trans [OF _ ygt]) |
|
650 |
have iyless: "inverse y < x" |
|
651 |
by (simp add: inverse_less_imp_less [of x] ygt) |
|
652 |
have "inverse y \<in> A" |
|
653 |
by (simp add: preal_downwards_closed [OF A x] iyless)} |
|
654 |
thus ?thesis by (auto simp add: inverse_set_def) |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
655 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
656 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
657 |
qed |
14335 | 658 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
659 |
lemma inverse_set_not_rat_set: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
660 |
assumes A: "A \<in> preal" shows "inverse_set A < {r. 0 < r}" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
661 |
proof |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
662 |
show "inverse_set A \<subseteq> {r. 0 < r}" by (force simp add: inverse_set_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
663 |
next |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
664 |
show "inverse_set A \<noteq> {r. 0 < r}" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
665 |
by (insert preal_not_mem_inverse_set_Ex [OF A], blast) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
666 |
qed |
14335 | 667 |
|
668 |
text{*Part 3 of Dedekind sections definition*} |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
669 |
lemma inverse_set_lemma3: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
670 |
"[|A \<in> preal; u \<in> inverse_set A; 0 < z; z < u|] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
671 |
==> z \<in> inverse_set A" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
672 |
apply (auto simp add: inverse_set_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
673 |
apply (auto intro: order_less_trans) |
14335 | 674 |
done |
675 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
676 |
text{*Part 4 of Dedekind sections definition*} |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
677 |
lemma inverse_set_lemma4: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
678 |
"[|A \<in> preal; y \<in> inverse_set A|] ==> \<exists>u \<in> inverse_set A. y < u" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
679 |
apply (auto simp add: inverse_set_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
680 |
apply (drule dense [of y]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
681 |
apply (blast intro: order_less_trans) |
14335 | 682 |
done |
683 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
684 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
685 |
lemma mem_inverse_set: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
686 |
"A \<in> preal ==> inverse_set A \<in> preal" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
687 |
apply (simp (no_asm_simp) add: preal_def cut_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
688 |
apply (blast intro!: inverse_set_not_empty inverse_set_not_rat_set |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
689 |
inverse_set_lemma3 inverse_set_lemma4) |
14335 | 690 |
done |
691 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
692 |
|
14335 | 693 |
subsection{*Gleason's Lemma 9-3.4, page 122*} |
694 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
695 |
lemma Gleason9_34_exists: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
696 |
assumes A: "A \<in> preal" |
14369 | 697 |
and "\<forall>x\<in>A. x + u \<in> A" |
698 |
and "0 \<le> z" |
|
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14377
diff
changeset
|
699 |
shows "\<exists>b\<in>A. b + (of_int z) * u \<in> A" |
14369 | 700 |
proof (cases z rule: int_cases) |
701 |
case (nonneg n) |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
702 |
show ?thesis |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
703 |
proof (simp add: prems, induct n) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
704 |
case 0 |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
705 |
from preal_nonempty [OF A] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
706 |
show ?case by force |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
707 |
case (Suc k) |
15013 | 708 |
from this obtain b where "b \<in> A" "b + of_nat k * u \<in> A" .. |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14377
diff
changeset
|
709 |
hence "b + of_int (int k)*u + u \<in> A" by (simp add: prems) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
710 |
thus ?case by (force simp add: left_distrib add_ac prems) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
711 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
712 |
next |
14369 | 713 |
case (neg n) |
714 |
with prems show ?thesis by simp |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
715 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
716 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
717 |
lemma Gleason9_34_contra: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
718 |
assumes A: "A \<in> preal" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
719 |
shows "[|\<forall>x\<in>A. x + u \<in> A; 0 < u; 0 < y; y \<notin> A|] ==> False" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
720 |
proof (induct u, induct y) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
721 |
fix a::int and b::int |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
722 |
fix c::int and d::int |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
723 |
assume bpos [simp]: "0 < b" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
724 |
and dpos [simp]: "0 < d" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
725 |
and closed: "\<forall>x\<in>A. x + (Fract c d) \<in> A" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
726 |
and upos: "0 < Fract c d" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
727 |
and ypos: "0 < Fract a b" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
728 |
and notin: "Fract a b \<notin> A" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
729 |
have cpos [simp]: "0 < c" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
730 |
by (simp add: zero_less_Fract_iff [OF dpos, symmetric] upos) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
731 |
have apos [simp]: "0 < a" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
732 |
by (simp add: zero_less_Fract_iff [OF bpos, symmetric] ypos) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
733 |
let ?k = "a*d" |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14377
diff
changeset
|
734 |
have frle: "Fract a b \<le> Fract ?k 1 * (Fract c d)" |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
735 |
proof - |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
736 |
have "?thesis = ((a * d * b * d) \<le> c * b * (a * d * b * d))" |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14377
diff
changeset
|
737 |
by (simp add: mult_rat le_rat order_less_imp_not_eq2 mult_ac) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
738 |
moreover |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
739 |
have "(1 * (a * d * b * d)) \<le> c * b * (a * d * b * d)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
740 |
by (rule mult_mono, |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
741 |
simp_all add: int_one_le_iff_zero_less zero_less_mult_iff |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
742 |
order_less_imp_le) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
743 |
ultimately |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
744 |
show ?thesis by simp |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
745 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
746 |
have k: "0 \<le> ?k" by (simp add: order_less_imp_le zero_less_mult_iff) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
747 |
from Gleason9_34_exists [OF A closed k] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
748 |
obtain z where z: "z \<in> A" |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14377
diff
changeset
|
749 |
and mem: "z + of_int ?k * Fract c d \<in> A" .. |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14377
diff
changeset
|
750 |
have less: "z + of_int ?k * Fract c d < Fract a b" |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
751 |
by (rule not_in_preal_ub [OF A notin mem ypos]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
752 |
have "0<z" by (rule preal_imp_pos [OF A z]) |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14377
diff
changeset
|
753 |
with frle and less show False by (simp add: Fract_of_int_eq) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
754 |
qed |
14335 | 755 |
|
756 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
757 |
lemma Gleason9_34: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
758 |
assumes A: "A \<in> preal" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
759 |
and upos: "0 < u" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
760 |
shows "\<exists>r \<in> A. r + u \<notin> A" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
761 |
proof (rule ccontr, simp) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
762 |
assume closed: "\<forall>r\<in>A. r + u \<in> A" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
763 |
from preal_exists_bound [OF A] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
764 |
obtain y where y: "y \<notin> A" and ypos: "0 < y" by blast |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
765 |
show False |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
766 |
by (rule Gleason9_34_contra [OF A closed upos ypos y]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
767 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
768 |
|
14335 | 769 |
|
770 |
||
771 |
subsection{*Gleason's Lemma 9-3.6*} |
|
772 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
773 |
lemma lemma_gleason9_36: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
774 |
assumes A: "A \<in> preal" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
775 |
and x: "1 < x" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
776 |
shows "\<exists>r \<in> A. r*x \<notin> A" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
777 |
proof - |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
778 |
from preal_nonempty [OF A] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
779 |
obtain y where y: "y \<in> A" and ypos: "0<y" .. |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
780 |
show ?thesis |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
781 |
proof (rule classical) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
782 |
assume "~(\<exists>r\<in>A. r * x \<notin> A)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
783 |
with y have ymem: "y * x \<in> A" by blast |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
784 |
from ypos mult_strict_left_mono [OF x] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
785 |
have yless: "y < y*x" by simp |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
786 |
let ?d = "y*x - y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
787 |
from yless have dpos: "0 < ?d" and eq: "y + ?d = y*x" by auto |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
788 |
from Gleason9_34 [OF A dpos] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
789 |
obtain r where r: "r\<in>A" and notin: "r + ?d \<notin> A" .. |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
790 |
have rpos: "0<r" by (rule preal_imp_pos [OF A r]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
791 |
with dpos have rdpos: "0 < r + ?d" by arith |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
792 |
have "~ (r + ?d \<le> y + ?d)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
793 |
proof |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
794 |
assume le: "r + ?d \<le> y + ?d" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
795 |
from ymem have yd: "y + ?d \<in> A" by (simp add: eq) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
796 |
have "r + ?d \<in> A" by (rule preal_downwards_closed' [OF A yd rdpos le]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
797 |
with notin show False by simp |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
798 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
799 |
hence "y < r" by simp |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
800 |
with ypos have dless: "?d < (r * ?d)/y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
801 |
by (simp add: pos_less_divide_eq mult_commute [of ?d] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
802 |
mult_strict_right_mono dpos) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
803 |
have "r + ?d < r*x" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
804 |
proof - |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
805 |
have "r + ?d < r + (r * ?d)/y" by (simp add: dless) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
806 |
also with ypos have "... = (r/y) * (y + ?d)" |
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14387
diff
changeset
|
807 |
by (simp only: right_distrib divide_inverse mult_ac, simp) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
808 |
also have "... = r*x" using ypos |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
809 |
by simp |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
810 |
finally show "r + ?d < r*x" . |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
811 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
812 |
with r notin rdpos |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
813 |
show "\<exists>r\<in>A. r * x \<notin> A" by (blast dest: preal_downwards_closed [OF A]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
814 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
815 |
qed |
14335 | 816 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
817 |
subsection{*Existence of Inverse: Part 2*} |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
818 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
819 |
lemma mem_Rep_preal_inverse_iff: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
820 |
"(z \<in> Rep_preal(inverse R)) = |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
821 |
(0 < z \<and> (\<exists>y. z < y \<and> inverse y \<notin> Rep_preal R))" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
822 |
apply (simp add: preal_inverse_def mem_inverse_set Rep_preal) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
823 |
apply (simp add: inverse_set_def) |
14335 | 824 |
done |
825 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
826 |
lemma Rep_preal_of_rat: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
827 |
"0 < q ==> Rep_preal (preal_of_rat q) = {x. 0 < x \<and> x < q}" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
828 |
by (simp add: preal_of_rat_def rat_mem_preal) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
829 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
830 |
lemma subset_inverse_mult_lemma: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
831 |
assumes xpos: "0 < x" and xless: "x < 1" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
832 |
shows "\<exists>r u y. 0 < r & r < y & inverse y \<notin> Rep_preal R & |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
833 |
u \<in> Rep_preal R & x = r * u" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
834 |
proof - |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
835 |
from xpos and xless have "1 < inverse x" by (simp add: one_less_inverse_iff) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
836 |
from lemma_gleason9_36 [OF Rep_preal this] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
837 |
obtain r where r: "r \<in> Rep_preal R" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
838 |
and notin: "r * (inverse x) \<notin> Rep_preal R" .. |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
839 |
have rpos: "0<r" by (rule preal_imp_pos [OF Rep_preal r]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
840 |
from preal_exists_greater [OF Rep_preal r] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
841 |
obtain u where u: "u \<in> Rep_preal R" and rless: "r < u" .. |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
842 |
have upos: "0<u" by (rule preal_imp_pos [OF Rep_preal u]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
843 |
show ?thesis |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
844 |
proof (intro exI conjI) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
845 |
show "0 < x/u" using xpos upos |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
846 |
by (simp add: zero_less_divide_iff) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
847 |
show "x/u < x/r" using xpos upos rpos |
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14387
diff
changeset
|
848 |
by (simp add: divide_inverse mult_less_cancel_left rless) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
849 |
show "inverse (x / r) \<notin> Rep_preal R" using notin |
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14387
diff
changeset
|
850 |
by (simp add: divide_inverse mult_commute) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
851 |
show "u \<in> Rep_preal R" by (rule u) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
852 |
show "x = x / u * u" using upos |
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14387
diff
changeset
|
853 |
by (simp add: divide_inverse mult_commute) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
854 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
855 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
856 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
857 |
lemma subset_inverse_mult: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
858 |
"Rep_preal(preal_of_rat 1) \<subseteq> Rep_preal(inverse R * R)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
859 |
apply (auto simp add: Bex_def Rep_preal_of_rat mem_Rep_preal_inverse_iff |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
860 |
mem_Rep_preal_mult_iff) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
861 |
apply (blast dest: subset_inverse_mult_lemma) |
14335 | 862 |
done |
863 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
864 |
lemma inverse_mult_subset_lemma: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
865 |
assumes rpos: "0 < r" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
866 |
and rless: "r < y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
867 |
and notin: "inverse y \<notin> Rep_preal R" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
868 |
and q: "q \<in> Rep_preal R" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
869 |
shows "r*q < 1" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
870 |
proof - |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
871 |
have "q < inverse y" using rpos rless |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
872 |
by (simp add: not_in_preal_ub [OF Rep_preal notin] q) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
873 |
hence "r * q < r/y" using rpos |
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14387
diff
changeset
|
874 |
by (simp add: divide_inverse mult_less_cancel_left) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
875 |
also have "... \<le> 1" using rpos rless |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
876 |
by (simp add: pos_divide_le_eq) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
877 |
finally show ?thesis . |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
878 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
879 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
880 |
lemma inverse_mult_subset: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
881 |
"Rep_preal(inverse R * R) \<subseteq> Rep_preal(preal_of_rat 1)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
882 |
apply (auto simp add: Bex_def Rep_preal_of_rat mem_Rep_preal_inverse_iff |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
883 |
mem_Rep_preal_mult_iff) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
884 |
apply (simp add: zero_less_mult_iff preal_imp_pos [OF Rep_preal]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
885 |
apply (blast intro: inverse_mult_subset_lemma) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
886 |
done |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
887 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
888 |
lemma preal_mult_inverse: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
889 |
"inverse R * R = (preal_of_rat 1)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
890 |
apply (rule inj_Rep_preal [THEN injD]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
891 |
apply (rule equalityI [OF inverse_mult_subset subset_inverse_mult]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
892 |
done |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
893 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
894 |
lemma preal_mult_inverse_right: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
895 |
"R * inverse R = (preal_of_rat 1)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
896 |
apply (rule preal_mult_commute [THEN subst]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
897 |
apply (rule preal_mult_inverse) |
14335 | 898 |
done |
899 |
||
900 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
901 |
text{*Theorems needing @{text Gleason9_34}*} |
14335 | 902 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
903 |
lemma Rep_preal_self_subset: "Rep_preal (R) \<subseteq> Rep_preal(R + S)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
904 |
proof |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
905 |
fix r |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
906 |
assume r: "r \<in> Rep_preal R" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
907 |
have rpos: "0<r" by (rule preal_imp_pos [OF Rep_preal r]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
908 |
from mem_Rep_preal_Ex |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
909 |
obtain y where y: "y \<in> Rep_preal S" .. |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
910 |
have ypos: "0<y" by (rule preal_imp_pos [OF Rep_preal y]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
911 |
have ry: "r+y \<in> Rep_preal(R + S)" using r y |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
912 |
by (auto simp add: mem_Rep_preal_add_iff) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
913 |
show "r \<in> Rep_preal(R + S)" using r ypos rpos |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
914 |
by (simp add: preal_downwards_closed [OF Rep_preal ry]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
915 |
qed |
14335 | 916 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
917 |
lemma Rep_preal_sum_not_subset: "~ Rep_preal (R + S) \<subseteq> Rep_preal(R)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
918 |
proof - |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
919 |
from mem_Rep_preal_Ex |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
920 |
obtain y where y: "y \<in> Rep_preal S" .. |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
921 |
have ypos: "0<y" by (rule preal_imp_pos [OF Rep_preal y]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
922 |
from Gleason9_34 [OF Rep_preal ypos] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
923 |
obtain r where r: "r \<in> Rep_preal R" and notin: "r + y \<notin> Rep_preal R" .. |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
924 |
have "r + y \<in> Rep_preal (R + S)" using r y |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
925 |
by (auto simp add: mem_Rep_preal_add_iff) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
926 |
thus ?thesis using notin by blast |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
927 |
qed |
14335 | 928 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
929 |
lemma Rep_preal_sum_not_eq: "Rep_preal (R + S) \<noteq> Rep_preal(R)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
930 |
by (insert Rep_preal_sum_not_subset, blast) |
14335 | 931 |
|
932 |
text{*at last, Gleason prop. 9-3.5(iii) page 123*} |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
933 |
lemma preal_self_less_add_left: "(R::preal) < R + S" |
14335 | 934 |
apply (unfold preal_less_def psubset_def) |
935 |
apply (simp add: Rep_preal_self_subset Rep_preal_sum_not_eq [THEN not_sym]) |
|
936 |
done |
|
937 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
938 |
lemma preal_self_less_add_right: "(R::preal) < S + R" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
939 |
by (simp add: preal_add_commute preal_self_less_add_left) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
940 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
941 |
lemma preal_not_eq_self: "x \<noteq> x + (y::preal)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
942 |
by (insert preal_self_less_add_left [of x y], auto) |
14335 | 943 |
|
944 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
945 |
subsection{*Subtraction for Positive Reals*} |
14335 | 946 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
947 |
text{*Gleason prop. 9-3.5(iv), page 123: proving @{term "A < B ==> \<exists>D. A + D = |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
948 |
B"}. We define the claimed @{term D} and show that it is a positive real*} |
14335 | 949 |
|
950 |
text{*Part 1 of Dedekind sections definition*} |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
951 |
lemma diff_set_not_empty: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
952 |
"R < S ==> {} \<subset> diff_set (Rep_preal S) (Rep_preal R)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
953 |
apply (auto simp add: preal_less_def diff_set_def elim!: equalityE) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
954 |
apply (frule_tac x1 = S in Rep_preal [THEN preal_exists_greater]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
955 |
apply (drule preal_imp_pos [OF Rep_preal], clarify) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
956 |
apply (cut_tac a=x and b=u in add_eq_exists, force) |
14335 | 957 |
done |
958 |
||
959 |
text{*Part 2 of Dedekind sections definition*} |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
960 |
lemma diff_set_nonempty: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
961 |
"\<exists>q. 0 < q & q \<notin> diff_set (Rep_preal S) (Rep_preal R)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
962 |
apply (cut_tac X = S in Rep_preal_exists_bound) |
14335 | 963 |
apply (erule exE) |
964 |
apply (rule_tac x = x in exI, auto) |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
965 |
apply (simp add: diff_set_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
966 |
apply (auto dest: Rep_preal [THEN preal_downwards_closed]) |
14335 | 967 |
done |
968 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
969 |
lemma diff_set_not_rat_set: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
970 |
"diff_set (Rep_preal S) (Rep_preal R) < {r. 0 < r}" (is "?lhs < ?rhs") |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
971 |
proof |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
972 |
show "?lhs \<subseteq> ?rhs" by (auto simp add: diff_set_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
973 |
show "?lhs \<noteq> ?rhs" using diff_set_nonempty by blast |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
974 |
qed |
14335 | 975 |
|
976 |
text{*Part 3 of Dedekind sections definition*} |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
977 |
lemma diff_set_lemma3: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
978 |
"[|R < S; u \<in> diff_set (Rep_preal S) (Rep_preal R); 0 < z; z < u|] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
979 |
==> z \<in> diff_set (Rep_preal S) (Rep_preal R)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
980 |
apply (auto simp add: diff_set_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
981 |
apply (rule_tac x=x in exI) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
982 |
apply (drule Rep_preal [THEN preal_downwards_closed], auto) |
14335 | 983 |
done |
984 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
985 |
text{*Part 4 of Dedekind sections definition*} |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
986 |
lemma diff_set_lemma4: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
987 |
"[|R < S; y \<in> diff_set (Rep_preal S) (Rep_preal R)|] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
988 |
==> \<exists>u \<in> diff_set (Rep_preal S) (Rep_preal R). y < u" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
989 |
apply (auto simp add: diff_set_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
990 |
apply (drule Rep_preal [THEN preal_exists_greater], clarify) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
991 |
apply (cut_tac a="x+y" and b=u in add_eq_exists, clarify) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
992 |
apply (rule_tac x="y+xa" in exI) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
993 |
apply (auto simp add: add_ac) |
14335 | 994 |
done |
995 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
996 |
lemma mem_diff_set: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
997 |
"R < S ==> diff_set (Rep_preal S) (Rep_preal R) \<in> preal" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
998 |
apply (unfold preal_def cut_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
999 |
apply (blast intro!: diff_set_not_empty diff_set_not_rat_set |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1000 |
diff_set_lemma3 diff_set_lemma4) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1001 |
done |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1002 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1003 |
lemma mem_Rep_preal_diff_iff: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1004 |
"R < S ==> |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1005 |
(z \<in> Rep_preal(S-R)) = |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1006 |
(\<exists>x. 0 < x & 0 < z & x \<notin> Rep_preal R & x + z \<in> Rep_preal S)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1007 |
apply (simp add: preal_diff_def mem_diff_set Rep_preal) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1008 |
apply (force simp add: diff_set_def) |
14335 | 1009 |
done |
1010 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1011 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1012 |
text{*proving that @{term "R + D \<le> S"}*} |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1013 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1014 |
lemma less_add_left_lemma: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1015 |
assumes Rless: "R < S" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1016 |
and a: "a \<in> Rep_preal R" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1017 |
and cb: "c + b \<in> Rep_preal S" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1018 |
and "c \<notin> Rep_preal R" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1019 |
and "0 < b" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1020 |
and "0 < c" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1021 |
shows "a + b \<in> Rep_preal S" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1022 |
proof - |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1023 |
have "0<a" by (rule preal_imp_pos [OF Rep_preal a]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1024 |
moreover |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1025 |
have "a < c" using prems |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1026 |
by (blast intro: not_in_Rep_preal_ub ) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1027 |
ultimately show ?thesis using prems |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1028 |
by (simp add: preal_downwards_closed [OF Rep_preal cb]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1029 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1030 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1031 |
lemma less_add_left_le1: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1032 |
"R < (S::preal) ==> R + (S-R) \<le> S" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1033 |
apply (auto simp add: Bex_def preal_le_def mem_Rep_preal_add_iff |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1034 |
mem_Rep_preal_diff_iff) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1035 |
apply (blast intro: less_add_left_lemma) |
14335 | 1036 |
done |
1037 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1038 |
subsection{*proving that @{term "S \<le> R + D"} --- trickier*} |
14335 | 1039 |
|
1040 |
lemma lemma_sum_mem_Rep_preal_ex: |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1041 |
"x \<in> Rep_preal S ==> \<exists>e. 0 < e & x + e \<in> Rep_preal S" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1042 |
apply (drule Rep_preal [THEN preal_exists_greater], clarify) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1043 |
apply (cut_tac a=x and b=u in add_eq_exists, auto) |
14335 | 1044 |
done |
1045 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1046 |
lemma less_add_left_lemma2: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1047 |
assumes Rless: "R < S" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1048 |
and x: "x \<in> Rep_preal S" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1049 |
and xnot: "x \<notin> Rep_preal R" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1050 |
shows "\<exists>u v z. 0 < v & 0 < z & u \<in> Rep_preal R & z \<notin> Rep_preal R & |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1051 |
z + v \<in> Rep_preal S & x = u + v" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1052 |
proof - |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1053 |
have xpos: "0<x" by (rule preal_imp_pos [OF Rep_preal x]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1054 |
from lemma_sum_mem_Rep_preal_ex [OF x] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1055 |
obtain e where epos: "0 < e" and xe: "x + e \<in> Rep_preal S" by blast |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1056 |
from Gleason9_34 [OF Rep_preal epos] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1057 |
obtain r where r: "r \<in> Rep_preal R" and notin: "r + e \<notin> Rep_preal R" .. |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1058 |
with x xnot xpos have rless: "r < x" by (blast intro: not_in_Rep_preal_ub) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1059 |
from add_eq_exists [of r x] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1060 |
obtain y where eq: "x = r+y" by auto |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1061 |
show ?thesis |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1062 |
proof (intro exI conjI) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1063 |
show "r \<in> Rep_preal R" by (rule r) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1064 |
show "r + e \<notin> Rep_preal R" by (rule notin) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1065 |
show "r + e + y \<in> Rep_preal S" using xe eq by (simp add: add_ac) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1066 |
show "x = r + y" by (simp add: eq) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1067 |
show "0 < r + e" using epos preal_imp_pos [OF Rep_preal r] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1068 |
by simp |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1069 |
show "0 < y" using rless eq by arith |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1070 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1071 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1072 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1073 |
lemma less_add_left_le2: "R < (S::preal) ==> S \<le> R + (S-R)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1074 |
apply (auto simp add: preal_le_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1075 |
apply (case_tac "x \<in> Rep_preal R") |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1076 |
apply (cut_tac Rep_preal_self_subset [of R], force) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1077 |
apply (auto simp add: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_diff_iff) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1078 |
apply (blast dest: less_add_left_lemma2) |
14335 | 1079 |
done |
1080 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1081 |
lemma less_add_left: "R < (S::preal) ==> R + (S-R) = S" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1082 |
by (blast intro: preal_le_anti_sym [OF less_add_left_le1 less_add_left_le2]) |
14335 | 1083 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1084 |
lemma less_add_left_Ex: "R < (S::preal) ==> \<exists>D. R + D = S" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1085 |
by (fast dest: less_add_left) |
14335 | 1086 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1087 |
lemma preal_add_less2_mono1: "R < (S::preal) ==> R + T < S + T" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1088 |
apply (auto dest!: less_add_left_Ex simp add: preal_add_assoc) |
14335 | 1089 |
apply (rule_tac y1 = D in preal_add_commute [THEN subst]) |
1090 |
apply (auto intro: preal_self_less_add_left simp add: preal_add_assoc [symmetric]) |
|
1091 |
done |
|
1092 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1093 |
lemma preal_add_less2_mono2: "R < (S::preal) ==> T + R < T + S" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1094 |
by (auto intro: preal_add_less2_mono1 simp add: preal_add_commute [of T]) |
14335 | 1095 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1096 |
lemma preal_add_right_less_cancel: "R + T < S + T ==> R < (S::preal)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1097 |
apply (insert linorder_less_linear [of R S], auto) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1098 |
apply (drule_tac R = S and T = T in preal_add_less2_mono1) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1099 |
apply (blast dest: order_less_trans) |
14335 | 1100 |
done |
1101 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1102 |
lemma preal_add_left_less_cancel: "T + R < T + S ==> R < (S::preal)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1103 |
by (auto elim: preal_add_right_less_cancel simp add: preal_add_commute [of T]) |
14335 | 1104 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1105 |
lemma preal_add_less_cancel_right: "((R::preal) + T < S + T) = (R < S)" |
14335 | 1106 |
by (blast intro: preal_add_less2_mono1 preal_add_right_less_cancel) |
1107 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1108 |
lemma preal_add_less_cancel_left: "(T + (R::preal) < T + S) = (R < S)" |
14335 | 1109 |
by (blast intro: preal_add_less2_mono2 preal_add_left_less_cancel) |
1110 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1111 |
lemma preal_add_le_cancel_right: "((R::preal) + T \<le> S + T) = (R \<le> S)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1112 |
by (simp add: linorder_not_less [symmetric] preal_add_less_cancel_right) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1113 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1114 |
lemma preal_add_le_cancel_left: "(T + (R::preal) \<le> T + S) = (R \<le> S)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1115 |
by (simp add: linorder_not_less [symmetric] preal_add_less_cancel_left) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1116 |
|
14335 | 1117 |
lemma preal_add_less_mono: |
1118 |
"[| x1 < y1; x2 < y2 |] ==> x1 + x2 < y1 + (y2::preal)" |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1119 |
apply (auto dest!: less_add_left_Ex simp add: preal_add_ac) |
14335 | 1120 |
apply (rule preal_add_assoc [THEN subst]) |
1121 |
apply (rule preal_self_less_add_right) |
|
1122 |
done |
|
1123 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1124 |
lemma preal_add_right_cancel: "(R::preal) + T = S + T ==> R = S" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1125 |
apply (insert linorder_less_linear [of R S], safe) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1126 |
apply (drule_tac [!] T = T in preal_add_less2_mono1, auto) |
14335 | 1127 |
done |
1128 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1129 |
lemma preal_add_left_cancel: "C + A = C + B ==> A = (B::preal)" |
14335 | 1130 |
by (auto intro: preal_add_right_cancel simp add: preal_add_commute) |
1131 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1132 |
lemma preal_add_left_cancel_iff: "(C + A = C + B) = ((A::preal) = B)" |
14335 | 1133 |
by (fast intro: preal_add_left_cancel) |
1134 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1135 |
lemma preal_add_right_cancel_iff: "(A + C = B + C) = ((A::preal) = B)" |
14335 | 1136 |
by (fast intro: preal_add_right_cancel) |
1137 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1138 |
lemmas preal_cancels = |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1139 |
preal_add_less_cancel_right preal_add_less_cancel_left |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1140 |
preal_add_le_cancel_right preal_add_le_cancel_left |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1141 |
preal_add_left_cancel_iff preal_add_right_cancel_iff |
14335 | 1142 |
|
1143 |
||
1144 |
subsection{*Completeness of type @{typ preal}*} |
|
1145 |
||
1146 |
text{*Prove that supremum is a cut*} |
|
1147 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1148 |
text{*Part 1 of Dedekind sections definition*} |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1149 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1150 |
lemma preal_sup_set_not_empty: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1151 |
"P \<noteq> {} ==> {} \<subset> (\<Union>X \<in> P. Rep_preal(X))" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1152 |
apply auto |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1153 |
apply (cut_tac X = x in mem_Rep_preal_Ex, auto) |
14335 | 1154 |
done |
1155 |
||
1156 |
||
1157 |
text{*Part 2 of Dedekind sections definition*} |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1158 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1159 |
lemma preal_sup_not_exists: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1160 |
"\<forall>X \<in> P. X \<le> Y ==> \<exists>q. 0 < q & q \<notin> (\<Union>X \<in> P. Rep_preal(X))" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1161 |
apply (cut_tac X = Y in Rep_preal_exists_bound) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1162 |
apply (auto simp add: preal_le_def) |
14335 | 1163 |
done |
1164 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1165 |
lemma preal_sup_set_not_rat_set: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1166 |
"\<forall>X \<in> P. X \<le> Y ==> (\<Union>X \<in> P. Rep_preal(X)) < {r. 0 < r}" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1167 |
apply (drule preal_sup_not_exists) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1168 |
apply (blast intro: preal_imp_pos [OF Rep_preal]) |
14335 | 1169 |
done |
1170 |
||
1171 |
text{*Part 3 of Dedekind sections definition*} |
|
1172 |
lemma preal_sup_set_lemma3: |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1173 |
"[|P \<noteq> {}; \<forall>X \<in> P. X \<le> Y; u \<in> (\<Union>X \<in> P. Rep_preal(X)); 0 < z; z < u|] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1174 |
==> z \<in> (\<Union>X \<in> P. Rep_preal(X))" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1175 |
by (auto elim: Rep_preal [THEN preal_downwards_closed]) |
14335 | 1176 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1177 |
text{*Part 4 of Dedekind sections definition*} |
14335 | 1178 |
lemma preal_sup_set_lemma4: |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1179 |
"[|P \<noteq> {}; \<forall>X \<in> P. X \<le> Y; y \<in> (\<Union>X \<in> P. Rep_preal(X)) |] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1180 |
==> \<exists>u \<in> (\<Union>X \<in> P. Rep_preal(X)). y < u" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1181 |
by (blast dest: Rep_preal [THEN preal_exists_greater]) |
14335 | 1182 |
|
1183 |
lemma preal_sup: |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1184 |
"[|P \<noteq> {}; \<forall>X \<in> P. X \<le> Y|] ==> (\<Union>X \<in> P. Rep_preal(X)) \<in> preal" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1185 |
apply (unfold preal_def cut_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1186 |
apply (blast intro!: preal_sup_set_not_empty preal_sup_set_not_rat_set |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1187 |
preal_sup_set_lemma3 preal_sup_set_lemma4) |
14335 | 1188 |
done |
1189 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1190 |
lemma preal_psup_le: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1191 |
"[| \<forall>X \<in> P. X \<le> Y; x \<in> P |] ==> x \<le> psup P" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1192 |
apply (simp (no_asm_simp) add: preal_le_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1193 |
apply (subgoal_tac "P \<noteq> {}") |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1194 |
apply (auto simp add: psup_def preal_sup) |
14335 | 1195 |
done |
1196 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1197 |
lemma psup_le_ub: "[| P \<noteq> {}; \<forall>X \<in> P. X \<le> Y |] ==> psup P \<le> Y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1198 |
apply (simp (no_asm_simp) add: preal_le_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1199 |
apply (simp add: psup_def preal_sup) |
14335 | 1200 |
apply (auto simp add: preal_le_def) |
1201 |
done |
|
1202 |
||
1203 |
text{*Supremum property*} |
|
1204 |
lemma preal_complete: |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1205 |
"[| P \<noteq> {}; \<forall>X \<in> P. X \<le> Y |] ==> (\<exists>X \<in> P. Z < X) = (Z < psup P)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1206 |
apply (simp add: preal_less_def psup_def preal_sup) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1207 |
apply (auto simp add: preal_le_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1208 |
apply (rename_tac U) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1209 |
apply (cut_tac x = U and y = Z in linorder_less_linear) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1210 |
apply (auto simp add: preal_less_def) |
14335 | 1211 |
done |
1212 |
||
1213 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1214 |
subsection{*The Embadding from @{typ rat} into @{typ preal}*} |
14335 | 1215 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1216 |
lemma preal_of_rat_add_lemma1: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1217 |
"[|x < y + z; 0 < x; 0 < y|] ==> x * y * inverse (y + z) < (y::rat)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1218 |
apply (frule_tac c = "y * inverse (y + z) " in mult_strict_right_mono) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1219 |
apply (simp add: zero_less_mult_iff) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1220 |
apply (simp add: mult_ac) |
14335 | 1221 |
done |
1222 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1223 |
lemma preal_of_rat_add_lemma2: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1224 |
assumes "u < x + y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1225 |
and "0 < x" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1226 |
and "0 < y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1227 |
and "0 < u" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1228 |
shows "\<exists>v w::rat. w < y & 0 < v & v < x & 0 < w & u = v + w" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1229 |
proof (intro exI conjI) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1230 |
show "u * x * inverse(x+y) < x" using prems |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1231 |
by (simp add: preal_of_rat_add_lemma1) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1232 |
show "u * y * inverse(x+y) < y" using prems |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1233 |
by (simp add: preal_of_rat_add_lemma1 add_commute [of x]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1234 |
show "0 < u * x * inverse (x + y)" using prems |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1235 |
by (simp add: zero_less_mult_iff) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1236 |
show "0 < u * y * inverse (x + y)" using prems |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1237 |
by (simp add: zero_less_mult_iff) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1238 |
show "u = u * x * inverse (x + y) + u * y * inverse (x + y)" using prems |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1239 |
by (simp add: left_distrib [symmetric] right_distrib [symmetric] mult_ac) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1240 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1241 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1242 |
lemma preal_of_rat_add: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1243 |
"[| 0 < x; 0 < y|] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1244 |
==> preal_of_rat ((x::rat) + y) = preal_of_rat x + preal_of_rat y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1245 |
apply (unfold preal_of_rat_def preal_add_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1246 |
apply (simp add: rat_mem_preal) |
14335 | 1247 |
apply (rule_tac f = Abs_preal in arg_cong) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1248 |
apply (auto simp add: add_set_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1249 |
apply (blast dest: preal_of_rat_add_lemma2) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1250 |
done |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1251 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1252 |
lemma preal_of_rat_mult_lemma1: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1253 |
"[|x < y; 0 < x; 0 < z|] ==> x * z * inverse y < (z::rat)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1254 |
apply (frule_tac c = "z * inverse y" in mult_strict_right_mono) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1255 |
apply (simp add: zero_less_mult_iff) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1256 |
apply (subgoal_tac "y * (z * inverse y) = z * (y * inverse y)") |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1257 |
apply (simp_all add: mult_ac) |
14335 | 1258 |
done |
1259 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1260 |
lemma preal_of_rat_mult_lemma2: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1261 |
assumes xless: "x < y * z" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1262 |
and xpos: "0 < x" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1263 |
and ypos: "0 < y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1264 |
shows "x * z * inverse y * inverse z < (z::rat)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1265 |
proof - |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1266 |
have "0 < y * z" using prems by simp |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1267 |
hence zpos: "0 < z" using prems by (simp add: zero_less_mult_iff) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1268 |
have "x * z * inverse y * inverse z = x * inverse y * (z * inverse z)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1269 |
by (simp add: mult_ac) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1270 |
also have "... = x/y" using zpos |
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14387
diff
changeset
|
1271 |
by (simp add: divide_inverse) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1272 |
also have "... < z" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1273 |
by (simp add: pos_divide_less_eq [OF ypos] mult_commute) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1274 |
finally show ?thesis . |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1275 |
qed |
14335 | 1276 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1277 |
lemma preal_of_rat_mult_lemma3: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1278 |
assumes uless: "u < x * y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1279 |
and "0 < x" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1280 |
and "0 < y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1281 |
and "0 < u" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1282 |
shows "\<exists>v w::rat. v < x & w < y & 0 < v & 0 < w & u = v * w" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1283 |
proof - |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1284 |
from dense [OF uless] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1285 |
obtain r where "u < r" "r < x * y" by blast |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1286 |
thus ?thesis |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1287 |
proof (intro exI conjI) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1288 |
show "u * x * inverse r < x" using prems |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1289 |
by (simp add: preal_of_rat_mult_lemma1) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1290 |
show "r * y * inverse x * inverse y < y" using prems |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1291 |
by (simp add: preal_of_rat_mult_lemma2) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1292 |
show "0 < u * x * inverse r" using prems |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1293 |
by (simp add: zero_less_mult_iff) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1294 |
show "0 < r * y * inverse x * inverse y" using prems |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1295 |
by (simp add: zero_less_mult_iff) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1296 |
have "u * x * inverse r * (r * y * inverse x * inverse y) = |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1297 |
u * (r * inverse r) * (x * inverse x) * (y * inverse y)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1298 |
by (simp only: mult_ac) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1299 |
thus "u = u * x * inverse r * (r * y * inverse x * inverse y)" using prems |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1300 |
by simp |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1301 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1302 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1303 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1304 |
lemma preal_of_rat_mult: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1305 |
"[| 0 < x; 0 < y|] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1306 |
==> preal_of_rat ((x::rat) * y) = preal_of_rat x * preal_of_rat y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1307 |
apply (unfold preal_of_rat_def preal_mult_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1308 |
apply (simp add: rat_mem_preal) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1309 |
apply (rule_tac f = Abs_preal in arg_cong) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1310 |
apply (auto simp add: zero_less_mult_iff mult_strict_mono mult_set_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1311 |
apply (blast dest: preal_of_rat_mult_lemma3) |
14335 | 1312 |
done |
1313 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1314 |
lemma preal_of_rat_less_iff: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1315 |
"[| 0 < x; 0 < y|] ==> (preal_of_rat x < preal_of_rat y) = (x < y)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1316 |
by (force simp add: preal_of_rat_def preal_less_def rat_mem_preal) |
14335 | 1317 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1318 |
lemma preal_of_rat_le_iff: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1319 |
"[| 0 < x; 0 < y|] ==> (preal_of_rat x \<le> preal_of_rat y) = (x \<le> y)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1320 |
by (simp add: preal_of_rat_less_iff linorder_not_less [symmetric]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1321 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1322 |
lemma preal_of_rat_eq_iff: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1323 |
"[| 0 < x; 0 < y|] ==> (preal_of_rat x = preal_of_rat y) = (x = y)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1324 |
by (simp add: preal_of_rat_le_iff order_eq_iff) |
14335 | 1325 |
|
1326 |
||
1327 |
ML |
|
1328 |
{* |
|
1329 |
val inj_on_Abs_preal = thm"inj_on_Abs_preal"; |
|
1330 |
val inj_Rep_preal = thm"inj_Rep_preal"; |
|
1331 |
val mem_Rep_preal_Ex = thm"mem_Rep_preal_Ex"; |
|
1332 |
val preal_add_commute = thm"preal_add_commute"; |
|
1333 |
val preal_add_assoc = thm"preal_add_assoc"; |
|
1334 |
val preal_add_left_commute = thm"preal_add_left_commute"; |
|
1335 |
val preal_mult_commute = thm"preal_mult_commute"; |
|
1336 |
val preal_mult_assoc = thm"preal_mult_assoc"; |
|
1337 |
val preal_mult_left_commute = thm"preal_mult_left_commute"; |
|
1338 |
val preal_add_mult_distrib2 = thm"preal_add_mult_distrib2"; |
|
1339 |
val preal_add_mult_distrib = thm"preal_add_mult_distrib"; |
|
1340 |
val preal_self_less_add_left = thm"preal_self_less_add_left"; |
|
1341 |
val preal_self_less_add_right = thm"preal_self_less_add_right"; |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1342 |
val less_add_left = thm"less_add_left"; |
14335 | 1343 |
val preal_add_less2_mono1 = thm"preal_add_less2_mono1"; |
1344 |
val preal_add_less2_mono2 = thm"preal_add_less2_mono2"; |
|
1345 |
val preal_add_right_less_cancel = thm"preal_add_right_less_cancel"; |
|
1346 |
val preal_add_left_less_cancel = thm"preal_add_left_less_cancel"; |
|
1347 |
val preal_add_right_cancel = thm"preal_add_right_cancel"; |
|
1348 |
val preal_add_left_cancel = thm"preal_add_left_cancel"; |
|
1349 |
val preal_add_left_cancel_iff = thm"preal_add_left_cancel_iff"; |
|
1350 |
val preal_add_right_cancel_iff = thm"preal_add_right_cancel_iff"; |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1351 |
val preal_psup_le = thm"preal_psup_le"; |
14335 | 1352 |
val psup_le_ub = thm"psup_le_ub"; |
1353 |
val preal_complete = thm"preal_complete"; |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1354 |
val preal_of_rat_add = thm"preal_of_rat_add"; |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14335
diff
changeset
|
1355 |
val preal_of_rat_mult = thm"preal_of_rat_mult"; |
14335 | 1356 |
|
1357 |
val preal_add_ac = thms"preal_add_ac"; |
|
1358 |
val preal_mult_ac = thms"preal_mult_ac"; |
|
1359 |
*} |
|
1360 |
||
5078 | 1361 |
end |