changeset 76217 | 8655344f1cf6 |
parent 76216 | 9fc34f76b4e8 |
child 76219 | cf7db6353322 |
--- a/src/ZF/Induct/PropLog.thy Tue Sep 27 17:54:20 2022 +0100 +++ b/src/ZF/Induct/PropLog.thy Tue Sep 27 18:02:34 2022 +0100 @@ -107,7 +107,7 @@ subsection \<open>Proof theory of propositional logic\<close> lemma thms_mono: "G \<subseteq> H \<Longrightarrow> thms(G) \<subseteq> thms(H)" - apply (unfold thms.defs) + unfolding thms.defs apply (rule lfp_mono) apply (rule thms.bnd_mono)+ apply (assumption | rule univ_mono basic_monos)+