src/HOLCF/Eventual.thy
author wenzelm
Thu, 01 Oct 2009 16:09:47 +0200
changeset 32813 dac196e23093
parent 27408 22a515a55bf5
child 35771 2b75230f272f
permissions -rw-r--r--
handle Pattern.MATCH, not arbitrary exceptions; misc tuning, trying to reduce line length a bit;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
27408
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
     1
theory Eventual
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
     2
imports Infinite_Set
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
     3
begin
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
     4
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
     5
subsection {* Lemmas about MOST *}
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
     6
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
     7
lemma MOST_INFM:
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
     8
  assumes inf: "infinite (UNIV::'a set)"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
     9
  shows "MOST x::'a. P x \<Longrightarrow> INFM x::'a. P x"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    10
  unfolding Alm_all_def Inf_many_def
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    11
  apply (auto simp add: Collect_neg_eq)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    12
  apply (drule (1) finite_UnI)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    13
  apply (simp add: Compl_partition2 inf)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    14
  done
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    15
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    16
lemma MOST_comp: "\<lbrakk>inj f; MOST x. P x\<rbrakk> \<Longrightarrow> MOST x. P (f x)"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    17
unfolding MOST_iff_finiteNeg
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    18
by (drule (1) finite_vimageI, simp)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    19
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    20
lemma INFM_comp: "\<lbrakk>inj f; INFM x. P (f x)\<rbrakk> \<Longrightarrow> INFM x. P x"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    21
unfolding Inf_many_def
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    22
by (clarify, drule (1) finite_vimageI, simp)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    23
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    24
lemma MOST_SucI: "MOST n. P n \<Longrightarrow> MOST n. P (Suc n)"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    25
by (rule MOST_comp [OF inj_Suc])
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    26
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    27
lemma MOST_SucD: "MOST n. P (Suc n) \<Longrightarrow> MOST n. P n"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    28
unfolding MOST_nat
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    29
apply (clarify, rule_tac x="Suc m" in exI, clarify)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    30
apply (erule Suc_lessE, simp)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    31
done
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    32
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    33
lemma MOST_Suc_iff: "(MOST n. P (Suc n)) \<longleftrightarrow> (MOST n. P n)"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    34
by (rule iffI [OF MOST_SucD MOST_SucI])
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    35
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    36
lemma INFM_finite_Bex_distrib:
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    37
  "finite A \<Longrightarrow> (INFM y. \<exists>x\<in>A. P x y) \<longleftrightarrow> (\<exists>x\<in>A. INFM y. P x y)"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    38
by (induct set: finite, simp, simp add: INFM_disj_distrib)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    39
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    40
lemma MOST_finite_Ball_distrib:
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    41
  "finite A \<Longrightarrow> (MOST y. \<forall>x\<in>A. P x y) \<longleftrightarrow> (\<forall>x\<in>A. MOST y. P x y)"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    42
by (induct set: finite, simp, simp add: MOST_conj_distrib)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    43
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    44
lemma MOST_ge_nat: "MOST n::nat. m \<le> n"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    45
unfolding MOST_nat_le by fast
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    46
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    47
subsection {* Eventually constant sequences *}
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    48
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    49
definition
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    50
  eventually_constant :: "(nat \<Rightarrow> 'a) \<Rightarrow> bool"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    51
where
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    52
  "eventually_constant S = (\<exists>x. MOST i. S i = x)"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    53
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    54
lemma eventually_constant_MOST_MOST:
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    55
  "eventually_constant S \<longleftrightarrow> (MOST m. MOST n. S n = S m)"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    56
unfolding eventually_constant_def MOST_nat
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    57
apply safe
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    58
apply (rule_tac x=m in exI, clarify)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    59
apply (rule_tac x=m in exI, clarify)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    60
apply simp
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    61
apply fast
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    62
done
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    63
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    64
lemma eventually_constantI: "MOST i. S i = x \<Longrightarrow> eventually_constant S"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    65
unfolding eventually_constant_def by fast
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    66
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    67
lemma eventually_constant_comp:
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    68
  "eventually_constant (\<lambda>i. S i) \<Longrightarrow> eventually_constant (\<lambda>i. f (S i))"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    69
unfolding eventually_constant_def
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    70
apply (erule exE, rule_tac x="f x" in exI)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    71
apply (erule MOST_mono, simp)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    72
done
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    73
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    74
lemma eventually_constant_Suc_iff:
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    75
  "eventually_constant (\<lambda>i. S (Suc i)) \<longleftrightarrow> eventually_constant (\<lambda>i. S i)"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    76
unfolding eventually_constant_def
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    77
by (subst MOST_Suc_iff, rule refl)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    78
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    79
lemma eventually_constant_SucD:
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    80
  "eventually_constant (\<lambda>i. S (Suc i)) \<Longrightarrow> eventually_constant (\<lambda>i. S i)"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    81
by (rule eventually_constant_Suc_iff [THEN iffD1])
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    82
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    83
subsection {* Limits of eventually constant sequences *}
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    84
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    85
definition
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    86
  eventual :: "(nat \<Rightarrow> 'a) \<Rightarrow> 'a" where
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    87
  "eventual S = (THE x. MOST i. S i = x)"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    88
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    89
lemma eventual_eqI: "MOST i. S i = x \<Longrightarrow> eventual S = x"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    90
unfolding eventual_def
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    91
apply (rule the_equality, assumption)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    92
apply (rename_tac y)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    93
apply (subgoal_tac "MOST i::nat. y = x", simp)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    94
apply (erule MOST_rev_mp)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    95
apply (erule MOST_rev_mp)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    96
apply simp
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    97
done
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    98
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
    99
lemma MOST_eq_eventual:
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   100
  "eventually_constant S \<Longrightarrow> MOST i. S i = eventual S"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   101
unfolding eventually_constant_def
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   102
by (erule exE, simp add: eventual_eqI)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   103
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   104
lemma eventual_mem_range:
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   105
  "eventually_constant S \<Longrightarrow> eventual S \<in> range S"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   106
apply (drule MOST_eq_eventual)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   107
apply (simp only: MOST_nat_le, clarify)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   108
apply (drule spec, drule mp, rule order_refl)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   109
apply (erule range_eqI [OF sym])
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   110
done
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   111
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   112
lemma eventually_constant_MOST_iff:
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   113
  assumes S: "eventually_constant S"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   114
  shows "(MOST n. P (S n)) \<longleftrightarrow> P (eventual S)"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   115
apply (subgoal_tac "(MOST n. P (S n)) \<longleftrightarrow> (MOST n::nat. P (eventual S))")
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   116
apply simp
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   117
apply (rule iffI)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   118
apply (rule MOST_rev_mp [OF MOST_eq_eventual [OF S]])
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   119
apply (erule MOST_mono, force)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   120
apply (rule MOST_rev_mp [OF MOST_eq_eventual [OF S]])
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   121
apply (erule MOST_mono, simp)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   122
done
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   123
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   124
lemma MOST_eventual:
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   125
  "\<lbrakk>eventually_constant S; MOST n. P (S n)\<rbrakk> \<Longrightarrow> P (eventual S)"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   126
proof -
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   127
  assume "eventually_constant S"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   128
  hence "MOST n. S n = eventual S"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   129
    by (rule MOST_eq_eventual)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   130
  moreover assume "MOST n. P (S n)"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   131
  ultimately have "MOST n. S n = eventual S \<and> P (S n)"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   132
    by (rule MOST_conj_distrib [THEN iffD2, OF conjI])
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   133
  hence "MOST n::nat. P (eventual S)"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   134
    by (rule MOST_mono) auto
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   135
  thus ?thesis by simp
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   136
qed
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   137
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   138
lemma eventually_constant_MOST_Suc_eq:
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   139
  "eventually_constant S \<Longrightarrow> MOST n. S (Suc n) = S n"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   140
apply (drule MOST_eq_eventual)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   141
apply (frule MOST_Suc_iff [THEN iffD2])
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   142
apply (erule MOST_rev_mp)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   143
apply (erule MOST_rev_mp)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   144
apply simp
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   145
done
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   146
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   147
lemma eventual_comp:
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   148
  "eventually_constant S \<Longrightarrow> eventual (\<lambda>i. f (S i)) = f (eventual (\<lambda>i. S i))"
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   149
apply (rule eventual_eqI)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   150
apply (rule MOST_mono)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   151
apply (erule MOST_eq_eventual)
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   152
apply simp
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   153
done
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   154
22a515a55bf5 theory of eventually-constant sequences
huffman
parents:
diff changeset
   155
end