diff -r a642599ffdea -r 9fc34f76b4e8 src/ZF/UNITY/UNITY.thy --- a/src/ZF/UNITY/UNITY.thy Tue Sep 27 17:46:52 2022 +0100 +++ b/src/ZF/UNITY/UNITY.thy Tue Sep 27 17:54:20 2022 +0100 @@ -165,7 +165,7 @@ lemmas InitD = Init_type [THEN subsetD] lemma st_set_Init [iff]: "st_set(Init(F))" -apply (unfold st_set_def) + unfolding st_set_def apply (rule Init_type) done @@ -469,7 +469,7 @@ by (force simp add: unless_def constrains_def) lemma unlessI: "\F \ (A-B) co (A \ B)\ \ F \ A unless B" -apply (unfold unless_def) + unfolding unless_def apply (blast dest: constrainsD2) done @@ -515,34 +515,34 @@ lemma stable_Un: "\F \ stable(A); F \ stable(A')\ \ F \ stable(A \ A')" -apply (unfold stable_def) + unfolding stable_def apply (blast intro: constrains_Un) done lemma stable_UN: "\\i. i\I \ F \ stable(A(i)); F \ program\ \ F \ stable (\i \ I. A(i))" -apply (unfold stable_def) + unfolding stable_def apply (blast intro: constrains_UN) done lemma stable_Int: "\F \ stable(A); F \ stable(A')\ \ F \ stable (A \ A')" -apply (unfold stable_def) + unfolding stable_def apply (blast intro: constrains_Int) done lemma stable_INT: "\\i. i \ I \ F \ stable(A(i)); F \ program\ \ F \ stable (\i \ I. A(i))" -apply (unfold stable_def) + unfolding stable_def apply (blast intro: constrains_INT) done lemma stable_All: "\\z. F \ stable({s \ state. P(s, z)}); F \ program\ \ F \ stable({s \ state. \z. P(s, z)})" -apply (unfold stable_def) + unfolding stable_def apply (rule constrains_All, auto) done @@ -562,7 +562,7 @@ subsection\The Operator \<^term>\invariant\\ lemma invariant_type: "invariant(A) \ program" -apply (unfold invariant_def) + unfolding invariant_def apply (blast dest: stable_type [THEN subsetD]) done @@ -575,7 +575,7 @@ by (unfold invariant_def initially_def, auto) lemma invariantD2: "F \ invariant(A) \ F \ program \ st_set(A)" -apply (unfold invariant_def) + unfolding invariant_def apply (blast dest: stableD2) done