misc tuning;
authorwenzelm
Tue, 21 Feb 2012 16:28:18 +0100
changeset 474448c4c5c8dcf7a
parent 47443 3074685ab7ed
child 47445 41701fce8ac7
misc tuning;
more indentation;
src/HOL/Library/Float.thy
src/HOL/Library/Fraction_Field.thy
     1.1 --- a/src/HOL/Library/Float.thy	Tue Feb 21 16:04:58 2012 +0100
     1.2 +++ b/src/HOL/Library/Float.thy	Tue Feb 21 16:28:18 2012 +0100
     1.3 @@ -9,8 +9,7 @@
     1.4  imports Complex_Main Lattice_Algebras
     1.5  begin
     1.6  
     1.7 -definition
     1.8 -  pow2 :: "int \<Rightarrow> real" where
     1.9 +definition pow2 :: "int \<Rightarrow> real" where
    1.10    [simp]: "pow2 a = (if (0 <= a) then (2^(nat a)) else (inverse (2^(nat (-a)))))"
    1.11  
    1.12  datatype float = Float int int
    1.13 @@ -30,17 +29,20 @@
    1.14  primrec scale :: "float \<Rightarrow> int" where
    1.15    "scale (Float a b) = b"
    1.16  
    1.17 -instantiation float :: zero begin
    1.18 +instantiation float :: zero
    1.19 +begin
    1.20  definition zero_float where "0 = Float 0 0"
    1.21  instance ..
    1.22  end
    1.23  
    1.24 -instantiation float :: one begin
    1.25 +instantiation float :: one
    1.26 +begin
    1.27  definition one_float where "1 = Float 1 0"
    1.28  instance ..
    1.29  end
    1.30  
    1.31 -instantiation float :: number begin
    1.32 +instantiation float :: number
    1.33 +begin
    1.34  definition number_of_float where "number_of n = Float n 0"
    1.35  instance ..
    1.36  end
    1.37 @@ -124,13 +126,13 @@
    1.38    by (auto simp add: pow2_def)
    1.39  
    1.40  lemma pow2_int: "pow2 (int c) = 2^c"
    1.41 -by (simp add: pow2_def)
    1.42 +  by (simp add: pow2_def)
    1.43  
    1.44 -lemma zero_less_pow2[simp]:
    1.45 -  "0 < pow2 x"
    1.46 +lemma zero_less_pow2[simp]: "0 < pow2 x"
    1.47    by (simp add: pow2_powr)
    1.48  
    1.49 -lemma normfloat_imp_odd_or_zero: "normfloat f = Float a b \<Longrightarrow> odd a \<or> (a = 0 \<and> b = 0)"
    1.50 +lemma normfloat_imp_odd_or_zero:
    1.51 +  "normfloat f = Float a b \<Longrightarrow> odd a \<or> (a = 0 \<and> b = 0)"
    1.52  proof (induct f rule: normfloat.induct)
    1.53    case (1 u v)
    1.54    from 1 have ab: "normfloat (Float u v) = Float a b" by auto
    1.55 @@ -169,7 +171,7 @@
    1.56  
    1.57  lemma float_eq_odd_helper: 
    1.58    assumes odd: "odd a'"
    1.59 -  and floateq: "real (Float a b) = real (Float a' b')"
    1.60 +    and floateq: "real (Float a b) = real (Float a' b')"
    1.61    shows "b \<le> b'"
    1.62  proof - 
    1.63    from odd have "a' \<noteq> 0" by auto
    1.64 @@ -191,8 +193,8 @@
    1.65  
    1.66  lemma float_eq_odd: 
    1.67    assumes odd1: "odd a"
    1.68 -  and odd2: "odd a'"
    1.69 -  and floateq: "real (Float a b) = real (Float a' b')"
    1.70 +    and odd2: "odd a'"
    1.71 +    and floateq: "real (Float a b) = real (Float a' b')"
    1.72    shows "a = a' \<and> b = b'"
    1.73  proof -
    1.74    from 
    1.75 @@ -216,7 +218,7 @@
    1.76    have ab': "odd a' \<or> (a' = 0 \<and> b' = 0)" by (rule normfloat_imp_odd_or_zero[OF normg])
    1.77    {
    1.78      assume odd: "odd a"
    1.79 -    then have "a \<noteq> 0" by (simp add: even_def, arith)
    1.80 +    then have "a \<noteq> 0" by (simp add: even_def) arith
    1.81      with float_eq have "a' \<noteq> 0" by auto
    1.82      with ab' have "odd a'" by simp
    1.83      from odd this float_eq have "a = a' \<and> b = b'" by (rule float_eq_odd)
    1.84 @@ -236,7 +238,8 @@
    1.85      done
    1.86  qed
    1.87  
    1.88 -instantiation float :: plus begin
    1.89 +instantiation float :: plus
    1.90 +begin
    1.91  fun plus_float where
    1.92  [simp del]: "(Float a_m a_e) + (Float b_m b_e) = 
    1.93       (if a_e \<le> b_e then Float (a_m + b_m * 2^(nat(b_e - a_e))) a_e 
    1.94 @@ -244,17 +247,20 @@
    1.95  instance ..
    1.96  end
    1.97  
    1.98 -instantiation float :: uminus begin
    1.99 +instantiation float :: uminus
   1.100 +begin
   1.101  primrec uminus_float where [simp del]: "uminus_float (Float m e) = Float (-m) e"
   1.102  instance ..
   1.103  end
   1.104  
   1.105 -instantiation float :: minus begin
   1.106 -definition minus_float where [simp del]: "(z::float) - w = z + (- w)"
   1.107 +instantiation float :: minus
   1.108 +begin
   1.109 +definition minus_float where "(z::float) - w = z + (- w)"
   1.110  instance ..
   1.111  end
   1.112  
   1.113 -instantiation float :: times begin
   1.114 +instantiation float :: times
   1.115 +begin
   1.116  fun times_float where [simp del]: "(Float a_m a_e) * (Float b_m b_e) = Float (a_m * b_m) (a_e + b_e)"
   1.117  instance ..
   1.118  end
   1.119 @@ -265,7 +271,8 @@
   1.120  primrec float_nprt :: "float \<Rightarrow> float" where
   1.121    "float_nprt (Float a e) = (if 0 <= a then 0 else (Float a e))" 
   1.122  
   1.123 -instantiation float :: ord begin
   1.124 +instantiation float :: ord
   1.125 +begin
   1.126  definition le_float_def: "z \<le> (w :: float) \<equiv> real z \<le> real w"
   1.127  definition less_float_def: "z < (w :: float) \<equiv> real z < real w"
   1.128  instance ..
   1.129 @@ -276,7 +283,7 @@
   1.130        auto simp add: pow2_int[symmetric] pow2_add[symmetric])
   1.131  
   1.132  lemma real_of_float_minus[simp]: "real (- a) = - real (a :: float)"
   1.133 -  by (cases a) (simp add: uminus_float.simps)
   1.134 +  by (cases a) simp
   1.135  
   1.136  lemma real_of_float_sub[simp]: "real (a - b) = real a - real (b :: float)"
   1.137    by (cases a, cases b) (simp add: minus_float_def)
   1.138 @@ -285,7 +292,7 @@
   1.139    by (cases a, cases b) (simp add: times_float.simps pow2_add)
   1.140  
   1.141  lemma real_of_float_0[simp]: "real (0 :: float) = 0"
   1.142 -  by (auto simp add: zero_float_def float_zero)
   1.143 +  by (auto simp add: zero_float_def)
   1.144  
   1.145  lemma real_of_float_1[simp]: "real (1 :: float) = 1"
   1.146    by (auto simp add: one_float_def)
   1.147 @@ -1161,16 +1168,20 @@
   1.148  qed
   1.149  
   1.150  definition lb_mult :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" where
   1.151 -"lb_mult prec x y = (case normfloat (x * y) of Float m e \<Rightarrow> let
   1.152 -    l = bitlen m - int prec
   1.153 -  in if l > 0 then Float (m div (2^nat l)) (e + l)
   1.154 -              else Float m e)"
   1.155 +  "lb_mult prec x y =
   1.156 +    (case normfloat (x * y) of Float m e \<Rightarrow>
   1.157 +      let
   1.158 +        l = bitlen m - int prec
   1.159 +      in if l > 0 then Float (m div (2^nat l)) (e + l)
   1.160 +                  else Float m e)"
   1.161  
   1.162  definition ub_mult :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" where
   1.163 -"ub_mult prec x y = (case normfloat (x * y) of Float m e \<Rightarrow> let
   1.164 -    l = bitlen m - int prec
   1.165 -  in if l > 0 then Float (m div (2^nat l) + 1) (e + l)
   1.166 -              else Float m e)"
   1.167 +  "ub_mult prec x y =
   1.168 +    (case normfloat (x * y) of Float m e \<Rightarrow>
   1.169 +      let
   1.170 +        l = bitlen m - int prec
   1.171 +      in if l > 0 then Float (m div (2^nat l) + 1) (e + l)
   1.172 +                  else Float m e)"
   1.173  
   1.174  lemma lb_mult: "real (lb_mult prec x y) \<le> real (x * y)"
   1.175  proof (cases "normfloat (x * y)")
   1.176 @@ -1225,7 +1236,8 @@
   1.177  primrec float_abs :: "float \<Rightarrow> float" where
   1.178    "float_abs (Float m e) = Float \<bar>m\<bar> e"
   1.179  
   1.180 -instantiation float :: abs begin
   1.181 +instantiation float :: abs
   1.182 +begin
   1.183  definition abs_float_def: "\<bar>x\<bar> = float_abs x"
   1.184  instance ..
   1.185  end
   1.186 @@ -1290,10 +1302,10 @@
   1.187  declare ceiling_fl.simps[simp del]
   1.188  
   1.189  definition lb_mod :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" where
   1.190 -"lb_mod prec x ub lb = x - ceiling_fl (float_divr prec x lb) * ub"
   1.191 +  "lb_mod prec x ub lb = x - ceiling_fl (float_divr prec x lb) * ub"
   1.192  
   1.193  definition ub_mod :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" where
   1.194 -"ub_mod prec x ub lb = x - floor_fl (float_divl prec x ub) * lb"
   1.195 +  "ub_mod prec x ub lb = x - floor_fl (float_divl prec x ub) * lb"
   1.196  
   1.197  lemma lb_mod: fixes k :: int assumes "0 \<le> real x" and "real k * y \<le> real x" (is "?k * y \<le> ?x")
   1.198    assumes "0 < real lb" "real lb \<le> y" (is "?lb \<le> y") "y \<le> real ub" (is "y \<le> ?ub")
   1.199 @@ -1307,7 +1319,9 @@
   1.200    finally show ?thesis unfolding lb_mod_def real_of_float_sub real_of_float_mult by auto
   1.201  qed
   1.202  
   1.203 -lemma ub_mod: fixes k :: int and x :: float assumes "0 \<le> real x" and "real x \<le> real k * y" (is "?x \<le> ?k * y")
   1.204 +lemma ub_mod:
   1.205 +  fixes k :: int and x :: float
   1.206 +  assumes "0 \<le> real x" and "real x \<le> real k * y" (is "?x \<le> ?k * y")
   1.207    assumes "0 < real lb" "real lb \<le> y" (is "?lb \<le> y") "y \<le> real ub" (is "y \<le> ?ub")
   1.208    shows "?x - ?k * y \<le> real (ub_mod prec x ub lb)"
   1.209  proof -
     2.1 --- a/src/HOL/Library/Fraction_Field.thy	Tue Feb 21 16:04:58 2012 +0100
     2.2 +++ b/src/HOL/Library/Fraction_Field.thy	Tue Feb 21 16:28:18 2012 +0100
     2.3 @@ -2,8 +2,8 @@
     2.4      Author:     Amine Chaieb, University of Cambridge
     2.5  *)
     2.6  
     2.7 -header{* A formalization of the fraction field of any integral domain 
     2.8 -         A generalization of Rat.thy from int to any integral domain *}
     2.9 +header{* A formalization of the fraction field of any integral domain;
    2.10 +         generalization of theory Rat from int to any integral domain *}
    2.11  
    2.12  theory Fraction_Field
    2.13  imports Main
    2.14 @@ -14,7 +14,7 @@
    2.15  subsubsection {* Construction of the type of fractions *}
    2.16  
    2.17  definition fractrel :: "(('a::idom * 'a ) * ('a * 'a)) set" where
    2.18 -  "fractrel == {(x, y). snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x}"
    2.19 +  "fractrel = {(x, y). snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x}"
    2.20  
    2.21  lemma fractrel_iff [simp]:
    2.22    "(x, y) \<in> fractrel \<longleftrightarrow> snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x"
    2.23 @@ -70,8 +70,7 @@
    2.24  
    2.25  subsubsection {* Representation and basic operations *}
    2.26  
    2.27 -definition
    2.28 -  Fract :: "'a::idom \<Rightarrow> 'a \<Rightarrow> 'a fract" where
    2.29 +definition Fract :: "'a::idom \<Rightarrow> 'a \<Rightarrow> 'a fract" where
    2.30    "Fract a b = Abs_fract (fractrel `` {if b = 0 then (0, 1) else (a, b)})"
    2.31  
    2.32  code_datatype Fract
    2.33 @@ -95,14 +94,11 @@
    2.34  instantiation fract :: (idom) "{comm_ring_1, power}"
    2.35  begin
    2.36  
    2.37 -definition
    2.38 -  Zero_fract_def [code_unfold]: "0 = Fract 0 1"
    2.39 +definition Zero_fract_def [code_unfold]: "0 = Fract 0 1"
    2.40  
    2.41 -definition
    2.42 -  One_fract_def [code_unfold]: "1 = Fract 1 1"
    2.43 +definition One_fract_def [code_unfold]: "1 = Fract 1 1"
    2.44  
    2.45 -definition
    2.46 -  add_fract_def:
    2.47 +definition add_fract_def:
    2.48    "q + r = Abs_fract (\<Union>x \<in> Rep_fract q. \<Union>y \<in> Rep_fract r.
    2.49      fractrel `` {(fst x * snd y + fst y * snd x, snd x * snd y)})"
    2.50  
    2.51 @@ -117,8 +113,7 @@
    2.52    with assms show ?thesis by (simp add: Fract_def add_fract_def UN_fractrel2)
    2.53  qed
    2.54  
    2.55 -definition
    2.56 -  minus_fract_def:
    2.57 +definition minus_fract_def:
    2.58    "- q = Abs_fract (\<Union>x \<in> Rep_fract q. fractrel `` {(- fst x, snd x)})"
    2.59  
    2.60  lemma minus_fract [simp, code]: "- Fract a b = Fract (- a) (b::'a::idom)"
    2.61 @@ -131,16 +126,14 @@
    2.62  lemma minus_fract_cancel [simp]: "Fract (- a) (- b) = Fract a b"
    2.63    by (cases "b = 0") (simp_all add: eq_fract)
    2.64  
    2.65 -definition
    2.66 -  diff_fract_def: "q - r = q + - (r::'a fract)"
    2.67 +definition diff_fract_def: "q - r = q + - (r::'a fract)"
    2.68  
    2.69  lemma diff_fract [simp]:
    2.70    assumes "b \<noteq> 0" and "d \<noteq> 0"
    2.71    shows "Fract a b - Fract c d = Fract (a * d - c * b) (b * d)"
    2.72    using assms by (simp add: diff_fract_def diff_minus)
    2.73  
    2.74 -definition
    2.75 -  mult_fract_def:
    2.76 +definition mult_fract_def:
    2.77    "q * r = Abs_fract (\<Union>x \<in> Rep_fract q. \<Union>y \<in> Rep_fract r.
    2.78      fractrel``{(fst x * fst y, snd x * snd y)})"
    2.79  
    2.80 @@ -238,8 +231,7 @@
    2.81  instantiation fract :: (idom) field_inverse_zero
    2.82  begin
    2.83  
    2.84 -definition
    2.85 -  inverse_fract_def:
    2.86 +definition inverse_fract_def:
    2.87    "inverse q = Abs_fract (\<Union>x \<in> Rep_fract q.
    2.88       fractrel `` {if fst x = 0 then (0, 1) else (snd x, fst x)})"
    2.89  
    2.90 @@ -252,8 +244,7 @@
    2.91    then show ?thesis by (simp add: Fract_def inverse_fract_def UN_fractrel)
    2.92  qed
    2.93  
    2.94 -definition
    2.95 -  divide_fract_def: "q / r = q * inverse (r:: 'a fract)"
    2.96 +definition divide_fract_def: "q / r = q * inverse (r:: 'a fract)"
    2.97  
    2.98  lemma divide_fract [simp]: "Fract a b / Fract c d = Fract (a * d) (b * c)"
    2.99    by (simp add: divide_fract_def)
   2.100 @@ -261,14 +252,15 @@
   2.101  instance proof
   2.102    fix q :: "'a fract"
   2.103    assume "q \<noteq> 0"
   2.104 -  then show "inverse q * q = 1" apply (cases q rule: Fract_cases_nonzero)
   2.105 -    by (simp_all add: mult_fract  inverse_fract fract_expand eq_fract mult_commute)
   2.106 +  then show "inverse q * q = 1"
   2.107 +    by (cases q rule: Fract_cases_nonzero)
   2.108 +      (simp_all add: fract_expand eq_fract mult_commute)
   2.109  next
   2.110    fix q r :: "'a fract"
   2.111    show "q / r = q * inverse r" by (simp add: divide_fract_def)
   2.112  next
   2.113 -  show "inverse 0 = (0:: 'a fract)" by (simp add: fract_expand)
   2.114 -    (simp add: fract_collapse)
   2.115 +  show "inverse 0 = (0:: 'a fract)"
   2.116 +    by (simp add: fract_expand) (simp add: fract_collapse)
   2.117  qed
   2.118  
   2.119  end
   2.120 @@ -292,7 +284,7 @@
   2.121      have "?le a b c d = ?le (a * x) (b * x) c d"
   2.122      proof -
   2.123        from x have "0 < x * x" by (auto simp add: zero_less_mult_iff)
   2.124 -      hence "?le a b c d =
   2.125 +      then have "?le a b c d =
   2.126            ((a * d) * (b * d) * (x * x) \<le> (c * b) * (b * d) * (x * x))"
   2.127          by (simp add: mult_le_cancel_right)
   2.128        also have "... = ?le (a * x) (b * x) c d"
   2.129 @@ -304,7 +296,7 @@
   2.130    let ?D = "b * d" and ?D' = "b' * d'"
   2.131    from neq have D: "?D \<noteq> 0" by simp
   2.132    from neq have "?D' \<noteq> 0" by simp
   2.133 -  hence "?le a b c d = ?le (a * ?D') (b * ?D') c d"
   2.134 +  then have "?le a b c d = ?le (a * ?D') (b * ?D') c d"
   2.135      by (rule le_factor)
   2.136    also have "... = ((a * b') * ?D * ?D' * d * d' \<le> (c * d') * ?D * ?D' * b * b')"
   2.137      by (simp add: mult_ac)
   2.138 @@ -320,13 +312,11 @@
   2.139  instantiation fract :: (linordered_idom) linorder
   2.140  begin
   2.141  
   2.142 -definition
   2.143 -  le_fract_def:
   2.144 +definition le_fract_def:
   2.145     "q \<le> r \<longleftrightarrow> the_elem (\<Union>x \<in> Rep_fract q. \<Union>y \<in> Rep_fract r.
   2.146        {(fst x * snd y)*(snd x * snd y) \<le> (fst y * snd x)*(snd x * snd y)})"
   2.147  
   2.148 -definition
   2.149 -  less_fract_def: "z < (w::'a fract) \<longleftrightarrow> z \<le> w \<and> \<not> w \<le> z"
   2.150 +definition less_fract_def: "z < (w::'a fract) \<longleftrightarrow> z \<le> w \<and> \<not> w \<le> z"
   2.151  
   2.152  lemma le_fract [simp]:
   2.153    assumes "b \<noteq> 0" and "d \<noteq> 0"
   2.154 @@ -409,28 +399,25 @@
   2.155  instantiation fract :: (linordered_idom) "{distrib_lattice, abs_if, sgn_if}"
   2.156  begin
   2.157  
   2.158 -definition
   2.159 -  abs_fract_def: "\<bar>q\<bar> = (if q < 0 then -q else (q::'a fract))"
   2.160 +definition abs_fract_def: "\<bar>q\<bar> = (if q < 0 then -q else (q::'a fract))"
   2.161  
   2.162 -definition
   2.163 -  sgn_fract_def:
   2.164 -    "sgn (q::'a fract) = (if q=0 then 0 else if 0<q then 1 else - 1)"
   2.165 +definition sgn_fract_def:
   2.166 +  "sgn (q::'a fract) = (if q=0 then 0 else if 0<q then 1 else - 1)"
   2.167  
   2.168  theorem abs_fract [simp]: "\<bar>Fract a b\<bar> = Fract \<bar>a\<bar> \<bar>b\<bar>"
   2.169    by (auto simp add: abs_fract_def Zero_fract_def le_less
   2.170        eq_fract zero_less_mult_iff mult_less_0_iff split: abs_split)
   2.171  
   2.172 -definition
   2.173 -  inf_fract_def:
   2.174 -    "(inf \<Colon> 'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract) = min"
   2.175 +definition inf_fract_def:
   2.176 +  "(inf \<Colon> 'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract) = min"
   2.177  
   2.178 -definition
   2.179 -  sup_fract_def:
   2.180 -    "(sup \<Colon> 'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract) = max"
   2.181 +definition sup_fract_def:
   2.182 +  "(sup \<Colon> 'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract) = max"
   2.183  
   2.184 -instance by intro_classes
   2.185 -  (auto simp add: abs_fract_def sgn_fract_def
   2.186 -    min_max.sup_inf_distrib1 inf_fract_def sup_fract_def)
   2.187 +instance
   2.188 +  by intro_classes
   2.189 +    (auto simp add: abs_fract_def sgn_fract_def
   2.190 +      min_max.sup_inf_distrib1 inf_fract_def sup_fract_def)
   2.191  
   2.192  end
   2.193  
   2.194 @@ -485,8 +472,8 @@
   2.195    proof -
   2.196      fix a::'a and b::'a
   2.197      assume b: "b < 0"
   2.198 -    hence "0 < -b" by simp
   2.199 -    hence "P (Fract (-a) (-b))" by (rule step)
   2.200 +    then have "0 < -b" by simp
   2.201 +    then have "P (Fract (-a) (-b))" by (rule step)
   2.202      thus "P (Fract a b)" by (simp add: order_less_imp_not_eq [OF b])
   2.203    qed
   2.204    case (Fract a b)