Theory Fix

theory Fix
imports Cfun
(*  Title:      HOL/HOLCF/Fix.thy
    Author:     Franz Regensburger
    Author:     Brian Huffman
*)

section ‹Fixed point operator and admissibility›

theory Fix
imports Cfun
begin

default_sort pcpo

subsection ‹Iteration›

primrec iterate :: "nat ⇒ ('a::cpo → 'a) → ('a → 'a)" where
    "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.›

lemma chain_iterate [simp]: "chain (λi. iterate i⋅F⋅⊥)"
by (rule chainI, unfold iterate_Suc2, rule monofun_cfun_arg, rule minimal)


subsection ‹Least fixed point operator›

definition
  "fix" :: "('a → 'a) → 'a" where
  "fix = (Λ F. ⨆i. iterate i⋅F⋅⊥)"

text ‹Binder syntax for @{term fix}›

abbreviation
  fix_syn :: "('a ⇒ 'a) ⇒ 'a"  (binder "μ " 10) where
  "fix_syn (λx. f x) ≡ fix⋅(Λ x. f x)"

notation (ASCII)
  fix_syn  (binder "FIX " 10)

text ‹Properties of @{term fix}›

text ‹direct connection between @{term fix} and iteration›

lemma fix_def2: "fix⋅F = (⨆i. iterate i⋅F⋅⊥)"
unfolding fix_def by simp

lemma iterate_below_fix: "iterate n⋅f⋅⊥ ⊑ fix⋅f"
  unfolding fix_def2
  using chain_iterate by (rule is_ub_thelub)

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_below: "F⋅x ⊑ x ⟹ fix⋅F ⊑ x"
apply (simp add: fix_def2)
apply (rule lub_below)
apply (rule chain_iterate)
apply (induct_tac i)
apply simp
apply simp
apply (erule rev_below_trans)
apply (erule monofun_cfun_arg)
done

lemma fix_least: "F⋅x = x ⟹ fix⋅F ⊑ x"
by (rule fix_least_below, simp)

lemma fix_eqI:
  assumes fixed: "F⋅x = x" and least: "⋀z. F⋅z = z ⟹ x ⊑ z"
  shows "fix⋅F = x"
apply (rule below_antisym)
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_bottom_iff: "(fix⋅F = ⊥) = (F⋅⊥ = ⊥)"
apply (rule iffI)
apply (erule subst)
apply (rule fix_eq [symmetric])
apply (erule fix_least [THEN bottomI])
done

lemma fix_strict: "F⋅⊥ = ⊥ ⟹ fix⋅F = ⊥"
by (simp add: fix_bottom_iff)

lemma fix_defined: "F⋅⊥ ≠ ⊥ ⟹ fix⋅F ≠ ⊥"
by (simp add: fix_bottom_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 cont_fix_ind:
  "⟦cont F; adm P; P ⊥; ⋀x. P x ⟹ P (F x)⟧ ⟹ P (fix⋅(Abs_cfun F))"
by (simp add: fix_ind)

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

lemma parallel_fix_ind:
  assumes adm: "adm (λx. P (fst x) (snd x))"
  assumes base: "P ⊥ ⊥"
  assumes step: "⋀x y. P x y ⟹ P (F⋅x) (G⋅y)"
  shows "P (fix⋅F) (fix⋅G)"
proof -
  from adm have adm': "adm (case_prod P)"
    unfolding split_def .
  have "⋀i. P (iterate i⋅F⋅⊥) (iterate i⋅G⋅⊥)"
    by (induct_tac i, simp add: base, simp add: step)
  hence "⋀i. case_prod P (iterate i⋅F⋅⊥, iterate i⋅G⋅⊥)"
    by simp
  hence "case_prod P (⨆i. (iterate i⋅F⋅⊥, iterate i⋅G⋅⊥))"
    by - (rule admD [OF adm'], simp, assumption)
  hence "case_prod P (⨆i. iterate i⋅F⋅⊥, ⨆i. iterate i⋅G⋅⊥)"
    by (simp add: lub_Pair)
  hence "P (⨆i. iterate i⋅F⋅⊥) (⨆i. iterate i⋅G⋅⊥)"
    by simp
  thus "P (fix⋅F) (fix⋅G)"
    by (simp add: fix_def2)
qed

lemma cont_parallel_fix_ind:
  assumes "cont F" and "cont G"
  assumes "adm (λx. P (fst x) (snd x))"
  assumes "P ⊥ ⊥"
  assumes "⋀x y. P x y ⟹ P (F x) (G y)"
  shows "P (fix⋅(Abs_cfun F)) (fix⋅(Abs_cfun G))"
by (rule parallel_fix_ind, simp_all add: assms)

subsection ‹Fixed-points on product types›

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. fst (F⋅(x, μ y. snd (F⋅(x, y)))),
    μ y. snd (F⋅(μ x. fst (F⋅(x, μ y. snd (F⋅(x, y)))), y)))"
  (is "fix⋅F = (?x, ?y)")
proof (rule fix_eqI)
  have 1: "fst (F⋅(?x, ?y)) = ?x"
    by (rule trans [symmetric, OF fix_eq], simp)
  have 2: "snd (F⋅(?x, ?y)) = ?y"
    by (rule trans [symmetric, OF fix_eq], simp)
  from 1 2 show "F⋅(?x, ?y) = (?x, ?y)" by (simp add: prod_eq_iff)
next
  fix z assume F_z: "F⋅z = z"
  obtain x y where z: "z = (x,y)" by (rule prod.exhaust)
  from F_z z have F_x: "fst (F⋅(x, y)) = x" by simp
  from F_z z have F_y: "snd (F⋅(x, y)) = y" by simp
  let ?y1 = "μ y. snd (F⋅(x, y))"
  have "?y1 ⊑ y" by (rule fix_least, simp add: F_y)
  hence "fst (F⋅(x, ?y1)) ⊑ fst (F⋅(x, y))"
    by (simp add: fst_monofun monofun_cfun)
  hence "fst (F⋅(x, ?y1)) ⊑ x" using F_x by simp
  hence 1: "?x ⊑ x" by (simp add: fix_least_below)
  hence "snd (F⋅(?x, y)) ⊑ snd (F⋅(x, y))"
    by (simp add: snd_monofun monofun_cfun)
  hence "snd (F⋅(?x, y)) ⊑ y" using F_y by simp
  hence 2: "?y ⊑ y" by (simp add: fix_least_below)
  show "(?x, ?y) ⊑ z" using z 1 2 by simp
qed

end