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