header {* Least Fixed Point Operation on Bounded Natural Functors *}
theory BNF_LFP
imports BNF_FP_Base
keywords
"datatype_new" :: thy_decl and
"datatype_compat" :: thy_decl
begin
lemma subset_emptyI: "(!!x. x ∈ A ==> False) ==> A ⊆ {}"
by blast
lemma image_Collect_subsetI: "(!!x. P x ==> f x ∈ B) ==> f ` {x. P x} ⊆ B"
by blast
lemma Collect_restrict: "{x. x ∈ X ∧ P x} ⊆ X"
by auto
lemma prop_restrict: "[|x ∈ Z; Z ⊆ {x. x ∈ X ∧ P x}|] ==> P x"
by auto
lemma underS_I: "[|i ≠ j; (i, j) ∈ R|] ==> i ∈ underS R j"
unfolding underS_def by simp
lemma underS_E: "i ∈ underS R j ==> i ≠ j ∧ (i, j) ∈ R"
unfolding underS_def by simp
lemma underS_Field: "i ∈ underS R j ==> i ∈ Field R"
unfolding underS_def Field_def by auto
lemma FieldI2: "(i, j) ∈ R ==> j ∈ Field R"
unfolding Field_def by auto
lemma fst_convol': "fst (〈f, g〉 x) = f x"
using fst_convol unfolding convol_def by simp
lemma snd_convol': "snd (〈f, g〉 x) = g x"
using snd_convol unfolding convol_def by simp
lemma convol_expand_snd: "fst o f = g ==> 〈g, snd o f〉 = f"
unfolding convol_def by auto
lemma convol_expand_snd':
assumes "(fst o f = g)"
shows "h = snd o f <-> 〈g, h〉 = f"
proof -
from assms have *: "〈g, snd o f〉 = f" by (rule convol_expand_snd)
then have "h = snd o f <-> h = snd o 〈g, snd o f〉" by simp
moreover have "… <-> h = snd o f" by (simp add: snd_convol)
moreover have "… <-> 〈g, h〉 = f" by (subst (2) *[symmetric]) (auto simp: convol_def fun_eq_iff)
ultimately show ?thesis by simp
qed
lemma bij_betwE: "bij_betw f A B ==> ∀a∈A. f a ∈ B"
unfolding bij_betw_def by auto
lemma bij_betw_imageE: "bij_betw f A B ==> f ` A = B"
unfolding bij_betw_def by auto
lemma f_the_inv_into_f_bij_betw: "bij_betw f A B ==>
(bij_betw f A B ==> x ∈ B) ==> f (the_inv_into A f x) = x"
unfolding bij_betw_def by (blast intro: f_the_inv_into_f)
lemma ex_bij_betw: "|A| ≤o (r :: 'b rel) ==> ∃f B :: 'b set. bij_betw f B A"
by (subst (asm) internalize_card_of_ordLeq)
(auto dest!: iffD2[OF card_of_ordIso ordIso_symmetric])
lemma bij_betwI':
"[|!!x y. [|x ∈ X; y ∈ X|] ==> (f x = f y) = (x = y);
!!x. x ∈ X ==> f x ∈ Y;
!!y. y ∈ Y ==> ∃x ∈ X. y = f x|] ==> bij_betw f X Y"
unfolding bij_betw_def inj_on_def by blast
lemma surj_fun_eq:
assumes surj_on: "f ` X = UNIV" and eq_on: "∀x ∈ X. (g1 o f) x = (g2 o f) x"
shows "g1 = g2"
proof (rule ext)
fix y
from surj_on obtain x where "x ∈ X" and "y = f x" by blast
thus "g1 y = g2 y" using eq_on by simp
qed
lemma Card_order_wo_rel: "Card_order r ==> wo_rel r"
unfolding wo_rel_def card_order_on_def by blast
lemma Cinfinite_limit: "[|x ∈ Field r; Cinfinite r|] ==>
∃y ∈ Field r. x ≠ y ∧ (x, y) ∈ r"
unfolding cinfinite_def by (auto simp add: infinite_Card_order_limit)
lemma Card_order_trans:
"[|Card_order r; x ≠ y; (x, y) ∈ r; y ≠ z; (y, z) ∈ r|] ==> x ≠ z ∧ (x, z) ∈ r"
unfolding card_order_on_def well_order_on_def linear_order_on_def
partial_order_on_def preorder_on_def trans_def antisym_def by blast
lemma Cinfinite_limit2:
assumes x1: "x1 ∈ Field r" and x2: "x2 ∈ Field r" and r: "Cinfinite r"
shows "∃y ∈ Field r. (x1 ≠ y ∧ (x1, y) ∈ r) ∧ (x2 ≠ y ∧ (x2, y) ∈ r)"
proof -
from r have trans: "trans r" and total: "Total r" and antisym: "antisym r"
unfolding card_order_on_def well_order_on_def linear_order_on_def
partial_order_on_def preorder_on_def by auto
obtain y1 where y1: "y1 ∈ Field r" "x1 ≠ y1" "(x1, y1) ∈ r"
using Cinfinite_limit[OF x1 r] by blast
obtain y2 where y2: "y2 ∈ Field r" "x2 ≠ y2" "(x2, y2) ∈ r"
using Cinfinite_limit[OF x2 r] by blast
show ?thesis
proof (cases "y1 = y2")
case True with y1 y2 show ?thesis by blast
next
case False
with y1(1) y2(1) total have "(y1, y2) ∈ r ∨ (y2, y1) ∈ r"
unfolding total_on_def by auto
thus ?thesis
proof
assume *: "(y1, y2) ∈ r"
with trans y1(3) have "(x1, y2) ∈ r" unfolding trans_def by blast
with False y1 y2 * antisym show ?thesis by (cases "x1 = y2") (auto simp: antisym_def)
next
assume *: "(y2, y1) ∈ r"
with trans y2(3) have "(x2, y1) ∈ r" unfolding trans_def by blast
with False y1 y2 * antisym show ?thesis by (cases "x2 = y1") (auto simp: antisym_def)
qed
qed
qed
lemma Cinfinite_limit_finite: "[|finite X; X ⊆ Field r; Cinfinite r|]
==> ∃y ∈ Field r. ∀x ∈ X. (x ≠ y ∧ (x, y) ∈ r)"
proof (induct X rule: finite_induct)
case empty thus ?case unfolding cinfinite_def using ex_in_conv[of "Field r"] finite.emptyI by auto
next
case (insert x X)
then obtain y where y: "y ∈ Field r" "∀x ∈ X. (x ≠ y ∧ (x, y) ∈ r)" by blast
then obtain z where z: "z ∈ Field r" "x ≠ z ∧ (x, z) ∈ r" "y ≠ z ∧ (y, z) ∈ r"
using Cinfinite_limit2[OF _ y(1) insert(5), of x] insert(4) by blast
show ?case
apply (intro bexI ballI)
apply (erule insertE)
apply hypsubst
apply (rule z(2))
using Card_order_trans[OF insert(5)[THEN conjunct2]] y(2) z(3)
apply blast
apply (rule z(1))
done
qed
lemma insert_subsetI: "[|x ∈ A; X ⊆ A|] ==> insert x X ⊆ A"
by auto
lemma well_order_induct_imp:
"wo_rel r ==> (!!x. ∀y. y ≠ x ∧ (y, x) ∈ r --> y ∈ Field r --> P y ==> x ∈ Field r --> P x) ==>
x ∈ Field r --> P x"
by (erule wo_rel.well_order_induct)
lemma meta_spec2:
assumes "(!!x y. PROP P x y)"
shows "PROP P x y"
by (rule assms)
lemma nchotomy_relcomppE:
assumes "!!y. ∃x. y = f x" "(r OO s) a c" "!!b. r a (f b) ==> s (f b) c ==> P"
shows P
proof (rule relcompp.cases[OF assms(2)], hypsubst)
fix b assume "r a b" "s b c"
moreover from assms(1) obtain b' where "b = f b'" by blast
ultimately show P by (blast intro: assms(3))
qed
lemma vimage2p_rel_fun: "rel_fun (vimage2p f g R) R f g"
unfolding rel_fun_def vimage2p_def by auto
lemma predicate2D_vimage2p: "[|R ≤ vimage2p f g S; R x y|] ==> S (f x) (g y)"
unfolding vimage2p_def by auto
lemma id_transfer: "rel_fun A A id id"
unfolding rel_fun_def by simp
lemma ssubst_Pair_rhs: "[|(r, s) ∈ R; s' = s|] ==> (r, s') ∈ R"
by (rule ssubst)
ML_file "Tools/BNF/bnf_lfp_util.ML"
ML_file "Tools/BNF/bnf_lfp_tactics.ML"
ML_file "Tools/BNF/bnf_lfp.ML"
ML_file "Tools/BNF/bnf_lfp_compat.ML"
ML_file "Tools/BNF/bnf_lfp_rec_sugar_more.ML"
hide_fact (open) id_transfer
end