author | haftmann |
Sun, 16 Oct 2011 14:48:00 +0200 | |
changeset 45153 | 93e290c11b0f |
parent 45152 | e877b76c72bd |
child 45154 | 66e8a5812f41 |
--- a/src/HOL/Transitive_Closure.thy Sun Oct 16 14:48:00 2011 +0200 +++ b/src/HOL/Transitive_Closure.thy Sun Oct 16 14:48:00 2011 +0200 @@ -779,8 +779,8 @@ by (induct n) (simp, simp add: O_assoc [symmetric]) lemma rel_pow_empty: - "0 < n ==> ({} :: 'a * 'a => bool) ^^ n = {}" -by (cases n) auto + "0 < n \<Longrightarrow> ({} :: ('a \<times> 'a) set) ^^ n = {}" + by (cases n) auto lemma rtrancl_imp_UN_rel_pow: assumes "p \<in> R^*"