src/ZF/UNITY/Distributor.thy
changeset 76216 9fc34f76b4e8
parent 76215 a642599ffdea
--- a/src/ZF/UNITY/Distributor.thy	Tue Sep 27 17:46:52 2022 +0100
+++ b/src/ZF/UNITY/Distributor.thy	Tue Sep 27 17:54:20 2022 +0100
@@ -55,19 +55,19 @@
 
 
 lemma (in distr) In_value_type [simp,TC]: "s \<in> state \<Longrightarrow> s`In \<in> list(A)"
-apply (unfold state_def)
+  unfolding state_def
 apply (drule_tac a = In in apply_type, auto)
 done
 
 lemma (in distr) iIn_value_type [simp,TC]:
      "s \<in> state \<Longrightarrow> s`iIn \<in> list(nat)"
-apply (unfold state_def)
+  unfolding state_def
 apply (drule_tac a = iIn in apply_type, auto)
 done
 
 lemma (in distr) Out_value_type [simp,TC]:
      "s \<in> state \<Longrightarrow> s`Out(n):list(A)"
-apply (unfold state_def)
+  unfolding state_def
 apply (drule_tac a = "Out (n)" in apply_type)
 apply auto
 done