tuned proof (no induction needed); removed unused lemma and fuzzy comment
authorkrauss
Tue, 06 Apr 2010 09:27:03 +0200
changeset 36075 2e0370c03066
parent 36074 6301046146b6
child 36077 13f0e77128f8
tuned proof (no induction needed); removed unused lemma and fuzzy comment
src/HOLCF/Fix.thy
--- a/src/HOLCF/Fix.thy	Fri Apr 02 17:20:43 2010 +0200
+++ b/src/HOLCF/Fix.thy	Tue Apr 06 09:27:03 2010 +0200
@@ -34,16 +34,10 @@
   "iterate m\<cdot>F\<cdot>(iterate n\<cdot>F\<cdot>x) = iterate (m + n)\<cdot>F\<cdot>x"
 by (induct m) simp_all
 
-text {*
-  The sequence of function iterations is a chain.
-  This property is essential since monotonicity of iterate makes no sense.
-*}
-
-lemma chain_iterate2: "x \<sqsubseteq> F\<cdot>x \<Longrightarrow> chain (\<lambda>i. iterate i\<cdot>F\<cdot>x)"
-by (rule chainI, induct_tac i, auto elim: monofun_cfun_arg)
+text {* The sequence of function iterations is a chain. *}
 
 lemma chain_iterate [simp]: "chain (\<lambda>i. iterate i\<cdot>F\<cdot>\<bottom>)"
-by (rule chain_iterate2 [OF minimal])
+by (rule chainI, unfold iterate_Suc2, rule monofun_cfun_arg, rule minimal)
 
 
 subsection {* Least fixed point operator *}