| author | wenzelm | 
| Sun, 28 Apr 2019 22:22:29 +0200 | |
| changeset 70207 | 511352b4d5d3 | 
| parent 70201 | 2e496190039d | 
| child 73648 | 1bd3463e30b8 | 
| permissions | -rw-r--r-- | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1  | 
section \<open>The Reals as Dedekind Sections of Positive Rationals\<close>  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
2  | 
|
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
3  | 
text \<open>Fundamentals of Abstract Analysis [Gleason- p. 121] provides some of the definitions.\<close>  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
4  | 
|
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
5  | 
(* Title: HOL/ex/Dedekind_Real.thy  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
6  | 
Author: Jacques D. Fleuriot, University of Cambridge  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
7  | 
Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4; 2019  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
8  | 
*)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
9  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
10  | 
theory Dedekind_Real  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
11  | 
imports Complex_Main  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
12  | 
begin  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
13  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
14  | 
text\<open>Could be moved to \<open>Groups\<close>\<close>  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
15  | 
lemma add_eq_exists: "\<exists>x. a+x = (b::'a::ab_group_add)"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
16  | 
by (rule_tac x="b-a" in exI, simp)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
17  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
18  | 
subsection \<open>Dedekind cuts or sections\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
19  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
20  | 
definition  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
21  | 
cut :: "rat set \<Rightarrow> bool" where  | 
| 70200 | 22  | 
  "cut A \<equiv> {} \<subset> A \<and> A \<subset> {0<..} \<and>
 | 
23  | 
(\<forall>y \<in> A. ((\<forall>z. 0<z \<and> z < y \<longrightarrow> z \<in> A) \<and> (\<exists>u \<in> A. y < u)))"  | 
|
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
24  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
25  | 
lemma cut_of_rat:  | 
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
26  | 
  assumes q: "0 < q" shows "cut {r::rat. 0 < r \<and> r < q}" (is "cut ?A")
 | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
27  | 
proof -  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
28  | 
  from q have pos: "?A \<subset> {0<..}" by force
 | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
29  | 
  have nonempty: "{} \<subset> ?A"
 | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
30  | 
proof  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
31  | 
    show "{} \<subseteq> ?A" by simp
 | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
32  | 
    show "{} \<noteq> ?A"
 | 
| 70200 | 33  | 
using field_lbound_gt_zero q by auto  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
34  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
35  | 
show ?thesis  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
36  | 
by (simp add: cut_def pos nonempty,  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
37  | 
blast dest: dense intro: order_less_trans)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
38  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
39  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
40  | 
|
| 59814 | 41  | 
typedef preal = "Collect cut"  | 
42  | 
by (blast intro: cut_of_rat [OF zero_less_one])  | 
|
| 
45694
 
4a8743618257
prefer typedef without extra definition and alternative name;
 
wenzelm 
parents: 
41541 
diff
changeset
 | 
43  | 
|
| 59814 | 44  | 
lemma Abs_preal_induct [induct type: preal]:  | 
45  | 
"(\<And>x. cut x \<Longrightarrow> P (Abs_preal x)) \<Longrightarrow> P x"  | 
|
46  | 
using Abs_preal_induct [of P x] by simp  | 
|
47  | 
||
| 70200 | 48  | 
lemma cut_Rep_preal [simp]: "cut (Rep_preal x)"  | 
| 59814 | 49  | 
using Rep_preal [of x] by simp  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
50  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
51  | 
definition  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
52  | 
psup :: "preal set \<Rightarrow> preal" where  | 
| 37765 | 53  | 
"psup P = Abs_preal (\<Union>X \<in> P. Rep_preal X)"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
54  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
55  | 
definition  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
56  | 
add_set :: "[rat set,rat set] \<Rightarrow> rat set" where  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
57  | 
  "add_set A B = {w. \<exists>x \<in> A. \<exists>y \<in> B. w = x + y}"
 | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
58  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
59  | 
definition  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
60  | 
diff_set :: "[rat set,rat set] \<Rightarrow> rat set" where  | 
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
61  | 
  "diff_set A B = {w. \<exists>x. 0 < w \<and> 0 < x \<and> x \<notin> B \<and> x + w \<in> A}"
 | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
62  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
63  | 
definition  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
64  | 
mult_set :: "[rat set,rat set] \<Rightarrow> rat set" where  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
65  | 
  "mult_set A B = {w. \<exists>x \<in> A. \<exists>y \<in> B. w = x * y}"
 | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
66  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
67  | 
definition  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
68  | 
inverse_set :: "rat set \<Rightarrow> rat set" where  | 
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
69  | 
  "inverse_set A \<equiv> {x. \<exists>y. 0 < x \<and> x < y \<and> inverse y \<notin> A}"
 | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
70  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
71  | 
instantiation preal :: "{ord, plus, minus, times, inverse, one}"
 | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
72  | 
begin  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
73  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
74  | 
definition  | 
| 37765 | 75  | 
preal_less_def:  | 
| 70201 | 76  | 
"r < s \<equiv> Rep_preal r < Rep_preal s"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
77  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
78  | 
definition  | 
| 37765 | 79  | 
preal_le_def:  | 
| 70201 | 80  | 
"r \<le> s \<equiv> Rep_preal r \<subseteq> Rep_preal s"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
81  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
82  | 
definition  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
83  | 
preal_add_def:  | 
| 70201 | 84  | 
"r + s \<equiv> Abs_preal (add_set (Rep_preal r) (Rep_preal s))"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
85  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
86  | 
definition  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
87  | 
preal_diff_def:  | 
| 70201 | 88  | 
"r - s \<equiv> Abs_preal (diff_set (Rep_preal r) (Rep_preal s))"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
89  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
90  | 
definition  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
91  | 
preal_mult_def:  | 
| 70201 | 92  | 
"r * s \<equiv> Abs_preal (mult_set (Rep_preal r) (Rep_preal s))"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
93  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
94  | 
definition  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
95  | 
preal_inverse_def:  | 
| 70201 | 96  | 
"inverse r \<equiv> Abs_preal (inverse_set (Rep_preal r))"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
97  | 
|
| 70201 | 98  | 
definition "r div s = r * inverse (s::preal)"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
99  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
100  | 
definition  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
101  | 
preal_one_def:  | 
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
102  | 
    "1 \<equiv> Abs_preal {x. 0 < x \<and> x < 1}"
 | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
103  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
104  | 
instance ..  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
105  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
106  | 
end  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
107  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
108  | 
|
| 61343 | 109  | 
text\<open>Reduces equality on abstractions to equality on representatives\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
110  | 
declare Abs_preal_inject [simp]  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
111  | 
declare Abs_preal_inverse [simp]  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
112  | 
|
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
113  | 
lemma rat_mem_preal: "0 < q \<Longrightarrow> cut {r::rat. 0 < r \<and> r < q}"
 | 
| 59814 | 114  | 
by (simp add: cut_of_rat)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
115  | 
|
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
116  | 
lemma preal_nonempty: "cut A \<Longrightarrow> \<exists>x\<in>A. 0 < x"  | 
| 59814 | 117  | 
unfolding cut_def [abs_def] by blast  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
118  | 
|
| 59814 | 119  | 
lemma preal_Ex_mem: "cut A \<Longrightarrow> \<exists>x. x \<in> A"  | 
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
120  | 
using preal_nonempty by blast  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
121  | 
|
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
122  | 
lemma preal_exists_bound: "cut A \<Longrightarrow> \<exists>x. 0 < x \<and> x \<notin> A"  | 
| 
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
123  | 
using Dedekind_Real.cut_def by fastforce  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
124  | 
|
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
125  | 
lemma preal_exists_greater: "\<lbrakk>cut A; y \<in> A\<rbrakk> \<Longrightarrow> \<exists>u \<in> A. y < u"  | 
| 59814 | 126  | 
unfolding cut_def [abs_def] by blast  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
127  | 
|
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
128  | 
lemma preal_downwards_closed: "\<lbrakk>cut A; y \<in> A; 0 < z; z < y\<rbrakk> \<Longrightarrow> z \<in> A"  | 
| 59814 | 129  | 
unfolding cut_def [abs_def] by blast  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
130  | 
|
| 61343 | 131  | 
text\<open>Relaxing the final premise\<close>  | 
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
132  | 
lemma preal_downwards_closed': "\<lbrakk>cut A; y \<in> A; 0 < z; z \<le> y\<rbrakk> \<Longrightarrow> z \<in> A"  | 
| 
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
133  | 
using less_eq_rat_def preal_downwards_closed by blast  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
134  | 
|
| 61343 | 135  | 
text\<open>A positive fraction not in a positive real is an upper bound.  | 
136  | 
Gleason p. 122 - Remark (1)\<close>  | 
|
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
137  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
138  | 
lemma not_in_preal_ub:  | 
| 59814 | 139  | 
assumes A: "cut A"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
140  | 
and notx: "x \<notin> A"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
141  | 
and y: "y \<in> A"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
142  | 
and pos: "0 < x"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
143  | 
shows "y < x"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
144  | 
proof (cases rule: linorder_cases)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
145  | 
assume "x<y"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
146  | 
with notx show ?thesis  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
147  | 
by (simp add: preal_downwards_closed [OF A y] pos)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
148  | 
next  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
149  | 
assume "x=y"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
150  | 
with notx and y show ?thesis by simp  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
151  | 
next  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
152  | 
assume "y<x"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
153  | 
thus ?thesis .  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
154  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
155  | 
|
| 69597 | 156  | 
text \<open>preal lemmas instantiated to \<^term>\<open>Rep_preal X\<close>\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
157  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
158  | 
lemma mem_Rep_preal_Ex: "\<exists>x. x \<in> Rep_preal X"  | 
| 59814 | 159  | 
thm preal_Ex_mem  | 
| 70200 | 160  | 
by (rule preal_Ex_mem [OF cut_Rep_preal])  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
161  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
162  | 
lemma Rep_preal_exists_bound: "\<exists>x>0. x \<notin> Rep_preal X"  | 
| 70200 | 163  | 
by (rule preal_exists_bound [OF cut_Rep_preal])  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
164  | 
|
| 70200 | 165  | 
lemmas not_in_Rep_preal_ub = not_in_preal_ub [OF cut_Rep_preal]  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
166  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
167  | 
|
| 61343 | 168  | 
subsection\<open>Properties of Ordering\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
169  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
170  | 
instance preal :: order  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
171  | 
proof  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
172  | 
fix w :: preal  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
173  | 
show "w \<le> w" by (simp add: preal_le_def)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
174  | 
next  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
175  | 
fix i j k :: preal  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
176  | 
assume "i \<le> j" and "j \<le> k"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
177  | 
then show "i \<le> k" by (simp add: preal_le_def)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
178  | 
next  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
179  | 
fix z w :: preal  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
180  | 
assume "z \<le> w" and "w \<le> z"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
181  | 
then show "z = w" by (simp add: preal_le_def Rep_preal_inject)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
182  | 
next  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
183  | 
fix z w :: preal  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
184  | 
show "z < w \<longleftrightarrow> z \<le> w \<and> \<not> w \<le> z"  | 
| 70200 | 185  | 
by (auto simp: preal_le_def preal_less_def Rep_preal_inject)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
186  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
187  | 
|
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
188  | 
lemma preal_imp_pos: "\<lbrakk>cut A; r \<in> A\<rbrakk> \<Longrightarrow> 0 < r"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
189  | 
by (auto simp: cut_def)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
190  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
191  | 
instance preal :: linorder  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
192  | 
proof  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
193  | 
fix x y :: preal  | 
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
194  | 
show "x \<le> y \<or> y \<le> x"  | 
| 
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
195  | 
unfolding preal_le_def  | 
| 70200 | 196  | 
by (meson cut_Rep_preal not_in_preal_ub preal_downwards_closed preal_imp_pos subsetI)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
197  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
198  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
199  | 
instantiation preal :: distrib_lattice  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
200  | 
begin  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
201  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
202  | 
definition  | 
| 61076 | 203  | 
"(inf :: preal \<Rightarrow> preal \<Rightarrow> preal) = min"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
204  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
205  | 
definition  | 
| 61076 | 206  | 
"(sup :: preal \<Rightarrow> preal \<Rightarrow> preal) = max"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
207  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
208  | 
instance  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
209  | 
by intro_classes  | 
| 70200 | 210  | 
(auto simp: inf_preal_def sup_preal_def max_min_distrib2)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
211  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
212  | 
end  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
213  | 
|
| 61343 | 214  | 
subsection\<open>Properties of Addition\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
215  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
216  | 
lemma preal_add_commute: "(x::preal) + y = y + x"  | 
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
217  | 
unfolding preal_add_def add_set_def  | 
| 
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
218  | 
by (metis (no_types, hide_lams) add.commute)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
219  | 
|
| 61343 | 220  | 
text\<open>Lemmas for proving that addition of two positive reals gives  | 
221  | 
a positive real\<close>  | 
|
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
222  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
223  | 
lemma mem_add_set:  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
224  | 
assumes "cut A" "cut B"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
225  | 
shows "cut (add_set A B)"  | 
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
226  | 
proof -  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
227  | 
  have "{} \<subset> add_set A B"
 | 
| 70200 | 228  | 
using assms by (force simp: add_set_def dest: preal_nonempty)  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
229  | 
moreover  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
230  | 
obtain q where "q > 0" "q \<notin> add_set A B"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
231  | 
proof -  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
232  | 
obtain a b where "a > 0" "a \<notin> A" "b > 0" "b \<notin> B" "\<And>x. x \<in> A \<Longrightarrow> x < a" "\<And>y. y \<in> B \<Longrightarrow> y < b"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
233  | 
by (meson assms preal_exists_bound not_in_preal_ub)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
234  | 
with assms have "a+b \<notin> add_set A B"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
235  | 
by (fastforce simp add: add_set_def)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
236  | 
then show thesis  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
237  | 
using \<open>0 < a\<close> \<open>0 < b\<close> add_pos_pos that by blast  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
238  | 
qed  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
239  | 
  then have "add_set A B \<subset> {0<..}"
 | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
240  | 
unfolding add_set_def  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
241  | 
using preal_imp_pos [OF \<open>cut A\<close>] preal_imp_pos [OF \<open>cut B\<close>] by fastforce  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
242  | 
moreover have "z \<in> add_set A B"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
243  | 
if u: "u \<in> add_set A B" and "0 < z" "z < u" for u z  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
244  | 
using u unfolding add_set_def  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
245  | 
proof (clarify)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
246  | 
fix x::rat and y::rat  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
247  | 
assume ueq: "u = x + y" and x: "x \<in> A" and y:"y \<in> B"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
248  | 
have xpos [simp]: "x > 0" and ypos [simp]: "y > 0"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
249  | 
using assms preal_imp_pos x y by blast+  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
250  | 
have xypos [simp]: "x+y > 0" by (simp add: pos_add_strict)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
251  | 
let ?f = "z/(x+y)"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
252  | 
have fless: "?f < 1"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
253  | 
using divide_less_eq_1_pos \<open>z < u\<close> ueq xypos by blast  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
254  | 
show "\<exists>x' \<in> A. \<exists>y'\<in>B. z = x' + y'"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
255  | 
proof (intro bexI)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
256  | 
show "z = x*?f + y*?f"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
257  | 
by (simp add: distrib_right [symmetric] divide_inverse ac_simps order_less_imp_not_eq2)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
258  | 
next  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
259  | 
show "y * ?f \<in> B"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
260  | 
proof (rule preal_downwards_closed [OF \<open>cut B\<close> y])  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
261  | 
show "0 < y * ?f"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
262  | 
by (simp add: \<open>0 < z\<close>)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
263  | 
next  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
264  | 
show "y * ?f < y"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
265  | 
by (insert mult_strict_left_mono [OF fless ypos], simp)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
266  | 
qed  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
267  | 
next  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
268  | 
show "x * ?f \<in> A"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
269  | 
proof (rule preal_downwards_closed [OF \<open>cut A\<close> x])  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
270  | 
show "0 < x * ?f"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
271  | 
by (simp add: \<open>0 < z\<close>)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
272  | 
next  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
273  | 
show "x * ?f < x"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
274  | 
by (insert mult_strict_left_mono [OF fless xpos], simp)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
275  | 
qed  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
276  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
277  | 
qed  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
278  | 
moreover  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
279  | 
have "\<And>y. y \<in> add_set A B \<Longrightarrow> \<exists>u \<in> add_set A B. y < u"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
280  | 
unfolding add_set_def using preal_exists_greater assms by fastforce  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
281  | 
ultimately show ?thesis  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
282  | 
by (simp add: Dedekind_Real.cut_def)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
283  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
284  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
285  | 
lemma preal_add_assoc: "((x::preal) + y) + z = x + (y + z)"  | 
| 70201 | 286  | 
apply (simp add: preal_add_def mem_add_set)  | 
| 70200 | 287  | 
apply (force simp: add_set_def ac_simps)  | 
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
288  | 
done  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
289  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
290  | 
instance preal :: ab_semigroup_add  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
291  | 
proof  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
292  | 
fix a b c :: preal  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
293  | 
show "(a + b) + c = a + (b + c)" by (rule preal_add_assoc)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
294  | 
show "a + b = b + a" by (rule preal_add_commute)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
295  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
296  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
297  | 
|
| 61343 | 298  | 
subsection\<open>Properties of Multiplication\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
299  | 
|
| 61343 | 300  | 
text\<open>Proofs essentially same as for addition\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
301  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
302  | 
lemma preal_mult_commute: "(x::preal) * y = y * x"  | 
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
303  | 
unfolding preal_mult_def mult_set_def  | 
| 
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
304  | 
by (metis (no_types, hide_lams) mult.commute)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
305  | 
|
| 61343 | 306  | 
text\<open>Multiplication of two positive reals gives a positive real.\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
307  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
308  | 
lemma mem_mult_set:  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
309  | 
assumes "cut A" "cut B"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
310  | 
shows "cut (mult_set A B)"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
311  | 
proof -  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
312  | 
  have "{} \<subset> mult_set A B"
 | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
313  | 
using assms  | 
| 70200 | 314  | 
by (force simp: mult_set_def dest: preal_nonempty)  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
315  | 
moreover  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
316  | 
obtain q where "q > 0" "q \<notin> mult_set A B"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
317  | 
proof -  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
318  | 
obtain x y where x [simp]: "0 < x" "x \<notin> A" and y [simp]: "0 < y" "y \<notin> B"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
319  | 
using preal_exists_bound assms by blast  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
320  | 
show thesis  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
321  | 
proof  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
322  | 
show "0 < x*y" by simp  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
323  | 
show "x * y \<notin> mult_set A B"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
324  | 
proof -  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
325  | 
          {
 | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
326  | 
fix u::rat and v::rat  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
327  | 
assume u: "u \<in> A" and v: "v \<in> B" and xy: "x*y = u*v"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
328  | 
moreover have "u<x" and "v<y" using assms x y u v by (blast dest: not_in_preal_ub)+  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
329  | 
moreover have "0\<le>v"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
330  | 
using less_imp_le preal_imp_pos assms x y u v by blast  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
331  | 
moreover have "u*v < x*y"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
332  | 
using assms x \<open>u < x\<close> \<open>v < y\<close> \<open>0 \<le> v\<close> by (blast intro: mult_strict_mono)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
333  | 
ultimately have False by force  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
334  | 
}  | 
| 70200 | 335  | 
thus ?thesis by (auto simp: mult_set_def)  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
336  | 
qed  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
337  | 
qed  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
338  | 
qed  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
339  | 
  then have "mult_set A B \<subset> {0<..}"
 | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
340  | 
unfolding mult_set_def  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
341  | 
using preal_imp_pos [OF \<open>cut A\<close>] preal_imp_pos [OF \<open>cut B\<close>] by fastforce  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
342  | 
moreover have "z \<in> mult_set A B"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
343  | 
if u: "u \<in> mult_set A B" and "0 < z" "z < u" for u z  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
344  | 
using u unfolding mult_set_def  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
345  | 
proof (clarify)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
346  | 
fix x::rat and y::rat  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
347  | 
assume ueq: "u = x * y" and x: "x \<in> A" and y: "y \<in> B"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
348  | 
have [simp]: "y > 0"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
349  | 
using \<open>cut B\<close> preal_imp_pos y by blast  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
350  | 
show "\<exists>x' \<in> A. \<exists>y' \<in> B. z = x' * y'"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
351  | 
proof  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
352  | 
have "z = (z/y)*y"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
353  | 
by (simp add: divide_inverse mult.commute [of y] mult.assoc order_less_imp_not_eq2)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
354  | 
then show "\<exists>y'\<in>B. z = (z/y) * y'"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
355  | 
using y by blast  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
356  | 
next  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
357  | 
show "z/y \<in> A"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
358  | 
proof (rule preal_downwards_closed [OF \<open>cut A\<close> x])  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
359  | 
show "0 < z/y"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
360  | 
by (simp add: \<open>0 < z\<close>)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
361  | 
show "z/y < x"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
362  | 
using \<open>0 < y\<close> pos_divide_less_eq \<open>z < u\<close> ueq by blast  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
363  | 
qed  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
364  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
365  | 
qed  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
366  | 
moreover have "\<And>y. y \<in> mult_set A B \<Longrightarrow> \<exists>u \<in> mult_set A B. y < u"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
367  | 
apply (simp add: mult_set_def)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
368  | 
by (metis preal_exists_greater mult_strict_right_mono preal_imp_pos assms)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
369  | 
ultimately show ?thesis  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
370  | 
by (simp add: Dedekind_Real.cut_def)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
371  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
372  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
373  | 
lemma preal_mult_assoc: "((x::preal) * y) * z = x * (y * z)"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
374  | 
apply (simp add: preal_mult_def mem_mult_set Rep_preal)  | 
| 70200 | 375  | 
apply (force simp: mult_set_def ac_simps)  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
376  | 
done  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
377  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
378  | 
instance preal :: ab_semigroup_mult  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
379  | 
proof  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
380  | 
fix a b c :: preal  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
381  | 
show "(a * b) * c = a * (b * c)" by (rule preal_mult_assoc)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
382  | 
show "a * b = b * a" by (rule preal_mult_commute)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
383  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
384  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
385  | 
|
| 61343 | 386  | 
text\<open>Positive real 1 is the multiplicative identity element\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
387  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
388  | 
lemma preal_mult_1: "(1::preal) * z = z"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
389  | 
proof (induct z)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
390  | 
fix A :: "rat set"  | 
| 59814 | 391  | 
assume A: "cut A"  | 
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
392  | 
  have "{w. \<exists>u. 0 < u \<and> u < 1 \<and> (\<exists>v \<in> A. w = u * v)} = A" (is "?lhs = A")
 | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
393  | 
proof  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
394  | 
show "?lhs \<subseteq> A"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
395  | 
proof clarify  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
396  | 
fix x::rat and u::rat and v::rat  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
397  | 
assume upos: "0<u" and "u<1" and v: "v \<in> A"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
398  | 
have vpos: "0<v" by (rule preal_imp_pos [OF A v])  | 
| 61343 | 399  | 
hence "u*v < 1*v" by (simp only: mult_strict_right_mono upos \<open>u < 1\<close> v)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
400  | 
thus "u * v \<in> A"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
401  | 
by (force intro: preal_downwards_closed [OF A v] mult_pos_pos upos vpos)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
402  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
403  | 
next  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
404  | 
show "A \<subseteq> ?lhs"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
405  | 
proof clarify  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
406  | 
fix x::rat  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
407  | 
assume x: "x \<in> A"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
408  | 
have xpos: "0<x" by (rule preal_imp_pos [OF A x])  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
409  | 
from preal_exists_greater [OF A x]  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
410  | 
obtain v where v: "v \<in> A" and xlessv: "x < v" ..  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
411  | 
have vpos: "0<v" by (rule preal_imp_pos [OF A v])  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
412  | 
show "\<exists>u. 0 < u \<and> u < 1 \<and> (\<exists>v\<in>A. x = u * v)"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
413  | 
proof (intro exI conjI)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
414  | 
show "0 < x/v"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
415  | 
by (simp add: zero_less_divide_iff xpos vpos)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
416  | 
show "x / v < 1"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
417  | 
by (simp add: pos_divide_less_eq vpos xlessv)  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
418  | 
have "x = (x/v)*v"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
419  | 
by (simp add: divide_inverse mult.assoc vpos order_less_imp_not_eq2)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
420  | 
then show "\<exists>v'\<in>A. x = (x / v) * v'"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
421  | 
using v by blast  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
422  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
423  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
424  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
425  | 
thus "1 * Abs_preal A = Abs_preal A"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
426  | 
by (simp add: preal_one_def preal_mult_def mult_set_def rat_mem_preal A)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
427  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
428  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
429  | 
instance preal :: comm_monoid_mult  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
430  | 
by intro_classes (rule preal_mult_1)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
431  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
432  | 
|
| 61343 | 433  | 
subsection\<open>Distribution of Multiplication across Addition\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
434  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
435  | 
lemma mem_Rep_preal_add_iff:  | 
| 70201 | 436  | 
"(z \<in> Rep_preal(r+s)) = (\<exists>x \<in> Rep_preal r. \<exists>y \<in> Rep_preal s. z = x + y)"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
437  | 
apply (simp add: preal_add_def mem_add_set Rep_preal)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
438  | 
apply (simp add: add_set_def)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
439  | 
done  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
440  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
441  | 
lemma mem_Rep_preal_mult_iff:  | 
| 70201 | 442  | 
"(z \<in> Rep_preal(r*s)) = (\<exists>x \<in> Rep_preal r. \<exists>y \<in> Rep_preal s. z = x * y)"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
443  | 
apply (simp add: preal_mult_def mem_mult_set Rep_preal)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
444  | 
apply (simp add: mult_set_def)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
445  | 
done  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
446  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
447  | 
lemma distrib_subset1:  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
448  | 
"Rep_preal (w * (x + y)) \<subseteq> Rep_preal (w * x + w * y)"  | 
| 70200 | 449  | 
by (force simp: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_mult_iff distrib_left)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
450  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
451  | 
lemma preal_add_mult_distrib_mean:  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
452  | 
assumes a: "a \<in> Rep_preal w"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
453  | 
and b: "b \<in> Rep_preal w"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
454  | 
and d: "d \<in> Rep_preal x"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
455  | 
and e: "e \<in> Rep_preal y"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
456  | 
shows "\<exists>c \<in> Rep_preal w. a * d + b * e = c * (d + e)"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
457  | 
proof  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
458  | 
let ?c = "(a*d + b*e)/(d+e)"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
459  | 
have [simp]: "0<a" "0<b" "0<d" "0<e" "0<d+e"  | 
| 70200 | 460  | 
by (blast intro: preal_imp_pos [OF cut_Rep_preal] a b d e pos_add_strict)+  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
461  | 
have cpos: "0 < ?c"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
462  | 
by (simp add: zero_less_divide_iff zero_less_mult_iff pos_add_strict)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
463  | 
show "a * d + b * e = ?c * (d + e)"  | 
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
57492 
diff
changeset
 | 
464  | 
by (simp add: divide_inverse mult.assoc order_less_imp_not_eq2)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
465  | 
show "?c \<in> Rep_preal w"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
466  | 
proof (cases rule: linorder_le_cases)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
467  | 
assume "a \<le> b"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
468  | 
hence "?c \<le> b"  | 
| 
49962
 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 
webertj 
parents: 
49834 
diff
changeset
 | 
469  | 
by (simp add: pos_divide_le_eq distrib_left mult_right_mono  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
470  | 
order_less_imp_le)  | 
| 70200 | 471  | 
thus ?thesis by (rule preal_downwards_closed' [OF cut_Rep_preal b cpos])  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
472  | 
next  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
473  | 
assume "b \<le> a"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
474  | 
hence "?c \<le> a"  | 
| 
49962
 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 
webertj 
parents: 
49834 
diff
changeset
 | 
475  | 
by (simp add: pos_divide_le_eq distrib_left mult_right_mono  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
476  | 
order_less_imp_le)  | 
| 70200 | 477  | 
thus ?thesis by (rule preal_downwards_closed' [OF cut_Rep_preal a cpos])  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
478  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
479  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
480  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
481  | 
lemma distrib_subset2:  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
482  | 
"Rep_preal (w * x + w * y) \<subseteq> Rep_preal (w * (x + y))"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
483  | 
apply (clarsimp simp: mem_Rep_preal_add_iff mem_Rep_preal_mult_iff)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
484  | 
using mem_Rep_preal_add_iff preal_add_mult_distrib_mean by blast  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
485  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
486  | 
lemma preal_add_mult_distrib2: "(w * ((x::preal) + y)) = (w * x) + (w * y)"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
487  | 
by (metis Rep_preal_inverse distrib_subset1 distrib_subset2 subset_antisym)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
488  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
489  | 
lemma preal_add_mult_distrib: "(((x::preal) + y) * w) = (x * w) + (y * w)"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
490  | 
by (simp add: preal_mult_commute preal_add_mult_distrib2)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
491  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
492  | 
instance preal :: comm_semiring  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
493  | 
by intro_classes (rule preal_add_mult_distrib)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
494  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
495  | 
|
| 61343 | 496  | 
subsection\<open>Existence of Inverse, a Positive Real\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
497  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
498  | 
lemma mem_inverse_set:  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
499  | 
assumes "cut A" shows "cut (inverse_set A)"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
500  | 
proof -  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
501  | 
have "\<exists>x y. 0 < x \<and> x < y \<and> inverse y \<notin> A"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
502  | 
proof -  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
503  | 
from preal_exists_bound [OF \<open>cut A\<close>]  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
504  | 
obtain x where [simp]: "0<x" "x \<notin> A" by blast  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
505  | 
show ?thesis  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
506  | 
proof (intro exI conjI)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
507  | 
show "0 < inverse (x+1)"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
508  | 
by (simp add: order_less_trans [OF _ less_add_one])  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
509  | 
show "inverse(x+1) < inverse x"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
510  | 
by (simp add: less_imp_inverse_less less_add_one)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
511  | 
show "inverse (inverse x) \<notin> A"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
512  | 
by (simp add: order_less_imp_not_eq2)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
513  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
514  | 
qed  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
515  | 
  then have "{} \<subset> inverse_set A"
 | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
516  | 
using inverse_set_def by fastforce  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
517  | 
moreover obtain q where "q > 0" "q \<notin> inverse_set A"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
518  | 
proof -  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
519  | 
from preal_nonempty [OF \<open>cut A\<close>]  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
520  | 
obtain x where x: "x \<in> A" and xpos [simp]: "0<x" ..  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
521  | 
show ?thesis  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
522  | 
proof  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
523  | 
show "0 < inverse x" by simp  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
524  | 
show "inverse x \<notin> inverse_set A"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
525  | 
proof -  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
526  | 
        { fix y::rat 
 | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
527  | 
assume ygt: "inverse x < y"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
528  | 
have [simp]: "0 < y" by (simp add: order_less_trans [OF _ ygt])  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
529  | 
have iyless: "inverse y < x"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
530  | 
by (simp add: inverse_less_imp_less [of x] ygt)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
531  | 
have "inverse y \<in> A"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
532  | 
by (simp add: preal_downwards_closed [OF \<open>cut A\<close> x] iyless)}  | 
| 70200 | 533  | 
thus ?thesis by (auto simp: inverse_set_def)  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
534  | 
qed  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
535  | 
qed  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
536  | 
qed  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
537  | 
  moreover have "inverse_set A \<subset> {0<..}"
 | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
538  | 
using calculation inverse_set_def by blast  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
539  | 
moreover have "z \<in> inverse_set A"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
540  | 
if u: "u \<in> inverse_set A" and "0 < z" "z < u" for u z  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
541  | 
using u that less_trans unfolding inverse_set_def by auto  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
542  | 
moreover have "\<And>y. y \<in> inverse_set A \<Longrightarrow> \<exists>u \<in> inverse_set A. y < u"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
543  | 
by (simp add: inverse_set_def) (meson dense less_trans)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
544  | 
ultimately show ?thesis  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
545  | 
by (simp add: Dedekind_Real.cut_def)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
546  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
547  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
548  | 
|
| 61343 | 549  | 
subsection\<open>Gleason's Lemma 9-3.4, page 122\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
550  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
551  | 
lemma Gleason9_34_exists:  | 
| 59814 | 552  | 
assumes A: "cut A"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
553  | 
and "\<forall>x\<in>A. x + u \<in> A"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
554  | 
and "0 \<le> z"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
555  | 
shows "\<exists>b\<in>A. b + (of_int z) * u \<in> A"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
556  | 
proof (cases z rule: int_cases)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
557  | 
case (nonneg n)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
558  | 
show ?thesis  | 
| 41541 | 559  | 
proof (simp add: nonneg, induct n)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
560  | 
case 0  | 
| 41541 | 561  | 
from preal_nonempty [OF A]  | 
562  | 
show ?case by force  | 
|
563  | 
next  | 
|
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
564  | 
case (Suc k)  | 
| 41541 | 565  | 
then obtain b where b: "b \<in> A" "b + of_nat k * u \<in> A" ..  | 
566  | 
hence "b + of_int (int k)*u + u \<in> A" by (simp add: assms)  | 
|
| 70200 | 567  | 
thus ?case by (force simp: algebra_simps b)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
568  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
569  | 
next  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
570  | 
case (neg n)  | 
| 41541 | 571  | 
with assms show ?thesis by simp  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
572  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
573  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
574  | 
lemma Gleason9_34_contra:  | 
| 59814 | 575  | 
assumes A: "cut A"  | 
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
576  | 
shows "\<lbrakk>\<forall>x\<in>A. x + u \<in> A; 0 < u; 0 < y; y \<notin> A\<rbrakk> \<Longrightarrow> False"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
577  | 
proof (induct u, induct y)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
578  | 
fix a::int and b::int  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
579  | 
fix c::int and d::int  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
580  | 
assume bpos [simp]: "0 < b"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
581  | 
and dpos [simp]: "0 < d"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
582  | 
and closed: "\<forall>x\<in>A. x + (Fract c d) \<in> A"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
583  | 
and upos: "0 < Fract c d"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
584  | 
and ypos: "0 < Fract a b"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
585  | 
and notin: "Fract a b \<notin> A"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
586  | 
have cpos [simp]: "0 < c"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
587  | 
by (simp add: zero_less_Fract_iff [OF dpos, symmetric] upos)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
588  | 
have apos [simp]: "0 < a"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
589  | 
by (simp add: zero_less_Fract_iff [OF bpos, symmetric] ypos)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
590  | 
let ?k = "a*d"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
591  | 
have frle: "Fract a b \<le> Fract ?k 1 * (Fract c d)"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
592  | 
proof -  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
593  | 
have "?thesis = ((a * d * b * d) \<le> c * b * (a * d * b * d))"  | 
| 
57514
 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 
haftmann 
parents: 
57512 
diff
changeset
 | 
594  | 
by (simp add: order_less_imp_not_eq2 ac_simps)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
595  | 
moreover  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
596  | 
have "(1 * (a * d * b * d)) \<le> c * b * (a * d * b * d)"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
597  | 
by (rule mult_mono,  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
598  | 
simp_all add: int_one_le_iff_zero_less zero_less_mult_iff  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
599  | 
order_less_imp_le)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
600  | 
ultimately  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
601  | 
show ?thesis by simp  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
602  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
603  | 
have k: "0 \<le> ?k" by (simp add: order_less_imp_le zero_less_mult_iff)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
604  | 
from Gleason9_34_exists [OF A closed k]  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
605  | 
obtain z where z: "z \<in> A"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
606  | 
and mem: "z + of_int ?k * Fract c d \<in> A" ..  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
607  | 
have less: "z + of_int ?k * Fract c d < Fract a b"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
608  | 
by (rule not_in_preal_ub [OF A notin mem ypos])  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
609  | 
have "0<z" by (rule preal_imp_pos [OF A z])  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
610  | 
with frle and less show False by (simp add: Fract_of_int_eq)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
611  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
612  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
613  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
614  | 
lemma Gleason9_34:  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
615  | 
assumes "cut A" "0 < u"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
616  | 
shows "\<exists>r \<in> A. r + u \<notin> A"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
617  | 
using assms Gleason9_34_contra preal_exists_bound by blast  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
618  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
619  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
620  | 
|
| 61343 | 621  | 
subsection\<open>Gleason's Lemma 9-3.6\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
622  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
623  | 
lemma lemma_gleason9_36:  | 
| 59814 | 624  | 
assumes A: "cut A"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
625  | 
and x: "1 < x"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
626  | 
shows "\<exists>r \<in> A. r*x \<notin> A"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
627  | 
proof -  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
628  | 
from preal_nonempty [OF A]  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
629  | 
obtain y where y: "y \<in> A" and ypos: "0<y" ..  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
630  | 
show ?thesis  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
631  | 
proof (rule classical)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
632  | 
assume "~(\<exists>r\<in>A. r * x \<notin> A)"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
633  | 
with y have ymem: "y * x \<in> A" by blast  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
634  | 
from ypos mult_strict_left_mono [OF x]  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
635  | 
have yless: "y < y*x" by simp  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
636  | 
let ?d = "y*x - y"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
637  | 
from yless have dpos: "0 < ?d" and eq: "y + ?d = y*x" by auto  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
638  | 
from Gleason9_34 [OF A dpos]  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
639  | 
obtain r where r: "r\<in>A" and notin: "r + ?d \<notin> A" ..  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
640  | 
have rpos: "0<r" by (rule preal_imp_pos [OF A r])  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
641  | 
with dpos have rdpos: "0 < r + ?d" by arith  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
642  | 
have "~ (r + ?d \<le> y + ?d)"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
643  | 
proof  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
644  | 
assume le: "r + ?d \<le> y + ?d"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
645  | 
from ymem have yd: "y + ?d \<in> A" by (simp add: eq)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
646  | 
have "r + ?d \<in> A" by (rule preal_downwards_closed' [OF A yd rdpos le])  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
647  | 
with notin show False by simp  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
648  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
649  | 
hence "y < r" by simp  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
650  | 
with ypos have dless: "?d < (r * ?d)/y"  | 
| 
61762
 
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
 
paulson <lp15@cam.ac.uk> 
parents: 
61343 
diff
changeset
 | 
651  | 
using dpos less_divide_eq_1 by fastforce  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
652  | 
have "r + ?d < r*x"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
653  | 
proof -  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
654  | 
have "r + ?d < r + (r * ?d)/y" by (simp add: dless)  | 
| 70200 | 655  | 
also from ypos have "\<dots> = (r/y) * (y + ?d)"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
656  | 
by (simp only: algebra_simps divide_inverse, simp)  | 
| 70200 | 657  | 
also have "\<dots> = r*x" using ypos  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
658  | 
by simp  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
659  | 
finally show "r + ?d < r*x" .  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
660  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
661  | 
with r notin rdpos  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
662  | 
show "\<exists>r\<in>A. r * x \<notin> A" by (blast dest: preal_downwards_closed [OF A])  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
663  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
664  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
665  | 
|
| 61343 | 666  | 
subsection\<open>Existence of Inverse: Part 2\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
667  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
668  | 
lemma mem_Rep_preal_inverse_iff:  | 
| 70201 | 669  | 
"(z \<in> Rep_preal(inverse r)) \<longleftrightarrow> (0 < z \<and> (\<exists>y. z < y \<and> inverse y \<notin> Rep_preal r))"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
670  | 
apply (simp add: preal_inverse_def mem_inverse_set Rep_preal)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
671  | 
apply (simp add: inverse_set_def)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
672  | 
done  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
673  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
674  | 
lemma Rep_preal_one:  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
675  | 
     "Rep_preal 1 = {x. 0 < x \<and> x < 1}"
 | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
676  | 
by (simp add: preal_one_def rat_mem_preal)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
677  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
678  | 
lemma subset_inverse_mult_lemma:  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
679  | 
assumes xpos: "0 < x" and xless: "x < 1"  | 
| 70201 | 680  | 
shows "\<exists>v u y. 0 < v \<and> v < y \<and> inverse y \<notin> Rep_preal R \<and>  | 
681  | 
u \<in> Rep_preal R \<and> x = v * u"  | 
|
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
682  | 
proof -  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
683  | 
from xpos and xless have "1 < inverse x" by (simp add: one_less_inverse_iff)  | 
| 70200 | 684  | 
from lemma_gleason9_36 [OF cut_Rep_preal this]  | 
| 70201 | 685  | 
obtain t where t: "t \<in> Rep_preal R"  | 
686  | 
and notin: "t * (inverse x) \<notin> Rep_preal R" ..  | 
|
687  | 
have rpos: "0<t" by (rule preal_imp_pos [OF cut_Rep_preal t])  | 
|
688  | 
from preal_exists_greater [OF cut_Rep_preal t]  | 
|
689  | 
obtain u where u: "u \<in> Rep_preal R" and rless: "t < u" ..  | 
|
| 70200 | 690  | 
have upos: "0<u" by (rule preal_imp_pos [OF cut_Rep_preal u])  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
691  | 
show ?thesis  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
692  | 
proof (intro exI conjI)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
693  | 
show "0 < x/u" using xpos upos  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
694  | 
by (simp add: zero_less_divide_iff)  | 
| 70201 | 695  | 
show "x/u < x/t" using xpos upos rpos  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
696  | 
by (simp add: divide_inverse mult_less_cancel_left rless)  | 
| 70201 | 697  | 
show "inverse (x / t) \<notin> Rep_preal R" using notin  | 
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
57492 
diff
changeset
 | 
698  | 
by (simp add: divide_inverse mult.commute)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
699  | 
show "u \<in> Rep_preal R" by (rule u)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
700  | 
show "x = x / u * u" using upos  | 
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
57492 
diff
changeset
 | 
701  | 
by (simp add: divide_inverse mult.commute)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
702  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
703  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
704  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
705  | 
lemma subset_inverse_mult:  | 
| 70201 | 706  | 
"Rep_preal 1 \<subseteq> Rep_preal(inverse r * r)"  | 
| 70200 | 707  | 
by (force simp: Rep_preal_one mem_Rep_preal_inverse_iff mem_Rep_preal_mult_iff dest: subset_inverse_mult_lemma)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
708  | 
|
| 70201 | 709  | 
lemma inverse_mult_subset: "Rep_preal(inverse r * r) \<subseteq> Rep_preal 1"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
710  | 
proof -  | 
| 70201 | 711  | 
have "0 < u * v" if "v \<in> Rep_preal r" "0 < u" "u < t" for u v t :: rat  | 
| 70200 | 712  | 
using that by (simp add: zero_less_mult_iff preal_imp_pos [OF cut_Rep_preal])  | 
| 70201 | 713  | 
moreover have "t * q < 1"  | 
714  | 
if "q \<in> Rep_preal r" "0 < t" "t < y" "inverse y \<notin> Rep_preal r"  | 
|
715  | 
for t q y :: rat  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
716  | 
proof -  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
717  | 
have "q < inverse y"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
718  | 
using not_in_Rep_preal_ub that by auto  | 
| 70201 | 719  | 
hence "t * q < t/y"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
720  | 
using that by (simp add: divide_inverse mult_less_cancel_left)  | 
| 70200 | 721  | 
also have "\<dots> \<le> 1"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
722  | 
using that by (simp add: pos_divide_le_eq)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
723  | 
finally show ?thesis .  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
724  | 
qed  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
725  | 
ultimately show ?thesis  | 
| 70200 | 726  | 
by (auto simp: Rep_preal_one mem_Rep_preal_inverse_iff mem_Rep_preal_mult_iff)  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
727  | 
qed  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
728  | 
|
| 70201 | 729  | 
lemma preal_mult_inverse: "inverse r * r = (1::preal)"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
730  | 
by (meson Rep_preal_inject inverse_mult_subset subset_antisym subset_inverse_mult)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
731  | 
|
| 70201 | 732  | 
lemma preal_mult_inverse_right: "r * inverse r = (1::preal)"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
733  | 
using preal_mult_commute preal_mult_inverse by auto  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
734  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
735  | 
|
| 61933 | 736  | 
text\<open>Theorems needing \<open>Gleason9_34\<close>\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
737  | 
|
| 70201 | 738  | 
lemma Rep_preal_self_subset: "Rep_preal (r) \<subseteq> Rep_preal(r + s)"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
739  | 
proof  | 
| 70201 | 740  | 
fix x  | 
741  | 
assume x: "x \<in> Rep_preal r"  | 
|
742  | 
obtain y where y: "y \<in> Rep_preal s" and "y > 0"  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
743  | 
using Rep_preal preal_nonempty by blast  | 
| 70201 | 744  | 
have ry: "x+y \<in> Rep_preal(r + s)" using x y  | 
| 70200 | 745  | 
by (auto simp: mem_Rep_preal_add_iff)  | 
| 70201 | 746  | 
then show "x \<in> Rep_preal(r + s)"  | 
747  | 
by (meson \<open>0 < y\<close> add_less_same_cancel1 not_in_Rep_preal_ub order.asym preal_imp_pos [OF cut_Rep_preal x])  | 
|
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
748  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
749  | 
|
| 70201 | 750  | 
lemma Rep_preal_sum_not_subset: "~ Rep_preal (r + s) \<subseteq> Rep_preal(r)"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
751  | 
proof -  | 
| 70201 | 752  | 
obtain y where y: "y \<in> Rep_preal s" and "y > 0"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
753  | 
using Rep_preal preal_nonempty by blast  | 
| 70201 | 754  | 
obtain x where "x \<in> Rep_preal r" and notin: "x + y \<notin> Rep_preal r"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
755  | 
using Dedekind_Real.Rep_preal Gleason9_34 \<open>0 < y\<close> by blast  | 
| 70201 | 756  | 
then have "x + y \<in> Rep_preal (r + s)" using y  | 
| 70200 | 757  | 
by (auto simp: mem_Rep_preal_add_iff)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
758  | 
thus ?thesis using notin by blast  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
759  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
760  | 
|
| 61343 | 761  | 
text\<open>at last, Gleason prop. 9-3.5(iii) page 123\<close>  | 
| 70201 | 762  | 
proposition preal_self_less_add_left: "(r::preal) < r + s"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
763  | 
by (meson Rep_preal_sum_not_subset not_less preal_le_def)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
764  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
765  | 
|
| 61343 | 766  | 
subsection\<open>Subtraction for Positive Reals\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
767  | 
|
| 70201 | 768  | 
text\<open>gleason prop. 9-3.5(iv), page 123: proving \<^prop>\<open>a < b \<Longrightarrow> \<exists>d. a + d = b\<close>.  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
769  | 
We define the claimed \<^term>\<open>D\<close> and show that it is a positive real\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
770  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
771  | 
lemma mem_diff_set:  | 
| 70201 | 772  | 
assumes "r < s"  | 
773  | 
shows "cut (diff_set (Rep_preal s) (Rep_preal r))"  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
774  | 
proof -  | 
| 70201 | 775  | 
obtain p where "Rep_preal r \<subseteq> Rep_preal s" "p \<in> Rep_preal s" "p \<notin> Rep_preal r"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
776  | 
using assms unfolding preal_less_def by auto  | 
| 70201 | 777  | 
  then have "{} \<subset> diff_set (Rep_preal s) (Rep_preal r)"
 | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
778  | 
apply (simp add: diff_set_def psubset_eq)  | 
| 70200 | 779  | 
by (metis cut_Rep_preal add_eq_exists less_add_same_cancel1 preal_exists_greater preal_imp_pos)  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
780  | 
moreover  | 
| 70201 | 781  | 
obtain q where "q > 0" "q \<notin> Rep_preal s"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
782  | 
using Rep_preal_exists_bound by blast  | 
| 70201 | 783  | 
then have qnot: "q \<notin> diff_set (Rep_preal s) (Rep_preal r)"  | 
| 70200 | 784  | 
by (auto simp: diff_set_def dest: cut_Rep_preal [THEN preal_downwards_closed])  | 
| 70201 | 785  | 
  moreover have "diff_set (Rep_preal s) (Rep_preal r) \<subset> {0<..}" (is "?lhs < ?rhs")
 | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
786  | 
using \<open>0 < q\<close> diff_set_def qnot by blast  | 
| 70201 | 787  | 
moreover have "z \<in> diff_set (Rep_preal s) (Rep_preal r)"  | 
788  | 
if u: "u \<in> diff_set (Rep_preal s) (Rep_preal r)" and "0 < z" "z < u" for u z  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
789  | 
using u that less_trans Rep_preal unfolding diff_set_def Dedekind_Real.cut_def by auto  | 
| 70201 | 790  | 
moreover have "\<exists>u \<in> diff_set (Rep_preal s) (Rep_preal r). y < u"  | 
791  | 
if y: "y \<in> diff_set (Rep_preal s) (Rep_preal r)" for y  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
792  | 
proof -  | 
| 70201 | 793  | 
obtain a b where "0 < a" "0 < b" "a \<notin> Rep_preal r" "a + y + b \<in> Rep_preal s"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
794  | 
using y  | 
| 70200 | 795  | 
by (simp add: diff_set_def) (metis cut_Rep_preal add_eq_exists less_add_same_cancel1 preal_exists_greater)  | 
| 70201 | 796  | 
then have "a + (y + b) \<in> Rep_preal s"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
797  | 
by (simp add: add.assoc)  | 
| 70201 | 798  | 
then have "y + b \<in> diff_set (Rep_preal s) (Rep_preal r)"  | 
799  | 
using \<open>0 < a\<close> \<open>0 < b\<close> \<open>a \<notin> Rep_preal r\<close> y  | 
|
| 70200 | 800  | 
by (auto simp: diff_set_def)  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
801  | 
then show ?thesis  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
802  | 
using \<open>0 < b\<close> less_add_same_cancel1 by blast  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
803  | 
qed  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
804  | 
ultimately show ?thesis  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
805  | 
by (simp add: Dedekind_Real.cut_def)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
806  | 
qed  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
807  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
808  | 
lemma mem_Rep_preal_diff_iff:  | 
| 70201 | 809  | 
"r < s \<Longrightarrow>  | 
810  | 
(z \<in> Rep_preal (s - r)) \<longleftrightarrow>  | 
|
811  | 
(\<exists>x. 0 < x \<and> 0 < z \<and> x \<notin> Rep_preal r \<and> x + z \<in> Rep_preal s)"  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
812  | 
apply (simp add: preal_diff_def mem_diff_set Rep_preal)  | 
| 70200 | 813  | 
apply (force simp: diff_set_def)  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
814  | 
done  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
815  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
816  | 
proposition less_add_left:  | 
| 70201 | 817  | 
fixes r::preal  | 
818  | 
assumes "r < s"  | 
|
819  | 
shows "r + (s-r) = s"  | 
|
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
820  | 
proof -  | 
| 70201 | 821  | 
have "a + b \<in> Rep_preal s"  | 
822  | 
if "a \<in> Rep_preal r" "c + b \<in> Rep_preal s" "c \<notin> Rep_preal r"  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
823  | 
and "0 < b" "0 < c" for a b c  | 
| 70200 | 824  | 
by (meson cut_Rep_preal add_less_imp_less_right add_pos_pos not_in_Rep_preal_ub preal_downwards_closed preal_imp_pos that)  | 
| 70201 | 825  | 
then have "r + (s-r) \<le> s"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
826  | 
using assms mem_Rep_preal_add_iff mem_Rep_preal_diff_iff preal_le_def by auto  | 
| 70201 | 827  | 
have "x \<in> Rep_preal (r + (s - r))" if "x \<in> Rep_preal s" for x  | 
828  | 
proof (cases "x \<in> Rep_preal r")  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
829  | 
case True  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
830  | 
then show ?thesis  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
831  | 
using Rep_preal_self_subset by blast  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
832  | 
next  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
833  | 
case False  | 
| 70201 | 834  | 
have "\<exists>u v z. 0 < v \<and> 0 < z \<and> u \<in> Rep_preal r \<and> z \<notin> Rep_preal r \<and> z + v \<in> Rep_preal s \<and> x = u + v"  | 
835  | 
if x: "x \<in> Rep_preal s"  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
836  | 
proof -  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
837  | 
have xpos: "x > 0"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
838  | 
using Rep_preal preal_imp_pos that by blast  | 
| 70201 | 839  | 
obtain e where epos: "0 < e" and xe: "x + e \<in> Rep_preal s"  | 
| 70200 | 840  | 
by (metis cut_Rep_preal x add_eq_exists less_add_same_cancel1 preal_exists_greater)  | 
841  | 
from Gleason9_34 [OF cut_Rep_preal epos]  | 
|
| 70201 | 842  | 
obtain u where r: "u \<in> Rep_preal r" and notin: "u + e \<notin> Rep_preal r" ..  | 
843  | 
with x False xpos have rless: "u < x" by (blast intro: not_in_Rep_preal_ub)  | 
|
844  | 
from add_eq_exists [of u x]  | 
|
845  | 
obtain y where eq: "x = u+y" by auto  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
846  | 
show ?thesis  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
847  | 
proof (intro exI conjI)  | 
| 70201 | 848  | 
show "u + e \<notin> Rep_preal r" by (rule notin)  | 
849  | 
show "u + e + y \<in> Rep_preal s" using xe eq by (simp add: ac_simps)  | 
|
850  | 
show "0 < u + e"  | 
|
| 70200 | 851  | 
using epos preal_imp_pos [OF cut_Rep_preal r] by simp  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
852  | 
qed (use r rless eq in auto)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
853  | 
qed  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
854  | 
then show ?thesis  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
855  | 
using assms mem_Rep_preal_add_iff mem_Rep_preal_diff_iff that by blast  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
856  | 
qed  | 
| 70201 | 857  | 
then have "s \<le> r + (s-r)"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
858  | 
by (auto simp: preal_le_def)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
859  | 
then show ?thesis  | 
| 70201 | 860  | 
by (simp add: \<open>r + (s - r) \<le> s\<close> antisym)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
861  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
862  | 
|
| 70201 | 863  | 
lemma preal_add_less2_mono1: "r < (s::preal) \<Longrightarrow> r + t < s + t"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
864  | 
by (metis add.assoc add.commute less_add_left preal_self_less_add_left)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
865  | 
|
| 70201 | 866  | 
lemma preal_add_less2_mono2: "r < (s::preal) \<Longrightarrow> t + r < t + s"  | 
867  | 
by (auto intro: preal_add_less2_mono1 simp add: preal_add_commute [of t])  | 
|
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
868  | 
|
| 70201 | 869  | 
lemma preal_add_right_less_cancel: "r + t < s + t \<Longrightarrow> r < (s::preal)"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
870  | 
by (metis linorder_cases order.asym preal_add_less2_mono1)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
871  | 
|
| 70201 | 872  | 
lemma preal_add_left_less_cancel: "t + r < t + s \<Longrightarrow> r < (s::preal)"  | 
873  | 
by (auto elim: preal_add_right_less_cancel simp add: preal_add_commute [of t])  | 
|
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
874  | 
|
| 70201 | 875  | 
lemma preal_add_less_cancel_left [simp]: "(t + (r::preal) < t + s) \<longleftrightarrow> (r < s)"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
876  | 
by (blast intro: preal_add_less2_mono2 preal_add_left_less_cancel)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
877  | 
|
| 70201 | 878  | 
lemma preal_add_less_cancel_right [simp]: "((r::preal) + t < s + t) = (r < s)"  | 
879  | 
using preal_add_less_cancel_left [symmetric, of r s t] by (simp add: ac_simps)  | 
|
| 
59815
 
cce82e360c2f
explicit commutative additive inverse operation;
 
haftmann 
parents: 
59814 
diff
changeset
 | 
880  | 
|
| 70201 | 881  | 
lemma preal_add_le_cancel_left [simp]: "(t + (r::preal) \<le> t + s) = (r \<le> s)"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
882  | 
by (simp add: linorder_not_less [symmetric])  | 
| 
59815
 
cce82e360c2f
explicit commutative additive inverse operation;
 
haftmann 
parents: 
59814 
diff
changeset
 | 
883  | 
|
| 70201 | 884  | 
lemma preal_add_le_cancel_right [simp]: "((r::preal) + t \<le> s + t) = (r \<le> s)"  | 
885  | 
using preal_add_le_cancel_left [symmetric, of r s t] by (simp add: ac_simps)  | 
|
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
886  | 
|
| 70201 | 887  | 
lemma preal_add_right_cancel: "(r::preal) + t = s + t \<Longrightarrow> r = s"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
888  | 
by (metis less_irrefl linorder_cases preal_add_less_cancel_right)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
889  | 
|
| 70201 | 890  | 
lemma preal_add_left_cancel: "c + a = c + b \<Longrightarrow> a = (b::preal)"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
891  | 
by (auto intro: preal_add_right_cancel simp add: preal_add_commute)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
892  | 
|
| 
59815
 
cce82e360c2f
explicit commutative additive inverse operation;
 
haftmann 
parents: 
59814 
diff
changeset
 | 
893  | 
instance preal :: linordered_ab_semigroup_add  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
894  | 
proof  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
895  | 
fix a b c :: preal  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
896  | 
show "a \<le> b \<Longrightarrow> c + a \<le> c + b" by (simp only: preal_add_le_cancel_left)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
897  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
898  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
899  | 
|
| 69597 | 900  | 
subsection\<open>Completeness of type \<^typ>\<open>preal\<close>\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
901  | 
|
| 61343 | 902  | 
text\<open>Prove that supremum is a cut\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
903  | 
|
| 61343 | 904  | 
text\<open>Part 1 of Dedekind sections definition\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
905  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
906  | 
lemma preal_sup:  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
907  | 
  assumes le: "\<And>X. X \<in> P \<Longrightarrow> X \<le> Y" and "P \<noteq> {}" 
 | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
908  | 
shows "cut (\<Union>X \<in> P. Rep_preal(X))"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
909  | 
proof -  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
910  | 
  have "{} \<subset> (\<Union>X \<in> P. Rep_preal(X))"
 | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
911  | 
    using \<open>P \<noteq> {}\<close> mem_Rep_preal_Ex by fastforce
 | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
912  | 
moreover  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
913  | 
obtain q where "q > 0" and "q \<notin> (\<Union>X \<in> P. Rep_preal(X))"  | 
| 70200 | 914  | 
using Rep_preal_exists_bound [of Y] le by (auto simp: preal_le_def)  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
915  | 
  then have "(\<Union>X \<in> P. Rep_preal(X)) \<subset> {0<..}"
 | 
| 70200 | 916  | 
using cut_Rep_preal preal_imp_pos by force  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
917  | 
moreover  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
918  | 
have "\<And>u z. \<lbrakk>u \<in> (\<Union>X \<in> P. Rep_preal(X)); 0 < z; z < u\<rbrakk> \<Longrightarrow> z \<in> (\<Union>X \<in> P. Rep_preal(X))"  | 
| 70200 | 919  | 
by (auto elim: cut_Rep_preal [THEN preal_downwards_closed])  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
920  | 
moreover  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
921  | 
have "\<And>y. y \<in> (\<Union>X \<in> P. Rep_preal(X)) \<Longrightarrow> \<exists>u \<in> (\<Union>X \<in> P. Rep_preal(X)). y < u"  | 
| 70200 | 922  | 
by (blast dest: cut_Rep_preal [THEN preal_exists_greater])  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
923  | 
ultimately show ?thesis  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
924  | 
by (simp add: Dedekind_Real.cut_def)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
925  | 
qed  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
926  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
927  | 
lemma preal_psup_le:  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
928  | 
"\<lbrakk>\<And>X. X \<in> P \<Longrightarrow> X \<le> Y; x \<in> P\<rbrakk> \<Longrightarrow> x \<le> psup P"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
929  | 
using preal_sup [of P Y] unfolding preal_le_def psup_def by fastforce  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
930  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
931  | 
lemma psup_le_ub: "\<lbrakk>\<And>X. X \<in> P \<Longrightarrow> X \<le> Y; P \<noteq> {}\<rbrakk> \<Longrightarrow> psup P \<le> Y"
 | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
932  | 
using preal_sup [of P Y] by (simp add: SUP_least preal_le_def psup_def)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
933  | 
|
| 61343 | 934  | 
text\<open>Supremum property\<close>  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
935  | 
proposition preal_complete:  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
936  | 
  assumes le: "\<And>X. X \<in> P \<Longrightarrow> X \<le> Y" and "P \<noteq> {}" 
 | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
937  | 
shows "(\<exists>X \<in> P. Z < X) \<longleftrightarrow> (Z < psup P)" (is "?lhs = ?rhs")  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
938  | 
proof  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
939  | 
assume ?lhs  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
940  | 
then show ?rhs  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
941  | 
using preal_sup [OF assms] preal_less_def psup_def by auto  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
942  | 
next  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
943  | 
assume ?rhs  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
944  | 
then show ?lhs  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
945  | 
    by (meson \<open>P \<noteq> {}\<close> not_less psup_le_ub) 
 | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
946  | 
qed  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
947  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
948  | 
subsection \<open>Defining the Reals from the Positive Reals\<close>  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
949  | 
|
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
950  | 
text \<open>Here we do quotients the old-fashioned way\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
951  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
952  | 
definition  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
953  | 
realrel :: "((preal * preal) * (preal * preal)) set" where  | 
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
954  | 
  "realrel = {p. \<exists>x1 y1 x2 y2. p = ((x1,y1),(x2,y2)) \<and> x1+y2 = x2+y1}"
 | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
955  | 
|
| 
45694
 
4a8743618257
prefer typedef without extra definition and alternative name;
 
wenzelm 
parents: 
41541 
diff
changeset
 | 
956  | 
definition "Real = UNIV//realrel"  | 
| 
 
4a8743618257
prefer typedef without extra definition and alternative name;
 
wenzelm 
parents: 
41541 
diff
changeset
 | 
957  | 
|
| 49834 | 958  | 
typedef real = Real  | 
| 
45694
 
4a8743618257
prefer typedef without extra definition and alternative name;
 
wenzelm 
parents: 
41541 
diff
changeset
 | 
959  | 
morphisms Rep_Real Abs_Real  | 
| 70200 | 960  | 
unfolding Real_def by (auto simp: quotient_def)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
961  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
962  | 
text \<open>This doesn't involve the overloaded "real" function: users don't see it\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
963  | 
definition  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
964  | 
real_of_preal :: "preal \<Rightarrow> real" where  | 
| 37765 | 965  | 
  "real_of_preal m = Abs_Real (realrel `` {(m + 1, 1)})"
 | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
966  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
967  | 
instantiation real :: "{zero, one, plus, minus, uminus, times, inverse, ord, abs, sgn}"
 | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
968  | 
begin  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
969  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
970  | 
definition  | 
| 37765 | 971  | 
  real_zero_def: "0 = Abs_Real(realrel``{(1, 1)})"
 | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
972  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
973  | 
definition  | 
| 37765 | 974  | 
  real_one_def: "1 = Abs_Real(realrel``{(1 + 1, 1)})"
 | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
975  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
976  | 
definition  | 
| 37765 | 977  | 
real_add_def: "z + w =  | 
| 39910 | 978  | 
the_elem (\<Union>(x,y) \<in> Rep_Real(z). \<Union>(u,v) \<in> Rep_Real(w).  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
979  | 
                 { Abs_Real(realrel``{(x+u, y+v)}) })"
 | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
980  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
981  | 
definition  | 
| 39910 | 982  | 
  real_minus_def: "- r =  the_elem (\<Union>(x,y) \<in> Rep_Real(r). { Abs_Real(realrel``{(y,x)}) })"
 | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
983  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
984  | 
definition  | 
| 37765 | 985  | 
real_diff_def: "r - (s::real) = r + - s"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
986  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
987  | 
definition  | 
| 37765 | 988  | 
real_mult_def:  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
989  | 
"z * w =  | 
| 39910 | 990  | 
the_elem (\<Union>(x,y) \<in> Rep_Real(z). \<Union>(u,v) \<in> Rep_Real(w).  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
991  | 
                 { Abs_Real(realrel``{(x*u + y*v, x*v + y*u)}) })"
 | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
992  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
993  | 
definition  | 
| 70201 | 994  | 
real_inverse_def: "inverse (r::real) = (THE s. (r = 0 \<and> s = 0) \<or> s * r = 1)"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
995  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
996  | 
definition  | 
| 70201 | 997  | 
real_divide_def: "r div (s::real) = r * inverse s"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
998  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
999  | 
definition  | 
| 37765 | 1000  | 
real_le_def: "z \<le> (w::real) \<longleftrightarrow>  | 
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1001  | 
(\<exists>x y u v. x+v \<le> u+y \<and> (x,y) \<in> Rep_Real z \<and> (u,v) \<in> Rep_Real w)"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1002  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1003  | 
definition  | 
| 61076 | 1004  | 
real_less_def: "x < (y::real) \<longleftrightarrow> x \<le> y \<and> x \<noteq> y"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1005  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1006  | 
definition  | 
| 61945 | 1007  | 
real_abs_def: "\<bar>r::real\<bar> = (if r < 0 then - r else r)"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1008  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1009  | 
definition  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1010  | 
real_sgn_def: "sgn (x::real) = (if x=0 then 0 else if 0<x then 1 else - 1)"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1011  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1012  | 
instance ..  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1013  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1014  | 
end  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1015  | 
|
| 61343 | 1016  | 
subsection \<open>Equivalence relation over positive reals\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1017  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1018  | 
lemma realrel_iff [simp]: "(((x1,y1),(x2,y2)) \<in> realrel) = (x1 + y2 = x2 + y1)"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1019  | 
by (simp add: realrel_def)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1020  | 
|
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1021  | 
lemma preal_trans_lemma:  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1022  | 
assumes "x + y1 = x1 + y" and "x + y2 = x2 + y"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1023  | 
shows "x1 + y2 = x2 + (y1::preal)"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1024  | 
by (metis add.left_commute assms preal_add_left_cancel)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1025  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1026  | 
lemma equiv_realrel: "equiv UNIV realrel"  | 
| 70200 | 1027  | 
by (auto simp: equiv_def refl_on_def sym_def trans_def realrel_def intro: dest: preal_trans_lemma)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1028  | 
|
| 69597 | 1029  | 
text\<open>Reduces equality of equivalence classes to the \<^term>\<open>realrel\<close> relation:  | 
1030  | 
  \<^term>\<open>(realrel `` {x} = realrel `` {y}) = ((x,y) \<in> realrel)\<close>\<close>
 | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1031  | 
lemmas equiv_realrel_iff [simp] =  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1032  | 
eq_equiv_class_iff [OF equiv_realrel UNIV_I UNIV_I]  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1033  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1034  | 
lemma realrel_in_real [simp]: "realrel``{(x,y)} \<in> Real"
 | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1035  | 
by (simp add: Real_def realrel_def quotient_def, blast)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1036  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1037  | 
declare Abs_Real_inject [simp] Abs_Real_inverse [simp]  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1038  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1039  | 
|
| 61343 | 1040  | 
text\<open>Case analysis on the representation of a real number as an equivalence  | 
1041  | 
class of pairs of positive reals.\<close>  | 
|
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1042  | 
lemma eq_Abs_Real [case_names Abs_Real, cases type: real]:  | 
| 67613 | 1043  | 
     "(\<And>x y. z = Abs_Real(realrel``{(x,y)}) \<Longrightarrow> P) \<Longrightarrow> P"
 | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1044  | 
by (metis Rep_Real_inverse prod.exhaust Rep_Real [of z, unfolded Real_def, THEN quotientE])  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1045  | 
|
| 61343 | 1046  | 
subsection \<open>Addition and Subtraction\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1047  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1048  | 
lemma real_add:  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1049  | 
     "Abs_Real (realrel``{(x,y)}) + Abs_Real (realrel``{(u,v)}) =
 | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1050  | 
      Abs_Real (realrel``{(x+u, y+v)})"
 | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1051  | 
proof -  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1052  | 
  have "(\<lambda>z w. (\<lambda>(x,y). (\<lambda>(u,v). {Abs_Real (realrel `` {(x+u, y+v)})}) w) z)
 | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1053  | 
respects2 realrel"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1054  | 
by (clarsimp simp: congruent2_def) (metis add.left_commute preal_add_assoc)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1055  | 
thus ?thesis  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1056  | 
by (simp add: real_add_def UN_UN_split_split_eq UN_equiv_class2 [OF equiv_realrel equiv_realrel])  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1057  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1058  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1059  | 
lemma real_minus: "- Abs_Real(realrel``{(x,y)}) = Abs_Real(realrel `` {(y,x)})"
 | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1060  | 
proof -  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1061  | 
  have "(\<lambda>(x,y). {Abs_Real (realrel``{(y,x)})}) respects realrel"
 | 
| 70200 | 1062  | 
by (auto simp: congruent_def add.commute)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1063  | 
thus ?thesis  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1064  | 
by (simp add: real_minus_def UN_equiv_class [OF equiv_realrel])  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1065  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1066  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1067  | 
instance real :: ab_group_add  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1068  | 
proof  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1069  | 
fix x y z :: real  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1070  | 
show "(x + y) + z = x + (y + z)"  | 
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
57492 
diff
changeset
 | 
1071  | 
by (cases x, cases y, cases z, simp add: real_add add.assoc)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1072  | 
show "x + y = y + x"  | 
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
57492 
diff
changeset
 | 
1073  | 
by (cases x, cases y, simp add: real_add add.commute)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1074  | 
show "0 + x = x"  | 
| 
57514
 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 
haftmann 
parents: 
57512 
diff
changeset
 | 
1075  | 
by (cases x, simp add: real_add real_zero_def ac_simps)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1076  | 
show "- x + x = 0"  | 
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
57492 
diff
changeset
 | 
1077  | 
by (cases x, simp add: real_minus real_add real_zero_def add.commute)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1078  | 
show "x - y = x + - y"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1079  | 
by (simp add: real_diff_def)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1080  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1081  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1082  | 
|
| 61343 | 1083  | 
subsection \<open>Multiplication\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1084  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1085  | 
lemma real_mult_congruent2_lemma:  | 
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1086  | 
"!!(x1::preal). \<lbrakk>x1 + y2 = x2 + y1\<rbrakk> \<Longrightarrow>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1087  | 
x * x1 + y * y1 + (x * y2 + y * x2) =  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1088  | 
x * x2 + y * y2 + (x * y1 + y * x1)"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1089  | 
by (metis (no_types, hide_lams) add.left_commute preal_add_commute preal_add_mult_distrib2)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1090  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1091  | 
lemma real_mult_congruent2:  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1092  | 
"(\<lambda>p1 p2.  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1093  | 
(\<lambda>(x1,y1). (\<lambda>(x2,y2).  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1094  | 
          { Abs_Real (realrel``{(x1*x2 + y1*y2, x1*y2+y1*x2)}) }) p2) p1)
 | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1095  | 
respects2 realrel"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1096  | 
apply (rule congruent2_commuteI [OF equiv_realrel])  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1097  | 
by (auto simp: mult.commute add.commute combine_common_factor preal_add_assoc preal_add_commute)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1098  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1099  | 
lemma real_mult:  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1100  | 
  "Abs_Real((realrel``{(x1,y1)})) * Abs_Real((realrel``{(x2,y2)})) =
 | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1101  | 
   Abs_Real(realrel `` {(x1*x2+y1*y2,x1*y2+y1*x2)})"
 | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1102  | 
by (simp add: real_mult_def UN_UN_split_split_eq  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1103  | 
UN_equiv_class2 [OF equiv_realrel equiv_realrel real_mult_congruent2])  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1104  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1105  | 
lemma real_mult_commute: "(z::real) * w = w * z"  | 
| 
59815
 
cce82e360c2f
explicit commutative additive inverse operation;
 
haftmann 
parents: 
59814 
diff
changeset
 | 
1106  | 
by (cases z, cases w, simp add: real_mult ac_simps)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1107  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1108  | 
lemma real_mult_assoc: "((z1::real) * z2) * z3 = z1 * (z2 * z3)"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1109  | 
by (cases z1, cases z2, cases z3) (simp add: real_mult algebra_simps)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1110  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1111  | 
lemma real_mult_1: "(1::real) * z = z"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1112  | 
by (cases z) (simp add: real_mult real_one_def algebra_simps)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1113  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1114  | 
lemma real_add_mult_distrib: "((z1::real) + z2) * w = (z1 * w) + (z2 * w)"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1115  | 
by (cases z1, cases z2, cases w) (simp add: real_add real_mult algebra_simps)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1116  | 
|
| 61343 | 1117  | 
text\<open>one and zero are distinct\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1118  | 
lemma real_zero_not_eq_one: "0 \<noteq> (1::real)"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1119  | 
proof -  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1120  | 
have "(1::preal) < 1 + 1"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1121  | 
by (simp add: preal_self_less_add_left)  | 
| 
59815
 
cce82e360c2f
explicit commutative additive inverse operation;
 
haftmann 
parents: 
59814 
diff
changeset
 | 
1122  | 
then show ?thesis  | 
| 
 
cce82e360c2f
explicit commutative additive inverse operation;
 
haftmann 
parents: 
59814 
diff
changeset
 | 
1123  | 
by (simp add: real_zero_def real_one_def neq_iff)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1124  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1125  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1126  | 
instance real :: comm_ring_1  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1127  | 
proof  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1128  | 
fix x y z :: real  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1129  | 
show "(x * y) * z = x * (y * z)" by (rule real_mult_assoc)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1130  | 
show "x * y = y * x" by (rule real_mult_commute)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1131  | 
show "1 * x = x" by (rule real_mult_1)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1132  | 
show "(x + y) * z = x * z + y * z" by (rule real_add_mult_distrib)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1133  | 
show "0 \<noteq> (1::real)" by (rule real_zero_not_eq_one)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1134  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1135  | 
|
| 61343 | 1136  | 
subsection \<open>Inverse and Division\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1137  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1138  | 
lemma real_zero_iff: "Abs_Real (realrel `` {(x, x)}) = 0"
 | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1139  | 
by (simp add: real_zero_def add.commute)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1140  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1141  | 
lemma real_mult_inverse_left_ex:  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1142  | 
assumes "x \<noteq> 0" obtains y::real where "y*x = 1"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1143  | 
proof (cases x)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1144  | 
case (Abs_Real u v)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1145  | 
show ?thesis  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1146  | 
proof (cases u v rule: linorder_cases)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1147  | 
case less  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1148  | 
then have "v * inverse (v - u) = 1 + u * inverse (v - u)"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1149  | 
using less_add_left [of u v]  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1150  | 
by (metis preal_add_commute preal_add_mult_distrib preal_mult_inverse_right)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1151  | 
    then have "Abs_Real (realrel``{(1, inverse (v-u) + 1)}) * x - 1 = 0"
 | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1152  | 
by (simp add: Abs_Real real_mult preal_mult_inverse_right real_one_def) (simp add: algebra_simps)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1153  | 
with that show thesis by auto  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1154  | 
next  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1155  | 
case equal  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1156  | 
then show ?thesis  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1157  | 
using Abs_Real assms real_zero_iff by blast  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1158  | 
next  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1159  | 
case greater  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1160  | 
then have "u * inverse (u - v) = 1 + v * inverse (u - v)"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1161  | 
using less_add_left [of v u] by (metis add.commute distrib_right preal_mult_inverse_right)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1162  | 
    then have "Abs_Real (realrel``{(inverse (u-v) + 1, 1)}) * x - 1 = 0"
 | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1163  | 
by (simp add: Abs_Real real_mult preal_mult_inverse_right real_one_def) (simp add: algebra_simps)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1164  | 
with that show thesis by auto  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1165  | 
qed  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1166  | 
qed  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1167  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1168  | 
|
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1169  | 
lemma real_mult_inverse_left:  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1170  | 
fixes x :: real  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1171  | 
assumes "x \<noteq> 0" shows "inverse x * x = 1"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1172  | 
proof -  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1173  | 
obtain y where "y*x = 1"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1174  | 
using assms real_mult_inverse_left_ex by blast  | 
| 70201 | 1175  | 
then have "(THE s. s * x = 1) * x = 1"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1176  | 
proof (rule theI)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1177  | 
show "y' = y" if "y' * x = 1" for y'  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1178  | 
by (metis \<open>y * x = 1\<close> mult.left_commute mult.right_neutral that)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1179  | 
qed  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1180  | 
then show ?thesis  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1181  | 
using assms real_inverse_def by auto  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1182  | 
qed  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1183  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1184  | 
|
| 61343 | 1185  | 
subsection\<open>The Real Numbers form a Field\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1186  | 
|
| 
59867
 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 
haftmann 
parents: 
59815 
diff
changeset
 | 
1187  | 
instance real :: field  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1188  | 
proof  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1189  | 
fix x y z :: real  | 
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1190  | 
show "x \<noteq> 0 \<Longrightarrow> inverse x * x = 1" by (rule real_mult_inverse_left)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1191  | 
show "x / y = x * inverse y" by (simp add: real_divide_def)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1192  | 
show "inverse 0 = (0::real)" by (simp add: real_inverse_def)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1193  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1194  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1195  | 
|
| 61933 | 1196  | 
subsection\<open>The \<open>\<le>\<close> Ordering\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1197  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1198  | 
lemma real_le_refl: "w \<le> (w::real)"  | 
| 70200 | 1199  | 
by (cases w, force simp: real_le_def)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1200  | 
|
| 61343 | 1201  | 
text\<open>The arithmetic decision procedure is not set up for type preal.  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1202  | 
This lemma is currently unused, but it could simplify the proofs of the  | 
| 61343 | 1203  | 
following two lemmas.\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1204  | 
lemma preal_eq_le_imp_le:  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1205  | 
assumes eq: "a+b = c+d" and le: "c \<le> a"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1206  | 
shows "b \<le> (d::preal)"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1207  | 
proof -  | 
| 
59815
 
cce82e360c2f
explicit commutative additive inverse operation;
 
haftmann 
parents: 
59814 
diff
changeset
 | 
1208  | 
from le have "c+d \<le> a+d" by simp  | 
| 41541 | 1209  | 
hence "a+b \<le> a+d" by (simp add: eq)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1210  | 
thus "b \<le> d" by simp  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1211  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1212  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1213  | 
lemma real_le_lemma:  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1214  | 
assumes l: "u1 + v2 \<le> u2 + v1"  | 
| 41541 | 1215  | 
and "x1 + v1 = u1 + y1"  | 
1216  | 
and "x2 + v2 = u2 + y2"  | 
|
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1217  | 
shows "x1 + y2 \<le> x2 + (y1::preal)"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1218  | 
proof -  | 
| 41541 | 1219  | 
have "(x1+v1) + (u2+y2) = (u1+y1) + (x2+v2)" by (simp add: assms)  | 
| 
57514
 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 
haftmann 
parents: 
57512 
diff
changeset
 | 
1220  | 
hence "(x1+y2) + (u2+v1) = (x2+y1) + (u1+v2)" by (simp add: ac_simps)  | 
| 70200 | 1221  | 
also have "\<dots> \<le> (x2+y1) + (u2+v1)" by (simp add: assms)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1222  | 
finally show ?thesis by simp  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1223  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1224  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1225  | 
lemma real_le:  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1226  | 
  "Abs_Real(realrel``{(x1,y1)}) \<le> Abs_Real(realrel``{(x2,y2)})  \<longleftrightarrow>  x1 + y2 \<le> x2 + y1"
 | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1227  | 
unfolding real_le_def by (auto intro: real_le_lemma)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1228  | 
|
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1229  | 
lemma real_le_antisym: "\<lbrakk>z \<le> w; w \<le> z\<rbrakk> \<Longrightarrow> z = (w::real)"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1230  | 
by (cases z, cases w, simp add: real_le)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1231  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1232  | 
lemma real_trans_lemma:  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1233  | 
assumes "x + v \<le> u + y"  | 
| 41541 | 1234  | 
and "u + v' \<le> u' + v"  | 
1235  | 
and "x2 + v2 = u2 + y2"  | 
|
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1236  | 
shows "x + v' \<le> u' + (y::preal)"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1237  | 
proof -  | 
| 
57514
 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 
haftmann 
parents: 
57512 
diff
changeset
 | 
1238  | 
have "(x+v') + (u+v) = (x+v) + (u+v')" by (simp add: ac_simps)  | 
| 70200 | 1239  | 
also have "\<dots> \<le> (u+y) + (u+v')" by (simp add: assms)  | 
1240  | 
also have "\<dots> \<le> (u+y) + (u'+v)" by (simp add: assms)  | 
|
1241  | 
also have "\<dots> = (u'+y) + (u+v)" by (simp add: ac_simps)  | 
|
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1242  | 
finally show ?thesis by simp  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1243  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1244  | 
|
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1245  | 
lemma real_le_trans: "\<lbrakk>i \<le> j; j \<le> k\<rbrakk> \<Longrightarrow> i \<le> (k::real)"  | 
| 70200 | 1246  | 
by (cases i, cases j, cases k) (auto simp: real_le intro: real_trans_lemma)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1247  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1248  | 
instance real :: order  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1249  | 
proof  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1250  | 
show "u < v \<longleftrightarrow> u \<le> v \<and> \<not> v \<le> u" for u v::real  | 
| 70200 | 1251  | 
by (auto simp: real_less_def intro: real_le_antisym)  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1252  | 
qed (auto intro: real_le_refl real_le_trans real_le_antisym)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1253  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1254  | 
instance real :: linorder  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1255  | 
proof  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1256  | 
show "x \<le> y \<or> y \<le> x" for x y :: real  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1257  | 
by (meson eq_refl le_cases real_le_def)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1258  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1259  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1260  | 
instantiation real :: distrib_lattice  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1261  | 
begin  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1262  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1263  | 
definition  | 
| 61076 | 1264  | 
"(inf :: real \<Rightarrow> real \<Rightarrow> real) = min"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1265  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1266  | 
definition  | 
| 61076 | 1267  | 
"(sup :: real \<Rightarrow> real \<Rightarrow> real) = max"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1268  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1269  | 
instance  | 
| 70200 | 1270  | 
by standard (auto simp: inf_real_def sup_real_def max_min_distrib2)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1271  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1272  | 
end  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1273  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1274  | 
subsection\<open>The Reals Form an Ordered Field\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1275  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1276  | 
lemma real_le_eq_diff: "(x \<le> y) \<longleftrightarrow> (x-y \<le> (0::real))"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1277  | 
by (cases x, cases y) (simp add: real_le real_zero_def real_diff_def real_add real_minus preal_add_commute)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1278  | 
|
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1279  | 
lemma real_add_left_mono:  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1280  | 
assumes le: "x \<le> y" shows "z + x \<le> z + (y::real)"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1281  | 
proof -  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1282  | 
have "z + x - (z + y) = (z + -z) + (x - y)"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1283  | 
by (simp add: algebra_simps)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1284  | 
with le show ?thesis  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1285  | 
by (simp add: real_le_eq_diff[of x] real_le_eq_diff[of "z+x"])  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1286  | 
qed  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1287  | 
|
| 70201 | 1288  | 
lemma real_sum_gt_zero_less: "(0 < s + (-w::real)) \<Longrightarrow> (w < s)"  | 
1289  | 
by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of s])  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1290  | 
|
| 70201 | 1291  | 
lemma real_less_sum_gt_zero: "(w < s) \<Longrightarrow> (0 < s + (-w::real))"  | 
1292  | 
by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of s])  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1293  | 
|
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1294  | 
lemma real_mult_order:  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1295  | 
fixes x y::real  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1296  | 
assumes "0 < x" "0 < y"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1297  | 
shows "0 < x * y"  | 
| 70200 | 1298  | 
proof (cases x, cases y)  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1299  | 
show "0 < x * y"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1300  | 
    if x: "x = Abs_Real (Dedekind_Real.realrel `` {(x1, x2)})"
 | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1301  | 
      and y: "y = Abs_Real (Dedekind_Real.realrel `` {(y1, y2)})"
 | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1302  | 
for x1 x2 y1 y2  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1303  | 
proof -  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1304  | 
have "x2 < x1" "y2 < y1"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1305  | 
using assms not_le real_zero_def real_le x y  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1306  | 
by (metis preal_add_le_cancel_left real_zero_iff)+  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1307  | 
then obtain xd yd where "x1 = x2 + xd" "y1 = y2 + yd"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1308  | 
using less_add_left by metis  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1309  | 
then have "\<not> (x * y \<le> 0)"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1310  | 
apply (simp add: x y real_mult real_zero_def real_le)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1311  | 
apply (simp add: not_le algebra_simps preal_self_less_add_left)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1312  | 
done  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1313  | 
then show ?thesis  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1314  | 
by auto  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1315  | 
qed  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1316  | 
qed  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1317  | 
|
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1318  | 
lemma real_mult_less_mono2: "\<lbrakk>(0::real) < z; x < y\<rbrakk> \<Longrightarrow> z * x < z * y"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1319  | 
by (metis add_uminus_conv_diff real_less_sum_gt_zero real_mult_order real_sum_gt_zero_less right_diff_distrib')  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1320  | 
|
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1321  | 
|
| 
59867
 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 
haftmann 
parents: 
59815 
diff
changeset
 | 
1322  | 
instance real :: linordered_field  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1323  | 
proof  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1324  | 
fix x y z :: real  | 
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1325  | 
show "x \<le> y \<Longrightarrow> z + x \<le> z + y" by (rule real_add_left_mono)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1326  | 
show "\<bar>x\<bar> = (if x < 0 then -x else x)" by (simp only: real_abs_def)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1327  | 
show "sgn x = (if x=0 then 0 else if 0<x then 1 else - 1)"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1328  | 
by (simp only: real_sgn_def)  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1329  | 
show "z * x < z * y" if "x < y" "0 < z"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1330  | 
by (simp add: real_mult_less_mono2 that)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1331  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1332  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1333  | 
|
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1334  | 
subsection \<open>Completeness of the reals\<close>  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1335  | 
|
| 69597 | 1336  | 
text\<open>The function \<^term>\<open>real_of_preal\<close> requires many proofs, but it seems  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1337  | 
to be essential for proving completeness of the reals from that of the  | 
| 61343 | 1338  | 
positive reals.\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1339  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1340  | 
lemma real_of_preal_add:  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1341  | 
"real_of_preal ((x::preal) + y) = real_of_preal x + real_of_preal y"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1342  | 
by (simp add: real_of_preal_def real_add algebra_simps)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1343  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1344  | 
lemma real_of_preal_mult:  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1345  | 
"real_of_preal ((x::preal) * y) = real_of_preal x * real_of_preal y"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1346  | 
by (simp add: real_of_preal_def real_mult algebra_simps)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1347  | 
|
| 61343 | 1348  | 
text\<open>Gleason prop 9-4.4 p 127\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1349  | 
lemma real_of_preal_trichotomy:  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1350  | 
"\<exists>m. (x::real) = real_of_preal m \<or> x = 0 \<or> x = -(real_of_preal m)"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1351  | 
proof (cases x)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1352  | 
case (Abs_Real u v)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1353  | 
show ?thesis  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1354  | 
proof (cases u v rule: linorder_cases)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1355  | 
case less  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1356  | 
then show ?thesis  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1357  | 
using less_add_left  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1358  | 
apply (simp add: Abs_Real real_of_preal_def real_minus real_zero_def)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1359  | 
by (metis preal_add_assoc preal_add_commute)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1360  | 
next  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1361  | 
case equal  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1362  | 
then show ?thesis  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1363  | 
using Abs_Real real_zero_iff by blast  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1364  | 
next  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1365  | 
case greater  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1366  | 
then show ?thesis  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1367  | 
using less_add_left  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1368  | 
apply (simp add: Abs_Real real_of_preal_def real_minus real_zero_def)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1369  | 
by (metis preal_add_assoc preal_add_commute)  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1370  | 
qed  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1371  | 
qed  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1372  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1373  | 
lemma real_of_preal_less_iff [simp]:  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1374  | 
"(real_of_preal m1 < real_of_preal m2) = (m1 < m2)"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1375  | 
by (metis not_less preal_add_less_cancel_right real_le real_of_preal_def)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1376  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1377  | 
lemma real_of_preal_le_iff [simp]:  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1378  | 
"(real_of_preal m1 \<le> real_of_preal m2) = (m1 \<le> m2)"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1379  | 
by (simp add: linorder_not_less [symmetric])  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1380  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1381  | 
lemma real_of_preal_zero_less [simp]: "0 < real_of_preal m"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1382  | 
by (metis less_add_same_cancel2 preal_self_less_add_left real_of_preal_add real_of_preal_less_iff)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1383  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1384  | 
|
| 61343 | 1385  | 
subsection\<open>Theorems About the Ordering\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1386  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1387  | 
lemma real_gt_zero_preal_Ex: "(0 < x) \<longleftrightarrow> (\<exists>y. x = real_of_preal y)"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1388  | 
using order.asym real_of_preal_trichotomy by fastforce  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1389  | 
|
| 61343 | 1390  | 
subsection \<open>Completeness of Positive Reals\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1391  | 
|
| 61343 | 1392  | 
text \<open>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1393  | 
Supremum property for the set of positive reals  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1394  | 
|
| 61933 | 1395  | 
Let \<open>P\<close> be a non-empty set of positive reals, with an upper  | 
1396  | 
bound \<open>y\<close>. Then \<open>P\<close> has a least upper bound  | 
|
1397  | 
(written \<open>S\<close>).  | 
|
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1398  | 
|
| 61933 | 1399  | 
FIXME: Can the premise be weakened to \<open>\<forall>x \<in> P. x\<le> y\<close>?  | 
| 61343 | 1400  | 
\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1401  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1402  | 
lemma posreal_complete:  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1403  | 
assumes positive_P: "\<forall>x \<in> P. (0::real) < x"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1404  | 
and not_empty_P: "\<exists>x. x \<in> P"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1405  | 
and upper_bound_Ex: "\<exists>y. \<forall>x \<in> P. x<y"  | 
| 70201 | 1406  | 
shows "\<exists>s. \<forall>y. (\<exists>x \<in> P. y < x) = (y < s)"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1407  | 
proof (rule exI, rule allI)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1408  | 
fix y  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1409  | 
  let ?pP = "{w. real_of_preal w \<in> P}"
 | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1410  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1411  | 
show "(\<exists>x\<in>P. y < x) = (y < real_of_preal (psup ?pP))"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1412  | 
proof (cases "0 < y")  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1413  | 
assume neg_y: "\<not> 0 < y"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1414  | 
show ?thesis  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1415  | 
proof  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1416  | 
assume "\<exists>x\<in>P. y < x"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1417  | 
thus "y < real_of_preal (psup ?pP)"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1418  | 
by (metis dual_order.strict_trans neg_y not_less_iff_gr_or_eq real_of_preal_zero_less)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1419  | 
next  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1420  | 
assume "y < real_of_preal (psup ?pP)"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1421  | 
obtain "x" where x_in_P: "x \<in> P" using not_empty_P ..  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1422  | 
thus "\<exists>x \<in> P. y < x" using x_in_P  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1423  | 
using neg_y not_less_iff_gr_or_eq positive_P by fastforce  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1424  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1425  | 
next  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1426  | 
assume pos_y: "0 < y"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1427  | 
then obtain py where y_is_py: "y = real_of_preal py"  | 
| 70200 | 1428  | 
by (auto simp: real_gt_zero_preal_Ex)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1429  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1430  | 
obtain a where "a \<in> P" using not_empty_P ..  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1431  | 
with positive_P have a_pos: "0 < a" ..  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1432  | 
then obtain pa where "a = real_of_preal pa"  | 
| 70200 | 1433  | 
by (auto simp: real_gt_zero_preal_Ex)  | 
| 61343 | 1434  | 
hence "pa \<in> ?pP" using \<open>a \<in> P\<close> by auto  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1435  | 
    hence pP_not_empty: "?pP \<noteq> {}" by auto
 | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1436  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1437  | 
obtain sup where sup: "\<forall>x \<in> P. x < sup"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1438  | 
using upper_bound_Ex ..  | 
| 61343 | 1439  | 
from this and \<open>a \<in> P\<close> have "a < sup" ..  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1440  | 
hence "0 < sup" using a_pos by arith  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1441  | 
then obtain possup where "sup = real_of_preal possup"  | 
| 70200 | 1442  | 
by (auto simp: real_gt_zero_preal_Ex)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1443  | 
hence "\<forall>X \<in> ?pP. X \<le> possup"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1444  | 
using sup by auto  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1445  | 
with pP_not_empty have psup: "\<And>Z. (\<exists>X \<in> ?pP. Z < X) = (Z < psup ?pP)"  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1446  | 
by (meson preal_complete)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1447  | 
show ?thesis  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1448  | 
proof  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1449  | 
assume "\<exists>x \<in> P. y < x"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1450  | 
then obtain x where x_in_P: "x \<in> P" and y_less_x: "y < x" ..  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1451  | 
hence "0 < x" using pos_y by arith  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1452  | 
then obtain px where x_is_px: "x = real_of_preal px"  | 
| 70200 | 1453  | 
by (auto simp: real_gt_zero_preal_Ex)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1454  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1455  | 
have py_less_X: "\<exists>X \<in> ?pP. py < X"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1456  | 
proof  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1457  | 
show "py < px" using y_is_py and x_is_px and y_less_x  | 
| 70200 | 1458  | 
by simp  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1459  | 
show "px \<in> ?pP" using x_in_P and x_is_px by simp  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1460  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1461  | 
|
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1462  | 
have "(\<exists>X \<in> ?pP. py < X) \<Longrightarrow> (py < psup ?pP)"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1463  | 
using psup by simp  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1464  | 
hence "py < psup ?pP" using py_less_X by simp  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1465  | 
      thus "y < real_of_preal (psup {w. real_of_preal w \<in> P})"
 | 
| 70200 | 1466  | 
using y_is_py and pos_y by simp  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1467  | 
next  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1468  | 
assume y_less_psup: "y < real_of_preal (psup ?pP)"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1469  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1470  | 
hence "py < psup ?pP" using y_is_py  | 
| 70200 | 1471  | 
by simp  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1472  | 
then obtain "X" where py_less_X: "py < X" and X_in_pP: "X \<in> ?pP"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1473  | 
using psup by auto  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1474  | 
then obtain x where x_is_X: "x = real_of_preal X"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1475  | 
by (simp add: real_gt_zero_preal_Ex)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1476  | 
hence "y < x" using py_less_X and y_is_py  | 
| 70200 | 1477  | 
by simp  | 
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1478  | 
moreover have "x \<in> P"  | 
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1479  | 
using x_is_X and X_in_pP by simp  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1480  | 
ultimately show "\<exists> x \<in> P. y < x" ..  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1481  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1482  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1483  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1484  | 
|
| 
70199
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1485  | 
|
| 
 
b3630f5cc403
Massive restructuring; deleting unused theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
70197 
diff
changeset
 | 
1486  | 
subsection \<open>Completeness\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1487  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1488  | 
lemma reals_complete:  | 
| 
54263
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1489  | 
fixes S :: "real set"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1490  | 
assumes notempty_S: "\<exists>X. X \<in> S"  | 
| 
54263
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1491  | 
and exists_Ub: "bdd_above S"  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1492  | 
shows "\<exists>x. (\<forall>s\<in>S. s \<le> x) \<and> (\<forall>y. (\<forall>s\<in>S. s \<le> y) \<longrightarrow> x \<le> y)"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1493  | 
proof -  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1494  | 
obtain X where X_in_S: "X \<in> S" using notempty_S ..  | 
| 
54263
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1495  | 
obtain Y where Y_isUb: "\<forall>s\<in>S. s \<le> Y"  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1496  | 
using exists_Ub by (auto simp: bdd_above_def)  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1497  | 
  let ?SHIFT = "{z. \<exists>x \<in>S. z = x + (-X) + 1} \<inter> {x. 0 < x}"
 | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1498  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1499  | 
  {
 | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1500  | 
fix x  | 
| 
54263
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1501  | 
assume S_le_x: "\<forall>s\<in>S. s \<le> x"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1502  | 
    {
 | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1503  | 
fix s  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1504  | 
      assume "s \<in> {z. \<exists>x\<in>S. z = x + - X + 1}"
 | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1505  | 
hence "\<exists> x \<in> S. s = x + -X + 1" ..  | 
| 53373 | 1506  | 
then obtain x1 where x1: "x1 \<in> S" "s = x1 + (-X) + 1" ..  | 
1507  | 
then have "x1 \<le> x" using S_le_x by simp  | 
|
1508  | 
with x1 have "s \<le> x + - X + 1" by arith  | 
|
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1509  | 
}  | 
| 
54263
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1510  | 
then have "\<forall>s\<in>?SHIFT. s \<le> x + (-X) + 1"  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1511  | 
by auto  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1512  | 
} note S_Ub_is_SHIFT_Ub = this  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1513  | 
|
| 
54263
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1514  | 
have *: "\<forall>s\<in>?SHIFT. s \<le> Y + (-X) + 1" using Y_isUb by (rule S_Ub_is_SHIFT_Ub)  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1515  | 
have "\<forall>s\<in>?SHIFT. s < Y + (-X) + 2"  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1516  | 
proof  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1517  | 
fix s assume "s\<in>?SHIFT"  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1518  | 
with * have "s \<le> Y + (-X) + 1" by simp  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1519  | 
also have "\<dots> < Y + (-X) + 2" by simp  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1520  | 
finally show "s < Y + (-X) + 2" .  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1521  | 
qed  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1522  | 
moreover have "\<forall>y \<in> ?SHIFT. 0 < y" by auto  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1523  | 
moreover have shifted_not_empty: "\<exists>u. u \<in> ?SHIFT"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1524  | 
using X_in_S and Y_isUb by auto  | 
| 
54263
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1525  | 
ultimately obtain t where t_is_Lub: "\<forall>y. (\<exists>x\<in>?SHIFT. y < x) = (y < t)"  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1526  | 
using posreal_complete [of ?SHIFT] unfolding bdd_above_def by blast  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1527  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1528  | 
show ?thesis  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1529  | 
proof  | 
| 
54263
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1530  | 
show "(\<forall>s\<in>S. s \<le> (t + X + (-1))) \<and> (\<forall>y. (\<forall>s\<in>S. s \<le> y) \<longrightarrow> (t + X + (-1)) \<le> y)"  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1531  | 
proof safe  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1532  | 
fix x  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1533  | 
assume "\<forall>s\<in>S. s \<le> x"  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1534  | 
hence "\<forall>s\<in>?SHIFT. s \<le> x + (-X) + 1"  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1535  | 
using S_Ub_is_SHIFT_Ub by simp  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1536  | 
then have "\<not> x + (-X) + 1 < t"  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1537  | 
by (subst t_is_Lub[rule_format, symmetric]) (simp add: not_less)  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1538  | 
thus "t + X + -1 \<le> x" by arith  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1539  | 
next  | 
| 
54263
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1540  | 
fix y  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1541  | 
assume y_in_S: "y \<in> S"  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1542  | 
obtain "u" where u_in_shift: "u \<in> ?SHIFT" using shifted_not_empty ..  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1543  | 
hence "\<exists> x \<in> S. u = x + - X + 1" by simp  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1544  | 
then obtain "x" where x_and_u: "u = x + - X + 1" ..  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1545  | 
have u_le_t: "u \<le> t"  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1546  | 
proof (rule dense_le)  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1547  | 
fix x assume "x < u" then have "x < t"  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1548  | 
using u_in_shift t_is_Lub by auto  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1549  | 
then show "x \<le> t" by simp  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1550  | 
qed  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1551  | 
|
| 
54263
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1552  | 
show "y \<le> t + X + -1"  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1553  | 
proof cases  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1554  | 
assume "y \<le> x"  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1555  | 
moreover have "x = u + X + - 1" using x_and_u by arith  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1556  | 
moreover have "u + X + - 1 \<le> t + X + -1" using u_le_t by arith  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1557  | 
ultimately show "y \<le> t + X + -1" by arith  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1558  | 
next  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1559  | 
assume "~(y \<le> x)"  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1560  | 
hence x_less_y: "x < y" by arith  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1561  | 
|
| 
54263
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1562  | 
have "x + (-X) + 1 \<in> ?SHIFT" using x_and_u and u_in_shift by simp  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1563  | 
hence "0 < x + (-X) + 1" by simp  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1564  | 
hence "0 < y + (-X) + 1" using x_less_y by arith  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1565  | 
hence *: "y + (-X) + 1 \<in> ?SHIFT" using y_in_S by simp  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1566  | 
have "y + (-X) + 1 \<le> t"  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1567  | 
proof (rule dense_le)  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1568  | 
fix x assume "x < y + (-X) + 1" then have "x < t"  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1569  | 
using * t_is_Lub by auto  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1570  | 
then show "x \<le> t" by simp  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1571  | 
qed  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1572  | 
thus ?thesis by simp  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1573  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1574  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1575  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1576  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1577  | 
|
| 61343 | 1578  | 
subsection \<open>The Archimedean Property of the Reals\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1579  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1580  | 
theorem reals_Archimedean:  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1581  | 
fixes x :: real  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1582  | 
assumes x_pos: "0 < x"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1583  | 
shows "\<exists>n. inverse (of_nat (Suc n)) < x"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1584  | 
proof (rule ccontr)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1585  | 
assume contr: "\<not> ?thesis"  | 
| 
70197
 
e383580ffc35
partial updating to eliminate ASCII style and some applys
 
paulson <lp15@cam.ac.uk> 
parents: 
69597 
diff
changeset
 | 
1586  | 
have "\<forall>n. x * of_nat (Suc n) \<le> 1"  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1587  | 
proof  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1588  | 
fix n  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1589  | 
from contr have "x \<le> inverse (of_nat (Suc n))"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1590  | 
by (simp add: linorder_not_less)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1591  | 
hence "x \<le> (1 / (of_nat (Suc n)))"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1592  | 
by (simp add: inverse_eq_divide)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1593  | 
moreover have "(0::real) \<le> of_nat (Suc n)"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1594  | 
by (rule of_nat_0_le_iff)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1595  | 
ultimately have "x * of_nat (Suc n) \<le> (1 / of_nat (Suc n)) * of_nat (Suc n)"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1596  | 
by (rule mult_right_mono)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1597  | 
thus "x * of_nat (Suc n) \<le> 1" by (simp del: of_nat_Suc)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1598  | 
qed  | 
| 
54263
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1599  | 
  hence 2: "bdd_above {z. \<exists>n. z = x * (of_nat (Suc n))}"
 | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1600  | 
by (auto intro!: bdd_aboveI[of _ 1])  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1601  | 
  have 1: "\<exists>X. X \<in> {z. \<exists>n. z = x* (of_nat (Suc n))}" by auto
 | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1602  | 
obtain t where  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1603  | 
    upper: "\<And>z. z \<in> {z. \<exists>n. z = x * of_nat (Suc n)} \<Longrightarrow> z \<le> t" and
 | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1604  | 
    least: "\<And>y. (\<And>a. a \<in> {z. \<exists>n. z = x * of_nat (Suc n)} \<Longrightarrow> a \<le> y) \<Longrightarrow> t \<le> y"
 | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1605  | 
using reals_complete[OF 1 2] by auto  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1606  | 
|
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1607  | 
have "t \<le> t + - x"  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1608  | 
proof (rule least)  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1609  | 
    fix a assume a: "a \<in> {z. \<exists>n. z = x * (of_nat (Suc n))}"
 | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1610  | 
have "\<forall>n::nat. x * of_nat n \<le> t + - x"  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1611  | 
proof  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1612  | 
fix n  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1613  | 
have "x * of_nat (Suc n) \<le> t"  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1614  | 
by (simp add: upper)  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1615  | 
hence "x * (of_nat n) + x \<le> t"  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1616  | 
by (simp add: distrib_left)  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1617  | 
thus "x * (of_nat n) \<le> t + - x" by arith  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1618  | 
qed hence "\<forall>m. x * of_nat (Suc m) \<le> t + - x" by (simp del: of_nat_Suc)  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1619  | 
with a show "a \<le> t + - x"  | 
| 
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1620  | 
by auto  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1621  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1622  | 
thus False using x_pos by arith  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1623  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1624  | 
|
| 61343 | 1625  | 
text \<open>  | 
| 61933 | 1626  | 
There must be other proofs, e.g. \<open>Suc\<close> of the largest  | 
1627  | 
integer in the cut representing \<open>x\<close>.  | 
|
| 61343 | 1628  | 
\<close>  | 
| 
36793
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1629  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1630  | 
lemma reals_Archimedean2: "\<exists>n. (x::real) < of_nat (n::nat)"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1631  | 
proof cases  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1632  | 
assume "x \<le> 0"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1633  | 
hence "x < of_nat (1::nat)" by simp  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1634  | 
thus ?thesis ..  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1635  | 
next  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1636  | 
assume "\<not> x \<le> 0"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1637  | 
hence x_greater_zero: "0 < x" by simp  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1638  | 
hence "0 < inverse x" by simp  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1639  | 
then obtain n where "inverse (of_nat (Suc n)) < inverse x"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1640  | 
using reals_Archimedean by blast  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1641  | 
hence "inverse (of_nat (Suc n)) * x < inverse x * x"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1642  | 
using x_greater_zero by (rule mult_strict_right_mono)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1643  | 
hence "inverse (of_nat (Suc n)) * x < 1"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1644  | 
using x_greater_zero by simp  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1645  | 
hence "of_nat (Suc n) * (inverse (of_nat (Suc n)) * x) < of_nat (Suc n) * 1"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1646  | 
by (rule mult_strict_left_mono) (simp del: of_nat_Suc)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1647  | 
hence "x < of_nat (Suc n)"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1648  | 
by (simp add: algebra_simps del: of_nat_Suc)  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1649  | 
thus "\<exists>(n::nat). x < of_nat n" ..  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1650  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1651  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1652  | 
instance real :: archimedean_field  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1653  | 
proof  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1654  | 
fix r :: real  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1655  | 
obtain n :: nat where "r < of_nat n"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1656  | 
using reals_Archimedean2 ..  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1657  | 
then have "r \<le> of_int (int n)"  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1658  | 
by simp  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1659  | 
then show "\<exists>z. r \<le> of_int z" ..  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1660  | 
qed  | 
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1661  | 
|
| 
 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 
huffman 
parents:  
diff
changeset
 | 
1662  | 
end  |