# HG changeset patch # User huffman # Date 1214868316 -7200 # Node ID 22a515a55bf5f36addb3cecf39ec21c72455f86e # Parent 68e111812b832354965ffe9c87785d3528fe61fb theory of eventually-constant sequences diff -r 68e111812b83 -r 22a515a55bf5 src/HOLCF/Eventual.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/Eventual.thy Tue Jul 01 01:25:16 2008 +0200 @@ -0,0 +1,155 @@ +theory Eventual +imports Infinite_Set +begin + +subsection {* Lemmas about MOST *} + +lemma MOST_INFM: + assumes inf: "infinite (UNIV::'a set)" + shows "MOST x::'a. P x \ INFM x::'a. P x" + unfolding Alm_all_def Inf_many_def + apply (auto simp add: Collect_neg_eq) + apply (drule (1) finite_UnI) + apply (simp add: Compl_partition2 inf) + done + +lemma MOST_comp: "\inj f; MOST x. P x\ \ MOST x. P (f x)" +unfolding MOST_iff_finiteNeg +by (drule (1) finite_vimageI, simp) + +lemma INFM_comp: "\inj f; INFM x. P (f x)\ \ INFM x. P x" +unfolding Inf_many_def +by (clarify, drule (1) finite_vimageI, simp) + +lemma MOST_SucI: "MOST n. P n \ MOST n. P (Suc n)" +by (rule MOST_comp [OF inj_Suc]) + +lemma MOST_SucD: "MOST n. P (Suc n) \ MOST n. P n" +unfolding MOST_nat +apply (clarify, rule_tac x="Suc m" in exI, clarify) +apply (erule Suc_lessE, simp) +done + +lemma MOST_Suc_iff: "(MOST n. P (Suc n)) \ (MOST n. P n)" +by (rule iffI [OF MOST_SucD MOST_SucI]) + +lemma INFM_finite_Bex_distrib: + "finite A \ (INFM y. \x\A. P x y) \ (\x\A. INFM y. P x y)" +by (induct set: finite, simp, simp add: INFM_disj_distrib) + +lemma MOST_finite_Ball_distrib: + "finite A \ (MOST y. \x\A. P x y) \ (\x\A. MOST y. P x y)" +by (induct set: finite, simp, simp add: MOST_conj_distrib) + +lemma MOST_ge_nat: "MOST n::nat. m \ n" +unfolding MOST_nat_le by fast + +subsection {* Eventually constant sequences *} + +definition + eventually_constant :: "(nat \ 'a) \ bool" +where + "eventually_constant S = (\x. MOST i. S i = x)" + +lemma eventually_constant_MOST_MOST: + "eventually_constant S \ (MOST m. MOST n. S n = S m)" +unfolding eventually_constant_def MOST_nat +apply safe +apply (rule_tac x=m in exI, clarify) +apply (rule_tac x=m in exI, clarify) +apply simp +apply fast +done + +lemma eventually_constantI: "MOST i. S i = x \ eventually_constant S" +unfolding eventually_constant_def by fast + +lemma eventually_constant_comp: + "eventually_constant (\i. S i) \ eventually_constant (\i. f (S i))" +unfolding eventually_constant_def +apply (erule exE, rule_tac x="f x" in exI) +apply (erule MOST_mono, simp) +done + +lemma eventually_constant_Suc_iff: + "eventually_constant (\i. S (Suc i)) \ eventually_constant (\i. S i)" +unfolding eventually_constant_def +by (subst MOST_Suc_iff, rule refl) + +lemma eventually_constant_SucD: + "eventually_constant (\i. S (Suc i)) \ eventually_constant (\i. S i)" +by (rule eventually_constant_Suc_iff [THEN iffD1]) + +subsection {* Limits of eventually constant sequences *} + +definition + eventual :: "(nat \ 'a) \ 'a" where + "eventual S = (THE x. MOST i. S i = x)" + +lemma eventual_eqI: "MOST i. S i = x \ eventual S = x" +unfolding eventual_def +apply (rule the_equality, assumption) +apply (rename_tac y) +apply (subgoal_tac "MOST i::nat. y = x", simp) +apply (erule MOST_rev_mp) +apply (erule MOST_rev_mp) +apply simp +done + +lemma MOST_eq_eventual: + "eventually_constant S \ MOST i. S i = eventual S" +unfolding eventually_constant_def +by (erule exE, simp add: eventual_eqI) + +lemma eventual_mem_range: + "eventually_constant S \ eventual S \ range S" +apply (drule MOST_eq_eventual) +apply (simp only: MOST_nat_le, clarify) +apply (drule spec, drule mp, rule order_refl) +apply (erule range_eqI [OF sym]) +done + +lemma eventually_constant_MOST_iff: + assumes S: "eventually_constant S" + shows "(MOST n. P (S n)) \ P (eventual S)" +apply (subgoal_tac "(MOST n. P (S n)) \ (MOST n::nat. P (eventual S))") +apply simp +apply (rule iffI) +apply (rule MOST_rev_mp [OF MOST_eq_eventual [OF S]]) +apply (erule MOST_mono, force) +apply (rule MOST_rev_mp [OF MOST_eq_eventual [OF S]]) +apply (erule MOST_mono, simp) +done + +lemma MOST_eventual: + "\eventually_constant S; MOST n. P (S n)\ \ P (eventual S)" +proof - + assume "eventually_constant S" + hence "MOST n. S n = eventual S" + by (rule MOST_eq_eventual) + moreover assume "MOST n. P (S n)" + ultimately have "MOST n. S n = eventual S \ P (S n)" + by (rule MOST_conj_distrib [THEN iffD2, OF conjI]) + hence "MOST n::nat. P (eventual S)" + by (rule MOST_mono) auto + thus ?thesis by simp +qed + +lemma eventually_constant_MOST_Suc_eq: + "eventually_constant S \ MOST n. S (Suc n) = S n" +apply (drule MOST_eq_eventual) +apply (frule MOST_Suc_iff [THEN iffD2]) +apply (erule MOST_rev_mp) +apply (erule MOST_rev_mp) +apply simp +done + +lemma eventual_comp: + "eventually_constant S \ eventual (\i. f (S i)) = f (eventual (\i. S i))" +apply (rule eventual_eqI) +apply (rule MOST_mono) +apply (erule MOST_eq_eventual) +apply simp +done + +end