author | haftmann |
Mon, 02 Dec 2019 17:15:17 +0000 | |
changeset 71196 | 29e7c6d11cf1 |
parent 71195 | d50a718ccf35 |
child 71199 | 0aaf16a0f7cc |
--- a/src/HOL/ex/Word.thy Mon Dec 02 17:15:16 2019 +0000 +++ b/src/HOL/ex/Word.thy Mon Dec 02 17:15:17 2019 +0000 @@ -267,6 +267,10 @@ subsubsection \<open>Properties\<close> +lemma exp_eq_zero_iff: + \<open>(2 :: 'a::len word) ^ n = 0 \<longleftrightarrow> LENGTH('a) \<le> n\<close> + by transfer simp + subsubsection \<open>Division\<close>