author | huffman |
Thu, 23 Aug 2007 18:47:08 +0200 | |
changeset 24412 | 9c7bb416f344 |
parent 24411 | 0c9af72fb303 |
child 24413 | 5073729e5c12 |
--- a/src/HOL/Word/BinInduct.thy Thu Aug 23 16:47:16 2007 +0200 +++ b/src/HOL/Word/BinInduct.thy Thu Aug 23 18:47:08 2007 +0200 @@ -126,7 +126,7 @@ lemma bin_last_Min [simp]: "bin_last Numeral.Min = bit.B1" by (subst Min_1_eq [symmetric], rule bin_last_BIT) -lemma bin_rest_BIT_bin_last: "(bin_rest x) BIT (bin_last x) = x" +lemma bin_rest_BIT_bin_last [simp]: "(bin_rest x) BIT (bin_last x) = x" by (cases x rule: BIT_cases) simp lemma wf_bin_rest: