| author | wenzelm | 
| Tue, 27 Jul 2021 15:20:20 +0200 | |
| changeset 74071 | b25b7c264a93 | 
| parent 67443 | 3abf6a722518 | 
| permissions | -rw-r--r-- | 
| 17456 | 1 | (* Title: CCL/Lfp.thy | 
| 1474 | 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 0 | 3 | Copyright 1992 University of Cambridge | 
| 4 | *) | |
| 5 | ||
| 60770 | 6 | section \<open>The Knaster-Tarski Theorem\<close> | 
| 17456 | 7 | |
| 8 | theory Lfp | |
| 9 | imports Set | |
| 10 | begin | |
| 11 | ||
| 20140 | 12 | definition | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
62020diff
changeset | 13 | lfp :: "['a set\<Rightarrow>'a set] \<Rightarrow> 'a set" where \<comment> \<open>least fixed point\<close> | 
| 17456 | 14 |   "lfp(f) == Inter({u. f(u) <= u})"
 | 
| 15 | ||
| 20140 | 16 | (* lfp(f) is the greatest lower bound of {u. f(u) <= u} *)
 | 
| 17 | ||
| 58977 | 18 | lemma lfp_lowerbound: "f(A) <= A \<Longrightarrow> lfp(f) <= A" | 
| 20140 | 19 | unfolding lfp_def by blast | 
| 20 | ||
| 58977 | 21 | lemma lfp_greatest: "(\<And>u. f(u) <= u \<Longrightarrow> A<=u) \<Longrightarrow> A <= lfp(f)" | 
| 20140 | 22 | unfolding lfp_def by blast | 
| 23 | ||
| 58977 | 24 | lemma lfp_lemma2: "mono(f) \<Longrightarrow> f(lfp(f)) <= lfp(f)" | 
| 20140 | 25 | by (rule lfp_greatest, rule subset_trans, drule monoD, rule lfp_lowerbound, assumption+) | 
| 26 | ||
| 58977 | 27 | lemma lfp_lemma3: "mono(f) \<Longrightarrow> lfp(f) <= f(lfp(f))" | 
| 20140 | 28 | by (rule lfp_lowerbound, frule monoD, drule lfp_lemma2, assumption+) | 
| 29 | ||
| 58977 | 30 | lemma lfp_Tarski: "mono(f) \<Longrightarrow> lfp(f) = f(lfp(f))" | 
| 20140 | 31 | by (rule equalityI lfp_lemma2 lfp_lemma3 | assumption)+ | 
| 32 | ||
| 33 | ||
| 34 | (*** General induction rule for least fixed points ***) | |
| 35 | ||
| 36 | lemma induct: | |
| 37 | assumes lfp: "a: lfp(f)" | |
| 38 | and mono: "mono(f)" | |
| 58977 | 39 |     and indhyp: "\<And>x. \<lbrakk>x: f(lfp(f) Int {x. P(x)})\<rbrakk> \<Longrightarrow> P(x)"
 | 
| 20140 | 40 | shows "P(a)" | 
| 41 | apply (rule_tac a = a in Int_lower2 [THEN subsetD, THEN CollectD]) | |
| 42 | apply (rule lfp [THEN [2] lfp_lowerbound [THEN subsetD]]) | |
| 43 | apply (rule Int_greatest, rule subset_trans, rule Int_lower1 [THEN mono [THEN monoD]], | |
| 44 | rule mono [THEN lfp_lemma2], rule CollectI [THEN subsetI], rule indhyp, assumption) | |
| 45 | done | |
| 46 | ||
| 47 | (** Definition forms of lfp_Tarski and induct, to control unfolding **) | |
| 48 | ||
| 58977 | 49 | lemma def_lfp_Tarski: "\<lbrakk>h == lfp(f); mono(f)\<rbrakk> \<Longrightarrow> h = f(h)" | 
| 20140 | 50 | apply unfold | 
| 51 | apply (drule lfp_Tarski) | |
| 52 | apply assumption | |
| 53 | done | |
| 54 | ||
| 58977 | 55 | lemma def_induct: "\<lbrakk>A == lfp(f);  a:A;  mono(f); \<And>x. x: f(A Int {x. P(x)}) \<Longrightarrow> P(x)\<rbrakk> \<Longrightarrow> P(a)"
 | 
| 20140 | 56 | apply (rule induct [of concl: P a]) | 
| 57 | apply simp | |
| 58 | apply assumption | |
| 59 | apply blast | |
| 60 | done | |
| 61 | ||
| 62 | (*Monotonicity of lfp!*) | |
| 58977 | 63 | lemma lfp_mono: "\<lbrakk>mono(g); \<And>Z. f(Z) <= g(Z)\<rbrakk> \<Longrightarrow> lfp(f) <= lfp(g)" | 
| 20140 | 64 | apply (rule lfp_lowerbound) | 
| 65 | apply (rule subset_trans) | |
| 66 | apply (erule meta_spec) | |
| 67 | apply (erule lfp_lemma2) | |
| 68 | done | |
| 17456 | 69 | |
| 0 | 70 | end |