| author | wenzelm | 
| Sat, 18 Jun 2011 12:13:42 +0200 | |
| changeset 43431 | f3d5cecfecdc | 
| parent 41541 | 1fa4725c4656 | 
| child 45694 | 4a8743618257 | 
| permissions | -rw-r--r-- | 
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1 | (* Title: HOL/ex/Dedekind_Real.thy | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2 | Author: Jacques D. Fleuriot, University of Cambridge | 
| 36794 | 3 | Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4 | 
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 4 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 5 | The positive reals as Dedekind sections of positive | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 6 | rationals. Fundamentals of Abstract Analysis [Gleason- p. 121] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 7 | provides some of the definitions. | 
| 
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 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 11 | imports Rat Lubs | 
| 
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 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 14 | section {* Positive real numbers *}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 15 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 16 | text{*Could be generalized and moved to @{text Groups}*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 17 | lemma add_eq_exists: "\<exists>x. a+x = (b::rat)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 18 | by (rule_tac x="b-a" in exI, simp) | 
| 
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 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 21 | cut :: "rat set => bool" where | 
| 37765 | 22 |   "cut A = ({} \<subset> A &
 | 
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 23 |             A < {r. 0 < r} &
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 24 | (\<forall>y \<in> A. ((\<forall>z. 0<z & z < y --> z \<in> A) & (\<exists>u \<in> A. y < u))))" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 25 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 26 | lemma interval_empty_iff: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 27 |   "{y. (x::'a::dense_linorder) < y \<and> y < z} = {} \<longleftrightarrow> \<not> x < z"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 28 | by (auto dest: dense) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 29 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 30 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 31 | lemma cut_of_rat: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 32 |   assumes q: "0 < q" shows "cut {r::rat. 0 < r & r < q}" (is "cut ?A")
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 33 | proof - | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 34 |   from q have pos: "?A < {r. 0 < r}" by force
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 35 |   have nonempty: "{} \<subset> ?A"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 36 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 37 |     show "{} \<subseteq> ?A" by simp
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 38 |     show "{} \<noteq> ?A"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 39 |       by (force simp only: q eq_commute [of "{}"] interval_empty_iff)
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 40 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 41 | show ?thesis | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 42 | by (simp add: cut_def pos nonempty, | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 43 | blast dest: dense intro: order_less_trans) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 44 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 45 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 46 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 47 | typedef preal = "{A. cut A}"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 48 | by (blast intro: cut_of_rat [OF zero_less_one]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 49 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 50 | definition | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 51 | psup :: "preal set => preal" where | 
| 37765 | 52 | "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 | 53 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 54 | definition | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 55 | add_set :: "[rat set,rat set] => rat set" where | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 56 |   "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 | 57 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 58 | definition | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 59 | diff_set :: "[rat set,rat set] => rat set" where | 
| 37765 | 60 |   "diff_set A B = {w. \<exists>x. 0 < w & 0 < x & x \<notin> B & x + w \<in> A}"
 | 
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 61 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 62 | definition | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 63 | mult_set :: "[rat set,rat set] => rat set" where | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 64 |   "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 | 65 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 66 | definition | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 67 | inverse_set :: "rat set => rat set" where | 
| 37765 | 68 |   "inverse_set A = {x. \<exists>y. 0 < x & x < y & inverse y \<notin> A}"
 | 
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 69 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 70 | instantiation preal :: "{ord, plus, minus, times, inverse, one}"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 71 | begin | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 72 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 73 | definition | 
| 37765 | 74 | preal_less_def: | 
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 75 | "R < S == Rep_preal R < Rep_preal S" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 76 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 77 | definition | 
| 37765 | 78 | preal_le_def: | 
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 79 | "R \<le> S == Rep_preal R \<subseteq> Rep_preal S" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 80 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 81 | definition | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 82 | preal_add_def: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 83 | "R + S == Abs_preal (add_set (Rep_preal R) (Rep_preal S))" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 84 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 85 | definition | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 86 | preal_diff_def: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 87 | "R - S == Abs_preal (diff_set (Rep_preal R) (Rep_preal S))" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 88 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 89 | definition | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 90 | preal_mult_def: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 91 | "R * S == Abs_preal (mult_set (Rep_preal R) (Rep_preal S))" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 92 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 93 | definition | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 94 | preal_inverse_def: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 95 | "inverse R == Abs_preal (inverse_set (Rep_preal R))" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 96 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 97 | definition "R / S = R * inverse (S\<Colon>preal)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 98 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 99 | definition | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 100 | preal_one_def: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 101 |     "1 == Abs_preal {x. 0 < x & x < 1}"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 102 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 103 | instance .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 104 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 105 | end | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 106 | |
| 
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 | text{*Reduces equality on abstractions to equality on representatives*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 109 | declare Abs_preal_inject [simp] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 110 | declare Abs_preal_inverse [simp] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 111 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 112 | lemma rat_mem_preal: "0 < q ==> {r::rat. 0 < r & r < q} \<in> preal"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 113 | by (simp add: preal_def cut_of_rat) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 114 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 115 | lemma preal_nonempty: "A \<in> preal ==> \<exists>x\<in>A. 0 < x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 116 | by (unfold preal_def cut_def, blast) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 117 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 118 | lemma preal_Ex_mem: "A \<in> preal \<Longrightarrow> \<exists>x. x \<in> A" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 119 | by (drule preal_nonempty, fast) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 120 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 121 | lemma preal_imp_psubset_positives: "A \<in> preal ==> A < {r. 0 < r}"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 122 | by (force simp add: preal_def cut_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 123 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 124 | lemma preal_exists_bound: "A \<in> preal ==> \<exists>x. 0 < x & x \<notin> A" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 125 | by (drule preal_imp_psubset_positives, auto) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 126 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 127 | lemma preal_exists_greater: "[| A \<in> preal; y \<in> A |] ==> \<exists>u \<in> A. y < u" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 128 | by (unfold preal_def cut_def, blast) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 129 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 130 | lemma preal_downwards_closed: "[| A \<in> preal; y \<in> A; 0 < z; z < y |] ==> z \<in> A" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 131 | by (unfold preal_def cut_def, blast) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 132 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 133 | text{*Relaxing the final premise*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 134 | lemma preal_downwards_closed': | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 135 | "[| A \<in> preal; y \<in> A; 0 < z; z \<le> y |] ==> z \<in> A" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 136 | apply (simp add: order_le_less) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 137 | apply (blast intro: preal_downwards_closed) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 138 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 139 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 140 | text{*A positive fraction not in a positive real is an upper bound.
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 141 | Gleason p. 122 - Remark (1)*} | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 142 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 143 | lemma not_in_preal_ub: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 144 | assumes A: "A \<in> preal" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 145 | and notx: "x \<notin> A" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 146 | and y: "y \<in> A" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 147 | and pos: "0 < x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 148 | shows "y < x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 149 | proof (cases rule: linorder_cases) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 150 | assume "x<y" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 151 | with notx show ?thesis | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 152 | 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 | 153 | next | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 154 | assume "x=y" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 155 | with notx and y show ?thesis by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 156 | next | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 157 | assume "y<x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 158 | thus ?thesis . | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 159 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 160 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 161 | text {* preal lemmas instantiated to @{term "Rep_preal X"} *}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 162 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 163 | lemma mem_Rep_preal_Ex: "\<exists>x. x \<in> Rep_preal X" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 164 | by (rule preal_Ex_mem [OF Rep_preal]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 165 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 166 | lemma Rep_preal_exists_bound: "\<exists>x>0. x \<notin> Rep_preal X" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 167 | by (rule preal_exists_bound [OF Rep_preal]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 168 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 169 | lemmas not_in_Rep_preal_ub = not_in_preal_ub [OF Rep_preal] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 170 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 171 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 172 | subsection{*Properties of Ordering*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 173 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 174 | instance preal :: order | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 175 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 176 | fix w :: preal | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 177 | 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 | 178 | next | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 179 | fix i j k :: preal | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 180 | assume "i \<le> j" and "j \<le> k" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 181 | 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 | 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 | assume "z \<le> w" and "w \<le> z" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 185 | 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 | 186 | next | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 187 | fix z w :: preal | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 188 | show "z < w \<longleftrightarrow> z \<le> w \<and> \<not> w \<le> z" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 189 | by (auto simp add: preal_le_def preal_less_def Rep_preal_inject) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 190 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 191 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 192 | lemma preal_imp_pos: "[|A \<in> preal; r \<in> A|] ==> 0 < r" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 193 | by (insert preal_imp_psubset_positives, blast) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 194 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 195 | instance preal :: linorder | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 196 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 197 | fix x y :: preal | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 198 | show "x <= y | y <= x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 199 | apply (auto simp add: preal_le_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 200 | apply (rule ccontr) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 201 | apply (blast dest: not_in_Rep_preal_ub intro: preal_imp_pos [OF Rep_preal] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 202 | elim: order_less_asym) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 203 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 204 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 205 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 206 | instantiation preal :: distrib_lattice | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 207 | begin | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 208 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 209 | definition | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 210 | "(inf \<Colon> preal \<Rightarrow> preal \<Rightarrow> preal) = min" | 
| 
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 | definition | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 213 | "(sup \<Colon> preal \<Rightarrow> preal \<Rightarrow> preal) = max" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 214 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 215 | instance | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 216 | by intro_classes | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 217 | (auto simp add: inf_preal_def sup_preal_def min_max.sup_inf_distrib1) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 218 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 219 | end | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 220 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 221 | subsection{*Properties of Addition*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 222 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 223 | lemma preal_add_commute: "(x::preal) + y = y + x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 224 | apply (unfold preal_add_def add_set_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 225 | apply (rule_tac f = Abs_preal in arg_cong) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 226 | apply (force simp add: add_commute) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 227 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 228 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 229 | text{*Lemmas for proving that addition of two positive reals gives
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 230 | a positive real*} | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 231 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 232 | text{*Part 1 of Dedekind sections definition*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 233 | lemma add_set_not_empty: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 234 |      "[|A \<in> preal; B \<in> preal|] ==> {} \<subset> add_set A B"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 235 | apply (drule preal_nonempty)+ | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 236 | apply (auto simp add: add_set_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 237 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 238 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 239 | text{*Part 2 of Dedekind sections definition.  A structured version of
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 240 | this proof is @{text preal_not_mem_mult_set_Ex} below.*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 241 | lemma preal_not_mem_add_set_Ex: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 242 | "[|A \<in> preal; B \<in> preal|] ==> \<exists>q>0. q \<notin> add_set A B" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 243 | apply (insert preal_exists_bound [of A] preal_exists_bound [of B], auto) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 244 | apply (rule_tac x = "x+xa" in exI) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 245 | apply (simp add: add_set_def, clarify) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 246 | apply (drule (3) not_in_preal_ub)+ | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 247 | apply (force dest: add_strict_mono) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 248 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 249 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 250 | lemma add_set_not_rat_set: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 251 | assumes A: "A \<in> preal" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 252 | and B: "B \<in> preal" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 253 |      shows "add_set A B < {r. 0 < r}"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 254 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 255 | from preal_imp_pos [OF A] preal_imp_pos [OF B] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 256 |   show "add_set A B \<subseteq> {r. 0 < r}" by (force simp add: add_set_def) 
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 257 | next | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 258 |   show "add_set A B \<noteq> {r. 0 < r}"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 259 | by (insert preal_not_mem_add_set_Ex [OF A B], blast) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 260 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 261 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 262 | text{*Part 3 of Dedekind sections definition*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 263 | lemma add_set_lemma3: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 264 | "[|A \<in> preal; B \<in> preal; u \<in> add_set A B; 0 < z; z < u|] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 265 | ==> z \<in> add_set A B" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 266 | proof (unfold add_set_def, clarify) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 267 | fix x::rat and y::rat | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 268 | assume A: "A \<in> preal" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 269 | and B: "B \<in> preal" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 270 | and [simp]: "0 < z" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 271 | and zless: "z < x + y" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 272 | and x: "x \<in> A" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 273 | and y: "y \<in> B" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 274 | have xpos [simp]: "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 | 275 | have ypos [simp]: "0<y" by (rule preal_imp_pos [OF B y]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 276 | have xypos [simp]: "0 < x+y" by (simp add: pos_add_strict) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 277 | let ?f = "z/(x+y)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 278 | have fless: "?f < 1" by (simp add: zless pos_divide_less_eq) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 279 | show "\<exists>x' \<in> A. \<exists>y'\<in>B. z = x' + y'" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 280 | proof (intro bexI) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 281 | show "z = x*?f + y*?f" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 282 | by (simp add: left_distrib [symmetric] divide_inverse mult_ac | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 283 | order_less_imp_not_eq2) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 284 | next | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 285 | show "y * ?f \<in> B" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 286 | proof (rule preal_downwards_closed [OF B y]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 287 | show "0 < y * ?f" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 288 | by (simp add: divide_inverse zero_less_mult_iff) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 289 | next | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 290 | show "y * ?f < y" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 291 | by (insert mult_strict_left_mono [OF fless ypos], simp) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 292 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 293 | next | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 294 | show "x * ?f \<in> A" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 295 | proof (rule preal_downwards_closed [OF A x]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 296 | show "0 < x * ?f" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 297 | by (simp add: divide_inverse zero_less_mult_iff) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 298 | next | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 299 | show "x * ?f < x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 300 | by (insert mult_strict_left_mono [OF fless xpos], simp) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 301 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 302 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 303 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 304 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 305 | text{*Part 4 of Dedekind sections definition*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 306 | lemma add_set_lemma4: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 307 | "[|A \<in> preal; B \<in> preal; y \<in> add_set A B|] ==> \<exists>u \<in> add_set A B. y < u" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 308 | apply (auto simp add: add_set_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 309 | apply (frule preal_exists_greater [of A], auto) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 310 | apply (rule_tac x="u + y" in exI) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 311 | apply (auto intro: add_strict_left_mono) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 312 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 313 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 314 | lemma mem_add_set: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 315 | "[|A \<in> preal; B \<in> preal|] ==> add_set A B \<in> preal" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 316 | apply (simp (no_asm_simp) add: preal_def cut_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 317 | apply (blast intro!: add_set_not_empty add_set_not_rat_set | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 318 | add_set_lemma3 add_set_lemma4) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 319 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 320 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 321 | lemma preal_add_assoc: "((x::preal) + y) + z = x + (y + z)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 322 | apply (simp add: preal_add_def mem_add_set Rep_preal) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 323 | apply (force simp add: add_set_def add_ac) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 324 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 325 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 326 | instance preal :: ab_semigroup_add | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 327 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 328 | fix a b c :: preal | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 329 | 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 | 330 | 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 | 331 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 332 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 333 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 334 | subsection{*Properties of Multiplication*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 335 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 336 | text{*Proofs essentially same as for addition*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 337 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 338 | lemma preal_mult_commute: "(x::preal) * y = y * x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 339 | apply (unfold preal_mult_def mult_set_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 340 | apply (rule_tac f = Abs_preal in arg_cong) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 341 | apply (force simp add: mult_commute) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 342 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 343 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 344 | text{*Multiplication of two positive reals gives a positive real.*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 345 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 346 | text{*Lemmas for proving positive reals multiplication set in @{typ preal}*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 347 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 348 | text{*Part 1 of Dedekind sections definition*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 349 | lemma mult_set_not_empty: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 350 |      "[|A \<in> preal; B \<in> preal|] ==> {} \<subset> mult_set A B"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 351 | apply (insert preal_nonempty [of A] preal_nonempty [of B]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 352 | apply (auto simp add: mult_set_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 353 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 354 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 355 | text{*Part 2 of Dedekind sections definition*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 356 | lemma preal_not_mem_mult_set_Ex: | 
| 41541 | 357 | assumes A: "A \<in> preal" | 
| 358 | and B: "B \<in> preal" | |
| 359 | shows "\<exists>q. 0 < q & q \<notin> mult_set A B" | |
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 360 | proof - | 
| 41541 | 361 | from preal_exists_bound [OF A] obtain x where 1 [simp]: "0 < x" "x \<notin> A" by blast | 
| 362 | from preal_exists_bound [OF B] obtain y where 2 [simp]: "0 < y" "y \<notin> B" by blast | |
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 363 | show ?thesis | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 364 | proof (intro exI conjI) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 365 | show "0 < x*y" by (simp add: mult_pos_pos) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 366 | show "x * y \<notin> mult_set A B" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 367 | proof - | 
| 41541 | 368 |       {
 | 
| 369 | fix u::rat and v::rat | |
| 370 | assume u: "u \<in> A" and v: "v \<in> B" and xy: "x*y = u*v" | |
| 371 | moreover from A B 1 2 u v have "u<x" and "v<y" by (blast dest: not_in_preal_ub)+ | |
| 372 | moreover | |
| 373 | from A B 1 2 u v have "0\<le>v" | |
| 374 | by (blast intro: preal_imp_pos [OF B] order_less_imp_le) | |
| 375 | moreover | |
| 376 | from A B 1 `u < x` `v < y` `0 \<le> v` | |
| 377 | have "u*v < x*y" by (blast intro: mult_strict_mono) | |
| 378 | ultimately have False by force | |
| 379 | } | |
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 380 | thus ?thesis by (auto simp add: mult_set_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 381 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 382 | qed | 
| 
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 | lemma mult_set_not_rat_set: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 386 | assumes A: "A \<in> preal" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 387 | and B: "B \<in> preal" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 388 |   shows "mult_set A B < {r. 0 < r}"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 389 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 390 |   show "mult_set A B \<subseteq> {r. 0 < r}"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 391 | by (force simp add: mult_set_def | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 392 | intro: preal_imp_pos [OF A] preal_imp_pos [OF B] mult_pos_pos) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 393 |   show "mult_set A B \<noteq> {r. 0 < r}"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 394 | using preal_not_mem_mult_set_Ex [OF A B] by blast | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 395 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 396 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 397 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 398 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 399 | text{*Part 3 of Dedekind sections definition*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 400 | lemma mult_set_lemma3: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 401 | "[|A \<in> preal; B \<in> preal; u \<in> mult_set A B; 0 < z; z < u|] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 402 | ==> z \<in> mult_set A B" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 403 | proof (unfold mult_set_def, clarify) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 404 | fix x::rat and y::rat | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 405 | assume A: "A \<in> preal" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 406 | and B: "B \<in> preal" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 407 | and [simp]: "0 < z" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 408 | and zless: "z < x * y" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 409 | and x: "x \<in> A" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 410 | and y: "y \<in> B" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 411 | have [simp]: "0<y" by (rule preal_imp_pos [OF B y]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 412 | show "\<exists>x' \<in> A. \<exists>y' \<in> B. z = x' * y'" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 413 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 414 | show "\<exists>y'\<in>B. z = (z/y) * y'" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 415 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 416 | show "z = (z/y)*y" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 417 | by (simp add: divide_inverse mult_commute [of y] mult_assoc | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 418 | order_less_imp_not_eq2) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 419 | show "y \<in> B" by fact | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 420 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 421 | next | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 422 | show "z/y \<in> A" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 423 | proof (rule preal_downwards_closed [OF A x]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 424 | show "0 < z/y" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 425 | by (simp add: zero_less_divide_iff) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 426 | show "z/y < x" by (simp add: pos_divide_less_eq zless) | 
| 
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 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 429 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 430 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 431 | text{*Part 4 of Dedekind sections definition*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 432 | lemma mult_set_lemma4: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 433 | "[|A \<in> preal; B \<in> preal; y \<in> mult_set A B|] ==> \<exists>u \<in> mult_set A B. y < u" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 434 | apply (auto simp add: mult_set_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 435 | apply (frule preal_exists_greater [of A], auto) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 436 | apply (rule_tac x="u * y" in exI) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 437 | apply (auto intro: preal_imp_pos [of A] preal_imp_pos [of B] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 438 | mult_strict_right_mono) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 439 | done | 
| 
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 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 442 | lemma mem_mult_set: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 443 | "[|A \<in> preal; B \<in> preal|] ==> mult_set A B \<in> preal" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 444 | apply (simp (no_asm_simp) add: preal_def cut_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 445 | apply (blast intro!: mult_set_not_empty mult_set_not_rat_set | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 446 | mult_set_lemma3 mult_set_lemma4) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 447 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 448 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 449 | lemma preal_mult_assoc: "((x::preal) * y) * z = x * (y * z)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 450 | apply (simp add: preal_mult_def mem_mult_set Rep_preal) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 451 | apply (force simp add: mult_set_def mult_ac) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 452 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 453 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 454 | instance preal :: ab_semigroup_mult | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 455 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 456 | fix a b c :: preal | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 457 | 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 | 458 | 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 | 459 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 460 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 461 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 462 | text{* Positive real 1 is the multiplicative identity element *}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 463 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 464 | lemma preal_mult_1: "(1::preal) * z = z" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 465 | proof (induct z) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 466 | fix A :: "rat set" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 467 | assume A: "A \<in> preal" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 468 |   have "{w. \<exists>u. 0 < u \<and> u < 1 & (\<exists>v \<in> A. w = u * v)} = A" (is "?lhs = A")
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 469 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 470 | show "?lhs \<subseteq> A" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 471 | proof clarify | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 472 | fix x::rat and u::rat and v::rat | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 473 | 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 | 474 | have vpos: "0<v" by (rule preal_imp_pos [OF A v]) | 
| 41541 | 475 | hence "u*v < 1*v" by (simp only: mult_strict_right_mono upos `u < 1` v) | 
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 476 | thus "u * v \<in> A" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 477 | by (force intro: preal_downwards_closed [OF A v] mult_pos_pos | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 478 | upos vpos) | 
| 
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 | next | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 481 | show "A \<subseteq> ?lhs" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 482 | proof clarify | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 483 | fix x::rat | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 484 | assume x: "x \<in> A" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 485 | 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 | 486 | from preal_exists_greater [OF A x] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 487 | 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 | 488 | 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 | 489 | 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 | 490 | proof (intro exI conjI) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 491 | show "0 < x/v" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 492 | by (simp add: zero_less_divide_iff xpos vpos) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 493 | show "x / v < 1" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 494 | by (simp add: pos_divide_less_eq vpos xlessv) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 495 | show "\<exists>v'\<in>A. x = (x / v) * v'" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 496 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 497 | show "x = (x/v)*v" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 498 | by (simp add: divide_inverse mult_assoc vpos | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 499 | order_less_imp_not_eq2) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 500 | show "v \<in> A" by fact | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 501 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 502 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 503 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 504 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 505 | thus "1 * Abs_preal A = Abs_preal A" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 506 | by (simp add: preal_one_def preal_mult_def mult_set_def | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 507 | rat_mem_preal A) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 508 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 509 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 510 | instance preal :: comm_monoid_mult | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 511 | by intro_classes (rule preal_mult_1) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 512 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 513 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 514 | subsection{*Distribution of Multiplication across Addition*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 515 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 516 | lemma mem_Rep_preal_add_iff: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 517 | "(z \<in> Rep_preal(R+S)) = (\<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. z = x + y)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 518 | apply (simp add: preal_add_def mem_add_set Rep_preal) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 519 | apply (simp add: add_set_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 520 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 521 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 522 | lemma mem_Rep_preal_mult_iff: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 523 | "(z \<in> Rep_preal(R*S)) = (\<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. z = x * y)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 524 | apply (simp add: preal_mult_def mem_mult_set Rep_preal) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 525 | apply (simp add: mult_set_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 526 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 527 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 528 | lemma distrib_subset1: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 529 | "Rep_preal (w * (x + y)) \<subseteq> Rep_preal (w * x + w * y)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 530 | apply (auto simp add: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_mult_iff) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 531 | apply (force simp add: right_distrib) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 532 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 533 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 534 | lemma preal_add_mult_distrib_mean: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 535 | assumes a: "a \<in> Rep_preal w" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 536 | and b: "b \<in> Rep_preal w" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 537 | and d: "d \<in> Rep_preal x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 538 | and e: "e \<in> Rep_preal y" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 539 | 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 | 540 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 541 | let ?c = "(a*d + b*e)/(d+e)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 542 | have [simp]: "0<a" "0<b" "0<d" "0<e" "0<d+e" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 543 | by (blast intro: preal_imp_pos [OF Rep_preal] a b d e pos_add_strict)+ | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 544 | have cpos: "0 < ?c" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 545 | 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 | 546 | show "a * d + b * e = ?c * (d + e)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 547 | by (simp add: divide_inverse mult_assoc order_less_imp_not_eq2) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 548 | show "?c \<in> Rep_preal w" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 549 | proof (cases rule: linorder_le_cases) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 550 | assume "a \<le> b" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 551 | hence "?c \<le> b" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 552 | by (simp add: pos_divide_le_eq right_distrib mult_right_mono | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 553 | order_less_imp_le) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 554 | thus ?thesis by (rule preal_downwards_closed' [OF Rep_preal b cpos]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 555 | next | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 556 | assume "b \<le> a" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 557 | hence "?c \<le> a" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 558 | by (simp add: pos_divide_le_eq right_distrib mult_right_mono | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 559 | order_less_imp_le) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 560 | thus ?thesis by (rule preal_downwards_closed' [OF Rep_preal a cpos]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 561 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 562 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 563 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 564 | lemma distrib_subset2: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 565 | "Rep_preal (w * x + w * y) \<subseteq> Rep_preal (w * (x + y))" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 566 | apply (auto simp add: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_mult_iff) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 567 | apply (drule_tac w=w and x=x and y=y in preal_add_mult_distrib_mean, auto) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 568 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 569 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 570 | lemma preal_add_mult_distrib2: "(w * ((x::preal) + y)) = (w * x) + (w * y)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 571 | apply (rule Rep_preal_inject [THEN iffD1]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 572 | apply (rule equalityI [OF distrib_subset1 distrib_subset2]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 573 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 574 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 575 | lemma preal_add_mult_distrib: "(((x::preal) + y) * w) = (x * w) + (y * w)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 576 | by (simp add: preal_mult_commute preal_add_mult_distrib2) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 577 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 578 | instance preal :: comm_semiring | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 579 | by intro_classes (rule preal_add_mult_distrib) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 580 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 581 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 582 | subsection{*Existence of Inverse, a Positive Real*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 583 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 584 | lemma mem_inv_set_ex: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 585 | assumes A: "A \<in> preal" shows "\<exists>x y. 0 < x & x < y & inverse y \<notin> A" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 586 | proof - | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 587 | from preal_exists_bound [OF A] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 588 | obtain x where [simp]: "0<x" "x \<notin> A" by blast | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 589 | show ?thesis | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 590 | proof (intro exI conjI) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 591 | show "0 < inverse (x+1)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 592 | by (simp add: order_less_trans [OF _ less_add_one]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 593 | show "inverse(x+1) < inverse x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 594 | by (simp add: less_imp_inverse_less less_add_one) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 595 | show "inverse (inverse x) \<notin> A" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 596 | by (simp add: order_less_imp_not_eq2) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 597 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 598 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 599 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 600 | text{*Part 1 of Dedekind sections definition*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 601 | lemma inverse_set_not_empty: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 602 |      "A \<in> preal ==> {} \<subset> inverse_set A"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 603 | apply (insert mem_inv_set_ex [of A]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 604 | apply (auto simp add: inverse_set_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 605 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 606 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 607 | text{*Part 2 of Dedekind sections definition*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 608 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 609 | lemma preal_not_mem_inverse_set_Ex: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 610 | assumes A: "A \<in> preal" shows "\<exists>q. 0 < q & q \<notin> inverse_set A" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 611 | proof - | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 612 | from preal_nonempty [OF A] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 613 | obtain x where x: "x \<in> A" and xpos [simp]: "0<x" .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 614 | show ?thesis | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 615 | proof (intro exI conjI) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 616 | show "0 < inverse x" by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 617 | show "inverse x \<notin> inverse_set A" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 618 | proof - | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 619 |       { fix y::rat 
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 620 | assume ygt: "inverse x < y" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 621 | have [simp]: "0 < y" by (simp add: order_less_trans [OF _ ygt]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 622 | have iyless: "inverse y < x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 623 | by (simp add: inverse_less_imp_less [of x] ygt) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 624 | have "inverse y \<in> A" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 625 | by (simp add: preal_downwards_closed [OF A x] iyless)} | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 626 | thus ?thesis by (auto simp add: inverse_set_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 627 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 628 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 629 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 630 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 631 | lemma inverse_set_not_rat_set: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 632 |    assumes A: "A \<in> preal"  shows "inverse_set A < {r. 0 < r}"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 633 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 634 |   show "inverse_set A \<subseteq> {r. 0 < r}"  by (force simp add: inverse_set_def)
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 635 | next | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 636 |   show "inverse_set A \<noteq> {r. 0 < r}"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 637 | by (insert preal_not_mem_inverse_set_Ex [OF A], blast) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 638 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 639 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 640 | text{*Part 3 of Dedekind sections definition*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 641 | lemma inverse_set_lemma3: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 642 | "[|A \<in> preal; u \<in> inverse_set A; 0 < z; z < u|] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 643 | ==> z \<in> inverse_set A" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 644 | apply (auto simp add: inverse_set_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 645 | apply (auto intro: order_less_trans) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 646 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 647 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 648 | text{*Part 4 of Dedekind sections definition*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 649 | lemma inverse_set_lemma4: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 650 | "[|A \<in> preal; y \<in> inverse_set A|] ==> \<exists>u \<in> inverse_set A. y < u" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 651 | apply (auto simp add: inverse_set_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 652 | apply (drule dense [of y]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 653 | apply (blast intro: order_less_trans) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 654 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 655 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 656 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 657 | lemma mem_inverse_set: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 658 | "A \<in> preal ==> inverse_set A \<in> preal" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 659 | apply (simp (no_asm_simp) add: preal_def cut_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 660 | apply (blast intro!: inverse_set_not_empty inverse_set_not_rat_set | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 661 | inverse_set_lemma3 inverse_set_lemma4) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 662 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 663 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 664 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 665 | subsection{*Gleason's Lemma 9-3.4, page 122*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 666 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 667 | lemma Gleason9_34_exists: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 668 | assumes A: "A \<in> preal" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 669 | and "\<forall>x\<in>A. x + u \<in> A" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 670 | and "0 \<le> z" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 671 | 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 | 672 | proof (cases z rule: int_cases) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 673 | case (nonneg n) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 674 | show ?thesis | 
| 41541 | 675 | proof (simp add: nonneg, induct n) | 
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 676 | case 0 | 
| 41541 | 677 | from preal_nonempty [OF A] | 
| 678 | show ?case by force | |
| 679 | next | |
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 680 | case (Suc k) | 
| 41541 | 681 | then obtain b where b: "b \<in> A" "b + of_nat k * u \<in> A" .. | 
| 682 | hence "b + of_int (int k)*u + u \<in> A" by (simp add: assms) | |
| 683 | thus ?case by (force simp add: algebra_simps b) | |
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 684 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 685 | next | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 686 | case (neg n) | 
| 41541 | 687 | with assms show ?thesis by simp | 
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 688 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 689 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 690 | lemma Gleason9_34_contra: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 691 | assumes A: "A \<in> preal" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 692 | shows "[|\<forall>x\<in>A. x + u \<in> A; 0 < u; 0 < y; y \<notin> A|] ==> False" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 693 | proof (induct u, induct y) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 694 | fix a::int and b::int | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 695 | fix c::int and d::int | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 696 | assume bpos [simp]: "0 < b" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 697 | and dpos [simp]: "0 < d" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 698 | 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 | 699 | and upos: "0 < Fract c d" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 700 | and ypos: "0 < Fract a b" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 701 | and notin: "Fract a b \<notin> A" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 702 | have cpos [simp]: "0 < c" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 703 | 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 | 704 | have apos [simp]: "0 < a" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 705 | 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 | 706 | let ?k = "a*d" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 707 | 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 | 708 | proof - | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 709 | have "?thesis = ((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 | 710 | by (simp add: order_less_imp_not_eq2 mult_ac) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 711 | moreover | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 712 | 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 | 713 | by (rule mult_mono, | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 714 | 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 | 715 | order_less_imp_le) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 716 | ultimately | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 717 | show ?thesis by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 718 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 719 | 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 | 720 | from Gleason9_34_exists [OF A closed k] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 721 | obtain z where z: "z \<in> A" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 722 | 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 | 723 | 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 | 724 | 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 | 725 | 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 | 726 | 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 | 727 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 728 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 729 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 730 | lemma Gleason9_34: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 731 | assumes A: "A \<in> preal" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 732 | and upos: "0 < u" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 733 | shows "\<exists>r \<in> A. r + u \<notin> A" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 734 | proof (rule ccontr, simp) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 735 | assume closed: "\<forall>r\<in>A. r + u \<in> A" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 736 | from preal_exists_bound [OF A] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 737 | obtain y where y: "y \<notin> A" and ypos: "0 < y" by blast | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 738 | show False | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 739 | by (rule Gleason9_34_contra [OF A closed upos ypos y]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 740 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 741 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 742 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 743 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 744 | subsection{*Gleason's Lemma 9-3.6*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 745 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 746 | lemma lemma_gleason9_36: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 747 | assumes A: "A \<in> preal" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 748 | and x: "1 < x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 749 | shows "\<exists>r \<in> A. r*x \<notin> A" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 750 | proof - | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 751 | from preal_nonempty [OF A] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 752 | 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 | 753 | show ?thesis | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 754 | proof (rule classical) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 755 | assume "~(\<exists>r\<in>A. r * x \<notin> A)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 756 | 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 | 757 | from ypos mult_strict_left_mono [OF x] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 758 | have yless: "y < y*x" by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 759 | let ?d = "y*x - y" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 760 | 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 | 761 | from Gleason9_34 [OF A dpos] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 762 | 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 | 763 | 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 | 764 | with dpos have rdpos: "0 < r + ?d" by arith | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 765 | have "~ (r + ?d \<le> y + ?d)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 766 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 767 | assume le: "r + ?d \<le> y + ?d" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 768 | 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 | 769 | 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 | 770 | with notin show False by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 771 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 772 | hence "y < r" by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 773 | with ypos have dless: "?d < (r * ?d)/y" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 774 | by (simp add: pos_less_divide_eq mult_commute [of ?d] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 775 | mult_strict_right_mono dpos) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 776 | have "r + ?d < r*x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 777 | proof - | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 778 | have "r + ?d < r + (r * ?d)/y" by (simp add: dless) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 779 | also with ypos have "... = (r/y) * (y + ?d)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 780 | by (simp only: algebra_simps divide_inverse, simp) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 781 | also have "... = r*x" using ypos | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 782 | by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 783 | finally show "r + ?d < r*x" . | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 784 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 785 | with r notin rdpos | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 786 | 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 | 787 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 788 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 789 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 790 | subsection{*Existence of Inverse: Part 2*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 791 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 792 | lemma mem_Rep_preal_inverse_iff: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 793 | "(z \<in> Rep_preal(inverse R)) = | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 794 | (0 < z \<and> (\<exists>y. z < y \<and> inverse y \<notin> Rep_preal R))" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 795 | apply (simp add: preal_inverse_def mem_inverse_set Rep_preal) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 796 | apply (simp add: inverse_set_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 797 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 798 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 799 | lemma Rep_preal_one: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 800 |      "Rep_preal 1 = {x. 0 < x \<and> x < 1}"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 801 | by (simp add: preal_one_def rat_mem_preal) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 802 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 803 | lemma subset_inverse_mult_lemma: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 804 | assumes xpos: "0 < x" and xless: "x < 1" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 805 | shows "\<exists>r u y. 0 < r & r < y & inverse y \<notin> Rep_preal R & | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 806 | u \<in> Rep_preal R & x = r * u" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 807 | proof - | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 808 | from xpos and xless have "1 < inverse x" by (simp add: one_less_inverse_iff) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 809 | from lemma_gleason9_36 [OF Rep_preal this] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 810 | obtain r where r: "r \<in> Rep_preal R" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 811 | and notin: "r * (inverse x) \<notin> Rep_preal R" .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 812 | have rpos: "0<r" by (rule preal_imp_pos [OF Rep_preal r]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 813 | from preal_exists_greater [OF Rep_preal r] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 814 | obtain u where u: "u \<in> Rep_preal R" and rless: "r < u" .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 815 | have upos: "0<u" by (rule preal_imp_pos [OF Rep_preal u]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 816 | show ?thesis | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 817 | proof (intro exI conjI) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 818 | show "0 < x/u" using xpos upos | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 819 | by (simp add: zero_less_divide_iff) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 820 | show "x/u < x/r" using xpos upos rpos | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 821 | by (simp add: divide_inverse mult_less_cancel_left rless) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 822 | show "inverse (x / r) \<notin> Rep_preal R" using notin | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 823 | by (simp add: divide_inverse mult_commute) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 824 | show "u \<in> Rep_preal R" by (rule u) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 825 | show "x = x / u * u" using upos | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 826 | by (simp add: divide_inverse mult_commute) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 827 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 828 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 829 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 830 | lemma subset_inverse_mult: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 831 | "Rep_preal 1 \<subseteq> Rep_preal(inverse R * R)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 832 | apply (auto simp add: Bex_def Rep_preal_one mem_Rep_preal_inverse_iff | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 833 | mem_Rep_preal_mult_iff) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 834 | apply (blast dest: subset_inverse_mult_lemma) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 835 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 836 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 837 | lemma inverse_mult_subset_lemma: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 838 | assumes rpos: "0 < r" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 839 | and rless: "r < y" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 840 | and notin: "inverse y \<notin> Rep_preal R" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 841 | and q: "q \<in> Rep_preal R" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 842 | shows "r*q < 1" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 843 | proof - | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 844 | have "q < inverse y" using rpos rless | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 845 | by (simp add: not_in_preal_ub [OF Rep_preal notin] q) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 846 | hence "r * q < r/y" using rpos | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 847 | by (simp add: divide_inverse mult_less_cancel_left) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 848 | also have "... \<le> 1" using rpos rless | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 849 | by (simp add: pos_divide_le_eq) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 850 | finally show ?thesis . | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 851 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 852 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 853 | lemma inverse_mult_subset: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 854 | "Rep_preal(inverse R * R) \<subseteq> Rep_preal 1" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 855 | apply (auto simp add: Bex_def Rep_preal_one mem_Rep_preal_inverse_iff | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 856 | mem_Rep_preal_mult_iff) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 857 | apply (simp add: zero_less_mult_iff preal_imp_pos [OF Rep_preal]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 858 | apply (blast intro: inverse_mult_subset_lemma) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 859 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 860 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 861 | lemma preal_mult_inverse: "inverse R * R = (1::preal)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 862 | apply (rule Rep_preal_inject [THEN iffD1]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 863 | apply (rule equalityI [OF inverse_mult_subset subset_inverse_mult]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 864 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 865 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 866 | lemma preal_mult_inverse_right: "R * inverse R = (1::preal)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 867 | apply (rule preal_mult_commute [THEN subst]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 868 | apply (rule preal_mult_inverse) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 869 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 870 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 871 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 872 | text{*Theorems needing @{text Gleason9_34}*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 873 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 874 | lemma Rep_preal_self_subset: "Rep_preal (R) \<subseteq> Rep_preal(R + S)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 875 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 876 | fix r | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 877 | assume r: "r \<in> Rep_preal R" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 878 | have rpos: "0<r" by (rule preal_imp_pos [OF Rep_preal r]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 879 | from mem_Rep_preal_Ex | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 880 | obtain y where y: "y \<in> Rep_preal S" .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 881 | have ypos: "0<y" by (rule preal_imp_pos [OF Rep_preal y]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 882 | have ry: "r+y \<in> Rep_preal(R + S)" using r y | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 883 | by (auto simp add: mem_Rep_preal_add_iff) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 884 | show "r \<in> Rep_preal(R + S)" using r ypos rpos | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 885 | by (simp add: preal_downwards_closed [OF Rep_preal ry]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 886 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 887 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 888 | lemma Rep_preal_sum_not_subset: "~ Rep_preal (R + S) \<subseteq> Rep_preal(R)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 889 | proof - | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 890 | from mem_Rep_preal_Ex | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 891 | obtain y where y: "y \<in> Rep_preal S" .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 892 | have ypos: "0<y" by (rule preal_imp_pos [OF Rep_preal y]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 893 | from Gleason9_34 [OF Rep_preal ypos] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 894 | obtain r where r: "r \<in> Rep_preal R" and notin: "r + y \<notin> Rep_preal R" .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 895 | have "r + y \<in> Rep_preal (R + S)" using r y | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 896 | by (auto simp add: mem_Rep_preal_add_iff) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 897 | thus ?thesis using notin by blast | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 898 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 899 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 900 | lemma Rep_preal_sum_not_eq: "Rep_preal (R + S) \<noteq> Rep_preal(R)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 901 | by (insert Rep_preal_sum_not_subset, blast) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 902 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 903 | text{*at last, Gleason prop. 9-3.5(iii) page 123*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 904 | lemma preal_self_less_add_left: "(R::preal) < R + S" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 905 | apply (unfold preal_less_def less_le) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 906 | apply (simp add: Rep_preal_self_subset Rep_preal_sum_not_eq [THEN not_sym]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 907 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 908 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 909 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 910 | subsection{*Subtraction for Positive Reals*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 911 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 912 | text{*Gleason prop. 9-3.5(iv), page 123: proving @{prop "A < B ==> \<exists>D. A + D =
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 913 | B"}. We define the claimed @{term D} and show that it is a positive real*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 914 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 915 | text{*Part 1 of Dedekind sections definition*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 916 | lemma diff_set_not_empty: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 917 |      "R < S ==> {} \<subset> diff_set (Rep_preal S) (Rep_preal R)"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 918 | apply (auto simp add: preal_less_def diff_set_def elim!: equalityE) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 919 | apply (frule_tac x1 = S in Rep_preal [THEN preal_exists_greater]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 920 | apply (drule preal_imp_pos [OF Rep_preal], clarify) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 921 | apply (cut_tac a=x and b=u in add_eq_exists, force) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 922 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 923 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 924 | text{*Part 2 of Dedekind sections definition*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 925 | lemma diff_set_nonempty: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 926 | "\<exists>q. 0 < q & q \<notin> diff_set (Rep_preal S) (Rep_preal R)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 927 | apply (cut_tac X = S in Rep_preal_exists_bound) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 928 | apply (erule exE) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 929 | apply (rule_tac x = x in exI, auto) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 930 | apply (simp add: diff_set_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 931 | apply (auto dest: Rep_preal [THEN preal_downwards_closed]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 932 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 933 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 934 | lemma diff_set_not_rat_set: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 935 |   "diff_set (Rep_preal S) (Rep_preal R) < {r. 0 < r}" (is "?lhs < ?rhs")
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 936 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 937 | show "?lhs \<subseteq> ?rhs" by (auto simp add: diff_set_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 938 | show "?lhs \<noteq> ?rhs" using diff_set_nonempty by blast | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 939 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 940 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 941 | text{*Part 3 of Dedekind sections definition*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 942 | lemma diff_set_lemma3: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 943 | "[|R < S; u \<in> diff_set (Rep_preal S) (Rep_preal R); 0 < z; z < u|] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 944 | ==> z \<in> diff_set (Rep_preal S) (Rep_preal R)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 945 | apply (auto simp add: diff_set_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 946 | apply (rule_tac x=x in exI) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 947 | apply (drule Rep_preal [THEN preal_downwards_closed], auto) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 948 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 949 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 950 | text{*Part 4 of Dedekind sections definition*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 951 | lemma diff_set_lemma4: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 952 | "[|R < S; y \<in> diff_set (Rep_preal S) (Rep_preal R)|] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 953 | ==> \<exists>u \<in> diff_set (Rep_preal S) (Rep_preal R). y < u" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 954 | apply (auto simp add: diff_set_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 955 | apply (drule Rep_preal [THEN preal_exists_greater], clarify) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 956 | apply (cut_tac a="x+y" and b=u in add_eq_exists, clarify) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 957 | apply (rule_tac x="y+xa" in exI) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 958 | apply (auto simp add: add_ac) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 959 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 960 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 961 | lemma mem_diff_set: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 962 | "R < S ==> diff_set (Rep_preal S) (Rep_preal R) \<in> preal" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 963 | apply (unfold preal_def cut_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 964 | apply (blast intro!: diff_set_not_empty diff_set_not_rat_set | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 965 | diff_set_lemma3 diff_set_lemma4) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 966 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 967 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 968 | lemma mem_Rep_preal_diff_iff: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 969 | "R < S ==> | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 970 | (z \<in> Rep_preal(S-R)) = | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 971 | (\<exists>x. 0 < x & 0 < z & x \<notin> Rep_preal R & x + z \<in> Rep_preal S)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 972 | apply (simp add: preal_diff_def mem_diff_set Rep_preal) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 973 | apply (force simp add: diff_set_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 974 | done | 
| 
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 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 977 | text{*proving that @{term "R + D \<le> S"}*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 978 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 979 | lemma less_add_left_lemma: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 980 | assumes Rless: "R < S" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 981 | and a: "a \<in> Rep_preal R" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 982 | and cb: "c + b \<in> Rep_preal S" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 983 | and "c \<notin> Rep_preal R" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 984 | and "0 < b" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 985 | and "0 < c" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 986 | shows "a + b \<in> Rep_preal S" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 987 | proof - | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 988 | have "0<a" by (rule preal_imp_pos [OF Rep_preal a]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 989 | moreover | 
| 41541 | 990 | have "a < c" using assms by (blast intro: not_in_Rep_preal_ub ) | 
| 991 | ultimately show ?thesis | |
| 992 | using assms by (simp add: preal_downwards_closed [OF Rep_preal cb]) | |
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 993 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 994 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 995 | lemma less_add_left_le1: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 996 | "R < (S::preal) ==> R + (S-R) \<le> S" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 997 | apply (auto simp add: Bex_def preal_le_def mem_Rep_preal_add_iff | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 998 | mem_Rep_preal_diff_iff) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 999 | apply (blast intro: less_add_left_lemma) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1000 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1001 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1002 | subsection{*proving that @{term "S \<le> R + D"} --- trickier*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1003 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1004 | lemma lemma_sum_mem_Rep_preal_ex: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1005 | "x \<in> Rep_preal S ==> \<exists>e. 0 < e & x + e \<in> Rep_preal S" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1006 | apply (drule Rep_preal [THEN preal_exists_greater], clarify) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1007 | apply (cut_tac a=x and b=u in add_eq_exists, auto) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1008 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1009 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1010 | lemma less_add_left_lemma2: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1011 | assumes Rless: "R < S" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1012 | and x: "x \<in> Rep_preal S" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1013 | and xnot: "x \<notin> Rep_preal R" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1014 | shows "\<exists>u v z. 0 < v & 0 < z & u \<in> Rep_preal R & z \<notin> Rep_preal R & | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1015 | z + v \<in> Rep_preal S & x = u + v" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1016 | proof - | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1017 | have xpos: "0<x" by (rule preal_imp_pos [OF Rep_preal x]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1018 | from lemma_sum_mem_Rep_preal_ex [OF x] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1019 | obtain e where epos: "0 < e" and xe: "x + e \<in> Rep_preal S" by blast | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1020 | from Gleason9_34 [OF Rep_preal epos] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1021 | obtain r where r: "r \<in> Rep_preal R" and notin: "r + e \<notin> Rep_preal R" .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1022 | with x xnot xpos have rless: "r < x" by (blast intro: not_in_Rep_preal_ub) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1023 | from add_eq_exists [of r x] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1024 | obtain y where eq: "x = r+y" by auto | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1025 | show ?thesis | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1026 | proof (intro exI conjI) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1027 | show "r \<in> Rep_preal R" by (rule r) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1028 | show "r + e \<notin> Rep_preal R" by (rule notin) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1029 | show "r + e + y \<in> Rep_preal S" using xe eq by (simp add: add_ac) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1030 | show "x = r + y" by (simp add: eq) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1031 | show "0 < r + e" using epos preal_imp_pos [OF Rep_preal r] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1032 | by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1033 | show "0 < y" using rless eq by arith | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1034 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1035 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1036 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1037 | lemma less_add_left_le2: "R < (S::preal) ==> S \<le> R + (S-R)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1038 | apply (auto simp add: preal_le_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1039 | apply (case_tac "x \<in> Rep_preal R") | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1040 | apply (cut_tac Rep_preal_self_subset [of R], force) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1041 | apply (auto simp add: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_diff_iff) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1042 | apply (blast dest: less_add_left_lemma2) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1043 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1044 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1045 | lemma less_add_left: "R < (S::preal) ==> R + (S-R) = S" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1046 | by (blast intro: antisym [OF less_add_left_le1 less_add_left_le2]) | 
| 
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 less_add_left_Ex: "R < (S::preal) ==> \<exists>D. R + D = S" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1049 | by (fast dest: less_add_left) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1050 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1051 | lemma preal_add_less2_mono1: "R < (S::preal) ==> R + T < S + T" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1052 | apply (auto dest!: less_add_left_Ex simp add: preal_add_assoc) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1053 | apply (rule_tac y1 = D in preal_add_commute [THEN subst]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1054 | apply (auto intro: preal_self_less_add_left simp add: preal_add_assoc [symmetric]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1055 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1056 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1057 | lemma preal_add_less2_mono2: "R < (S::preal) ==> T + R < T + S" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1058 | by (auto intro: preal_add_less2_mono1 simp add: preal_add_commute [of T]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1059 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1060 | lemma preal_add_right_less_cancel: "R + T < S + T ==> R < (S::preal)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1061 | apply (insert linorder_less_linear [of R S], auto) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1062 | apply (drule_tac R = S and T = T in preal_add_less2_mono1) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1063 | apply (blast dest: order_less_trans) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1064 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1065 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1066 | lemma preal_add_left_less_cancel: "T + R < T + S ==> R < (S::preal)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1067 | by (auto elim: preal_add_right_less_cancel simp add: preal_add_commute [of T]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1068 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1069 | lemma preal_add_less_cancel_left: "(T + (R::preal) < T + S) = (R < S)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1070 | by (blast intro: preal_add_less2_mono2 preal_add_left_less_cancel) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1071 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1072 | lemma preal_add_le_cancel_left: "(T + (R::preal) \<le> T + S) = (R \<le> S)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1073 | by (simp add: linorder_not_less [symmetric] preal_add_less_cancel_left) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1074 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1075 | lemma preal_add_right_cancel: "(R::preal) + T = S + T ==> R = S" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1076 | apply (insert linorder_less_linear [of R S], safe) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1077 | apply (drule_tac [!] T = T in preal_add_less2_mono1, auto) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1078 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1079 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1080 | lemma preal_add_left_cancel: "C + A = C + B ==> A = (B::preal)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1081 | by (auto intro: preal_add_right_cancel simp add: preal_add_commute) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1082 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1083 | instance preal :: linordered_cancel_ab_semigroup_add | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1084 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1085 | fix a b c :: preal | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1086 | show "a + b = a + c \<Longrightarrow> b = c" by (rule preal_add_left_cancel) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1087 | 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 | 1088 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1089 | |
| 
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 | subsection{*Completeness of type @{typ preal}*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1092 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1093 | text{*Prove that supremum is a cut*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1094 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1095 | text{*Part 1 of Dedekind sections definition*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1096 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1097 | lemma preal_sup_set_not_empty: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1098 |      "P \<noteq> {} ==> {} \<subset> (\<Union>X \<in> P. Rep_preal(X))"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1099 | apply auto | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1100 | apply (cut_tac X = x in mem_Rep_preal_Ex, auto) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1101 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1102 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1103 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1104 | text{*Part 2 of Dedekind sections definition*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1105 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1106 | lemma preal_sup_not_exists: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1107 | "\<forall>X \<in> P. X \<le> Y ==> \<exists>q. 0 < q & q \<notin> (\<Union>X \<in> P. Rep_preal(X))" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1108 | apply (cut_tac X = Y in Rep_preal_exists_bound) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1109 | apply (auto simp add: preal_le_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1110 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1111 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1112 | lemma preal_sup_set_not_rat_set: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1113 |      "\<forall>X \<in> P. X \<le> Y ==> (\<Union>X \<in> P. Rep_preal(X)) < {r. 0 < r}"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1114 | apply (drule preal_sup_not_exists) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1115 | apply (blast intro: preal_imp_pos [OF Rep_preal]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1116 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1117 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1118 | text{*Part 3 of Dedekind sections definition*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1119 | lemma preal_sup_set_lemma3: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1120 |      "[|P \<noteq> {}; \<forall>X \<in> P. X \<le> Y; u \<in> (\<Union>X \<in> P. Rep_preal(X)); 0 < z; z < u|]
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1121 | ==> z \<in> (\<Union>X \<in> P. Rep_preal(X))" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1122 | by (auto elim: Rep_preal [THEN preal_downwards_closed]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1123 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1124 | text{*Part 4 of Dedekind sections definition*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1125 | lemma preal_sup_set_lemma4: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1126 |      "[|P \<noteq> {}; \<forall>X \<in> P. X \<le> Y; y \<in> (\<Union>X \<in> P. Rep_preal(X)) |]
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1127 | ==> \<exists>u \<in> (\<Union>X \<in> P. Rep_preal(X)). y < u" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1128 | by (blast dest: Rep_preal [THEN preal_exists_greater]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1129 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1130 | lemma preal_sup: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1131 |      "[|P \<noteq> {}; \<forall>X \<in> P. X \<le> Y|] ==> (\<Union>X \<in> P. Rep_preal(X)) \<in> preal"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1132 | apply (unfold preal_def cut_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1133 | apply (blast intro!: preal_sup_set_not_empty preal_sup_set_not_rat_set | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1134 | preal_sup_set_lemma3 preal_sup_set_lemma4) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1135 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1136 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1137 | lemma preal_psup_le: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1138 | "[| \<forall>X \<in> P. X \<le> Y; x \<in> P |] ==> x \<le> psup P" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1139 | apply (simp (no_asm_simp) add: preal_le_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1140 | apply (subgoal_tac "P \<noteq> {}") 
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1141 | apply (auto simp add: psup_def preal_sup) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1142 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1143 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1144 | lemma psup_le_ub: "[| P \<noteq> {}; \<forall>X \<in> P. X \<le> Y |] ==> psup P \<le> Y"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1145 | apply (simp (no_asm_simp) add: preal_le_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1146 | apply (simp add: psup_def preal_sup) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1147 | apply (auto simp add: preal_le_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1148 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1149 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1150 | text{*Supremum property*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1151 | lemma preal_complete: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1152 |      "[| P \<noteq> {}; \<forall>X \<in> P. X \<le> Y |] ==> (\<exists>X \<in> P. Z < X) = (Z < psup P)"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1153 | apply (simp add: preal_less_def psup_def preal_sup) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1154 | apply (auto simp add: preal_le_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1155 | apply (rename_tac U) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1156 | apply (cut_tac x = U and y = Z in linorder_less_linear) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1157 | apply (auto simp add: preal_less_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1158 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1159 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1160 | section {*Defining the Reals from the Positive Reals*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1161 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1162 | definition | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1163 | realrel :: "((preal * preal) * (preal * preal)) set" where | 
| 37765 | 1164 |   "realrel = {p. \<exists>x1 y1 x2 y2. p = ((x1,y1),(x2,y2)) & x1+y2 = x2+y1}"
 | 
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1165 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1166 | typedef (Real) real = "UNIV//realrel" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1167 | by (auto simp add: quotient_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1168 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1169 | definition | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1170 | (** these don't use the overloaded "real" function: users don't see them **) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1171 | real_of_preal :: "preal => real" where | 
| 37765 | 1172 |   "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 | 1173 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1174 | 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 | 1175 | begin | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1176 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1177 | definition | 
| 37765 | 1178 |   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 | 1179 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1180 | definition | 
| 37765 | 1181 |   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 | 1182 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1183 | definition | 
| 37765 | 1184 | real_add_def: "z + w = | 
| 39910 | 1185 | 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 | 1186 |                  { Abs_Real(realrel``{(x+u, y+v)}) })"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1187 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1188 | definition | 
| 39910 | 1189 |   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 | 1190 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1191 | definition | 
| 37765 | 1192 | real_diff_def: "r - (s::real) = r + - s" | 
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1193 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1194 | definition | 
| 37765 | 1195 | real_mult_def: | 
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1196 | "z * w = | 
| 39910 | 1197 | 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 | 1198 |                  { 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 | 1199 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1200 | definition | 
| 37765 | 1201 | real_inverse_def: "inverse (R::real) = (THE S. (R = 0 & S = 0) | S * R = 1)" | 
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1202 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1203 | definition | 
| 37765 | 1204 | real_divide_def: "R / (S::real) = R * inverse S" | 
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1205 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1206 | definition | 
| 37765 | 1207 | real_le_def: "z \<le> (w::real) \<longleftrightarrow> | 
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1208 | (\<exists>x y u v. x+v \<le> u+y & (x,y) \<in> Rep_Real z & (u,v) \<in> Rep_Real w)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1209 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1210 | definition | 
| 37765 | 1211 | real_less_def: "x < (y\<Colon>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 | 1212 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1213 | definition | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1214 | real_abs_def: "abs (r::real) = (if r < 0 then - r else r)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1215 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1216 | definition | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1217 | 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 | 1218 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1219 | instance .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1220 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1221 | end | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1222 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1223 | subsection {* Equivalence relation over positive reals *}
 | 
| 
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 preal_trans_lemma: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1226 | assumes "x + y1 = x1 + y" | 
| 41541 | 1227 | and "x + y2 = x2 + y" | 
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1228 | shows "x1 + y2 = x2 + (y1::preal)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1229 | proof - | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1230 | have "(x1 + y2) + x = (x + y2) + x1" by (simp add: add_ac) | 
| 41541 | 1231 | also have "... = (x2 + y) + x1" by (simp add: assms) | 
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1232 | also have "... = x2 + (x1 + y)" by (simp add: add_ac) | 
| 41541 | 1233 | also have "... = x2 + (x + y1)" by (simp add: assms) | 
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1234 | also have "... = (x2 + y1) + x" by (simp add: add_ac) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1235 | finally have "(x1 + y2) + x = (x2 + y1) + x" . | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1236 | thus ?thesis by (rule add_right_imp_eq) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1237 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1238 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1239 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1240 | lemma realrel_iff [simp]: "(((x1,y1),(x2,y2)) \<in> realrel) = (x1 + y2 = x2 + y1)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1241 | by (simp add: realrel_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1242 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1243 | lemma equiv_realrel: "equiv UNIV realrel" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1244 | apply (auto simp add: equiv_def refl_on_def sym_def trans_def realrel_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1245 | apply (blast dest: preal_trans_lemma) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1246 | done | 
| 
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 | text{*Reduces equality of equivalence classes to the @{term realrel} relation:
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1249 |   @{term "(realrel `` {x} = realrel `` {y}) = ((x,y) \<in> realrel)"} *}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1250 | lemmas equiv_realrel_iff = | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1251 | 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 | 1252 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1253 | declare equiv_realrel_iff [simp] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1254 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1255 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1256 | lemma realrel_in_real [simp]: "realrel``{(x,y)}: Real"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1257 | by (simp add: Real_def realrel_def quotient_def, blast) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1258 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1259 | declare Abs_Real_inject [simp] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1260 | declare Abs_Real_inverse [simp] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1261 | |
| 
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 | text{*Case analysis on the representation of a real number as an equivalence
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1264 | class of pairs of positive reals.*} | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1265 | lemma eq_Abs_Real [case_names Abs_Real, cases type: real]: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1266 |      "(!!x y. z = Abs_Real(realrel``{(x,y)}) ==> P) ==> P"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1267 | apply (rule Rep_Real [of z, unfolded Real_def, THEN quotientE]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1268 | apply (drule arg_cong [where f=Abs_Real]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1269 | apply (auto simp add: Rep_Real_inverse) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1270 | done | 
| 
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 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1273 | subsection {* Addition and Subtraction *}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1274 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1275 | lemma real_add_congruent2_lemma: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1276 | "[|a + ba = aa + b; ab + bc = ac + bb|] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1277 | ==> a + ab + (ba + bc) = aa + ac + (b + (bb::preal))" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1278 | apply (simp add: add_assoc) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1279 | apply (rule add_left_commute [of ab, THEN ssubst]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1280 | apply (simp add: add_assoc [symmetric]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1281 | apply (simp add: add_ac) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1282 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1283 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1284 | lemma real_add: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1285 |      "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 | 1286 |       Abs_Real (realrel``{(x+u, y+v)})"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1287 | proof - | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1288 |   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 | 1289 | respects2 realrel" | 
| 40822 | 1290 | by (auto simp add: congruent2_def, blast intro: real_add_congruent2_lemma) | 
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1291 | thus ?thesis | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1292 | by (simp add: real_add_def UN_UN_split_split_eq | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1293 | UN_equiv_class2 [OF equiv_realrel equiv_realrel]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1294 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1295 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1296 | 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 | 1297 | proof - | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1298 |   have "(\<lambda>(x,y). {Abs_Real (realrel``{(y,x)})}) respects realrel"
 | 
| 40822 | 1299 | by (auto simp add: congruent_def add_commute) | 
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1300 | thus ?thesis | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1301 | 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 | 1302 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1303 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1304 | instance real :: ab_group_add | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1305 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1306 | fix x y z :: real | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1307 | show "(x + y) + z = x + (y + z)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1308 | by (cases x, cases y, cases z, simp add: real_add add_assoc) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1309 | show "x + y = y + x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1310 | by (cases x, cases y, simp add: real_add add_commute) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1311 | show "0 + x = x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1312 | by (cases x, simp add: real_add real_zero_def add_ac) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1313 | show "- x + x = 0" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1314 | by (cases x, simp add: real_minus real_add real_zero_def add_commute) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1315 | show "x - y = x + - y" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1316 | by (simp add: real_diff_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1317 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1318 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1319 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1320 | subsection {* Multiplication *}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1321 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1322 | lemma real_mult_congruent2_lemma: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1323 | "!!(x1::preal). [| x1 + y2 = x2 + y1 |] ==> | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1324 | x * x1 + y * y1 + (x * y2 + y * x2) = | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1325 | x * x2 + y * y2 + (x * y1 + y * x1)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1326 | apply (simp add: add_left_commute add_assoc [symmetric]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1327 | apply (simp add: add_assoc right_distrib [symmetric]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1328 | apply (simp add: add_commute) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1329 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1330 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1331 | lemma real_mult_congruent2: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1332 | "(%p1 p2. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1333 | (%(x1,y1). (%(x2,y2). | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1334 |           { 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 | 1335 | respects2 realrel" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1336 | apply (rule congruent2_commuteI [OF equiv_realrel], clarify) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1337 | apply (simp add: mult_commute add_commute) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1338 | apply (auto simp add: real_mult_congruent2_lemma) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1339 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1340 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1341 | lemma real_mult: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1342 |       "Abs_Real((realrel``{(x1,y1)})) * Abs_Real((realrel``{(x2,y2)})) =
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1343 |        Abs_Real(realrel `` {(x1*x2+y1*y2,x1*y2+y1*x2)})"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1344 | by (simp add: real_mult_def UN_UN_split_split_eq | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1345 | UN_equiv_class2 [OF equiv_realrel equiv_realrel real_mult_congruent2]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1346 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1347 | lemma real_mult_commute: "(z::real) * w = w * z" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1348 | by (cases z, cases w, simp add: real_mult add_ac mult_ac) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1349 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1350 | lemma real_mult_assoc: "((z1::real) * z2) * z3 = z1 * (z2 * z3)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1351 | apply (cases z1, cases z2, cases z3) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1352 | apply (simp add: real_mult algebra_simps) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1353 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1354 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1355 | lemma real_mult_1: "(1::real) * z = z" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1356 | apply (cases z) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1357 | apply (simp add: real_mult real_one_def algebra_simps) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1358 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1359 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1360 | lemma real_add_mult_distrib: "((z1::real) + z2) * w = (z1 * w) + (z2 * w)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1361 | apply (cases z1, cases z2, cases w) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1362 | apply (simp add: real_add real_mult algebra_simps) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1363 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1364 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1365 | text{*one and zero are distinct*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1366 | 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 | 1367 | proof - | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1368 | have "(1::preal) < 1 + 1" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1369 | by (simp add: preal_self_less_add_left) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1370 | thus ?thesis | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1371 | by (simp add: real_zero_def real_one_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1372 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1373 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1374 | instance real :: comm_ring_1 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1375 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1376 | fix x y z :: real | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1377 | 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 | 1378 | 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 | 1379 | show "1 * x = x" by (rule real_mult_1) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1380 | 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 | 1381 | 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 | 1382 | qed | 
| 
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 | subsection {* Inverse and Division *}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1385 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1386 | lemma real_zero_iff: "Abs_Real (realrel `` {(x, x)}) = 0"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1387 | by (simp add: real_zero_def add_commute) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1388 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1389 | text{*Instead of using an existential quantifier and constructing the inverse
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1390 | within the proof, we could define the inverse explicitly.*} | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1391 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1392 | lemma real_mult_inverse_left_ex: "x \<noteq> 0 ==> \<exists>y. y*x = (1::real)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1393 | apply (simp add: real_zero_def real_one_def, cases x) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1394 | apply (cut_tac x = xa and y = y in linorder_less_linear) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1395 | apply (auto dest!: less_add_left_Ex simp add: real_zero_iff) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1396 | apply (rule_tac | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1397 |         x = "Abs_Real (realrel``{(1, inverse (D) + 1)})"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1398 | in exI) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1399 | apply (rule_tac [2] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1400 |         x = "Abs_Real (realrel``{(inverse (D) + 1, 1)})" 
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1401 | in exI) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1402 | apply (auto simp add: real_mult preal_mult_inverse_right algebra_simps) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1403 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1404 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1405 | lemma real_mult_inverse_left: "x \<noteq> 0 ==> inverse(x)*x = (1::real)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1406 | apply (simp add: real_inverse_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1407 | apply (drule real_mult_inverse_left_ex, safe) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1408 | apply (rule theI, assumption, rename_tac z) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1409 | apply (subgoal_tac "(z * x) * y = z * (x * y)") | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1410 | apply (simp add: mult_commute) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1411 | apply (rule mult_assoc) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1412 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1413 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1414 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1415 | subsection{*The Real Numbers form a Field*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1416 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1417 | instance real :: field_inverse_zero | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1418 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1419 | fix x y z :: real | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1420 | show "x \<noteq> 0 ==> inverse x * x = 1" by (rule real_mult_inverse_left) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1421 | 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 | 1422 | 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 | 1423 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1424 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1425 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1426 | subsection{*The @{text "\<le>"} Ordering*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1427 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1428 | lemma real_le_refl: "w \<le> (w::real)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1429 | by (cases w, force simp add: real_le_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1430 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1431 | text{*The arithmetic decision procedure is not set up for type preal.
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1432 | This lemma is currently unused, but it could simplify the proofs of the | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1433 | following two lemmas.*} | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1434 | lemma preal_eq_le_imp_le: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1435 | 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 | 1436 | shows "b \<le> (d::preal)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1437 | proof - | 
| 41541 | 1438 | have "c+d \<le> a+d" by (simp add: le) | 
| 1439 | 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 | 1440 | thus "b \<le> d" by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1441 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1442 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1443 | lemma real_le_lemma: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1444 | assumes l: "u1 + v2 \<le> u2 + v1" | 
| 41541 | 1445 | and "x1 + v1 = u1 + y1" | 
| 1446 | and "x2 + v2 = u2 + y2" | |
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1447 | shows "x1 + y2 \<le> x2 + (y1::preal)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1448 | proof - | 
| 41541 | 1449 | have "(x1+v1) + (u2+y2) = (u1+y1) + (x2+v2)" by (simp add: assms) | 
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1450 | hence "(x1+y2) + (u2+v1) = (x2+y1) + (u1+v2)" by (simp add: add_ac) | 
| 41541 | 1451 | also have "... \<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 | 1452 | finally show ?thesis by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1453 | qed | 
| 
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 | lemma real_le: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1456 |      "(Abs_Real(realrel``{(x1,y1)}) \<le> Abs_Real(realrel``{(x2,y2)})) =  
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1457 | (x1 + y2 \<le> x2 + y1)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1458 | apply (simp add: real_le_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1459 | apply (auto intro: real_le_lemma) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1460 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1461 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1462 | lemma real_le_antisym: "[| z \<le> w; w \<le> z |] ==> z = (w::real)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1463 | by (cases z, cases w, simp add: real_le) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1464 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1465 | lemma real_trans_lemma: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1466 | assumes "x + v \<le> u + y" | 
| 41541 | 1467 | and "u + v' \<le> u' + v" | 
| 1468 | and "x2 + v2 = u2 + y2" | |
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1469 | shows "x + v' \<le> u' + (y::preal)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1470 | proof - | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1471 | have "(x+v') + (u+v) = (x+v) + (u+v')" by (simp add: add_ac) | 
| 41541 | 1472 | also have "... \<le> (u+y) + (u+v')" by (simp add: assms) | 
| 1473 | also have "... \<le> (u+y) + (u'+v)" by (simp add: assms) | |
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1474 | also have "... = (u'+y) + (u+v)" by (simp add: add_ac) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1475 | finally show ?thesis by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1476 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1477 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1478 | lemma real_le_trans: "[| i \<le> j; j \<le> k |] ==> i \<le> (k::real)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1479 | apply (cases i, cases j, cases k) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1480 | apply (simp add: real_le) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1481 | apply (blast intro: real_trans_lemma) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1482 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1483 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1484 | instance real :: order | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1485 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1486 | fix u v :: real | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1487 | show "u < v \<longleftrightarrow> u \<le> v \<and> \<not> v \<le> u" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1488 | by (auto simp add: real_less_def intro: real_le_antisym) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1489 | qed (assumption | rule real_le_refl real_le_trans real_le_antisym)+ | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1490 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1491 | (* Axiom 'linorder_linear' of class 'linorder': *) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1492 | lemma real_le_linear: "(z::real) \<le> w | w \<le> z" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1493 | apply (cases z, cases w) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1494 | apply (auto simp add: real_le real_zero_def add_ac) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1495 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1496 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1497 | instance real :: linorder | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1498 | by (intro_classes, rule real_le_linear) | 
| 
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 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1501 | lemma real_le_eq_diff: "(x \<le> y) = (x-y \<le> (0::real))" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1502 | apply (cases x, cases y) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1503 | apply (auto simp add: real_le real_zero_def real_diff_def real_add real_minus | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1504 | add_ac) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1505 | apply (simp_all add: add_assoc [symmetric]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1506 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1507 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1508 | lemma real_add_left_mono: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1509 | assumes le: "x \<le> y" shows "z + x \<le> z + (y::real)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1510 | proof - | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1511 | have "z + x - (z + y) = (z + -z) + (x - y)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1512 | by (simp add: algebra_simps) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1513 | with le show ?thesis | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1514 | by (simp add: real_le_eq_diff[of x] real_le_eq_diff[of "z+x"] diff_minus) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1515 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1516 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1517 | lemma real_sum_gt_zero_less: "(0 < S + (-W::real)) ==> (W < S)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1518 | by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S] diff_minus) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1519 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1520 | lemma real_less_sum_gt_zero: "(W < S) ==> (0 < S + (-W::real))" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1521 | by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S] diff_minus) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1522 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1523 | lemma real_mult_order: "[| 0 < x; 0 < y |] ==> (0::real) < x * y" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1524 | apply (cases x, cases y) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1525 | apply (simp add: linorder_not_le [where 'a = real, symmetric] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1526 | linorder_not_le [where 'a = preal] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1527 | real_zero_def real_le real_mult) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1528 |   --{*Reduce to the (simpler) @{text "\<le>"} relation *}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1529 | apply (auto dest!: less_add_left_Ex | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1530 | simp add: algebra_simps preal_self_less_add_left) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1531 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1532 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1533 | lemma real_mult_less_mono2: "[| (0::real) < z; x < y |] ==> z * x < z * y" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1534 | apply (rule real_sum_gt_zero_less) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1535 | apply (drule real_less_sum_gt_zero [of x y]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1536 | apply (drule real_mult_order, assumption) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1537 | apply (simp add: right_distrib) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1538 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1539 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1540 | instantiation real :: distrib_lattice | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1541 | begin | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1542 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1543 | definition | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1544 | "(inf \<Colon> real \<Rightarrow> real \<Rightarrow> real) = min" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1545 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1546 | definition | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1547 | "(sup \<Colon> real \<Rightarrow> real \<Rightarrow> real) = max" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1548 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1549 | instance | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1550 | by default (auto simp add: inf_real_def sup_real_def min_max.sup_inf_distrib1) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1551 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1552 | end | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1553 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1554 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1555 | subsection{*The Reals Form an Ordered Field*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1556 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1557 | instance real :: linordered_field_inverse_zero | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1558 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1559 | fix x y z :: real | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1560 | show "x \<le> y ==> z + x \<le> z + y" by (rule real_add_left_mono) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1561 | show "x < y ==> 0 < z ==> z * x < z * y" by (rule real_mult_less_mono2) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1562 | 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 | 1563 | 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 | 1564 | by (simp only: real_sgn_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1565 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1566 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1567 | text{*The function @{term real_of_preal} requires many proofs, but it seems
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1568 | to be essential for proving completeness of the reals from that of the | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1569 | positive reals.*} | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1570 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1571 | lemma real_of_preal_add: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1572 | "real_of_preal ((x::preal) + y) = real_of_preal x + real_of_preal y" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1573 | by (simp add: real_of_preal_def real_add algebra_simps) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1574 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1575 | lemma real_of_preal_mult: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1576 | "real_of_preal ((x::preal) * y) = real_of_preal x* real_of_preal y" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1577 | by (simp add: real_of_preal_def real_mult algebra_simps) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1578 | |
| 
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 | text{*Gleason prop 9-4.4 p 127*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1581 | lemma real_of_preal_trichotomy: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1582 | "\<exists>m. (x::real) = real_of_preal m | x = 0 | x = -(real_of_preal m)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1583 | apply (simp add: real_of_preal_def real_zero_def, cases x) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1584 | apply (auto simp add: real_minus add_ac) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1585 | apply (cut_tac x = x and y = y in linorder_less_linear) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1586 | apply (auto dest!: less_add_left_Ex simp add: add_assoc [symmetric]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1587 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1588 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1589 | lemma real_of_preal_leD: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1590 | "real_of_preal m1 \<le> real_of_preal m2 ==> m1 \<le> m2" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1591 | by (simp add: real_of_preal_def real_le) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1592 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1593 | lemma real_of_preal_lessI: "m1 < m2 ==> real_of_preal m1 < real_of_preal m2" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1594 | by (auto simp add: real_of_preal_leD linorder_not_le [symmetric]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1595 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1596 | lemma real_of_preal_lessD: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1597 | "real_of_preal m1 < real_of_preal m2 ==> m1 < m2" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1598 | by (simp add: real_of_preal_def real_le linorder_not_le [symmetric]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1599 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1600 | lemma real_of_preal_less_iff [simp]: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1601 | "(real_of_preal m1 < real_of_preal m2) = (m1 < m2)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1602 | by (blast intro: real_of_preal_lessI real_of_preal_lessD) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1603 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1604 | lemma real_of_preal_le_iff: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1605 | "(real_of_preal m1 \<le> real_of_preal m2) = (m1 \<le> m2)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1606 | by (simp add: linorder_not_less [symmetric]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1607 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1608 | lemma real_of_preal_zero_less: "0 < real_of_preal m" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1609 | apply (insert preal_self_less_add_left [of 1 m]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1610 | apply (auto simp add: real_zero_def real_of_preal_def | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1611 | real_less_def real_le_def add_ac) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1612 | apply (rule_tac x="m + 1" in exI, rule_tac x="1" in exI) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1613 | apply (simp add: add_ac) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1614 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1615 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1616 | lemma real_of_preal_minus_less_zero: "- real_of_preal m < 0" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1617 | by (simp add: real_of_preal_zero_less) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1618 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1619 | lemma real_of_preal_not_minus_gt_zero: "~ 0 < - real_of_preal m" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1620 | proof - | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1621 | from real_of_preal_minus_less_zero | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1622 | show ?thesis by (blast dest: order_less_trans) | 
| 
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 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1625 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1626 | subsection{*Theorems About the Ordering*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1627 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1628 | lemma real_gt_zero_preal_Ex: "(0 < x) = (\<exists>y. x = real_of_preal y)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1629 | apply (auto simp add: real_of_preal_zero_less) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1630 | apply (cut_tac x = x in real_of_preal_trichotomy) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1631 | apply (blast elim!: real_of_preal_not_minus_gt_zero [THEN notE]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1632 | done | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1633 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1634 | lemma real_gt_preal_preal_Ex: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1635 | "real_of_preal z < x ==> \<exists>y. x = real_of_preal y" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1636 | by (blast dest!: real_of_preal_zero_less [THEN order_less_trans] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1637 | intro: real_gt_zero_preal_Ex [THEN iffD1]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1638 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1639 | lemma real_ge_preal_preal_Ex: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1640 | "real_of_preal z \<le> x ==> \<exists>y. x = real_of_preal y" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1641 | by (blast dest: order_le_imp_less_or_eq real_gt_preal_preal_Ex) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1642 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1643 | lemma real_less_all_preal: "y \<le> 0 ==> \<forall>x. y < real_of_preal x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1644 | by (auto elim: order_le_imp_less_or_eq [THEN disjE] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1645 | intro: real_of_preal_zero_less [THEN [2] order_less_trans] | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1646 | simp add: real_of_preal_zero_less) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1647 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1648 | lemma real_less_all_real2: "~ 0 < y ==> \<forall>x. y < real_of_preal x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1649 | by (blast intro!: real_less_all_preal linorder_not_less [THEN iffD1]) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1650 | |
| 
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 | subsection{*Numerals and Arithmetic*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1653 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1654 | instantiation real :: number_ring | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1655 | begin | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1656 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1657 | definition | 
| 37765 | 1658 | real_number_of_def: "(number_of w :: real) = of_int w" | 
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1659 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1660 | instance | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1661 | by intro_classes (simp add: real_number_of_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1662 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1663 | end | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1664 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1665 | subsection {* Completeness of Positive Reals *}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1666 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1667 | text {*
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1668 | Supremum property for the set of positive reals | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1669 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1670 |   Let @{text "P"} be a non-empty set of positive reals, with an upper
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1671 |   bound @{text "y"}.  Then @{text "P"} has a least upper bound
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1672 |   (written @{text "S"}).
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1673 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1674 |   FIXME: Can the premise be weakened to @{text "\<forall>x \<in> P. x\<le> y"}?
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1675 | *} | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1676 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1677 | lemma posreal_complete: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1678 | 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 | 1679 | and not_empty_P: "\<exists>x. x \<in> P" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1680 | and upper_bound_Ex: "\<exists>y. \<forall>x \<in> P. x<y" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1681 | shows "\<exists>S. \<forall>y. (\<exists>x \<in> P. y < x) = (y < S)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1682 | proof (rule exI, rule allI) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1683 | fix y | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1684 |   let ?pP = "{w. real_of_preal w \<in> P}"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1685 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1686 | 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 | 1687 | proof (cases "0 < y") | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1688 | assume neg_y: "\<not> 0 < y" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1689 | show ?thesis | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1690 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1691 | assume "\<exists>x\<in>P. y < x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1692 | have "\<forall>x. y < real_of_preal x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1693 | using neg_y by (rule real_less_all_real2) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1694 | thus "y < real_of_preal (psup ?pP)" .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1695 | next | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1696 | assume "y < real_of_preal (psup ?pP)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1697 | obtain "x" where x_in_P: "x \<in> P" using not_empty_P .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1698 | hence "0 < x" using positive_P by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1699 | hence "y < x" using neg_y by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1700 | thus "\<exists>x \<in> P. y < x" using x_in_P .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1701 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1702 | next | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1703 | assume pos_y: "0 < y" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1704 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1705 | then obtain py where y_is_py: "y = real_of_preal py" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1706 | by (auto simp add: real_gt_zero_preal_Ex) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1707 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1708 | 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 | 1709 | with positive_P have a_pos: "0 < a" .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1710 | then obtain pa where "a = real_of_preal pa" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1711 | by (auto simp add: real_gt_zero_preal_Ex) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1712 | hence "pa \<in> ?pP" using `a \<in> P` by auto | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1713 |     hence pP_not_empty: "?pP \<noteq> {}" by auto
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1714 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1715 | 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 | 1716 | using upper_bound_Ex .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1717 | from this and `a \<in> P` have "a < sup" .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1718 | hence "0 < sup" using a_pos by arith | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1719 | then obtain possup where "sup = real_of_preal possup" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1720 | by (auto simp add: real_gt_zero_preal_Ex) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1721 | hence "\<forall>X \<in> ?pP. X \<le> possup" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1722 | using sup by (auto simp add: real_of_preal_lessI) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1723 | with pP_not_empty have psup: "\<And>Z. (\<exists>X \<in> ?pP. Z < X) = (Z < psup ?pP)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1724 | by (rule preal_complete) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1725 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1726 | show ?thesis | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1727 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1728 | assume "\<exists>x \<in> P. y < x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1729 | 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 | 1730 | hence "0 < x" using pos_y by arith | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1731 | then obtain px where x_is_px: "x = real_of_preal px" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1732 | by (auto simp add: real_gt_zero_preal_Ex) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1733 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1734 | 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 | 1735 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1736 | show "py < px" using y_is_py and x_is_px and y_less_x | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1737 | by (simp add: real_of_preal_lessI) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1738 | 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 | 1739 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1740 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1741 | have "(\<exists>X \<in> ?pP. py < X) ==> (py < psup ?pP)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1742 | using psup by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1743 | 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 | 1744 |       thus "y < real_of_preal (psup {w. real_of_preal w \<in> P})"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1745 | using y_is_py and pos_y by (simp add: real_of_preal_lessI) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1746 | next | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1747 | 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 | 1748 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1749 | hence "py < psup ?pP" using y_is_py | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1750 | by (simp add: real_of_preal_lessI) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1751 | 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 | 1752 | using psup by auto | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1753 | 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 | 1754 | by (simp add: real_gt_zero_preal_Ex) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1755 | hence "y < x" using py_less_X and y_is_py | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1756 | by (simp add: real_of_preal_lessI) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1757 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1758 | moreover have "x \<in> P" using x_is_X and X_in_pP by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1759 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1760 | ultimately show "\<exists> x \<in> P. y < x" .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1761 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1762 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1763 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1764 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1765 | text {*
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1766 |   \medskip Completeness properties using @{text "isUb"}, @{text "isLub"} etc.
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1767 | *} | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1768 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1769 | lemma posreals_complete: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1770 | assumes positive_S: "\<forall>x \<in> S. 0 < x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1771 | and not_empty_S: "\<exists>x. x \<in> S" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1772 | and upper_bound_Ex: "\<exists>u. isUb (UNIV::real set) S u" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1773 | shows "\<exists>t. isLub (UNIV::real set) S t" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1774 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1775 |   let ?pS = "{w. real_of_preal w \<in> S}"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1776 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1777 | obtain u where "isUb UNIV S u" using upper_bound_Ex .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1778 | hence sup: "\<forall>x \<in> S. x \<le> u" by (simp add: isUb_def setle_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1779 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1780 | obtain x where x_in_S: "x \<in> S" using not_empty_S .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1781 | hence x_gt_zero: "0 < x" using positive_S by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1782 | have "x \<le> u" using sup and x_in_S .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1783 | hence "0 < u" using x_gt_zero by arith | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1784 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1785 | then obtain pu where u_is_pu: "u = real_of_preal pu" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1786 | by (auto simp add: real_gt_zero_preal_Ex) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1787 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1788 | have pS_less_pu: "\<forall>pa \<in> ?pS. pa \<le> pu" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1789 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1790 | fix pa | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1791 | assume "pa \<in> ?pS" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1792 | then obtain a where "a \<in> S" and "a = real_of_preal pa" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1793 | by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1794 | moreover hence "a \<le> u" using sup by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1795 | ultimately show "pa \<le> pu" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1796 | using sup and u_is_pu by (simp add: real_of_preal_le_iff) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1797 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1798 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1799 | have "\<forall>y \<in> S. y \<le> real_of_preal (psup ?pS)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1800 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1801 | fix y | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1802 | assume y_in_S: "y \<in> S" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1803 | hence "0 < y" using positive_S by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1804 | then obtain py where y_is_py: "y = real_of_preal py" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1805 | by (auto simp add: real_gt_zero_preal_Ex) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1806 | hence py_in_pS: "py \<in> ?pS" using y_in_S by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1807 | with pS_less_pu have "py \<le> psup ?pS" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1808 | by (rule preal_psup_le) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1809 | thus "y \<le> real_of_preal (psup ?pS)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1810 | using y_is_py by (simp add: real_of_preal_le_iff) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1811 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1812 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1813 |   moreover {
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1814 | fix x | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1815 | assume x_ub_S: "\<forall>y\<in>S. y \<le> x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1816 | have "real_of_preal (psup ?pS) \<le> x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1817 | proof - | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1818 | obtain "s" where s_in_S: "s \<in> S" using not_empty_S .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1819 | hence s_pos: "0 < s" using positive_S by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1820 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1821 | hence "\<exists> ps. s = real_of_preal ps" by (simp add: real_gt_zero_preal_Ex) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1822 | then obtain "ps" where s_is_ps: "s = real_of_preal ps" .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1823 |       hence ps_in_pS: "ps \<in> {w. real_of_preal w \<in> S}" using s_in_S by simp
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1824 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1825 | from x_ub_S have "s \<le> x" using s_in_S .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1826 | hence "0 < x" using s_pos by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1827 | hence "\<exists> px. x = real_of_preal px" by (simp add: real_gt_zero_preal_Ex) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1828 | then obtain "px" where x_is_px: "x = real_of_preal px" .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1829 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1830 | have "\<forall>pe \<in> ?pS. pe \<le> px" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1831 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1832 | fix pe | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1833 | assume "pe \<in> ?pS" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1834 | hence "real_of_preal pe \<in> S" by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1835 | hence "real_of_preal pe \<le> x" using x_ub_S by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1836 | thus "pe \<le> px" using x_is_px by (simp add: real_of_preal_le_iff) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1837 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1838 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1839 |       moreover have "?pS \<noteq> {}" using ps_in_pS by auto
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1840 | ultimately have "(psup ?pS) \<le> px" by (simp add: psup_le_ub) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1841 | thus "real_of_preal (psup ?pS) \<le> x" using x_is_px by (simp add: real_of_preal_le_iff) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1842 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1843 | } | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1844 | ultimately show "isLub UNIV S (real_of_preal (psup ?pS))" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1845 | by (simp add: isLub_def leastP_def isUb_def setle_def setge_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1846 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1847 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1848 | text {*
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1849 | \medskip reals Completeness (again!) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1850 | *} | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1851 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1852 | lemma reals_complete: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1853 | assumes notempty_S: "\<exists>X. X \<in> S" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1854 | and exists_Ub: "\<exists>Y. isUb (UNIV::real set) S Y" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1855 | shows "\<exists>t. isLub (UNIV :: real set) S t" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1856 | proof - | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1857 | obtain X where X_in_S: "X \<in> S" using notempty_S .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1858 | obtain Y where Y_isUb: "isUb (UNIV::real set) S Y" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1859 | using exists_Ub .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1860 |   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 | 1861 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1862 |   {
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1863 | fix x | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1864 | assume "isUb (UNIV::real set) S x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1865 | hence S_le_x: "\<forall> y \<in> S. y <= x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1866 | by (simp add: isUb_def setle_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1867 |     {
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1868 | fix s | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1869 |       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 | 1870 | hence "\<exists> x \<in> S. s = x + -X + 1" .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1871 | then obtain x1 where "x1 \<in> S" and "s = x1 + (-X) + 1" .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1872 | moreover hence "x1 \<le> x" using S_le_x by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1873 | ultimately have "s \<le> x + - X + 1" by arith | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1874 | } | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1875 | then have "isUb (UNIV::real set) ?SHIFT (x + (-X) + 1)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1876 | by (auto simp add: isUb_def setle_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1877 | } note S_Ub_is_SHIFT_Ub = this | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1878 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1879 | hence "isUb UNIV ?SHIFT (Y + (-X) + 1)" using Y_isUb by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1880 | hence "\<exists>Z. isUb UNIV ?SHIFT Z" .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1881 | 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 | 1882 | 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 | 1883 | using X_in_S and Y_isUb by auto | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1884 | ultimately obtain t where t_is_Lub: "isLub UNIV ?SHIFT t" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1885 | using posreals_complete [of ?SHIFT] by blast | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1886 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1887 | show ?thesis | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1888 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1889 | show "isLub UNIV S (t + X + (-1))" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1890 | proof (rule isLubI2) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1891 |       {
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1892 | fix x | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1893 | assume "isUb (UNIV::real set) S x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1894 | hence "isUb (UNIV::real set) (?SHIFT) (x + (-X) + 1)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1895 | using S_Ub_is_SHIFT_Ub by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1896 | hence "t \<le> (x + (-X) + 1)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1897 | using t_is_Lub by (simp add: isLub_le_isUb) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1898 | hence "t + X + -1 \<le> x" by arith | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1899 | } | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1900 | then show "(t + X + -1) <=* Collect (isUb UNIV S)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1901 | by (simp add: setgeI) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1902 | next | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1903 | show "isUb UNIV S (t + X + -1)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1904 | proof - | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1905 |         {
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1906 | fix y | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1907 | assume y_in_S: "y \<in> S" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1908 | have "y \<le> t + X + -1" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1909 | proof - | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1910 | obtain "u" where u_in_shift: "u \<in> ?SHIFT" using shifted_not_empty .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1911 | hence "\<exists> x \<in> S. u = x + - X + 1" by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1912 | then obtain "x" where x_and_u: "u = x + - X + 1" .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1913 | have u_le_t: "u \<le> t" using u_in_shift and t_is_Lub by (simp add: isLubD2) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1914 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1915 | show ?thesis | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1916 | proof cases | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1917 | assume "y \<le> x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1918 | moreover have "x = u + X + - 1" using x_and_u by arith | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1919 | moreover have "u + X + - 1 \<le> t + X + -1" using u_le_t by arith | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1920 | ultimately show "y \<le> t + X + -1" by arith | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1921 | next | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1922 | assume "~(y \<le> x)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1923 | hence x_less_y: "x < y" by arith | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1924 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1925 | have "x + (-X) + 1 \<in> ?SHIFT" using x_and_u and u_in_shift by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1926 | hence "0 < x + (-X) + 1" by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1927 | hence "0 < y + (-X) + 1" using x_less_y by arith | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1928 | hence "y + (-X) + 1 \<in> ?SHIFT" using y_in_S by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1929 | hence "y + (-X) + 1 \<le> t" using t_is_Lub by (simp add: isLubD2) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1930 | thus ?thesis by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1931 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1932 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1933 | } | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1934 | then show ?thesis by (simp add: isUb_def setle_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1935 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1936 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1937 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1938 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1939 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1940 | text{*A version of the same theorem without all those predicates!*}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1941 | lemma reals_complete2: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1942 | fixes S :: "(real set)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1943 | assumes "\<exists>y. y\<in>S" and "\<exists>(x::real). \<forall>y\<in>S. y \<le> x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1944 | shows "\<exists>x. (\<forall>y\<in>S. y \<le> x) & | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1945 | (\<forall>z. ((\<forall>y\<in>S. y \<le> z) --> x \<le> z))" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1946 | proof - | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1947 | have "\<exists>x. isLub UNIV S x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1948 | by (rule reals_complete) | 
| 41541 | 1949 | (auto simp add: isLub_def isUb_def leastP_def setle_def setge_def assms) | 
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1950 | thus ?thesis | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1951 | by (metis UNIV_I isLub_isUb isLub_le_isUb isUbD isUb_def setleI) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1952 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1953 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1954 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1955 | subsection {* The Archimedean Property of the Reals *}
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1956 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1957 | theorem reals_Archimedean: | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1958 | fixes x :: real | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1959 | assumes x_pos: "0 < x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1960 | shows "\<exists>n. inverse (of_nat (Suc n)) < x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1961 | proof (rule ccontr) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1962 | assume contr: "\<not> ?thesis" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1963 | have "\<forall>n. x * of_nat (Suc n) <= 1" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1964 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1965 | fix n | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1966 | 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 | 1967 | by (simp add: linorder_not_less) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1968 | hence "x \<le> (1 / (of_nat (Suc n)))" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1969 | by (simp add: inverse_eq_divide) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1970 | moreover have "(0::real) \<le> of_nat (Suc n)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1971 | by (rule of_nat_0_le_iff) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1972 | 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 | 1973 | by (rule mult_right_mono) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1974 | 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 | 1975 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1976 |   hence "{z. \<exists>n. z = x * (of_nat (Suc n))} *<= 1"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1977 | by (simp add: setle_def del: of_nat_Suc, safe, rule spec) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1978 |   hence "isUb (UNIV::real set) {z. \<exists>n. z = x * (of_nat (Suc n))} 1"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1979 | by (simp add: isUbI) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1980 |   hence "\<exists>Y. isUb (UNIV::real set) {z. \<exists>n. z = x* (of_nat (Suc n))} Y" ..
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1981 |   moreover have "\<exists>X. X \<in> {z. \<exists>n. z = x* (of_nat (Suc n))}" by auto
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1982 |   ultimately have "\<exists>t. isLub UNIV {z. \<exists>n. z = x * of_nat (Suc n)} t"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1983 | by (simp add: reals_complete) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1984 | then obtain "t" where | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1985 |     t_is_Lub: "isLub UNIV {z. \<exists>n. z = x * of_nat (Suc n)} t" ..
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1986 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1987 | have "\<forall>n::nat. x * of_nat n \<le> t + - x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1988 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1989 | fix n | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1990 | from t_is_Lub have "x * of_nat (Suc n) \<le> t" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1991 | by (simp add: isLubD2) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1992 | hence "x * (of_nat n) + x \<le> t" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1993 | by (simp add: right_distrib) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1994 | thus "x * (of_nat n) \<le> t + - x" by arith | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1995 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1996 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1997 | hence "\<forall>m. x * of_nat (Suc m) \<le> t + - x" by (simp del: of_nat_Suc) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1998 |   hence "{z. \<exists>n. z = x * (of_nat (Suc n))}  *<= (t + - x)"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 1999 | by (auto simp add: setle_def) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2000 |   hence "isUb (UNIV::real set) {z. \<exists>n. z = x * (of_nat (Suc n))} (t + (-x))"
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2001 | by (simp add: isUbI) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2002 | hence "t \<le> t + - x" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2003 | using t_is_Lub by (simp add: isLub_le_isUb) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2004 | thus False using x_pos by arith | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2005 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2006 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2007 | text {*
 | 
| 37388 | 2008 |   There must be other proofs, e.g. @{text Suc} of the largest
 | 
| 36793 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2009 |   integer in the cut representing @{text "x"}.
 | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2010 | *} | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2011 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2012 | 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 | 2013 | proof cases | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2014 | assume "x \<le> 0" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2015 | hence "x < of_nat (1::nat)" by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2016 | thus ?thesis .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2017 | next | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2018 | assume "\<not> x \<le> 0" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2019 | hence x_greater_zero: "0 < x" by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2020 | hence "0 < inverse x" by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2021 | 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 | 2022 | using reals_Archimedean by blast | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2023 | 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 | 2024 | 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 | 2025 | hence "inverse (of_nat (Suc n)) * x < 1" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2026 | using x_greater_zero by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2027 | 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 | 2028 | 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 | 2029 | hence "x < of_nat (Suc n)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2030 | by (simp add: algebra_simps del: of_nat_Suc) | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2031 | thus "\<exists>(n::nat). x < of_nat n" .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2032 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2033 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2034 | instance real :: archimedean_field | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2035 | proof | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2036 | fix r :: real | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2037 | obtain n :: nat where "r < of_nat n" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2038 | using reals_Archimedean2 .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2039 | then have "r \<le> of_int (int n)" | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2040 | by simp | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2041 | then show "\<exists>z. r \<le> of_int z" .. | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2042 | qed | 
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2043 | |
| 
27da0a27b76f
put construction of reals using Dedekind cuts in HOL/ex
 huffman parents: diff
changeset | 2044 | end |