header {* Fixed point operator and admissibility *}
theory Fix
imports Cfun Cprod
begin
defaultsort pcpo
subsection {* Iteration *}
consts
iterate :: "nat => ('a::cpo -> 'a) -> ('a -> 'a)"
primrec
"iterate 0 = (Λ F x. x)"
"iterate (Suc n) = (Λ F x. F·(iterate n·F·x))"
text {* Derive inductive properties of iterate from primitive recursion *}
lemma iterate_0 [simp]: "iterate 0·F·x = x"
by simp
lemma iterate_Suc [simp]: "iterate (Suc n)·F·x = F·(iterate n·F·x)"
by simp
declare iterate.simps [simp del]
lemma iterate_Suc2: "iterate (Suc n)·F·x = iterate n·F·(F·x)"
by (induct n) simp_all
lemma iterate_iterate:
"iterate m·F·(iterate n·F·x) = iterate (m + n)·F·x"
by (induct m) simp_all
text {*
The sequence of function iterations is a chain.
This property is essential since monotonicity of iterate makes no sense.
*}
lemma chain_iterate2: "x \<sqsubseteq> F·x ==> chain (λi. iterate i·F·x)"
by (rule chainI, induct_tac i, auto elim: monofun_cfun_arg)
lemma chain_iterate [simp]: "chain (λi. iterate i·F·⊥)"
by (rule chain_iterate2 [OF minimal])
subsection {* Least fixed point operator *}
definition
"fix" :: "('a -> 'a) -> 'a" where
"fix = (Λ F. \<Squnion>i. iterate i·F·⊥)"
text {* Binder syntax for @{term fix} *}
abbreviation
fix_syn :: "('a => 'a) => 'a" (binder "FIX " 10) where
"fix_syn (λx. f x) ≡ fix·(Λ x. f x)"
notation (xsymbols)
fix_syn (binder "μ " 10)
text {* Properties of @{term fix} *}
text {* direct connection between @{term fix} and iteration *}
lemma fix_def2: "fix·F = (\<Squnion>i. iterate i·F·⊥)"
apply (unfold fix_def)
apply (rule beta_cfun)
apply (rule cont2cont_lub)
apply (rule ch2ch_lambda)
apply (rule chain_iterate)
apply simp
done
text {*
Kleene's fixed point theorems for continuous functions in pointed
omega cpo's
*}
lemma fix_eq: "fix·F = F·(fix·F)"
apply (simp add: fix_def2)
apply (subst lub_range_shift [of _ 1, symmetric])
apply (rule chain_iterate)
apply (subst contlub_cfun_arg)
apply (rule chain_iterate)
apply simp
done
lemma fix_least_less: "F·x \<sqsubseteq> x ==> fix·F \<sqsubseteq> x"
apply (simp add: fix_def2)
apply (rule is_lub_thelub)
apply (rule chain_iterate)
apply (rule ub_rangeI)
apply (induct_tac i)
apply simp
apply simp
apply (erule rev_trans_less)
apply (erule monofun_cfun_arg)
done
lemma fix_least: "F·x = x ==> fix·F \<sqsubseteq> x"
by (rule fix_least_less, simp)
lemma fix_eqI:
assumes fixed: "F·x = x" and least: "!!z. F·z = z ==> x \<sqsubseteq> z"
shows "fix·F = x"
apply (rule antisym_less)
apply (rule fix_least [OF fixed])
apply (rule least [OF fix_eq [symmetric]])
done
lemma fix_eq2: "f ≡ fix·F ==> f = F·f"
by (simp add: fix_eq [symmetric])
lemma fix_eq3: "f ≡ fix·F ==> f·x = F·f·x"
by (erule fix_eq2 [THEN cfun_fun_cong])
lemma fix_eq4: "f = fix·F ==> f = F·f"
apply (erule ssubst)
apply (rule fix_eq)
done
lemma fix_eq5: "f = fix·F ==> f·x = F·f·x"
by (erule fix_eq4 [THEN cfun_fun_cong])
text {* strictness of @{term fix} *}
lemma fix_defined_iff: "(fix·F = ⊥) = (F·⊥ = ⊥)"
apply (rule iffI)
apply (erule subst)
apply (rule fix_eq [symmetric])
apply (erule fix_least [THEN UU_I])
done
lemma fix_strict: "F·⊥ = ⊥ ==> fix·F = ⊥"
by (simp add: fix_defined_iff)
lemma fix_defined: "F·⊥ ≠ ⊥ ==> fix·F ≠ ⊥"
by (simp add: fix_defined_iff)
text {* @{term fix} applied to identity and constant functions *}
lemma fix_id: "(μ x. x) = ⊥"
by (simp add: fix_strict)
lemma fix_const: "(μ x. c) = c"
by (subst fix_eq, simp)
subsection {* Fixed point induction *}
lemma fix_ind: "[|adm P; P ⊥; !!x. P x ==> P (F·x)|] ==> P (fix·F)"
unfolding fix_def2
apply (erule admD)
apply (rule chain_iterate)
apply (rule nat_induct, simp_all)
done
lemma def_fix_ind:
"[|f ≡ fix·F; adm P; P ⊥; !!x. P x ==> P (F·x)|] ==> P f"
by (simp add: fix_ind)
lemma fix_ind2:
assumes adm: "adm P"
assumes 0: "P ⊥" and 1: "P (F·⊥)"
assumes step: "!!x. [|P x; P (F·x)|] ==> P (F·(F·x))"
shows "P (fix·F)"
unfolding fix_def2
apply (rule admD [OF adm chain_iterate])
apply (rule nat_less_induct)
apply (case_tac n)
apply (simp add: 0)
apply (case_tac nat)
apply (simp add: 1)
apply (frule_tac x=nat in spec)
apply (simp add: step)
done
subsection {* Recursive let bindings *}
definition
CLetrec :: "('a -> 'a × 'b) -> 'b" where
"CLetrec = (Λ F. csnd·(F·(μ x. cfst·(F·x))))"
nonterminals
recbinds recbindt recbind
syntax
"_recbind" :: "['a, 'a] => recbind" ("(2_ =/ _)" 10)
"" :: "recbind => recbindt" ("_")
"_recbindt" :: "[recbind, recbindt] => recbindt" ("_,/ _")
"" :: "recbindt => recbinds" ("_")
"_recbinds" :: "[recbindt, recbinds] => recbinds" ("_;/ _")
"_Letrec" :: "[recbinds, 'a] => 'a" ("(Letrec (_)/ in (_))" 10)
translations
(recbindt) "x = a, 〈y,ys〉 = 〈b,bs〉" == (recbindt) "〈x,y,ys〉 = 〈a,b,bs〉"
(recbindt) "x = a, y = b" == (recbindt) "〈x,y〉 = 〈a,b〉"
translations
"_Letrec (_recbinds b bs) e" == "_Letrec b (_Letrec bs e)"
"Letrec xs = a in 〈e,es〉" == "CONST CLetrec·(Λ xs. 〈a,e,es〉)"
"Letrec xs = a in e" == "CONST CLetrec·(Λ xs. 〈a,e〉)"
text {*
Bekic's Theorem: Simultaneous fixed points over pairs
can be written in terms of separate fixed points.
*}
lemma fix_cprod:
"fix·(F::'a × 'b -> 'a × 'b) =
〈μ x. cfst·(F·〈x, μ y. csnd·(F·〈x, y〉)〉),
μ y. csnd·(F·〈μ x. cfst·(F·〈x, μ y. csnd·(F·〈x, y〉)〉), y〉)〉"
(is "fix·F = 〈?x, ?y〉")
proof (rule fix_eqI)
have 1: "cfst·(F·〈?x, ?y〉) = ?x"
by (rule trans [symmetric, OF fix_eq], simp)
have 2: "csnd·(F·〈?x, ?y〉) = ?y"
by (rule trans [symmetric, OF fix_eq], simp)
from 1 2 show "F·〈?x, ?y〉 = 〈?x, ?y〉" by (simp add: eq_cprod)
next
fix z assume F_z: "F·z = z"
then obtain x y where z: "z = 〈x,y〉" by (rule_tac p=z in cprodE)
from F_z z have F_x: "cfst·(F·〈x, y〉) = x" by simp
from F_z z have F_y: "csnd·(F·〈x, y〉) = y" by simp
let ?y1 = "μ y. csnd·(F·〈x, y〉)"
have "?y1 \<sqsubseteq> y" by (rule fix_least, simp add: F_y)
hence "cfst·(F·〈x, ?y1〉) \<sqsubseteq> cfst·(F·〈x, y〉)" by (simp add: monofun_cfun)
hence "cfst·(F·〈x, ?y1〉) \<sqsubseteq> x" using F_x by simp
hence 1: "?x \<sqsubseteq> x" by (simp add: fix_least_less)
hence "csnd·(F·〈?x, y〉) \<sqsubseteq> csnd·(F·〈x, y〉)" by (simp add: monofun_cfun)
hence "csnd·(F·〈?x, y〉) \<sqsubseteq> y" using F_y by simp
hence 2: "?y \<sqsubseteq> y" by (simp add: fix_least_less)
show "〈?x, ?y〉 \<sqsubseteq> z" using z 1 2 by simp
qed
subsection {* Weak admissibility *}
definition
admw :: "('a => bool) => bool" where
"admw P = (∀F. (∀n. P (iterate n·F·⊥)) --> P (\<Squnion>i. iterate i·F·⊥))"
text {* an admissible formula is also weak admissible *}
lemma adm_impl_admw: "adm P ==> admw P"
apply (unfold admw_def)
apply (intro strip)
apply (erule admD)
apply (rule chain_iterate)
apply (erule spec)
done
text {* computational induction for weak admissible formulae *}
lemma wfix_ind: "[|admw P; ∀n. P (iterate n·F·⊥)|] ==> P (fix·F)"
by (simp add: fix_def2 admw_def)
lemma def_wfix_ind:
"[|f ≡ fix·F; admw P; ∀n. P (iterate n·F·⊥)|] ==> P f"
by (simp, rule wfix_ind)
end