src/HOL/IMP/Abs_Int1.thy
changeset 49353 023be49d7fb8
parent 49344 ce1ccb78ecda
child 49396 73fb17ed2e08
--- a/src/HOL/IMP/Abs_Int1.thy	Thu Sep 13 17:20:44 2012 +0200
+++ b/src/HOL/IMP/Abs_Int1.thy	Fri Sep 14 10:01:42 2012 +0200
@@ -28,8 +28,11 @@
    wt I X \<and> vars b \<subseteq> X \<and> wt P X \<and> wt C X \<and> wt Q X"
 by(auto simp add: wt_acom_def)
 
-lemma wt_post[simp]: "wt c  X \<Longrightarrow> wt (post c) X"
-by(induction c)(auto simp: wt_acom_def)
+lemma post_in_annos: "post C : set(annos C)"
+by(induction C) auto
+
+lemma wt_post[simp]: "wt C X \<Longrightarrow> wt (post C) X"
+by(simp add: wt_acom_def post_in_annos)
 
 lemma lpfp_inv:
 assumes "lpfp f x0 = Some x" and "\<And>x. P x \<Longrightarrow> P(f x)" and "P(bot x0)"