# HG changeset patch # User haftmann # Date 1201687064 -3600 # Node ID 8764a1f1253bf0527c3741c465d2aaad1371ac5c # Parent f6917792f8a4475481c203e49d79989de21002f8 Theorem Inductive.lfp_ordinal_induct generalized to complete lattices diff -r f6917792f8a4 -r 8764a1f1253b NEWS --- a/NEWS Tue Jan 29 18:00:12 2008 +0100 +++ b/NEWS Wed Jan 30 10:57:44 2008 +0100 @@ -35,6 +35,9 @@ *** HOL *** +* Theorem Inductive.lfp_ordinal_induct generalized to complete lattices. The +form set-specific version is available as Inductive.lfp_ordinal_induct_set. + * Theorems "power.simps" renamed to "power_int.simps". * New class semiring_div provides basic abstract properties of semirings diff -r f6917792f8a4 -r 8764a1f1253b src/HOL/Inductive.thy --- a/src/HOL/Inductive.thy Tue Jan 29 18:00:12 2008 +0100 +++ b/src/HOL/Inductive.thy Wed Jan 30 10:57:44 2008 +0100 @@ -24,12 +24,15 @@ subsection {* Least and greatest fixed points *} +context complete_lattice +begin + definition - lfp :: "('a\complete_lattice \ 'a) \ 'a" where + lfp :: "('a \ 'a) \ 'a" where "lfp f = Inf {u. f u \ u}" --{*least fixed point*} definition - gfp :: "('a\complete_lattice \ 'a) \ 'a" where + gfp :: "('a \ 'a) \ 'a" where "gfp f = Sup {u. u \ f u}" --{*greatest fixed point*} @@ -44,6 +47,8 @@ lemma lfp_greatest: "(!!u. f u \ u ==> A \ u) ==> A \ lfp f" by (auto simp add: lfp_def intro: Inf_greatest) +end + lemma lfp_lemma2: "mono f ==> f (lfp f) \ lfp f" by (iprover intro: lfp_greatest order_trans monoD lfp_lowerbound) @@ -81,25 +86,34 @@ by (rule lfp_induct [THEN subsetD, THEN CollectD, OF mono _ lfp]) (auto simp: inf_set_eq intro: indhyp) -lemma lfp_ordinal_induct: +lemma lfp_ordinal_induct: + fixes f :: "'a\complete_lattice \ 'a" + assumes mono: "mono f" + and P_f: "\S. P S \ P (f S)" + and P_Union: "\M. \S\M. P S \ P (Sup M)" + shows "P (lfp f)" +proof - + let ?M = "{S. S \ lfp f \ P S}" + have "P (Sup ?M)" using P_Union by simp + also have "Sup ?M = lfp f" + proof (rule antisym) + show "Sup ?M \ lfp f" by (blast intro: Sup_least) + hence "f (Sup ?M) \ f (lfp f)" by (rule mono [THEN monoD]) + hence "f (Sup ?M) \ lfp f" using mono [THEN lfp_unfold] by simp + hence "f (Sup ?M) \ ?M" using P_f P_Union by simp + hence "f (Sup ?M) \ Sup ?M" by (rule Sup_upper) + thus "lfp f \ Sup ?M" by (rule lfp_lowerbound) + qed + finally show ?thesis . +qed + +lemma lfp_ordinal_induct_set: assumes mono: "mono f" and P_f: "!!S. P S ==> P(f S)" and P_Union: "!!M. !S:M. P S ==> P(Union M)" shows "P(lfp f)" -proof - - let ?M = "{S. S \ lfp f & P S}" - have "P (Union ?M)" using P_Union by simp - also have "Union ?M = lfp f" - proof - show "Union ?M \ lfp f" by blast - hence "f (Union ?M) \ f (lfp f)" by (rule mono [THEN monoD]) - hence "f (Union ?M) \ lfp f" using mono [THEN lfp_unfold] by simp - hence "f (Union ?M) \ ?M" using P_f P_Union by simp - hence "f (Union ?M) \ Union ?M" by (rule Union_upper) - thus "lfp f \ Union ?M" by (rule lfp_lowerbound) - qed - finally show ?thesis . -qed + using assms unfolding Sup_set_def [symmetric] + by (rule lfp_ordinal_induct) text{*Definition forms of @{text lfp_unfold} and @{text lfp_induct},