diff -r 31781b2de73d -r 474ff28210c0 src/HOL/Library/Char_ord.thy --- a/src/HOL/Library/Char_ord.thy Thu Jun 14 23:04:36 2007 +0200 +++ b/src/HOL/Library/Char_ord.thy Thu Jun 14 23:04:39 2007 +0200 @@ -13,32 +13,36 @@ nibble_less_eq_def: "n \ m \ nat_of_nibble n \ nat_of_nibble m" nibble_less_def: "n < m \ nat_of_nibble n < nat_of_nibble m" proof - fix n :: nibble show "n \ n" unfolding nibble_less_eq_def nibble_less_def by auto + fix n :: nibble + show "n \ n" unfolding nibble_less_eq_def nibble_less_def by auto next fix n m q :: nibble assume "n \ m" - and "m \ q" + and "m \ q" then show "n \ q" unfolding nibble_less_eq_def nibble_less_def by auto next fix n m :: nibble assume "n \ m" - and "m \ n" - then show "n = m" unfolding nibble_less_eq_def nibble_less_def by (auto simp add: nat_of_nibble_eq) + and "m \ n" + then show "n = m" + unfolding nibble_less_eq_def nibble_less_def + by (auto simp add: nat_of_nibble_eq) next fix n m :: nibble show "n < m \ n \ m \ n \ m" - unfolding nibble_less_eq_def nibble_less_def less_le by (auto simp add: nat_of_nibble_eq) + unfolding nibble_less_eq_def nibble_less_def less_le + by (auto simp add: nat_of_nibble_eq) next fix n m :: nibble show "n \ m \ m \ n" - unfolding nibble_less_eq_def by auto + unfolding nibble_less_eq_def by auto qed instance nibble :: distrib_lattice - "inf \ min" - "sup \ max" - by default - (auto simp add: inf_nibble_def sup_nibble_def min_max.sup_inf_distrib1) + "inf \ min" + "sup \ max" + by default (auto simp add: + inf_nibble_def sup_nibble_def min_max.sup_inf_distrib1) instance char :: linorder char_less_eq_def: "c1 \ c2 \ case c1 of Char n1 m1 \ case c2 of Char n2 m2 \ @@ -50,10 +54,10 @@ lemmas [code func del] = char_less_eq_def char_less_def instance char :: distrib_lattice - "inf \ min" - "sup \ max" - by default - (auto simp add: inf_char_def sup_char_def min_max.sup_inf_distrib1) + "inf \ min" + "sup \ max" + by default (auto simp add: + inf_char_def sup_char_def min_max.sup_inf_distrib1) lemma [simp, code func]: shows char_less_eq_simp: "Char n1 m1 \ Char n2 m2 \ n1 < n2 \ n1 = n2 \ m1 \ m2"