misc tuning;
authorwenzelm
Tue, 21 Feb 2012 16:28:18 +0100
changeset 46573 8c4c5c8dcf7a
parent 46572 3074685ab7ed
child 46574 41701fce8ac7
misc tuning; more indentation;
src/HOL/Library/Float.thy
src/HOL/Library/Fraction_Field.thy
--- a/src/HOL/Library/Float.thy	Tue Feb 21 16:04:58 2012 +0100
+++ b/src/HOL/Library/Float.thy	Tue Feb 21 16:28:18 2012 +0100
@@ -9,8 +9,7 @@
 imports Complex_Main Lattice_Algebras
 begin
 
-definition
-  pow2 :: "int \<Rightarrow> real" where
+definition pow2 :: "int \<Rightarrow> real" where
   [simp]: "pow2 a = (if (0 <= a) then (2^(nat a)) else (inverse (2^(nat (-a)))))"
 
 datatype float = Float int int
@@ -30,17 +29,20 @@
 primrec scale :: "float \<Rightarrow> int" where
   "scale (Float a b) = b"
 
-instantiation float :: zero begin
+instantiation float :: zero
+begin
 definition zero_float where "0 = Float 0 0"
 instance ..
 end
 
-instantiation float :: one begin
+instantiation float :: one
+begin
 definition one_float where "1 = Float 1 0"
 instance ..
 end
 
-instantiation float :: number begin
+instantiation float :: number
+begin
 definition number_of_float where "number_of n = Float n 0"
 instance ..
 end
@@ -124,13 +126,13 @@
   by (auto simp add: pow2_def)
 
 lemma pow2_int: "pow2 (int c) = 2^c"
-by (simp add: pow2_def)
+  by (simp add: pow2_def)
 
-lemma zero_less_pow2[simp]:
-  "0 < pow2 x"
+lemma zero_less_pow2[simp]: "0 < pow2 x"
   by (simp add: pow2_powr)
 
-lemma normfloat_imp_odd_or_zero: "normfloat f = Float a b \<Longrightarrow> odd a \<or> (a = 0 \<and> b = 0)"
+lemma normfloat_imp_odd_or_zero:
+  "normfloat f = Float a b \<Longrightarrow> odd a \<or> (a = 0 \<and> b = 0)"
 proof (induct f rule: normfloat.induct)
   case (1 u v)
   from 1 have ab: "normfloat (Float u v) = Float a b" by auto
@@ -169,7 +171,7 @@
 
 lemma float_eq_odd_helper: 
   assumes odd: "odd a'"
-  and floateq: "real (Float a b) = real (Float a' b')"
+    and floateq: "real (Float a b) = real (Float a' b')"
   shows "b \<le> b'"
 proof - 
   from odd have "a' \<noteq> 0" by auto
@@ -191,8 +193,8 @@
 
 lemma float_eq_odd: 
   assumes odd1: "odd a"
-  and odd2: "odd a'"
-  and floateq: "real (Float a b) = real (Float a' b')"
+    and odd2: "odd a'"
+    and floateq: "real (Float a b) = real (Float a' b')"
   shows "a = a' \<and> b = b'"
 proof -
   from 
@@ -216,7 +218,7 @@
   have ab': "odd a' \<or> (a' = 0 \<and> b' = 0)" by (rule normfloat_imp_odd_or_zero[OF normg])
   {
     assume odd: "odd a"
-    then have "a \<noteq> 0" by (simp add: even_def, arith)
+    then have "a \<noteq> 0" by (simp add: even_def) arith
     with float_eq have "a' \<noteq> 0" by auto
     with ab' have "odd a'" by simp
     from odd this float_eq have "a = a' \<and> b = b'" by (rule float_eq_odd)
@@ -236,7 +238,8 @@
     done
 qed
 
-instantiation float :: plus begin
+instantiation float :: plus
+begin
 fun plus_float where
 [simp del]: "(Float a_m a_e) + (Float b_m b_e) = 
      (if a_e \<le> b_e then Float (a_m + b_m * 2^(nat(b_e - a_e))) a_e 
@@ -244,17 +247,20 @@
 instance ..
 end
 
-instantiation float :: uminus begin
+instantiation float :: uminus
+begin
 primrec uminus_float where [simp del]: "uminus_float (Float m e) = Float (-m) e"
 instance ..
 end
 
-instantiation float :: minus begin
-definition minus_float where [simp del]: "(z::float) - w = z + (- w)"
+instantiation float :: minus
+begin
+definition minus_float where "(z::float) - w = z + (- w)"
 instance ..
 end
 
-instantiation float :: times begin
+instantiation float :: times
+begin
 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)"
 instance ..
 end
@@ -265,7 +271,8 @@
 primrec float_nprt :: "float \<Rightarrow> float" where
   "float_nprt (Float a e) = (if 0 <= a then 0 else (Float a e))" 
 
-instantiation float :: ord begin
+instantiation float :: ord
+begin
 definition le_float_def: "z \<le> (w :: float) \<equiv> real z \<le> real w"
 definition less_float_def: "z < (w :: float) \<equiv> real z < real w"
 instance ..
@@ -276,7 +283,7 @@
       auto simp add: pow2_int[symmetric] pow2_add[symmetric])
 
 lemma real_of_float_minus[simp]: "real (- a) = - real (a :: float)"
-  by (cases a) (simp add: uminus_float.simps)
+  by (cases a) simp
 
 lemma real_of_float_sub[simp]: "real (a - b) = real a - real (b :: float)"
   by (cases a, cases b) (simp add: minus_float_def)
@@ -285,7 +292,7 @@
   by (cases a, cases b) (simp add: times_float.simps pow2_add)
 
 lemma real_of_float_0[simp]: "real (0 :: float) = 0"
-  by (auto simp add: zero_float_def float_zero)
+  by (auto simp add: zero_float_def)
 
 lemma real_of_float_1[simp]: "real (1 :: float) = 1"
   by (auto simp add: one_float_def)
@@ -1161,16 +1168,20 @@
 qed
 
 definition lb_mult :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" where
-"lb_mult prec x y = (case normfloat (x * y) of Float m e \<Rightarrow> let
-    l = bitlen m - int prec
-  in if l > 0 then Float (m div (2^nat l)) (e + l)
-              else Float m e)"
+  "lb_mult prec x y =
+    (case normfloat (x * y) of Float m e \<Rightarrow>
+      let
+        l = bitlen m - int prec
+      in if l > 0 then Float (m div (2^nat l)) (e + l)
+                  else Float m e)"
 
 definition ub_mult :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" where
-"ub_mult prec x y = (case normfloat (x * y) of Float m e \<Rightarrow> let
-    l = bitlen m - int prec
-  in if l > 0 then Float (m div (2^nat l) + 1) (e + l)
-              else Float m e)"
+  "ub_mult prec x y =
+    (case normfloat (x * y) of Float m e \<Rightarrow>
+      let
+        l = bitlen m - int prec
+      in if l > 0 then Float (m div (2^nat l) + 1) (e + l)
+                  else Float m e)"
 
 lemma lb_mult: "real (lb_mult prec x y) \<le> real (x * y)"
 proof (cases "normfloat (x * y)")
@@ -1225,7 +1236,8 @@
 primrec float_abs :: "float \<Rightarrow> float" where
   "float_abs (Float m e) = Float \<bar>m\<bar> e"
 
-instantiation float :: abs begin
+instantiation float :: abs
+begin
 definition abs_float_def: "\<bar>x\<bar> = float_abs x"
 instance ..
 end
@@ -1290,10 +1302,10 @@
 declare ceiling_fl.simps[simp del]
 
 definition lb_mod :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" where
-"lb_mod prec x ub lb = x - ceiling_fl (float_divr prec x lb) * ub"
+  "lb_mod prec x ub lb = x - ceiling_fl (float_divr prec x lb) * ub"
 
 definition ub_mod :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" where
-"ub_mod prec x ub lb = x - floor_fl (float_divl prec x ub) * lb"
+  "ub_mod prec x ub lb = x - floor_fl (float_divl prec x ub) * lb"
 
 lemma lb_mod: fixes k :: int assumes "0 \<le> real x" and "real k * y \<le> real x" (is "?k * y \<le> ?x")
   assumes "0 < real lb" "real lb \<le> y" (is "?lb \<le> y") "y \<le> real ub" (is "y \<le> ?ub")
@@ -1307,7 +1319,9 @@
   finally show ?thesis unfolding lb_mod_def real_of_float_sub real_of_float_mult by auto
 qed
 
-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")
+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")
   assumes "0 < real lb" "real lb \<le> y" (is "?lb \<le> y") "y \<le> real ub" (is "y \<le> ?ub")
   shows "?x - ?k * y \<le> real (ub_mod prec x ub lb)"
 proof -
--- a/src/HOL/Library/Fraction_Field.thy	Tue Feb 21 16:04:58 2012 +0100
+++ b/src/HOL/Library/Fraction_Field.thy	Tue Feb 21 16:28:18 2012 +0100
@@ -2,8 +2,8 @@
     Author:     Amine Chaieb, University of Cambridge
 *)
 
-header{* A formalization of the fraction field of any integral domain 
-         A generalization of Rat.thy from int to any integral domain *}
+header{* A formalization of the fraction field of any integral domain;
+         generalization of theory Rat from int to any integral domain *}
 
 theory Fraction_Field
 imports Main
@@ -14,7 +14,7 @@
 subsubsection {* Construction of the type of fractions *}
 
 definition fractrel :: "(('a::idom * 'a ) * ('a * 'a)) set" where
-  "fractrel == {(x, y). snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x}"
+  "fractrel = {(x, y). snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x}"
 
 lemma fractrel_iff [simp]:
   "(x, y) \<in> fractrel \<longleftrightarrow> snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x"
@@ -70,8 +70,7 @@
 
 subsubsection {* Representation and basic operations *}
 
-definition
-  Fract :: "'a::idom \<Rightarrow> 'a \<Rightarrow> 'a fract" where
+definition Fract :: "'a::idom \<Rightarrow> 'a \<Rightarrow> 'a fract" where
   "Fract a b = Abs_fract (fractrel `` {if b = 0 then (0, 1) else (a, b)})"
 
 code_datatype Fract
@@ -95,14 +94,11 @@
 instantiation fract :: (idom) "{comm_ring_1, power}"
 begin
 
-definition
-  Zero_fract_def [code_unfold]: "0 = Fract 0 1"
+definition Zero_fract_def [code_unfold]: "0 = Fract 0 1"
 
-definition
-  One_fract_def [code_unfold]: "1 = Fract 1 1"
+definition One_fract_def [code_unfold]: "1 = Fract 1 1"
 
-definition
-  add_fract_def:
+definition add_fract_def:
   "q + r = Abs_fract (\<Union>x \<in> Rep_fract q. \<Union>y \<in> Rep_fract r.
     fractrel `` {(fst x * snd y + fst y * snd x, snd x * snd y)})"
 
@@ -117,8 +113,7 @@
   with assms show ?thesis by (simp add: Fract_def add_fract_def UN_fractrel2)
 qed
 
-definition
-  minus_fract_def:
+definition minus_fract_def:
   "- q = Abs_fract (\<Union>x \<in> Rep_fract q. fractrel `` {(- fst x, snd x)})"
 
 lemma minus_fract [simp, code]: "- Fract a b = Fract (- a) (b::'a::idom)"
@@ -131,16 +126,14 @@
 lemma minus_fract_cancel [simp]: "Fract (- a) (- b) = Fract a b"
   by (cases "b = 0") (simp_all add: eq_fract)
 
-definition
-  diff_fract_def: "q - r = q + - (r::'a fract)"
+definition diff_fract_def: "q - r = q + - (r::'a fract)"
 
 lemma diff_fract [simp]:
   assumes "b \<noteq> 0" and "d \<noteq> 0"
   shows "Fract a b - Fract c d = Fract (a * d - c * b) (b * d)"
   using assms by (simp add: diff_fract_def diff_minus)
 
-definition
-  mult_fract_def:
+definition mult_fract_def:
   "q * r = Abs_fract (\<Union>x \<in> Rep_fract q. \<Union>y \<in> Rep_fract r.
     fractrel``{(fst x * fst y, snd x * snd y)})"
 
@@ -238,8 +231,7 @@
 instantiation fract :: (idom) field_inverse_zero
 begin
 
-definition
-  inverse_fract_def:
+definition inverse_fract_def:
   "inverse q = Abs_fract (\<Union>x \<in> Rep_fract q.
      fractrel `` {if fst x = 0 then (0, 1) else (snd x, fst x)})"
 
@@ -252,8 +244,7 @@
   then show ?thesis by (simp add: Fract_def inverse_fract_def UN_fractrel)
 qed
 
-definition
-  divide_fract_def: "q / r = q * inverse (r:: 'a fract)"
+definition divide_fract_def: "q / r = q * inverse (r:: 'a fract)"
 
 lemma divide_fract [simp]: "Fract a b / Fract c d = Fract (a * d) (b * c)"
   by (simp add: divide_fract_def)
@@ -261,14 +252,15 @@
 instance proof
   fix q :: "'a fract"
   assume "q \<noteq> 0"
-  then show "inverse q * q = 1" apply (cases q rule: Fract_cases_nonzero)
-    by (simp_all add: mult_fract  inverse_fract fract_expand eq_fract mult_commute)
+  then show "inverse q * q = 1"
+    by (cases q rule: Fract_cases_nonzero)
+      (simp_all add: fract_expand eq_fract mult_commute)
 next
   fix q r :: "'a fract"
   show "q / r = q * inverse r" by (simp add: divide_fract_def)
 next
-  show "inverse 0 = (0:: 'a fract)" by (simp add: fract_expand)
-    (simp add: fract_collapse)
+  show "inverse 0 = (0:: 'a fract)"
+    by (simp add: fract_expand) (simp add: fract_collapse)
 qed
 
 end
@@ -292,7 +284,7 @@
     have "?le a b c d = ?le (a * x) (b * x) c d"
     proof -
       from x have "0 < x * x" by (auto simp add: zero_less_mult_iff)
-      hence "?le a b c d =
+      then have "?le a b c d =
           ((a * d) * (b * d) * (x * x) \<le> (c * b) * (b * d) * (x * x))"
         by (simp add: mult_le_cancel_right)
       also have "... = ?le (a * x) (b * x) c d"
@@ -304,7 +296,7 @@
   let ?D = "b * d" and ?D' = "b' * d'"
   from neq have D: "?D \<noteq> 0" by simp
   from neq have "?D' \<noteq> 0" by simp
-  hence "?le a b c d = ?le (a * ?D') (b * ?D') c d"
+  then have "?le a b c d = ?le (a * ?D') (b * ?D') c d"
     by (rule le_factor)
   also have "... = ((a * b') * ?D * ?D' * d * d' \<le> (c * d') * ?D * ?D' * b * b')"
     by (simp add: mult_ac)
@@ -320,13 +312,11 @@
 instantiation fract :: (linordered_idom) linorder
 begin
 
-definition
-  le_fract_def:
+definition le_fract_def:
    "q \<le> r \<longleftrightarrow> the_elem (\<Union>x \<in> Rep_fract q. \<Union>y \<in> Rep_fract r.
       {(fst x * snd y)*(snd x * snd y) \<le> (fst y * snd x)*(snd x * snd y)})"
 
-definition
-  less_fract_def: "z < (w::'a fract) \<longleftrightarrow> z \<le> w \<and> \<not> w \<le> z"
+definition less_fract_def: "z < (w::'a fract) \<longleftrightarrow> z \<le> w \<and> \<not> w \<le> z"
 
 lemma le_fract [simp]:
   assumes "b \<noteq> 0" and "d \<noteq> 0"
@@ -409,28 +399,25 @@
 instantiation fract :: (linordered_idom) "{distrib_lattice, abs_if, sgn_if}"
 begin
 
-definition
-  abs_fract_def: "\<bar>q\<bar> = (if q < 0 then -q else (q::'a fract))"
+definition abs_fract_def: "\<bar>q\<bar> = (if q < 0 then -q else (q::'a fract))"
 
-definition
-  sgn_fract_def:
-    "sgn (q::'a fract) = (if q=0 then 0 else if 0<q then 1 else - 1)"
+definition sgn_fract_def:
+  "sgn (q::'a fract) = (if q=0 then 0 else if 0<q then 1 else - 1)"
 
 theorem abs_fract [simp]: "\<bar>Fract a b\<bar> = Fract \<bar>a\<bar> \<bar>b\<bar>"
   by (auto simp add: abs_fract_def Zero_fract_def le_less
       eq_fract zero_less_mult_iff mult_less_0_iff split: abs_split)
 
-definition
-  inf_fract_def:
-    "(inf \<Colon> 'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract) = min"
+definition inf_fract_def:
+  "(inf \<Colon> 'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract) = min"
 
-definition
-  sup_fract_def:
-    "(sup \<Colon> 'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract) = max"
+definition sup_fract_def:
+  "(sup \<Colon> 'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract) = max"
 
-instance by intro_classes
-  (auto simp add: abs_fract_def sgn_fract_def
-    min_max.sup_inf_distrib1 inf_fract_def sup_fract_def)
+instance
+  by intro_classes
+    (auto simp add: abs_fract_def sgn_fract_def
+      min_max.sup_inf_distrib1 inf_fract_def sup_fract_def)
 
 end
 
@@ -485,8 +472,8 @@
   proof -
     fix a::'a and b::'a
     assume b: "b < 0"
-    hence "0 < -b" by simp
-    hence "P (Fract (-a) (-b))" by (rule step)
+    then have "0 < -b" by simp
+    then have "P (Fract (-a) (-b))" by (rule step)
     thus "P (Fract a b)" by (simp add: order_less_imp_not_eq [OF b])
   qed
   case (Fract a b)