# HG changeset patch # User wenzelm # Date 1512323629 -3600 # Node ID 491fd7f0b5df7d42e76a3b7cd11d5617fabf6c89 # Parent acb0807ddb56c1aeca51afff91787ca41b6ec028 misc tuning and modernization; diff -r acb0807ddb56 -r 491fd7f0b5df src/HOL/Word/Bit_Comparison.thy --- a/src/HOL/Word/Bit_Comparison.thy Sun Dec 03 13:22:09 2017 +0100 +++ b/src/HOL/Word/Bit_Comparison.thy Sun Dec 03 18:53:49 2017 +0100 @@ -6,15 +6,21 @@ *) theory Bit_Comparison -imports Bits_Int + imports Bits_Int begin lemma AND_lower [simp]: - fixes x :: int and y :: int + fixes x y :: int assumes "0 \ x" shows "0 \ x AND y" using assms proof (induct x arbitrary: y rule: bin_induct) + case 1 + then show ?case by simp +next + case 2 + then show ?case by (simp only: Min_def) +next case (3 bin bit) show ?case proof (cases y rule: bin_exhaust) @@ -25,13 +31,10 @@ with 1 show ?thesis by simp (simp add: Bit_def) qed -next - case 2 - then show ?case by (simp only: Min_def) -qed simp +qed lemma OR_lower [simp]: - fixes x :: int and y :: int + fixes x y :: int assumes "0 \ x" "0 \ y" shows "0 \ x OR y" using assms @@ -51,7 +54,7 @@ qed simp_all lemma XOR_lower [simp]: - fixes x :: int and y :: int + fixes x y :: int assumes "0 \ x" "0 \ y" shows "0 \ x XOR y" using assms @@ -74,7 +77,7 @@ qed simp lemma AND_upper1 [simp]: - fixes x :: int and y :: int + fixes x y :: int assumes "0 \ x" shows "x AND y \ x" using assms @@ -98,11 +101,17 @@ lemmas AND_upper1'' [simp] = order_le_less_trans [OF AND_upper1] lemma AND_upper2 [simp]: - fixes x :: int and y :: int + fixes x y :: int assumes "0 \ y" shows "x AND y \ y" using assms proof (induct y arbitrary: x rule: bin_induct) + case 1 + then show ?case by simp +next + case 2 + then show ?case by (simp only: Min_def) +next case (3 bin bit) show ?case proof (cases x rule: bin_exhaust) @@ -113,16 +122,13 @@ with 1 show ?thesis by simp (simp add: Bit_def) qed -next - case 2 - then show ?case by (simp only: Min_def) -qed simp +qed lemmas AND_upper2' [simp] = order_trans [OF AND_upper2] lemmas AND_upper2'' [simp] = order_le_less_trans [OF AND_upper2] lemma OR_upper: - fixes x :: int and y :: int + fixes x y :: int assumes "0 \ x" "x < 2 ^ n" "y < 2 ^ n" shows "x OR y < 2 ^ n" using assms @@ -155,11 +161,17 @@ qed simp_all lemma XOR_upper: - fixes x :: int and y :: int + fixes x y :: int assumes "0 \ x" "x < 2 ^ n" "y < 2 ^ n" shows "x XOR y < 2 ^ n" using assms proof (induct x arbitrary: y n rule: bin_induct) + case 1 + then show ?case by simp +next + case 2 + then show ?case by (simp only: Min_def) +next case (3 bin bit) show ?case proof (cases y rule: bin_exhaust) @@ -185,9 +197,6 @@ by simp (simp add: Bit_def) qed qed -next - case 2 - then show ?case by (simp only: Min_def) -qed simp +qed end diff -r acb0807ddb56 -r 491fd7f0b5df src/HOL/Word/Bits_Bit.thy --- a/src/HOL/Word/Bits_Bit.thy Sun Dec 03 13:22:09 2017 +0100 +++ b/src/HOL/Word/Bits_Bit.thy Sun Dec 03 18:53:49 2017 +0100 @@ -5,7 +5,7 @@ section \Bit operations in $\cal Z_2$\ theory Bits_Bit -imports Bits "HOL-Library.Bit" + imports Bits "HOL-Library.Bit" begin instantiation bit :: bit @@ -46,21 +46,21 @@ "x XOR 1 = NOT x" "x XOR 0 = x" for x :: bit - by (cases x, auto)+ + by (cases x; auto)+ lemma bit_ops_comm: "x AND y = y AND x" "x OR y = y OR x" "x XOR y = y XOR x" for x :: bit - by (cases y, auto)+ + by (cases y; auto)+ lemma bit_ops_same [simp]: "x AND x = x" "x OR x = x" "x XOR x = 0" for x :: bit - by (cases x, auto)+ + by (cases x; auto)+ lemma bit_not_not [simp]: "NOT (NOT x) = x" for x :: bit diff -r acb0807ddb56 -r 491fd7f0b5df src/HOL/Word/Bits_Int.thy --- a/src/HOL/Word/Bits_Int.thy Sun Dec 03 13:22:09 2017 +0100 +++ b/src/HOL/Word/Bits_Int.thy Sun Dec 03 18:53:49 2017 +0100 @@ -32,11 +32,9 @@ declare bitAND_int.simps [simp del] -definition int_or_def: - "bitOR = (\x y::int. NOT (NOT x AND NOT y))" +definition int_or_def: "bitOR = (\x y::int. NOT (NOT x AND NOT y))" -definition int_xor_def: - "bitXOR = (\x y::int. (x AND NOT y) OR (NOT x AND y))" +definition int_xor_def: "bitXOR = (\x y::int. (x AND NOT y) OR (NOT x AND y))" instance .. @@ -44,9 +42,8 @@ subsubsection \Basic simplification rules\ -lemma int_not_BIT [simp]: - "NOT (w BIT b) = (NOT w) BIT (\ b)" - unfolding int_not_def Bit_def by (cases b, simp_all) +lemma int_not_BIT [simp]: "NOT (w BIT b) = (NOT w) BIT (\ b)" + by (cases b) (simp_all add: int_not_def Bit_def) lemma int_not_simps [simp]: "NOT (0::int) = -1" @@ -57,160 +54,172 @@ "NOT (- numeral (Num.Bit1 w)::int) = numeral (Num.Bit0 w)" unfolding int_not_def by simp_all -lemma int_not_not [simp]: "NOT (NOT (x::int)) = x" +lemma int_not_not [simp]: "NOT (NOT x) = x" + for x :: int unfolding int_not_def by simp -lemma int_and_0 [simp]: "(0::int) AND x = 0" +lemma int_and_0 [simp]: "0 AND x = 0" + for x :: int by (simp add: bitAND_int.simps) -lemma int_and_m1 [simp]: "(-1::int) AND x = x" +lemma int_and_m1 [simp]: "-1 AND x = x" + for x :: int by (simp add: bitAND_int.simps) -lemma int_and_Bits [simp]: - "(x BIT b) AND (y BIT c) = (x AND y) BIT (b \ c)" - by (subst bitAND_int.simps, simp add: Bit_eq_0_iff Bit_eq_m1_iff) +lemma int_and_Bits [simp]: "(x BIT b) AND (y BIT c) = (x AND y) BIT (b \ c)" + by (subst bitAND_int.simps) (simp add: Bit_eq_0_iff Bit_eq_m1_iff) -lemma int_or_zero [simp]: "(0::int) OR x = x" - unfolding int_or_def by simp +lemma int_or_zero [simp]: "0 OR x = x" + for x :: int + by (simp add: int_or_def) -lemma int_or_minus1 [simp]: "(-1::int) OR x = -1" - unfolding int_or_def by simp +lemma int_or_minus1 [simp]: "-1 OR x = -1" + for x :: int + by (simp add: int_or_def) -lemma int_or_Bits [simp]: - "(x BIT b) OR (y BIT c) = (x OR y) BIT (b \ c)" - unfolding int_or_def by simp +lemma int_or_Bits [simp]: "(x BIT b) OR (y BIT c) = (x OR y) BIT (b \ c)" + by (simp add: int_or_def) -lemma int_xor_zero [simp]: "(0::int) XOR x = x" - unfolding int_xor_def by simp +lemma int_xor_zero [simp]: "0 XOR x = x" + for x :: int + by (simp add: int_xor_def) -lemma int_xor_Bits [simp]: - "(x BIT b) XOR (y BIT c) = (x XOR y) BIT ((b \ c) \ \ (b \ c))" +lemma int_xor_Bits [simp]: "(x BIT b) XOR (y BIT c) = (x XOR y) BIT ((b \ c) \ \ (b \ c))" unfolding int_xor_def by auto + subsubsection \Binary destructors\ lemma bin_rest_NOT [simp]: "bin_rest (NOT x) = NOT (bin_rest x)" - by (cases x rule: bin_exhaust, simp) + by (cases x rule: bin_exhaust) simp lemma bin_last_NOT [simp]: "bin_last (NOT x) \ \ bin_last x" - by (cases x rule: bin_exhaust, simp) + by (cases x rule: bin_exhaust) simp lemma bin_rest_AND [simp]: "bin_rest (x AND y) = bin_rest x AND bin_rest y" - by (cases x rule: bin_exhaust, cases y rule: bin_exhaust, simp) + by (cases x rule: bin_exhaust, cases y rule: bin_exhaust) simp lemma bin_last_AND [simp]: "bin_last (x AND y) \ bin_last x \ bin_last y" - by (cases x rule: bin_exhaust, cases y rule: bin_exhaust, simp) + by (cases x rule: bin_exhaust, cases y rule: bin_exhaust) simp lemma bin_rest_OR [simp]: "bin_rest (x OR y) = bin_rest x OR bin_rest y" - by (cases x rule: bin_exhaust, cases y rule: bin_exhaust, simp) + by (cases x rule: bin_exhaust, cases y rule: bin_exhaust) simp lemma bin_last_OR [simp]: "bin_last (x OR y) \ bin_last x \ bin_last y" - by (cases x rule: bin_exhaust, cases y rule: bin_exhaust, simp) + by (cases x rule: bin_exhaust, cases y rule: bin_exhaust) simp lemma bin_rest_XOR [simp]: "bin_rest (x XOR y) = bin_rest x XOR bin_rest y" - by (cases x rule: bin_exhaust, cases y rule: bin_exhaust, simp) + by (cases x rule: bin_exhaust, cases y rule: bin_exhaust) simp -lemma bin_last_XOR [simp]: "bin_last (x XOR y) \ (bin_last x \ bin_last y) \ \ (bin_last x \ bin_last y)" - by (cases x rule: bin_exhaust, cases y rule: bin_exhaust, simp) +lemma bin_last_XOR [simp]: + "bin_last (x XOR y) \ (bin_last x \ bin_last y) \ \ (bin_last x \ bin_last y)" + by (cases x rule: bin_exhaust, cases y rule: bin_exhaust) simp lemma bin_nth_ops: - "!!x y. bin_nth (x AND y) n = (bin_nth x n & bin_nth y n)" - "!!x y. bin_nth (x OR y) n = (bin_nth x n | bin_nth y n)" - "!!x y. bin_nth (x XOR y) n = (bin_nth x n ~= bin_nth y n)" - "!!x. bin_nth (NOT x) n = (~ bin_nth x n)" + "\x y. bin_nth (x AND y) n \ bin_nth x n \ bin_nth y n" + "\x y. bin_nth (x OR y) n \ bin_nth x n \ bin_nth y n" + "\x y. bin_nth (x XOR y) n \ bin_nth x n \ bin_nth y n" + "\x. bin_nth (NOT x) n \ \ bin_nth x n" by (induct n) auto + subsubsection \Derived properties\ -lemma int_xor_minus1 [simp]: "(-1::int) XOR x = NOT x" +lemma int_xor_minus1 [simp]: "-1 XOR x = NOT x" + for x :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_xor_extra_simps [simp]: - "w XOR (0::int) = w" - "w XOR (-1::int) = NOT w" + "w XOR 0 = w" + "w XOR -1 = NOT w" + for w :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_or_extra_simps [simp]: - "w OR (0::int) = w" - "w OR (-1::int) = -1" + "w OR 0 = w" + "w OR -1 = -1" + for w :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_and_extra_simps [simp]: - "w AND (0::int) = 0" - "w AND (-1::int) = w" + "w AND 0 = 0" + "w AND -1 = w" + for w :: int by (auto simp add: bin_eq_iff bin_nth_ops) -(* commutativity of the above *) +text \Commutativity of the above.\ lemma bin_ops_comm: - shows - int_and_comm: "!!y::int. x AND y = y AND x" and - int_or_comm: "!!y::int. x OR y = y OR x" and - int_xor_comm: "!!y::int. x XOR y = y XOR x" + fixes x y :: int + shows int_and_comm: "x AND y = y AND x" + and int_or_comm: "x OR y = y OR x" + and int_xor_comm: "x XOR y = y XOR x" by (auto simp add: bin_eq_iff bin_nth_ops) lemma bin_ops_same [simp]: - "(x::int) AND x = x" - "(x::int) OR x = x" - "(x::int) XOR x = 0" + "x AND x = x" + "x OR x = x" + "x XOR x = 0" + for x :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemmas bin_log_esimps = int_and_extra_simps int_or_extra_simps int_xor_extra_simps int_and_0 int_and_m1 int_or_zero int_or_minus1 int_xor_zero int_xor_minus1 -(* basic properties of logical (bit-wise) operations *) + +subsubsection \Basic properties of logical (bit-wise) operations\ -lemma bbw_ao_absorb: - "!!y::int. x AND (y OR x) = x & x OR (y AND x) = x" +lemma bbw_ao_absorb: "x AND (y OR x) = x \ x OR (y AND x) = x" + for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_ao_absorbs_other: - "x AND (x OR y) = x \ (y AND x) OR x = (x::int)" - "(y OR x) AND x = x \ x OR (x AND y) = (x::int)" - "(x OR y) AND x = x \ (x AND y) OR x = (x::int)" + "x AND (x OR y) = x \ (y AND x) OR x = x" + "(y OR x) AND x = x \ x OR (x AND y) = x" + "(x OR y) AND x = x \ (x AND y) OR x = x" + for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemmas bbw_ao_absorbs [simp] = bbw_ao_absorb bbw_ao_absorbs_other -lemma int_xor_not: - "!!y::int. (NOT x) XOR y = NOT (x XOR y) & - x XOR (NOT y) = NOT (x XOR y)" +lemma int_xor_not: "(NOT x) XOR y = NOT (x XOR y) \ x XOR (NOT y) = NOT (x XOR y)" + for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) -lemma int_and_assoc: - "(x AND y) AND (z::int) = x AND (y AND z)" +lemma int_and_assoc: "(x AND y) AND z = x AND (y AND z)" + for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) -lemma int_or_assoc: - "(x OR y) OR (z::int) = x OR (y OR z)" +lemma int_or_assoc: "(x OR y) OR z = x OR (y OR z)" + for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) -lemma int_xor_assoc: - "(x XOR y) XOR (z::int) = x XOR (y XOR z)" +lemma int_xor_assoc: "(x XOR y) XOR z = x XOR (y XOR z)" + for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemmas bbw_assocs = int_and_assoc int_or_assoc int_xor_assoc (* BH: Why are these declared as simp rules??? *) lemma bbw_lcs [simp]: - "(y::int) AND (x AND z) = x AND (y AND z)" - "(y::int) OR (x OR z) = x OR (y OR z)" - "(y::int) XOR (x XOR z) = x XOR (y XOR z)" + "y AND (x AND z) = x AND (y AND z)" + "y OR (x OR z) = x OR (y OR z)" + "y XOR (x XOR z) = x XOR (y XOR z)" + for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_not_dist: - "!!y::int. NOT (x OR y) = (NOT x) AND (NOT y)" - "!!y::int. NOT (x AND y) = (NOT x) OR (NOT y)" + "NOT (x OR y) = (NOT x) AND (NOT y)" + "NOT (x AND y) = (NOT x) OR (NOT y)" + for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) -lemma bbw_oa_dist: - "!!y z::int. (x AND y) OR z = - (x OR z) AND (y OR z)" +lemma bbw_oa_dist: "(x AND y) OR z = (x OR z) AND (y OR z)" + for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) -lemma bbw_ao_dist: - "!!y z::int. (x OR y) AND z = - (x AND z) OR (y AND z)" +lemma bbw_ao_dist: "(x OR y) AND z = (x AND z) OR (y AND z)" + for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) (* @@ -218,10 +227,10 @@ declare bin_ops_comm [simp] bbw_assocs [simp] *) + subsubsection \Simplification with numerals\ -text \Cases for \0\ and \-1\ are already covered by - other simp rules.\ +text \Cases for \0\ and \-1\ are already covered by other simp rules.\ lemma bin_rl_eqI: "\bin_rest x = bin_rest y; bin_last x = bin_last y\ \ x = y" by (metis (mono_tags) BIT_eq_iff bin_ex_rl bin_last_BIT bin_rest_BIT) @@ -234,8 +243,8 @@ "bin_last (- numeral (Num.BitM w))" by (simp only: BIT_bin_simps [symmetric] bin_last_BIT) -text \FIXME: The rule sets below are very large (24 rules for each - operator). Is there a simpler way to do this?\ +(* FIXME: The rule sets below are very large (24 rules for each + operator). Is there a simpler way to do this? *) lemma int_and_numerals [simp]: "numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (numeral x AND numeral y) BIT False" @@ -262,7 +271,7 @@ "numeral (Num.Bit1 x) AND (1::int) = 1" "- numeral (Num.Bit0 x) AND (1::int) = 0" "- numeral (Num.Bit1 x) AND (1::int) = 1" - by (rule bin_rl_eqI, simp, simp)+ + by (rule bin_rl_eqI; simp)+ lemma int_or_numerals [simp]: "numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (numeral x OR numeral y) BIT False" @@ -289,7 +298,7 @@ "numeral (Num.Bit1 x) OR (1::int) = numeral (Num.Bit1 x)" "- numeral (Num.Bit0 x) OR (1::int) = - numeral (Num.BitM x)" "- numeral (Num.Bit1 x) OR (1::int) = - numeral (Num.Bit1 x)" - by (rule bin_rl_eqI, simp, simp)+ + by (rule bin_rl_eqI; simp)+ lemma int_xor_numerals [simp]: "numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (numeral x XOR numeral y) BIT False" @@ -316,12 +325,12 @@ "numeral (Num.Bit1 x) XOR (1::int) = numeral (Num.Bit0 x)" "- numeral (Num.Bit0 x) XOR (1::int) = - numeral (Num.BitM x)" "- numeral (Num.Bit1 x) XOR (1::int) = - numeral (Num.Bit0 (x + Num.One))" - by (rule bin_rl_eqI, simp, simp)+ + by (rule bin_rl_eqI; simp)+ + subsubsection \Interactions with arithmetic\ -lemma plus_and_or [rule_format]: - "ALL y::int. (x AND y) + (x OR y) = x + y" +lemma plus_and_or [rule_format]: "\y::int. (x AND y) + (x OR y) = x + y" apply (induct x rule: bin_induct) apply clarsimp apply clarsimp @@ -334,8 +343,8 @@ apply simp done -lemma le_int_or: - "bin_sign (y::int) = 0 ==> x <= x OR y" +lemma le_int_or: "bin_sign y = 0 \ x \ x OR y" + for x y :: int apply (induct y arbitrary: x rule: bin_induct) apply clarsimp apply clarsimp @@ -348,8 +357,7 @@ lemmas int_and_le = xtrans(3) [OF bbw_ao_absorbs (2) [THEN conjunct2, symmetric] le_int_or] -(* interaction between bit-wise and arithmetic *) -(* good example of bin_induction *) +text \Interaction between bit-wise and arithmetic: good example of \bin_induction\.\ lemma bin_add_not: "x + NOT x = (-1::int)" apply (induct x rule: bin_induct) apply clarsimp @@ -357,91 +365,77 @@ apply (case_tac bit, auto) done + subsubsection \Truncating results of bit-wise operations\ lemma bin_trunc_ao: - "!!x y. (bintrunc n x) AND (bintrunc n y) = bintrunc n (x AND y)" - "!!x y. (bintrunc n x) OR (bintrunc n y) = bintrunc n (x OR y)" + "bintrunc n x AND bintrunc n y = bintrunc n (x AND y)" + "bintrunc n x OR bintrunc n y = bintrunc n (x OR y)" by (auto simp add: bin_eq_iff bin_nth_ops nth_bintr) -lemma bin_trunc_xor: - "!!x y. bintrunc n (bintrunc n x XOR bintrunc n y) = - bintrunc n (x XOR y)" +lemma bin_trunc_xor: "bintrunc n (bintrunc n x XOR bintrunc n y) = bintrunc n (x XOR y)" by (auto simp add: bin_eq_iff bin_nth_ops nth_bintr) -lemma bin_trunc_not: - "!!x. bintrunc n (NOT (bintrunc n x)) = bintrunc n (NOT x)" +lemma bin_trunc_not: "bintrunc n (NOT (bintrunc n x)) = bintrunc n (NOT x)" by (auto simp add: bin_eq_iff bin_nth_ops nth_bintr) -(* want theorems of the form of bin_trunc_xor *) -lemma bintr_bintr_i: - "x = bintrunc n y ==> bintrunc n x = bintrunc n y" +text \Want theorems of the form of \bin_trunc_xor\.\ +lemma bintr_bintr_i: "x = bintrunc n y \ bintrunc n x = bintrunc n y" by auto lemmas bin_trunc_and = bin_trunc_ao(1) [THEN bintr_bintr_i] lemmas bin_trunc_or = bin_trunc_ao(2) [THEN bintr_bintr_i] + subsection \Setting and clearing bits\ -(** nth bit, set/clear **) +text \nth bit, set/clear\ -primrec - bin_sc :: "nat => bool => int => int" -where - Z: "bin_sc 0 b w = bin_rest w BIT b" +primrec bin_sc :: "nat \ bool \ int \ int" + where + Z: "bin_sc 0 b w = bin_rest w BIT b" | Suc: "bin_sc (Suc n) b w = bin_sc n b (bin_rest w) BIT bin_last w" -lemma bin_nth_sc [simp]: - "bin_nth (bin_sc n b w) n \ b" +lemma bin_nth_sc [simp]: "bin_nth (bin_sc n b w) n \ b" by (induct n arbitrary: w) auto -lemma bin_sc_sc_same [simp]: - "bin_sc n c (bin_sc n b w) = bin_sc n c w" +lemma bin_sc_sc_same [simp]: "bin_sc n c (bin_sc n b w) = bin_sc n c w" by (induct n arbitrary: w) auto -lemma bin_sc_sc_diff: - "m ~= n ==> - bin_sc m c (bin_sc n b w) = bin_sc n b (bin_sc m c w)" +lemma bin_sc_sc_diff: "m \ n \ bin_sc m c (bin_sc n b w) = bin_sc n b (bin_sc m c w)" apply (induct n arbitrary: w m) apply (case_tac [!] m) apply auto done -lemma bin_nth_sc_gen: - "bin_nth (bin_sc n b w) m = (if m = n then b else bin_nth w m)" +lemma bin_nth_sc_gen: "bin_nth (bin_sc n b w) m = (if m = n then b else bin_nth w m)" by (induct n arbitrary: w m) (case_tac [!] m, auto) -lemma bin_sc_nth [simp]: - "(bin_sc n (bin_nth w n) w) = w" +lemma bin_sc_nth [simp]: "bin_sc n (bin_nth w n) w = w" by (induct n arbitrary: w) auto -lemma bin_sign_sc [simp]: - "bin_sign (bin_sc n b w) = bin_sign w" +lemma bin_sign_sc [simp]: "bin_sign (bin_sc n b w) = bin_sign w" by (induct n arbitrary: w) auto -lemma bin_sc_bintr [simp]: - "bintrunc m (bin_sc n x (bintrunc m (w))) = bintrunc m (bin_sc n x w)" +lemma bin_sc_bintr [simp]: "bintrunc m (bin_sc n x (bintrunc m (w))) = bintrunc m (bin_sc n x w)" apply (induct n arbitrary: w m) apply (case_tac [!] w rule: bin_exhaust) apply (case_tac [!] m, auto) done -lemma bin_clr_le: - "bin_sc n False w <= w" +lemma bin_clr_le: "bin_sc n False w \ w" apply (induct n arbitrary: w) apply (case_tac [!] w rule: bin_exhaust) apply (auto simp: le_Bits) done -lemma bin_set_ge: - "bin_sc n True w >= w" +lemma bin_set_ge: "bin_sc n True w \ w" apply (induct n arbitrary: w) apply (case_tac [!] w rule: bin_exhaust) apply (auto simp: le_Bits) done -lemma bintr_bin_clr_le: - "bintrunc n (bin_sc m False w) <= bintrunc n w" +lemma bintr_bin_clr_le: "bintrunc n (bin_sc m False w) \ bintrunc n w" apply (induct n arbitrary: w m) apply simp apply (case_tac w rule: bin_exhaust) @@ -449,8 +443,7 @@ apply (auto simp: le_Bits) done -lemma bintr_bin_set_ge: - "bintrunc n (bin_sc m True w) >= bintrunc n w" +lemma bintr_bin_set_ge: "bintrunc n (bin_sc m True w) \ bintrunc n w" apply (induct n arbitrary: w m) apply simp apply (case_tac w rule: bin_exhaust) @@ -466,8 +459,7 @@ lemmas bin_sc_simps = bin_sc.Z bin_sc.Suc bin_sc_TM bin_sc_FP -lemma bin_sc_minus: - "0 < n ==> bin_sc (Suc (n - 1)) b w = bin_sc n b w" +lemma bin_sc_minus: "0 < n \ bin_sc (Suc (n - 1)) b w = bin_sc n b w" by auto lemmas bin_sc_Suc_minus = @@ -482,40 +474,35 @@ subsection \Splitting and concatenation\ definition bin_rcat :: "nat \ int list \ int" -where - "bin_rcat n = foldl (\u v. bin_cat u n v) 0" + where "bin_rcat n = foldl (\u v. bin_cat u n v) 0" fun bin_rsplit_aux :: "nat \ nat \ int \ int list \ int list" -where - "bin_rsplit_aux n m c bs = - (if m = 0 | n = 0 then bs else + where "bin_rsplit_aux n m c bs = + (if m = 0 \ n = 0 then bs + else let (a, b) = bin_split n c in bin_rsplit_aux n (m - n) a (b # bs))" definition bin_rsplit :: "nat \ nat \ int \ int list" -where - "bin_rsplit n w = bin_rsplit_aux n (fst w) (snd w) []" + where "bin_rsplit n w = bin_rsplit_aux n (fst w) (snd w) []" fun bin_rsplitl_aux :: "nat \ nat \ int \ int list \ int list" -where - "bin_rsplitl_aux n m c bs = - (if m = 0 | n = 0 then bs else + where "bin_rsplitl_aux n m c bs = + (if m = 0 \ n = 0 then bs + else let (a, b) = bin_split (min m n) c in bin_rsplitl_aux n (m - n) a (b # bs))" definition bin_rsplitl :: "nat \ nat \ int \ int list" -where - "bin_rsplitl n w = bin_rsplitl_aux n (fst w) (snd w) []" + where "bin_rsplitl n w = bin_rsplitl_aux n (fst w) (snd w) []" declare bin_rsplit_aux.simps [simp del] declare bin_rsplitl_aux.simps [simp del] -lemma bin_sign_cat: - "bin_sign (bin_cat x n y) = bin_sign x" +lemma bin_sign_cat: "bin_sign (bin_cat x n y) = bin_sign x" by (induct n arbitrary: y) auto -lemma bin_cat_Suc_Bit: - "bin_cat w (Suc n) (v BIT b) = bin_cat w n v BIT b" +lemma bin_cat_Suc_Bit: "bin_cat w (Suc n) (v BIT b) = bin_cat w n v BIT b" by auto lemma bin_nth_cat: @@ -527,9 +514,9 @@ done lemma bin_nth_split: - "bin_split n c = (a, b) ==> - (ALL k. bin_nth a k = bin_nth c (n + k)) & - (ALL k. bin_nth b k = (k < n & bin_nth c k))" + "bin_split n c = (a, b) \ + (\k. bin_nth a k = bin_nth c (n + k)) \ + (\k. bin_nth b k = (k < n \ bin_nth c k))" apply (induct n arbitrary: b c) apply clarsimp apply (clarsimp simp: Let_def split: prod.split_asm) @@ -537,45 +524,38 @@ apply auto done -lemma bin_cat_assoc: - "bin_cat (bin_cat x m y) n z = bin_cat x (m + n) (bin_cat y n z)" +lemma bin_cat_assoc: "bin_cat (bin_cat x m y) n z = bin_cat x (m + n) (bin_cat y n z)" by (induct n arbitrary: z) auto -lemma bin_cat_assoc_sym: - "bin_cat x m (bin_cat y n z) = bin_cat (bin_cat x (m - n) y) (min m n) z" - apply (induct n arbitrary: z m, clarsimp) +lemma bin_cat_assoc_sym: "bin_cat x m (bin_cat y n z) = bin_cat (bin_cat x (m - n) y) (min m n) z" + apply (induct n arbitrary: z m) + apply clarsimp apply (case_tac m, auto) done lemma bin_cat_zero [simp]: "bin_cat 0 n w = bintrunc n w" by (induct n arbitrary: w) auto -lemma bintr_cat1: - "bintrunc (k + n) (bin_cat a n b) = bin_cat (bintrunc k a) n b" +lemma bintr_cat1: "bintrunc (k + n) (bin_cat a n b) = bin_cat (bintrunc k a) n b" by (induct n arbitrary: b) auto lemma bintr_cat: "bintrunc m (bin_cat a n b) = bin_cat (bintrunc (m - n) a) n (bintrunc (min m n) b)" by (rule bin_eqI) (auto simp: bin_nth_cat nth_bintr) -lemma bintr_cat_same [simp]: - "bintrunc n (bin_cat a n b) = bintrunc n b" +lemma bintr_cat_same [simp]: "bintrunc n (bin_cat a n b) = bintrunc n b" by (auto simp add : bintr_cat) -lemma cat_bintr [simp]: - "bin_cat a n (bintrunc n b) = bin_cat a n b" +lemma cat_bintr [simp]: "bin_cat a n (bintrunc n b) = bin_cat a n b" by (induct n arbitrary: b) auto -lemma split_bintrunc: - "bin_split n c = (a, b) ==> b = bintrunc n c" +lemma split_bintrunc: "bin_split n c = (a, b) \ b = bintrunc n c" by (induct n arbitrary: b c) (auto simp: Let_def split: prod.split_asm) -lemma bin_cat_split: - "bin_split n w = (u, v) ==> w = bin_cat u n v" +lemma bin_cat_split: "bin_split n w = (u, v) \ w = bin_cat u n v" by (induct n arbitrary: v w) (auto simp: Let_def split: prod.split_asm) -lemma bin_split_cat: - "bin_split n (bin_cat v n w) = (v, bintrunc n w)" +lemma bin_split_cat: "bin_split n (bin_cat v n w) = (v, bintrunc n w)" by (induct n arbitrary: w) auto lemma bin_split_zero [simp]: "bin_split n 0 = (0, 0)" @@ -586,7 +566,7 @@ by (induct n) auto lemma bin_split_trunc: - "bin_split (min m n) c = (a, b) ==> + "bin_split (min m n) c = (a, b) \ bin_split n (bintrunc m c) = (bintrunc (m - n) a, b)" apply (induct n arbitrary: m b c, clarsimp) apply (simp add: bin_rest_trunc Let_def split: prod.split_asm) @@ -595,7 +575,7 @@ done lemma bin_split_trunc1: - "bin_split n c = (a, b) ==> + "bin_split n c = (a, b) \ bin_split n (bintrunc m c) = (bintrunc (m - n) a, bintrunc m b)" apply (induct n arbitrary: m b c, clarsimp) apply (simp add: bin_rest_trunc Let_def split: prod.split_asm) @@ -603,25 +583,25 @@ apply (auto simp: Let_def split: prod.split_asm) done -lemma bin_cat_num: - "bin_cat a n b = a * 2 ^ n + bintrunc n b" - apply (induct n arbitrary: b, clarsimp) +lemma bin_cat_num: "bin_cat a n b = a * 2 ^ n + bintrunc n b" + apply (induct n arbitrary: b) + apply clarsimp apply (simp add: Bit_def) done -lemma bin_split_num: - "bin_split n b = (b div 2 ^ n, b mod 2 ^ n)" - apply (induct n arbitrary: b, simp) +lemma bin_split_num: "bin_split n b = (b div 2 ^ n, b mod 2 ^ n)" + apply (induct n arbitrary: b) + apply simp apply (simp add: bin_rest_def zdiv_zmult2_eq) apply (case_tac b rule: bin_exhaust) apply simp apply (simp add: Bit_def mod_mult_mult1 p1mod22k) done + subsection \Miscellaneous lemmas\ -lemma nth_2p_bin: - "bin_nth (2 ^ n) m = (m = n)" +lemma nth_2p_bin: "bin_nth (2 ^ n) m = (m = n)" apply (induct n arbitrary: m) apply clarsimp apply safe @@ -629,18 +609,14 @@ apply (auto simp: Bit_B0_2t [symmetric]) done -(* for use when simplifying with bin_nth_Bit *) - -lemma ex_eq_or: - "(EX m. n = Suc m & (m = k | P m)) = (n = Suc k | (EX m. n = Suc m & P m))" +(*for use when simplifying with bin_nth_Bit*) +lemma ex_eq_or: "(\m. n = Suc m \ (m = k \ P m)) \ n = Suc k \ (\m. n = Suc m \ P m)" by auto lemma power_BIT: "2 ^ (Suc n) - 1 = (2 ^ n - 1) BIT True" - unfolding Bit_B1 - by (induct n) simp_all + by (induct n) (simp_all add: Bit_B1) -lemma mod_BIT: - "bin BIT bit mod 2 ^ Suc n = (bin mod 2 ^ n) BIT bit" +lemma mod_BIT: "bin BIT bit mod 2 ^ Suc n = (bin mod 2 ^ n) BIT bit" proof - have "2 * (bin mod 2 ^ n) + 1 = (2 * bin mod 2 ^ Suc n) + 1" by (simp add: mod_mult_mult1) @@ -652,9 +628,8 @@ by (auto simp add: Bit_def) qed -lemma AND_mod: - fixes x :: int - shows "x AND 2 ^ n - 1 = x mod 2 ^ n" +lemma AND_mod: "x AND 2 ^ n - 1 = x mod 2 ^ n" + for x :: int proof (induct x arbitrary: n rule: bin_induct) case 1 then show ?case diff -r acb0807ddb56 -r 491fd7f0b5df src/HOL/Word/Bool_List_Representation.thy --- a/src/HOL/Word/Bool_List_Representation.thy Sun Dec 03 13:22:09 2017 +0100 +++ b/src/HOL/Word/Bool_List_Representation.thy Sun Dec 03 18:53:49 2017 +0100 @@ -117,7 +117,7 @@ bin_to_bl_aux n (numeral (Num.Bit1 w)) bl = bin_to_bl_aux (n - 1) (numeral w) (True # bl)" by (cases n) auto -text \Link between bin and bool list.\ +text \Link between \bin\ and \bool list\.\ lemma bl_to_bin_aux_append: "bl_to_bin_aux (bs @ cs) w = bl_to_bin_aux cs (bl_to_bin_aux bs w)" by (induct bs arbitrary: w) auto @@ -238,7 +238,7 @@ lemma bin_nth_of_bl_aux: "bin_nth (bl_to_bin_aux bl w) n = - (n < size bl \ rev bl ! n | n \ length bl \ bin_nth w (n - size bl))" + (n < size bl \ rev bl ! n \ n \ length bl \ bin_nth w (n - size bl))" apply (induct bl arbitrary: w) apply clarsimp apply clarsimp @@ -298,8 +298,9 @@ case Nil then show ?case by simp next - case (Cons b bs) with bl_to_bin_lt2p_aux[where w=1] - show ?case unfolding bl_to_bin_def by simp + case (Cons b bs) + with bl_to_bin_lt2p_aux[where w=1] show ?case + by (simp add: bl_to_bin_def) qed lemma bl_to_bin_lt2p: "bl_to_bin bs < 2 ^ length bs" @@ -509,11 +510,11 @@ lemma length_takefill [simp]: "length (takefill fill n l) = n" by (simp add: takefill_alt) -lemma take_takefill': "\w n. n = k + m \ take k (takefill fill n w) = takefill fill k w" - by (induct k) (auto split: list.split) +lemma take_takefill': "n = k + m \ take k (takefill fill n w) = takefill fill k w" + by (induct k arbitrary: w n) (auto split: list.split) -lemma drop_takefill: "\w. drop k (takefill fill (m + k) w) = takefill fill m (drop k w)" - by (induct k) (auto split: list.split) +lemma drop_takefill: "drop k (takefill fill (m + k) w) = takefill fill m (drop k w)" + by (induct k arbitrary: w) (auto split: list.split) lemma takefill_le [simp]: "m \ n \ takefill x m (takefill x n l) = takefill x m l" by (auto simp: le_iff_add takefill_le') @@ -715,7 +716,7 @@ done lemma rbl_add_take2: - "length blb >= length bla ==> rbl_add bla (take (length bla) blb) = rbl_add bla blb" + "length blb \ length bla \ rbl_add bla (take (length bla) blb) = rbl_add bla blb" apply (induct bla arbitrary: blb) apply simp apply clarsimp @@ -1023,7 +1024,8 @@ with \length bs = length cs\ show ?thesis by simp next case False - from "1.hyps" \m \ 0\ \n \ 0\ have hyp: "\v bs. length bs = Suc (length cs) \ + from "1.hyps" \m \ 0\ \n \ 0\ + have hyp: "\v bs. length bs = Suc (length cs) \ length (bin_rsplit_aux n (m - n) v bs) = length (bin_rsplit_aux n (m - n) (fst (bin_split n w)) (snd (bin_split n w) # cs))" by auto diff -r acb0807ddb56 -r 491fd7f0b5df src/HOL/Word/Misc_Typedef.thy --- a/src/HOL/Word/Misc_Typedef.thy Sun Dec 03 13:22:09 2017 +0100 +++ b/src/HOL/Word/Misc_Typedef.thy Sun Dec 03 18:53:49 2017 +0100 @@ -7,19 +7,17 @@ section \Type Definition Theorems\ theory Misc_Typedef -imports Main + imports Main begin section "More lemmas about normal type definitions" -lemma - tdD1: "type_definition Rep Abs A \ \x. Rep x \ A" and - tdD2: "type_definition Rep Abs A \ \x. Abs (Rep x) = x" and - tdD3: "type_definition Rep Abs A \ \y. y \ A \ Rep (Abs y) = y" +lemma tdD1: "type_definition Rep Abs A \ \x. Rep x \ A" + and tdD2: "type_definition Rep Abs A \ \x. Abs (Rep x) = x" + and tdD3: "type_definition Rep Abs A \ \y. y \ A \ Rep (Abs y) = y" by (auto simp: type_definition_def) -lemma td_nat_int: - "type_definition int nat (Collect (op <= 0))" +lemma td_nat_int: "type_definition int nat (Collect (op \ 0))" unfolding type_definition_def by auto context type_definition @@ -27,39 +25,33 @@ declare Rep [iff] Rep_inverse [simp] Rep_inject [simp] -lemma Abs_eqD: "Abs x = Abs y ==> x \ A ==> y \ A ==> x = y" +lemma Abs_eqD: "Abs x = Abs y \ x \ A \ y \ A \ x = y" by (simp add: Abs_inject) -lemma Abs_inverse': - "r : A ==> Abs r = a ==> Rep a = r" +lemma Abs_inverse': "r \ A \ Abs r = a \ Rep a = r" by (safe elim!: Abs_inverse) -lemma Rep_comp_inverse: - "Rep \ f = g ==> Abs \ g = f" +lemma Rep_comp_inverse: "Rep \ f = g \ Abs \ g = f" using Rep_inverse by auto -lemma Rep_eqD [elim!]: "Rep x = Rep y ==> x = y" +lemma Rep_eqD [elim!]: "Rep x = Rep y \ x = y" by simp -lemma Rep_inverse': "Rep a = r ==> Abs r = a" +lemma Rep_inverse': "Rep a = r \ Abs r = a" by (safe intro!: Rep_inverse) -lemma comp_Abs_inverse: - "f \ Abs = g ==> g \ Rep = f" +lemma comp_Abs_inverse: "f \ Abs = g \ g \ Rep = f" using Rep_inverse by auto -lemma set_Rep: - "A = range Rep" +lemma set_Rep: "A = range Rep" proof (rule set_eqI) - fix x - show "(x \ A) = (x \ range Rep)" + show "x \ A \ x \ range Rep" for x by (auto dest: Abs_inverse [of x, symmetric]) qed lemma set_Rep_Abs: "A = range (Rep \ Abs)" proof (rule set_eqI) - fix x - show "(x \ A) = (x \ range (Rep \ Abs))" + show "x \ A \ x \ range (Rep \ Abs)" for x by (auto dest: Abs_inverse [of x, symmetric]) qed @@ -72,21 +64,18 @@ lemmas td_thm = type_definition_axioms -lemma fns1: - "Rep \ fa = fr \ Rep | fa \ Abs = Abs \ fr ==> Abs \ fr \ Rep = fa" +lemma fns1: "Rep \ fa = fr \ Rep \ fa \ Abs = Abs \ fr \ Abs \ fr \ Rep = fa" by (auto dest: Rep_comp_inverse elim: comp_Abs_inverse simp: o_assoc) lemmas fns1a = disjI1 [THEN fns1] lemmas fns1b = disjI2 [THEN fns1] -lemma fns4: - "Rep \ fa \ Abs = fr ==> - Rep \ fa = fr \ Rep & fa \ Abs = Abs \ fr" +lemma fns4: "Rep \ fa \ Abs = fr \ Rep \ fa = fr \ Rep \ fa \ Abs = Abs \ fr" by auto end -interpretation nat_int: type_definition int nat "Collect (op <= 0)" +interpretation nat_int: type_definition int nat "Collect (op \ 0)" by (rule td_nat_int) declare @@ -99,15 +88,14 @@ subsection "Extended form of type definition predicate" lemma td_conds: - "norm \ norm = norm ==> (fr \ norm = norm \ fr) = - (norm \ fr \ norm = fr \ norm & norm \ fr \ norm = norm \ fr)" + "norm \ norm = norm \ + fr \ norm = norm \ fr \ norm \ fr \ norm = fr \ norm \ norm \ fr \ norm = norm \ fr" apply safe apply (simp_all add: comp_assoc) apply (simp_all add: o_assoc) done -lemma fn_comm_power: - "fa \ tr = tr \ fr ==> fa ^^ n \ tr = tr \ fr ^^ n" +lemma fn_comm_power: "fa \ tr = tr \ fr \ fa ^^ n \ tr = tr \ fr ^^ n" apply (rule ext) apply (induct n) apply (auto dest: fun_cong) @@ -122,12 +110,10 @@ assumes eq_norm: "\x. Rep (Abs x) = norm x" begin -lemma Abs_norm [simp]: - "Abs (norm x) = Abs x" +lemma Abs_norm [simp]: "Abs (norm x) = Abs x" using eq_norm [of x] by (auto elim: Rep_inverse') -lemma td_th: - "g \ Abs = f ==> f (Rep x) = g x" +lemma td_th: "g \ Abs = f \ f (Rep x) = g x" by (drule comp_Abs_inverse [symmetric]) simp lemma eq_norm': "Rep \ Abs = norm" @@ -141,15 +127,13 @@ lemma set_iff_norm: "w : A \ w = norm w" by (auto simp: set_Rep_Abs eq_norm' eq_norm [symmetric]) -lemma inverse_norm: - "(Abs n = w) = (Rep w = norm n)" +lemma inverse_norm: "Abs n = w \ Rep w = norm n" apply (rule iffI) apply (clarsimp simp add: eq_norm) apply (simp add: eq_norm' [symmetric]) done -lemma norm_eq_iff: - "(norm x = norm y) = (Abs x = Abs y)" +lemma norm_eq_iff: "norm x = norm y \ Abs x = Abs y" by (simp add: eq_norm' [symmetric]) lemma norm_comps: @@ -160,9 +144,7 @@ lemmas norm_norm [simp] = norm_comps -lemma fns5: - "Rep \ fa \ Abs = fr ==> - fr \ norm = fr & norm \ fr = fr" +lemma fns5: "Rep \ fa \ Abs = fr \ fr \ norm = fr \ norm \ fr = fr" by (fold eq_norm') auto (* following give conditions for converses to td_fns1 @@ -174,9 +156,7 @@ takes norm-equivalent arguments to the same result, and (norm \ fr = fr) says that fr takes any argument to a normalised result *) -lemma fns2: - "Abs \ fr \ Rep = fa ==> - (norm \ fr \ norm = fr \ norm) = (Rep \ fa = fr \ Rep)" +lemma fns2: "Abs \ fr \ Rep = fa \ norm \ fr \ norm = fr \ norm \ Rep \ fa = fr \ Rep" apply (fold eq_norm') apply safe prefer 2 @@ -186,9 +166,7 @@ apply auto done -lemma fns3: - "Abs \ fr \ Rep = fa ==> - (norm \ fr \ norm = norm \ fr) = (fa \ Abs = Abs \ fr)" +lemma fns3: "Abs \ fr \ Rep = fa \ norm \ fr \ norm = norm \ fr \ fa \ Abs = Abs \ fr" apply (fold eq_norm') apply safe prefer 2 @@ -198,9 +176,7 @@ apply simp done -lemma fns: - "fr \ norm = norm \ fr ==> - (fa \ Abs = Abs \ fr) = (Rep \ fa = fr \ Rep)" +lemma fns: "fr \ norm = norm \ fr \ fa \ Abs = Abs \ fr \ Rep \ fa = fr \ Rep" apply safe apply (frule fns1b) prefer 2 @@ -212,8 +188,7 @@ apply (simp_all add: o_assoc) done -lemma range_norm: - "range (Rep \ Abs) = A" +lemma range_norm: "range (Rep \ Abs) = A" by (simp add: set_Rep_Abs) end diff -r acb0807ddb56 -r 491fd7f0b5df src/HOL/Word/WordBitwise.thy --- a/src/HOL/Word/WordBitwise.thy Sun Dec 03 13:22:09 2017 +0100 +++ b/src/HOL/Word/WordBitwise.thy Sun Dec 03 18:53:49 2017 +0100 @@ -415,9 +415,10 @@ val word_ss = simpset_of @{theory_context Word}; -fun mk_nat_clist ns = List.foldr - (uncurry (Thm.mk_binop @{cterm "Cons :: nat => _"})) - @{cterm "[] :: nat list"} ns; +fun mk_nat_clist ns = + List.foldr + (uncurry (Thm.mk_binop @{cterm "Cons :: nat \ _"})) + @{cterm "[] :: nat list"} ns; fun upt_conv ctxt ct = case Thm.term_of ct of @@ -426,8 +427,9 @@ val (i, j) = apply2 (snd o HOLogic.dest_number) (n, m); val ns = map (Numeral.mk_cnumber @{ctyp nat}) (i upto (j - 1)) |> mk_nat_clist; - val prop = Thm.mk_binop @{cterm "op = :: nat list => _"} ct ns - |> Thm.apply @{cterm Trueprop}; + val prop = + Thm.mk_binop @{cterm "op = :: nat list \ _"} ct ns + |> Thm.apply @{cterm Trueprop}; in try (fn () => Goal.prove_internal ctxt [] prop @@ -441,16 +443,19 @@ {lhss = [@{term "upt x y"}], proc = K upt_conv}; fun word_len_simproc_fn ctxt ct = - case Thm.term_of ct of - Const (@{const_name len_of}, _) $ t => (let + (case Thm.term_of ct of + Const (@{const_name len_of}, _) $ t => + (let val T = fastype_of t |> dest_Type |> snd |> the_single val n = Numeral.mk_cnumber @{ctyp nat} (Word_Lib.dest_binT T); - val prop = Thm.mk_binop @{cterm "op = :: nat => _"} ct n - |> Thm.apply @{cterm Trueprop}; - in Goal.prove_internal ctxt [] prop (K (simp_tac (put_simpset word_ss ctxt) 1)) - |> mk_meta_eq |> SOME end - handle TERM _ => NONE | TYPE _ => NONE) - | _ => NONE; + val prop = + Thm.mk_binop @{cterm "op = :: nat \ _"} ct n + |> Thm.apply @{cterm Trueprop}; + in + Goal.prove_internal ctxt [] prop (K (simp_tac (put_simpset word_ss ctxt) 1)) + |> mk_meta_eq |> SOME + end handle TERM _ => NONE | TYPE _ => NONE) + | _ => NONE); val word_len_simproc = Simplifier.make_simproc @{context} "word_len" @@ -462,21 +467,24 @@ fun nat_get_Suc_simproc_fn n_sucs ctxt ct = let val (f $ arg) = Thm.term_of ct; - val n = (case arg of @{term nat} $ n => n | n => n) + val n = + (case arg of @{term nat} $ n => n | n => n) |> HOLogic.dest_number |> snd; - val (i, j) = if n > n_sucs then (n_sucs, n - n_sucs) - else (n, 0); - val arg' = List.foldr (op $) (HOLogic.mk_number @{typ nat} j) - (replicate i @{term Suc}); + val (i, j) = if n > n_sucs then (n_sucs, n - n_sucs) else (n, 0); + val arg' = + List.foldr (op $) (HOLogic.mk_number @{typ nat} j) (replicate i @{term Suc}); val _ = if arg = arg' then raise TERM ("", []) else (); - fun propfn g = HOLogic.mk_eq (g arg, g arg') + fun propfn g = + HOLogic.mk_eq (g arg, g arg') |> HOLogic.mk_Trueprop |> Thm.cterm_of ctxt; - val eq1 = Goal.prove_internal ctxt [] (propfn I) - (K (simp_tac (put_simpset word_ss ctxt) 1)); - in Goal.prove_internal ctxt [] (propfn (curry (op $) f)) + val eq1 = + Goal.prove_internal ctxt [] (propfn I) + (K (simp_tac (put_simpset word_ss ctxt) 1)); + in + Goal.prove_internal ctxt [] (propfn (curry (op $) f)) (K (simp_tac (put_simpset HOL_ss ctxt addsimps [eq1]) 1)) - |> mk_meta_eq |> SOME end - handle TERM _ => NONE; + |> mk_meta_eq |> SOME + end handle TERM _ => NONE; fun nat_get_Suc_simproc n_sucs ts = Simplifier.make_simproc @{context} "nat_get_Suc" diff -r acb0807ddb56 -r 491fd7f0b5df src/HOL/Word/Word_Miscellaneous.thy --- a/src/HOL/Word/Word_Miscellaneous.thy Sun Dec 03 13:22:09 2017 +0100 +++ b/src/HOL/Word/Word_Miscellaneous.thy Sun Dec 03 18:53:49 2017 +0100 @@ -28,19 +28,19 @@ done lemma list_exhaust_size_gt0: - assumes y: "\a list. y = a # list \ P" + assumes "\a list. y = a # list \ P" shows "0 < length y \ P" apply (cases y) apply simp - apply (rule y) + apply (rule assms) apply fastforce done lemma list_exhaust_size_eq0: - assumes y: "y = [] \ P" + assumes "y = [] \ P" shows "length y = 0 \ P" apply (cases y) - apply (rule y) + apply (rule assms) apply simp apply simp done @@ -307,7 +307,7 @@ lemmas td_gal_lt = td_gal [simplified not_less [symmetric], simplified] -lemmas div_mult_le = div_times_less_eq_dividend +lemmas div_mult_le = div_times_less_eq_dividend lemmas sdl = div_nat_eqI