src/HOL/Nonstandard_Analysis/Free_Ultrafilter.thy
author wenzelm
Mon, 29 Feb 2016 22:34:36 +0100
changeset 62479 716336f19aa9
parent 61975 src/HOL/NSA/Free_Ultrafilter.thy@b4b11391c676
child 64435 c93b0e6131c3
permissions -rw-r--r--
clarified session; tuned headers;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
62479
716336f19aa9 clarified session;
wenzelm
parents: 61975
diff changeset
     1
(*  Title:      HOL/Nonstandard_Analysis/Free_Ultrafilter.thy
41589
bbd861837ebc tuned headers;
wenzelm
parents: 36176
diff changeset
     2
    Author:     Jacques D. Fleuriot, University of Cambridge
bbd861837ebc tuned headers;
wenzelm
parents: 36176
diff changeset
     3
    Author:     Lawrence C Paulson
bbd861837ebc tuned headers;
wenzelm
parents: 36176
diff changeset
     4
    Author:     Brian Huffman
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     5
*) 
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     6
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 60041
diff changeset
     7
section \<open>Filters and Ultrafilters\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     8
60036
218fcc645d22 move filters to their own theory
hoelzl
parents: 58878
diff changeset
     9
theory Free_Ultrafilter
55018
2a526bd279ed moved 'Zorn' into 'Main', since it's a BNF dependency
blanchet
parents: 52198
diff changeset
    10
imports "~~/src/HOL/Library/Infinite_Set"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    11
begin
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    12
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 60041
diff changeset
    13
subsection \<open>Definitions and basic properties\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    14
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 60041
diff changeset
    15
subsubsection \<open>Ultrafilters\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    16
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    17
locale ultrafilter =
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    18
  fixes F :: "'a filter"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    19
  assumes proper: "F \<noteq> bot"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    20
  assumes ultra: "eventually P F \<or> eventually (\<lambda>x. \<not> P x) F"
47486
4d49f3ffe97e replace locale 'UFT' with new un-named context block feature;
huffman
parents: 46008
diff changeset
    21
begin
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    22
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    23
lemma eventually_imp_frequently: "frequently P F \<Longrightarrow> eventually P F"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    24
  using ultra[of P] by (simp add: frequently_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    25
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    26
lemma frequently_eq_eventually: "frequently P F = eventually P F"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    27
  using eventually_imp_frequently eventually_frequently[OF proper] ..
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    28
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    29
lemma eventually_disj_iff: "eventually (\<lambda>x. P x \<or> Q x) F \<longleftrightarrow> eventually P F \<or> eventually Q F"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    30
  unfolding frequently_eq_eventually[symmetric] frequently_disj_iff ..
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    31
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    32
lemma eventually_all_iff: "eventually (\<lambda>x. \<forall>y. P x y) F = (\<forall>Y. eventually (\<lambda>x. P x (Y x)) F)"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    33
  using frequently_all[of P F] by (simp add: frequently_eq_eventually)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    34
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    35
lemma eventually_imp_iff: "eventually (\<lambda>x. P x \<longrightarrow> Q x) F \<longleftrightarrow> (eventually P F \<longrightarrow> eventually Q F)"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    36
  using frequently_imp_iff[of P Q F] by (simp add: frequently_eq_eventually)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    37
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    38
lemma eventually_iff_iff: "eventually (\<lambda>x. P x \<longleftrightarrow> Q x) F \<longleftrightarrow> (eventually P F \<longleftrightarrow> eventually Q F)"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    39
  unfolding iff_conv_conj_imp eventually_conj_iff eventually_imp_iff by simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    40
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    41
lemma eventually_not_iff: "eventually (\<lambda>x. \<not> P x) F \<longleftrightarrow> \<not> eventually P F"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    42
  unfolding not_eventually frequently_eq_eventually ..
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    43
47486
4d49f3ffe97e replace locale 'UFT' with new un-named context block feature;
huffman
parents: 46008
diff changeset
    44
end
4d49f3ffe97e replace locale 'UFT' with new un-named context block feature;
huffman
parents: 46008
diff changeset
    45
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 60041
diff changeset
    46
subsection \<open>Maximal filter = Ultrafilter\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    47
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 60041
diff changeset
    48
text \<open>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    49
   A filter F is an ultrafilter iff it is a maximal filter,
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    50
   i.e. whenever G is a filter and @{term "F \<subseteq> G"} then @{term "F = G"}
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 60041
diff changeset
    51
\<close>
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 60041
diff changeset
    52
text \<open>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    53
  Lemmas that shows existence of an extension to what was assumed to
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    54
  be a maximal filter. Will be used to derive contradiction in proof of
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    55
  property of ultrafilter.
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 60041
diff changeset
    56
\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    57
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    58
lemma extend_filter:
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    59
  "frequently P F \<Longrightarrow> inf F (principal {x. P x}) \<noteq> bot"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    60
  unfolding trivial_limit_def eventually_inf_principal by (simp add: not_eventually)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    61
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    62
lemma max_filter_ultrafilter:
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    63
  assumes proper: "F \<noteq> bot"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    64
  assumes max: "\<And>G. G \<noteq> bot \<Longrightarrow> G \<le> F \<Longrightarrow> F = G"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    65
  shows "ultrafilter F"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    66
proof
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    67
  fix P show "eventually P F \<or> (\<forall>\<^sub>Fx in F. \<not> P x)"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    68
  proof (rule disjCI)
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    69
    assume "\<not> (\<forall>\<^sub>Fx in F. \<not> P x)"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    70
    then have "inf F (principal {x. P x}) \<noteq> bot"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    71
      by (simp add: not_eventually extend_filter)
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    72
    then have F: "F = inf F (principal {x. P x})"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    73
      by (rule max) simp
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    74
    show "eventually P F"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    75
      by (subst F) (simp add: eventually_inf_principal)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    76
  qed
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    77
qed fact
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    78
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    79
lemma le_filter_frequently: "F \<le> G \<longleftrightarrow> (\<forall>P. frequently P F \<longrightarrow> frequently P G)"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    80
  unfolding frequently_def le_filter_def
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    81
  apply auto
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    82
  apply (erule_tac x="\<lambda>x. \<not> P x" in allE)
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    83
  apply auto
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    84
  done
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    85
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    86
lemma (in ultrafilter) max_filter:
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    87
  assumes G: "G \<noteq> bot" and sub: "G \<le> F" shows "F = G"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    88
proof (rule antisym)
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    89
  show "F \<le> G"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    90
    using sub
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    91
    by (auto simp: le_filter_frequently[of F] frequently_eq_eventually le_filter_def[of G]
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    92
             intro!: eventually_frequently G proper)
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    93
qed fact
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    94
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 60041
diff changeset
    95
subsection \<open>Ultrafilter Theorem\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    96
47486
4d49f3ffe97e replace locale 'UFT' with new un-named context block feature;
huffman
parents: 46008
diff changeset
    97
text "A local context makes proof of ultrafilter Theorem more modular"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    98
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
    99
lemma ex_max_ultrafilter:
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   100
  fixes F :: "'a filter"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   101
  assumes F: "F \<noteq> bot"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   102
  shows "\<exists>U\<le>F. ultrafilter U"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   103
proof -
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   104
  let ?X = "{G. G \<noteq> bot \<and> G \<le> F}"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   105
  let ?R = "{(b, a). a \<noteq> bot \<and> a \<le> b \<and> b \<le> F}"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   106
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   107
  have bot_notin_R: "\<And>c. c \<in> Chains ?R \<Longrightarrow> bot \<notin> c"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   108
    by (auto simp: Chains_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   109
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   110
  have [simp]: "Field ?R = ?X"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   111
    by (auto simp: Field_def bot_unique)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   112
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   113
  have "\<exists>m\<in>Field ?R. \<forall>a\<in>Field ?R. (m, a) \<in> ?R \<longrightarrow> a = m"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   114
  proof (rule Zorns_po_lemma)
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   115
    show "Partial_order ?R"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   116
      unfolding partial_order_on_def preorder_on_def
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   117
      by (auto simp: antisym_def refl_on_def trans_def Field_def bot_unique)
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   118
    show "\<forall>C\<in>Chains ?R. \<exists>u\<in>Field ?R. \<forall>a\<in>C. (a, u) \<in> ?R"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   119
    proof (safe intro!: bexI del: notI)
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   120
      fix c x assume c: "c \<in> Chains ?R"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   121
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   122
      { assume "c \<noteq> {}"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   123
        with c have "Inf c = bot \<longleftrightarrow> (\<exists>x\<in>c. x = bot)"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   124
          unfolding trivial_limit_def by (intro eventually_Inf_base) (auto simp: Chains_def)
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   125
        with c have 1: "Inf c \<noteq> bot"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   126
          by (simp add: bot_notin_R)
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 60041
diff changeset
   127
        from \<open>c \<noteq> {}\<close> obtain x where "x \<in> c" by auto
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   128
        with c have 2: "Inf c \<le> F"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   129
          by (auto intro!: Inf_lower2[of x] simp: Chains_def)
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   130
        note 1 2 }
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   131
      note Inf_c = this
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   132
      then have [simp]: "inf F (Inf c) = (if c = {} then F else Inf c)"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   133
        using c by (auto simp add: inf_absorb2)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   134
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   135
      show "inf F (Inf c) \<noteq> bot"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   136
        using c by (simp add: F Inf_c)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   137
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   138
      show "inf F (Inf c) \<in> Field ?R"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   139
        using c by (simp add: Chains_def Inf_c F)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   140
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   141
      assume x: "x \<in> c"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   142
      with c show "inf F (Inf c) \<le> x" "x \<le> F"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   143
        by (auto intro: Inf_lower simp: Chains_def)
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   144
    qed
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   145
  qed
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   146
  then guess U ..
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   147
  then show ?thesis
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   148
    by (intro exI[of _ U] conjI max_filter_ultrafilter) auto
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   149
qed
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   150
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 60041
diff changeset
   151
subsubsection \<open>Free Ultrafilters\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   152
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 60041
diff changeset
   153
text \<open>There exists a free ultrafilter on any infinite set\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   154
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   155
locale freeultrafilter = ultrafilter +
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   156
  assumes infinite: "eventually P F \<Longrightarrow> infinite {x. P x}"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   157
begin
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   158
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   159
lemma finite: "finite {x. P x} \<Longrightarrow> \<not> eventually P F"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   160
  by (erule contrapos_pn, erule infinite)
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   161
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   162
lemma finite': "finite {x. \<not> P x} \<Longrightarrow> eventually P F"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   163
  by (drule finite) (simp add: not_eventually frequently_eq_eventually)
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   164
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   165
lemma le_cofinite: "F \<le> cofinite"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   166
  by (intro filter_leI)
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   167
     (auto simp add: eventually_cofinite not_eventually frequently_eq_eventually dest!: finite)
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   168
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   169
lemma singleton: "\<not> eventually (\<lambda>x. x = a) F"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   170
by (rule finite, simp)
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   171
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   172
lemma singleton': "\<not> eventually (op = a) F"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   173
by (rule finite, simp)
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   174
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   175
lemma ultrafilter: "ultrafilter F" ..
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   176
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   177
end
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   178
47486
4d49f3ffe97e replace locale 'UFT' with new un-named context block feature;
huffman
parents: 46008
diff changeset
   179
lemma freeultrafilter_Ex:
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   180
  assumes [simp]: "infinite (UNIV :: 'a set)"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   181
  shows "\<exists>U::'a filter. freeultrafilter U"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   182
proof -
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   183
  from ex_max_ultrafilter[of "cofinite :: 'a filter"]
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   184
  obtain U :: "'a filter" where "U \<le> cofinite" "ultrafilter U"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   185
    by auto
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   186
  interpret ultrafilter U by fact
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   187
  have "freeultrafilter U"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   188
  proof
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   189
    fix P assume "eventually P U"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   190
    with proper have "frequently P U"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   191
      by (rule eventually_frequently)
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   192
    then have "frequently P cofinite"
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 60041
diff changeset
   193
      using \<open>U \<le> cofinite\<close> by (simp add: le_filter_frequently)
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   194
    then show "infinite {x. P x}"
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 60036
diff changeset
   195
      by (simp add: frequently_cofinite)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   196
  qed
46008
c296c75f4cf4 reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
wenzelm
parents: 41959
diff changeset
   197
  then show ?thesis ..
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   198
qed
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   199
47486
4d49f3ffe97e replace locale 'UFT' with new un-named context block feature;
huffman
parents: 46008
diff changeset
   200
end