author | wenzelm |
Sat, 01 Oct 2016 11:14:00 +0200 | |
changeset 63976 | c1a481bb82d3 |
parent 63975 | 6728b5007ad0 |
child 63977 | ec0fb01c6d50 |
--- a/src/HOL/Inductive.thy Fri Sep 30 17:12:50 2016 +0100 +++ b/src/HOL/Inductive.thy Sat Oct 01 11:14:00 2016 +0200 @@ -158,6 +158,10 @@ by (rule antisym) (simp_all add: gfp_upperbound gfp_unfold[symmetric]) +lemma lfp_le_gfp: "mono f \<Longrightarrow> lfp f \<le> gfp f" + by (iprover intro: gfp_upperbound lfp_lemma3) + + subsection \<open>Coinduction rules for greatest fixed points\<close> text \<open>Weak version.\<close>