src/HOL/Hilbert_Choice.thy
author haftmann
Sun, 10 Nov 2013 15:05:06 +0100
changeset 55747 45a5523d4a63
parent 53280 36ffe23b25f8
child 55951 9387251b6a46
permissions -rw-r--r--
qualifed popular user space names
     1 (*  Title:      HOL/Hilbert_Choice.thy
     2     Author:     Lawrence C Paulson, Tobias Nipkow
     3     Copyright   2001  University of Cambridge
     4 *)
     5 
     6 header {* Hilbert's Epsilon-Operator and the Axiom of Choice *}
     7 
     8 theory Hilbert_Choice
     9 imports Nat Wellfounded Big_Operators
    10 keywords "specification" "ax_specification" :: thy_goal
    11 begin
    12 
    13 subsection {* Hilbert's epsilon *}
    14 
    15 axiomatization Eps :: "('a => bool) => 'a" where
    16   someI: "P x ==> P (Eps P)"
    17 
    18 syntax (epsilon)
    19   "_Eps"        :: "[pttrn, bool] => 'a"    ("(3\<some>_./ _)" [0, 10] 10)
    20 syntax (HOL)
    21   "_Eps"        :: "[pttrn, bool] => 'a"    ("(3@ _./ _)" [0, 10] 10)
    22 syntax
    23   "_Eps"        :: "[pttrn, bool] => 'a"    ("(3SOME _./ _)" [0, 10] 10)
    24 translations
    25   "SOME x. P" == "CONST Eps (%x. P)"
    26 
    27 print_translation {*
    28   [(@{const_syntax Eps}, fn _ => fn [Abs abs] =>
    29       let val (x, t) = Syntax_Trans.atomic_abs_tr' abs
    30       in Syntax.const @{syntax_const "_Eps"} $ x $ t end)]
    31 *} -- {* to avoid eta-contraction of body *}
    32 
    33 definition inv_into :: "'a set => ('a => 'b) => ('b => 'a)" where
    34 "inv_into A f == %x. SOME y. y : A & f y = x"
    35 
    36 abbreviation inv :: "('a => 'b) => ('b => 'a)" where
    37 "inv == inv_into UNIV"
    38 
    39 
    40 subsection {*Hilbert's Epsilon-operator*}
    41 
    42 text{*Easier to apply than @{text someI} if the witness comes from an
    43 existential formula*}
    44 lemma someI_ex [elim?]: "\<exists>x. P x ==> P (SOME x. P x)"
    45 apply (erule exE)
    46 apply (erule someI)
    47 done
    48 
    49 text{*Easier to apply than @{text someI} because the conclusion has only one
    50 occurrence of @{term P}.*}
    51 lemma someI2: "[| P a;  !!x. P x ==> Q x |] ==> Q (SOME x. P x)"
    52 by (blast intro: someI)
    53 
    54 text{*Easier to apply than @{text someI2} if the witness comes from an
    55 existential formula*}
    56 lemma someI2_ex: "[| \<exists>a. P a; !!x. P x ==> Q x |] ==> Q (SOME x. P x)"
    57 by (blast intro: someI2)
    58 
    59 lemma some_equality [intro]:
    60      "[| P a;  !!x. P x ==> x=a |] ==> (SOME x. P x) = a"
    61 by (blast intro: someI2)
    62 
    63 lemma some1_equality: "[| EX!x. P x; P a |] ==> (SOME x. P x) = a"
    64 by blast
    65 
    66 lemma some_eq_ex: "P (SOME x. P x) =  (\<exists>x. P x)"
    67 by (blast intro: someI)
    68 
    69 lemma some_eq_trivial [simp]: "(SOME y. y=x) = x"
    70 apply (rule some_equality)
    71 apply (rule refl, assumption)
    72 done
    73 
    74 lemma some_sym_eq_trivial [simp]: "(SOME y. x=y) = x"
    75 apply (rule some_equality)
    76 apply (rule refl)
    77 apply (erule sym)
    78 done
    79 
    80 
    81 subsection{*Axiom of Choice, Proved Using the Description Operator*}
    82 
    83 lemma choice: "\<forall>x. \<exists>y. Q x y ==> \<exists>f. \<forall>x. Q x (f x)"
    84 by (fast elim: someI)
    85 
    86 lemma bchoice: "\<forall>x\<in>S. \<exists>y. Q x y ==> \<exists>f. \<forall>x\<in>S. Q x (f x)"
    87 by (fast elim: someI)
    88 
    89 lemma choice_iff: "(\<forall>x. \<exists>y. Q x y) \<longleftrightarrow> (\<exists>f. \<forall>x. Q x (f x))"
    90 by (fast elim: someI)
    91 
    92 lemma choice_iff': "(\<forall>x. P x \<longrightarrow> (\<exists>y. Q x y)) \<longleftrightarrow> (\<exists>f. \<forall>x. P x \<longrightarrow> Q x (f x))"
    93 by (fast elim: someI)
    94 
    95 lemma bchoice_iff: "(\<forall>x\<in>S. \<exists>y. Q x y) \<longleftrightarrow> (\<exists>f. \<forall>x\<in>S. Q x (f x))"
    96 by (fast elim: someI)
    97 
    98 lemma bchoice_iff': "(\<forall>x\<in>S. P x \<longrightarrow> (\<exists>y. Q x y)) \<longleftrightarrow> (\<exists>f. \<forall>x\<in>S. P x \<longrightarrow> Q x (f x))"
    99 by (fast elim: someI)
   100 
   101 subsection {*Function Inverse*}
   102 
   103 lemma inv_def: "inv f = (%y. SOME x. f x = y)"
   104 by(simp add: inv_into_def)
   105 
   106 lemma inv_into_into: "x : f ` A ==> inv_into A f x : A"
   107 apply (simp add: inv_into_def)
   108 apply (fast intro: someI2)
   109 done
   110 
   111 lemma inv_id [simp]: "inv id = id"
   112 by (simp add: inv_into_def id_def)
   113 
   114 lemma inv_into_f_f [simp]:
   115   "[| inj_on f A;  x : A |] ==> inv_into A f (f x) = x"
   116 apply (simp add: inv_into_def inj_on_def)
   117 apply (blast intro: someI2)
   118 done
   119 
   120 lemma inv_f_f: "inj f ==> inv f (f x) = x"
   121 by simp
   122 
   123 lemma f_inv_into_f: "y : f`A  ==> f (inv_into A f y) = y"
   124 apply (simp add: inv_into_def)
   125 apply (fast intro: someI2)
   126 done
   127 
   128 lemma inv_into_f_eq: "[| inj_on f A; x : A; f x = y |] ==> inv_into A f y = x"
   129 apply (erule subst)
   130 apply (fast intro: inv_into_f_f)
   131 done
   132 
   133 lemma inv_f_eq: "[| inj f; f x = y |] ==> inv f y = x"
   134 by (simp add:inv_into_f_eq)
   135 
   136 lemma inj_imp_inv_eq: "[| inj f; ALL x. f(g x) = x |] ==> inv f = g"
   137   by (blast intro: inv_into_f_eq)
   138 
   139 text{*But is it useful?*}
   140 lemma inj_transfer:
   141   assumes injf: "inj f" and minor: "!!y. y \<in> range(f) ==> P(inv f y)"
   142   shows "P x"
   143 proof -
   144   have "f x \<in> range f" by auto
   145   hence "P(inv f (f x))" by (rule minor)
   146   thus "P x" by (simp add: inv_into_f_f [OF injf])
   147 qed
   148 
   149 lemma inj_iff: "(inj f) = (inv f o f = id)"
   150 apply (simp add: o_def fun_eq_iff)
   151 apply (blast intro: inj_on_inverseI inv_into_f_f)
   152 done
   153 
   154 lemma inv_o_cancel[simp]: "inj f ==> inv f o f = id"
   155 by (simp add: inj_iff)
   156 
   157 lemma o_inv_o_cancel[simp]: "inj f ==> g o inv f o f = g"
   158 by (simp add: comp_assoc)
   159 
   160 lemma inv_into_image_cancel[simp]:
   161   "inj_on f A ==> S <= A ==> inv_into A f ` f ` S = S"
   162 by(fastforce simp: image_def)
   163 
   164 lemma inj_imp_surj_inv: "inj f ==> surj (inv f)"
   165 by (blast intro!: surjI inv_into_f_f)
   166 
   167 lemma surj_f_inv_f: "surj f ==> f(inv f y) = y"
   168 by (simp add: f_inv_into_f)
   169 
   170 lemma inv_into_injective:
   171   assumes eq: "inv_into A f x = inv_into A f y"
   172       and x: "x: f`A"
   173       and y: "y: f`A"
   174   shows "x=y"
   175 proof -
   176   have "f (inv_into A f x) = f (inv_into A f y)" using eq by simp
   177   thus ?thesis by (simp add: f_inv_into_f x y)
   178 qed
   179 
   180 lemma inj_on_inv_into: "B <= f`A ==> inj_on (inv_into A f) B"
   181 by (blast intro: inj_onI dest: inv_into_injective injD)
   182 
   183 lemma bij_betw_inv_into: "bij_betw f A B ==> bij_betw (inv_into A f) B A"
   184 by (auto simp add: bij_betw_def inj_on_inv_into)
   185 
   186 lemma surj_imp_inj_inv: "surj f ==> inj (inv f)"
   187 by (simp add: inj_on_inv_into)
   188 
   189 lemma surj_iff: "(surj f) = (f o inv f = id)"
   190 by (auto intro!: surjI simp: surj_f_inv_f fun_eq_iff[where 'b='a])
   191 
   192 lemma surj_iff_all: "surj f \<longleftrightarrow> (\<forall>x. f (inv f x) = x)"
   193   unfolding surj_iff by (simp add: o_def fun_eq_iff)
   194 
   195 lemma surj_imp_inv_eq: "[| surj f; \<forall>x. g(f x) = x |] ==> inv f = g"
   196 apply (rule ext)
   197 apply (drule_tac x = "inv f x" in spec)
   198 apply (simp add: surj_f_inv_f)
   199 done
   200 
   201 lemma bij_imp_bij_inv: "bij f ==> bij (inv f)"
   202 by (simp add: bij_def inj_imp_surj_inv surj_imp_inj_inv)
   203 
   204 lemma inv_equality: "[| !!x. g (f x) = x;  !!y. f (g y) = y |] ==> inv f = g"
   205 apply (rule ext)
   206 apply (auto simp add: inv_into_def)
   207 done
   208 
   209 lemma inv_inv_eq: "bij f ==> inv (inv f) = f"
   210 apply (rule inv_equality)
   211 apply (auto simp add: bij_def surj_f_inv_f)
   212 done
   213 
   214 (** bij(inv f) implies little about f.  Consider f::bool=>bool such that
   215     f(True)=f(False)=True.  Then it's consistent with axiom someI that
   216     inv f could be any function at all, including the identity function.
   217     If inv f=id then inv f is a bijection, but inj f, surj(f) and
   218     inv(inv f)=f all fail.
   219 **)
   220 
   221 lemma inv_into_comp:
   222   "[| inj_on f (g ` A); inj_on g A; x : f ` g ` A |] ==>
   223   inv_into A (f o g) x = (inv_into A g o inv_into (g ` A) f) x"
   224 apply (rule inv_into_f_eq)
   225   apply (fast intro: comp_inj_on)
   226  apply (simp add: inv_into_into)
   227 apply (simp add: f_inv_into_f inv_into_into)
   228 done
   229 
   230 lemma o_inv_distrib: "[| bij f; bij g |] ==> inv (f o g) = inv g o inv f"
   231 apply (rule inv_equality)
   232 apply (auto simp add: bij_def surj_f_inv_f)
   233 done
   234 
   235 lemma image_surj_f_inv_f: "surj f ==> f ` (inv f ` A) = A"
   236 by (simp add: image_eq_UN surj_f_inv_f)
   237 
   238 lemma image_inv_f_f: "inj f ==> (inv f) ` (f ` A) = A"
   239 by (simp add: image_eq_UN)
   240 
   241 lemma inv_image_comp: "inj f ==> inv f ` (f`X) = X"
   242 by (auto simp add: image_def)
   243 
   244 lemma bij_image_Collect_eq: "bij f ==> f ` Collect P = {y. P (inv f y)}"
   245 apply auto
   246 apply (force simp add: bij_is_inj)
   247 apply (blast intro: bij_is_surj [THEN surj_f_inv_f, symmetric])
   248 done
   249 
   250 lemma bij_vimage_eq_inv_image: "bij f ==> f -` A = inv f ` A" 
   251 apply (auto simp add: bij_is_surj [THEN surj_f_inv_f])
   252 apply (blast intro: bij_is_inj [THEN inv_into_f_f, symmetric])
   253 done
   254 
   255 lemma finite_fun_UNIVD1:
   256   assumes fin: "finite (UNIV :: ('a \<Rightarrow> 'b) set)"
   257   and card: "card (UNIV :: 'b set) \<noteq> Suc 0"
   258   shows "finite (UNIV :: 'a set)"
   259 proof -
   260   from fin have finb: "finite (UNIV :: 'b set)" by (rule finite_fun_UNIVD2)
   261   with card have "card (UNIV :: 'b set) \<ge> Suc (Suc 0)"
   262     by (cases "card (UNIV :: 'b set)") (auto simp add: card_eq_0_iff)
   263   then obtain n where "card (UNIV :: 'b set) = Suc (Suc n)" "n = card (UNIV :: 'b set) - Suc (Suc 0)" by auto
   264   then obtain b1 b2 where b1b2: "(b1 :: 'b) \<noteq> (b2 :: 'b)" by (auto simp add: card_Suc_eq)
   265   from fin have "finite (range (\<lambda>f :: 'a \<Rightarrow> 'b. inv f b1))" by (rule finite_imageI)
   266   moreover have "UNIV = range (\<lambda>f :: 'a \<Rightarrow> 'b. inv f b1)"
   267   proof (rule UNIV_eq_I)
   268     fix x :: 'a
   269     from b1b2 have "x = inv (\<lambda>y. if y = x then b1 else b2) b1" by (simp add: inv_into_def)
   270     thus "x \<in> range (\<lambda>f\<Colon>'a \<Rightarrow> 'b. inv f b1)" by blast
   271   qed
   272   ultimately show "finite (UNIV :: 'a set)" by simp
   273 qed
   274 
   275 lemma image_inv_into_cancel:
   276   assumes SURJ: "f`A=A'" and SUB: "B' \<le> A'"
   277   shows "f `((inv_into A f)`B') = B'"
   278   using assms
   279 proof (auto simp add: f_inv_into_f)
   280   let ?f' = "(inv_into A f)"
   281   fix a' assume *: "a' \<in> B'"
   282   then have "a' \<in> A'" using SUB by auto
   283   then have "a' = f (?f' a')"
   284     using SURJ by (auto simp add: f_inv_into_f)
   285   then show "a' \<in> f ` (?f' ` B')" using * by blast
   286 qed
   287 
   288 lemma inv_into_inv_into_eq:
   289   assumes "bij_betw f A A'" "a \<in> A"
   290   shows "inv_into A' (inv_into A f) a = f a"
   291 proof -
   292   let ?f' = "inv_into A f"   let ?f'' = "inv_into A' ?f'"
   293   have 1: "bij_betw ?f' A' A" using assms
   294   by (auto simp add: bij_betw_inv_into)
   295   obtain a' where 2: "a' \<in> A'" and 3: "?f' a' = a"
   296     using 1 `a \<in> A` unfolding bij_betw_def by force
   297   hence "?f'' a = a'"
   298     using `a \<in> A` 1 3 by (auto simp add: f_inv_into_f bij_betw_def)
   299   moreover have "f a = a'" using assms 2 3
   300     by (auto simp add: bij_betw_def)
   301   ultimately show "?f'' a = f a" by simp
   302 qed
   303 
   304 lemma inj_on_iff_surj:
   305   assumes "A \<noteq> {}"
   306   shows "(\<exists>f. inj_on f A \<and> f ` A \<le> A') \<longleftrightarrow> (\<exists>g. g ` A' = A)"
   307 proof safe
   308   fix f assume INJ: "inj_on f A" and INCL: "f ` A \<le> A'"
   309   let ?phi = "\<lambda>a' a. a \<in> A \<and> f a = a'"  let ?csi = "\<lambda>a. a \<in> A"
   310   let ?g = "\<lambda>a'. if a' \<in> f ` A then (SOME a. ?phi a' a) else (SOME a. ?csi a)"
   311   have "?g ` A' = A"
   312   proof
   313     show "?g ` A' \<le> A"
   314     proof clarify
   315       fix a' assume *: "a' \<in> A'"
   316       show "?g a' \<in> A"
   317       proof cases
   318         assume Case1: "a' \<in> f ` A"
   319         then obtain a where "?phi a' a" by blast
   320         hence "?phi a' (SOME a. ?phi a' a)" using someI[of "?phi a'" a] by blast
   321         with Case1 show ?thesis by auto
   322       next
   323         assume Case2: "a' \<notin> f ` A"
   324         hence "?csi (SOME a. ?csi a)" using assms someI_ex[of ?csi] by blast
   325         with Case2 show ?thesis by auto
   326       qed
   327     qed
   328   next
   329     show "A \<le> ?g ` A'"
   330     proof-
   331       {fix a assume *: "a \<in> A"
   332        let ?b = "SOME aa. ?phi (f a) aa"
   333        have "?phi (f a) a" using * by auto
   334        hence 1: "?phi (f a) ?b" using someI[of "?phi(f a)" a] by blast
   335        hence "?g(f a) = ?b" using * by auto
   336        moreover have "a = ?b" using 1 INJ * by (auto simp add: inj_on_def)
   337        ultimately have "?g(f a) = a" by simp
   338        with INCL * have "?g(f a) = a \<and> f a \<in> A'" by auto
   339       }
   340       thus ?thesis by force
   341     qed
   342   qed
   343   thus "\<exists>g. g ` A' = A" by blast
   344 next
   345   fix g  let ?f = "inv_into A' g"
   346   have "inj_on ?f (g ` A')"
   347     by (auto simp add: inj_on_inv_into)
   348   moreover
   349   {fix a' assume *: "a' \<in> A'"
   350    let ?phi = "\<lambda> b'. b' \<in> A' \<and> g b' = g a'"
   351    have "?phi a'" using * by auto
   352    hence "?phi(SOME b'. ?phi b')" using someI[of ?phi] by blast
   353    hence "?f(g a') \<in> A'" unfolding inv_into_def by auto
   354   }
   355   ultimately show "\<exists>f. inj_on f (g ` A') \<and> f ` g ` A' \<subseteq> A'" by auto
   356 qed
   357 
   358 lemma Ex_inj_on_UNION_Sigma:
   359   "\<exists>f. (inj_on f (\<Union> i \<in> I. A i) \<and> f ` (\<Union> i \<in> I. A i) \<le> (SIGMA i : I. A i))"
   360 proof
   361   let ?phi = "\<lambda> a i. i \<in> I \<and> a \<in> A i"
   362   let ?sm = "\<lambda> a. SOME i. ?phi a i"
   363   let ?f = "\<lambda>a. (?sm a, a)"
   364   have "inj_on ?f (\<Union> i \<in> I. A i)" unfolding inj_on_def by auto
   365   moreover
   366   { { fix i a assume "i \<in> I" and "a \<in> A i"
   367       hence "?sm a \<in> I \<and> a \<in> A(?sm a)" using someI[of "?phi a" i] by auto
   368     }
   369     hence "?f ` (\<Union> i \<in> I. A i) \<le> (SIGMA i : I. A i)" by auto
   370   }
   371   ultimately
   372   show "inj_on ?f (\<Union> i \<in> I. A i) \<and> ?f ` (\<Union> i \<in> I. A i) \<le> (SIGMA i : I. A i)"
   373   by auto
   374 qed
   375 
   376 subsection {* The Cantor-Bernstein Theorem *}
   377 
   378 lemma Cantor_Bernstein_aux:
   379   shows "\<exists>A' h. A' \<le> A \<and>
   380                 (\<forall>a \<in> A'. a \<notin> g`(B - f ` A')) \<and>
   381                 (\<forall>a \<in> A'. h a = f a) \<and>
   382                 (\<forall>a \<in> A - A'. h a \<in> B - (f ` A') \<and> a = g(h a))"
   383 proof-
   384   obtain H where H_def: "H = (\<lambda> A'. A - (g`(B - (f ` A'))))" by blast
   385   have 0: "mono H" unfolding mono_def H_def by blast
   386   then obtain A' where 1: "H A' = A'" using lfp_unfold by blast
   387   hence 2: "A' = A - (g`(B - (f ` A')))" unfolding H_def by simp
   388   hence 3: "A' \<le> A" by blast
   389   have 4: "\<forall>a \<in> A'.  a \<notin> g`(B - f ` A')"
   390   using 2 by blast
   391   have 5: "\<forall>a \<in> A - A'. \<exists>b \<in> B - (f ` A'). a = g b"
   392   using 2 by blast
   393   (*  *)
   394   obtain h where h_def:
   395   "h = (\<lambda> a. if a \<in> A' then f a else (SOME b. b \<in> B - (f ` A') \<and> a = g b))" by blast
   396   hence "\<forall>a \<in> A'. h a = f a" by auto
   397   moreover
   398   have "\<forall>a \<in> A - A'. h a \<in> B - (f ` A') \<and> a = g(h a)"
   399   proof
   400     fix a assume *: "a \<in> A - A'"
   401     let ?phi = "\<lambda> b. b \<in> B - (f ` A') \<and> a = g b"
   402     have "h a = (SOME b. ?phi b)" using h_def * by auto
   403     moreover have "\<exists>b. ?phi b" using 5 *  by auto
   404     ultimately show  "?phi (h a)" using someI_ex[of ?phi] by auto
   405   qed
   406   ultimately show ?thesis using 3 4 by blast
   407 qed
   408 
   409 theorem Cantor_Bernstein:
   410   assumes INJ1: "inj_on f A" and SUB1: "f ` A \<le> B" and
   411           INJ2: "inj_on g B" and SUB2: "g ` B \<le> A"
   412   shows "\<exists>h. bij_betw h A B"
   413 proof-
   414   obtain A' and h where 0: "A' \<le> A" and
   415   1: "\<forall>a \<in> A'. a \<notin> g`(B - f ` A')" and
   416   2: "\<forall>a \<in> A'. h a = f a" and
   417   3: "\<forall>a \<in> A - A'. h a \<in> B - (f ` A') \<and> a = g(h a)"
   418   using Cantor_Bernstein_aux[of A g B f] by blast
   419   have "inj_on h A"
   420   proof (intro inj_onI)
   421     fix a1 a2
   422     assume 4: "a1 \<in> A" and 5: "a2 \<in> A" and 6: "h a1 = h a2"
   423     show "a1 = a2"
   424     proof(cases "a1 \<in> A'")
   425       assume Case1: "a1 \<in> A'"
   426       show ?thesis
   427       proof(cases "a2 \<in> A'")
   428         assume Case11: "a2 \<in> A'"
   429         hence "f a1 = f a2" using Case1 2 6 by auto
   430         thus ?thesis using INJ1 Case1 Case11 0
   431         unfolding inj_on_def by blast
   432       next
   433         assume Case12: "a2 \<notin> A'"
   434         hence False using 3 5 2 6 Case1 by force
   435         thus ?thesis by simp
   436       qed
   437     next
   438     assume Case2: "a1 \<notin> A'"
   439       show ?thesis
   440       proof(cases "a2 \<in> A'")
   441         assume Case21: "a2 \<in> A'"
   442         hence False using 3 4 2 6 Case2 by auto
   443         thus ?thesis by simp
   444       next
   445         assume Case22: "a2 \<notin> A'"
   446         hence "a1 = g(h a1) \<and> a2 = g(h a2)" using Case2 4 5 3 by auto
   447         thus ?thesis using 6 by simp
   448       qed
   449     qed
   450   qed
   451   (*  *)
   452   moreover
   453   have "h ` A = B"
   454   proof safe
   455     fix a assume "a \<in> A"
   456     thus "h a \<in> B" using SUB1 2 3 by (cases "a \<in> A'") auto
   457   next
   458     fix b assume *: "b \<in> B"
   459     show "b \<in> h ` A"
   460     proof(cases "b \<in> f ` A'")
   461       assume Case1: "b \<in> f ` A'"
   462       then obtain a where "a \<in> A' \<and> b = f a" by blast
   463       thus ?thesis using 2 0 by force
   464     next
   465       assume Case2: "b \<notin> f ` A'"
   466       hence "g b \<notin> A'" using 1 * by auto
   467       hence 4: "g b \<in> A - A'" using * SUB2 by auto
   468       hence "h(g b) \<in> B \<and> g(h(g b)) = g b"
   469       using 3 by auto
   470       hence "h(g b) = b" using * INJ2 unfolding inj_on_def by auto
   471       thus ?thesis using 4 by force
   472     qed
   473   qed
   474   (*  *)
   475   ultimately show ?thesis unfolding bij_betw_def by auto
   476 qed
   477 
   478 subsection {*Other Consequences of Hilbert's Epsilon*}
   479 
   480 text {*Hilbert's Epsilon and the @{term split} Operator*}
   481 
   482 text{*Looping simprule*}
   483 lemma split_paired_Eps: "(SOME x. P x) = (SOME (a,b). P(a,b))"
   484   by simp
   485 
   486 lemma Eps_split: "Eps (split P) = (SOME xy. P (fst xy) (snd xy))"
   487   by (simp add: split_def)
   488 
   489 lemma Eps_split_eq [simp]: "(@(x',y'). x = x' & y = y') = (x,y)"
   490   by blast
   491 
   492 
   493 text{*A relation is wellfounded iff it has no infinite descending chain*}
   494 lemma wf_iff_no_infinite_down_chain:
   495   "wf r = (~(\<exists>f. \<forall>i. (f(Suc i),f i) \<in> r))"
   496 apply (simp only: wf_eq_minimal)
   497 apply (rule iffI)
   498  apply (rule notI)
   499  apply (erule exE)
   500  apply (erule_tac x = "{w. \<exists>i. w=f i}" in allE, blast)
   501 apply (erule contrapos_np, simp, clarify)
   502 apply (subgoal_tac "\<forall>n. nat_rec x (%i y. @z. z:Q & (z,y) :r) n \<in> Q")
   503  apply (rule_tac x = "nat_rec x (%i y. @z. z:Q & (z,y) :r)" in exI)
   504  apply (rule allI, simp)
   505  apply (rule someI2_ex, blast, blast)
   506 apply (rule allI)
   507 apply (induct_tac "n", simp_all)
   508 apply (rule someI2_ex, blast+)
   509 done
   510 
   511 lemma wf_no_infinite_down_chainE:
   512   assumes "wf r" obtains k where "(f (Suc k), f k) \<notin> r"
   513 using `wf r` wf_iff_no_infinite_down_chain[of r] by blast
   514 
   515 
   516 text{*A dynamically-scoped fact for TFL *}
   517 lemma tfl_some: "\<forall>P x. P x --> P (Eps P)"
   518   by (blast intro: someI)
   519 
   520 
   521 subsection {* Least value operator *}
   522 
   523 definition
   524   LeastM :: "['a => 'b::ord, 'a => bool] => 'a" where
   525   "LeastM m P == SOME x. P x & (\<forall>y. P y --> m x <= m y)"
   526 
   527 syntax
   528   "_LeastM" :: "[pttrn, 'a => 'b::ord, bool] => 'a"    ("LEAST _ WRT _. _" [0, 4, 10] 10)
   529 translations
   530   "LEAST x WRT m. P" == "CONST LeastM m (%x. P)"
   531 
   532 lemma LeastMI2:
   533   "P x ==> (!!y. P y ==> m x <= m y)
   534     ==> (!!x. P x ==> \<forall>y. P y --> m x \<le> m y ==> Q x)
   535     ==> Q (LeastM m P)"
   536   apply (simp add: LeastM_def)
   537   apply (rule someI2_ex, blast, blast)
   538   done
   539 
   540 lemma LeastM_equality:
   541   "P k ==> (!!x. P x ==> m k <= m x)
   542     ==> m (LEAST x WRT m. P x) = (m k::'a::order)"
   543   apply (rule LeastMI2, assumption, blast)
   544   apply (blast intro!: order_antisym)
   545   done
   546 
   547 lemma wf_linord_ex_has_least:
   548   "wf r ==> \<forall>x y. ((x,y):r^+) = ((y,x)~:r^*) ==> P k
   549     ==> \<exists>x. P x & (!y. P y --> (m x,m y):r^*)"
   550   apply (drule wf_trancl [THEN wf_eq_minimal [THEN iffD1]])
   551   apply (drule_tac x = "m`Collect P" in spec, force)
   552   done
   553 
   554 lemma ex_has_least_nat:
   555     "P k ==> \<exists>x. P x & (\<forall>y. P y --> m x <= (m y::nat))"
   556   apply (simp only: pred_nat_trancl_eq_le [symmetric])
   557   apply (rule wf_pred_nat [THEN wf_linord_ex_has_least])
   558    apply (simp add: less_eq linorder_not_le pred_nat_trancl_eq_le, assumption)
   559   done
   560 
   561 lemma LeastM_nat_lemma:
   562     "P k ==> P (LeastM m P) & (\<forall>y. P y --> m (LeastM m P) <= (m y::nat))"
   563   apply (simp add: LeastM_def)
   564   apply (rule someI_ex)
   565   apply (erule ex_has_least_nat)
   566   done
   567 
   568 lemmas LeastM_natI = LeastM_nat_lemma [THEN conjunct1]
   569 
   570 lemma LeastM_nat_le: "P x ==> m (LeastM m P) <= (m x::nat)"
   571 by (rule LeastM_nat_lemma [THEN conjunct2, THEN spec, THEN mp], assumption, assumption)
   572 
   573 
   574 subsection {* Greatest value operator *}
   575 
   576 definition
   577   GreatestM :: "['a => 'b::ord, 'a => bool] => 'a" where
   578   "GreatestM m P == SOME x. P x & (\<forall>y. P y --> m y <= m x)"
   579 
   580 definition
   581   Greatest :: "('a::ord => bool) => 'a" (binder "GREATEST " 10) where
   582   "Greatest == GreatestM (%x. x)"
   583 
   584 syntax
   585   "_GreatestM" :: "[pttrn, 'a => 'b::ord, bool] => 'a"
   586       ("GREATEST _ WRT _. _" [0, 4, 10] 10)
   587 translations
   588   "GREATEST x WRT m. P" == "CONST GreatestM m (%x. P)"
   589 
   590 lemma GreatestMI2:
   591   "P x ==> (!!y. P y ==> m y <= m x)
   592     ==> (!!x. P x ==> \<forall>y. P y --> m y \<le> m x ==> Q x)
   593     ==> Q (GreatestM m P)"
   594   apply (simp add: GreatestM_def)
   595   apply (rule someI2_ex, blast, blast)
   596   done
   597 
   598 lemma GreatestM_equality:
   599  "P k ==> (!!x. P x ==> m x <= m k)
   600     ==> m (GREATEST x WRT m. P x) = (m k::'a::order)"
   601   apply (rule_tac m = m in GreatestMI2, assumption, blast)
   602   apply (blast intro!: order_antisym)
   603   done
   604 
   605 lemma Greatest_equality:
   606   "P (k::'a::order) ==> (!!x. P x ==> x <= k) ==> (GREATEST x. P x) = k"
   607   apply (simp add: Greatest_def)
   608   apply (erule GreatestM_equality, blast)
   609   done
   610 
   611 lemma ex_has_greatest_nat_lemma:
   612   "P k ==> \<forall>x. P x --> (\<exists>y. P y & ~ ((m y::nat) <= m x))
   613     ==> \<exists>y. P y & ~ (m y < m k + n)"
   614   apply (induct n, force)
   615   apply (force simp add: le_Suc_eq)
   616   done
   617 
   618 lemma ex_has_greatest_nat:
   619   "P k ==> \<forall>y. P y --> m y < b
   620     ==> \<exists>x. P x & (\<forall>y. P y --> (m y::nat) <= m x)"
   621   apply (rule ccontr)
   622   apply (cut_tac P = P and n = "b - m k" in ex_has_greatest_nat_lemma)
   623     apply (subgoal_tac [3] "m k <= b", auto)
   624   done
   625 
   626 lemma GreatestM_nat_lemma:
   627   "P k ==> \<forall>y. P y --> m y < b
   628     ==> P (GreatestM m P) & (\<forall>y. P y --> (m y::nat) <= m (GreatestM m P))"
   629   apply (simp add: GreatestM_def)
   630   apply (rule someI_ex)
   631   apply (erule ex_has_greatest_nat, assumption)
   632   done
   633 
   634 lemmas GreatestM_natI = GreatestM_nat_lemma [THEN conjunct1]
   635 
   636 lemma GreatestM_nat_le:
   637   "P x ==> \<forall>y. P y --> m y < b
   638     ==> (m x::nat) <= m (GreatestM m P)"
   639   apply (blast dest: GreatestM_nat_lemma [THEN conjunct2, THEN spec, of P])
   640   done
   641 
   642 
   643 text {* \medskip Specialization to @{text GREATEST}. *}
   644 
   645 lemma GreatestI: "P (k::nat) ==> \<forall>y. P y --> y < b ==> P (GREATEST x. P x)"
   646   apply (simp add: Greatest_def)
   647   apply (rule GreatestM_natI, auto)
   648   done
   649 
   650 lemma Greatest_le:
   651     "P x ==> \<forall>y. P y --> y < b ==> (x::nat) <= (GREATEST x. P x)"
   652   apply (simp add: Greatest_def)
   653   apply (rule GreatestM_nat_le, auto)
   654   done
   655 
   656 
   657 subsection {* An aside: bounded accessible part *}
   658 
   659 text {* Finite monotone eventually stable sequences *}
   660 
   661 lemma finite_mono_remains_stable_implies_strict_prefix:
   662   fixes f :: "nat \<Rightarrow> 'a::order"
   663   assumes S: "finite (range f)" "mono f" and eq: "\<forall>n. f n = f (Suc n) \<longrightarrow> f (Suc n) = f (Suc (Suc n))"
   664   shows "\<exists>N. (\<forall>n\<le>N. \<forall>m\<le>N. m < n \<longrightarrow> f m < f n) \<and> (\<forall>n\<ge>N. f N = f n)"
   665   using assms
   666 proof -
   667   have "\<exists>n. f n = f (Suc n)"
   668   proof (rule ccontr)
   669     assume "\<not> ?thesis"
   670     then have "\<And>n. f n \<noteq> f (Suc n)" by auto
   671     then have "\<And>n. f n < f (Suc n)"
   672       using  `mono f` by (auto simp: le_less mono_iff_le_Suc)
   673     with lift_Suc_mono_less_iff[of f]
   674     have "\<And>n m. n < m \<Longrightarrow> f n < f m" by auto
   675     then have "inj f"
   676       by (auto simp: inj_on_def) (metis linorder_less_linear order_less_imp_not_eq)
   677     with `finite (range f)` have "finite (UNIV::nat set)"
   678       by (rule finite_imageD)
   679     then show False by simp
   680   qed
   681   then obtain n where n: "f n = f (Suc n)" ..
   682   def N \<equiv> "LEAST n. f n = f (Suc n)"
   683   have N: "f N = f (Suc N)"
   684     unfolding N_def using n by (rule LeastI)
   685   show ?thesis
   686   proof (intro exI[of _ N] conjI allI impI)
   687     fix n assume "N \<le> n"
   688     then have "\<And>m. N \<le> m \<Longrightarrow> m \<le> n \<Longrightarrow> f m = f N"
   689     proof (induct rule: dec_induct)
   690       case (step n) then show ?case
   691         using eq[rule_format, of "n - 1"] N
   692         by (cases n) (auto simp add: le_Suc_eq)
   693     qed simp
   694     from this[of n] `N \<le> n` show "f N = f n" by auto
   695   next
   696     fix n m :: nat assume "m < n" "n \<le> N"
   697     then show "f m < f n"
   698     proof (induct rule: less_Suc_induct[consumes 1])
   699       case (1 i)
   700       then have "i < N" by simp
   701       then have "f i \<noteq> f (Suc i)"
   702         unfolding N_def by (rule not_less_Least)
   703       with `mono f` show ?case by (simp add: mono_iff_le_Suc less_le)
   704     qed auto
   705   qed
   706 qed
   707 
   708 lemma finite_mono_strict_prefix_implies_finite_fixpoint:
   709   fixes f :: "nat \<Rightarrow> 'a set"
   710   assumes S: "\<And>i. f i \<subseteq> S" "finite S"
   711     and inj: "\<exists>N. (\<forall>n\<le>N. \<forall>m\<le>N. m < n \<longrightarrow> f m \<subset> f n) \<and> (\<forall>n\<ge>N. f N = f n)"
   712   shows "f (card S) = (\<Union>n. f n)"
   713 proof -
   714   from inj obtain N where inj: "(\<forall>n\<le>N. \<forall>m\<le>N. m < n \<longrightarrow> f m \<subset> f n)" and eq: "(\<forall>n\<ge>N. f N = f n)" by auto
   715 
   716   { fix i have "i \<le> N \<Longrightarrow> i \<le> card (f i)"
   717     proof (induct i)
   718       case 0 then show ?case by simp
   719     next
   720       case (Suc i)
   721       with inj[rule_format, of "Suc i" i]
   722       have "(f i) \<subset> (f (Suc i))" by auto
   723       moreover have "finite (f (Suc i))" using S by (rule finite_subset)
   724       ultimately have "card (f i) < card (f (Suc i))" by (intro psubset_card_mono)
   725       with Suc show ?case using inj by auto
   726     qed
   727   }
   728   then have "N \<le> card (f N)" by simp
   729   also have "\<dots> \<le> card S" using S by (intro card_mono)
   730   finally have "f (card S) = f N" using eq by auto
   731   then show ?thesis using eq inj[rule_format, of N]
   732     apply auto
   733     apply (case_tac "n < N")
   734     apply (auto simp: not_less)
   735     done
   736 qed
   737 
   738 primrec bacc :: "('a \<times> 'a) set \<Rightarrow> nat \<Rightarrow> 'a set" 
   739 where
   740   "bacc r 0 = {x. \<forall> y. (y, x) \<notin> r}"
   741 | "bacc r (Suc n) = (bacc r n \<union> {x. \<forall>y. (y, x) \<in> r \<longrightarrow> y \<in> bacc r n})"
   742 
   743 lemma bacc_subseteq_acc:
   744   "bacc r n \<subseteq> Wellfounded.acc r"
   745   by (induct n) (auto intro: acc.intros)
   746 
   747 lemma bacc_mono:
   748   "n \<le> m \<Longrightarrow> bacc r n \<subseteq> bacc r m"
   749   by (induct rule: dec_induct) auto
   750   
   751 lemma bacc_upper_bound:
   752   "bacc (r :: ('a \<times> 'a) set)  (card (UNIV :: 'a::finite set)) = (\<Union>n. bacc r n)"
   753 proof -
   754   have "mono (bacc r)" unfolding mono_def by (simp add: bacc_mono)
   755   moreover have "\<forall>n. bacc r n = bacc r (Suc n) \<longrightarrow> bacc r (Suc n) = bacc r (Suc (Suc n))" by auto
   756   moreover have "finite (range (bacc r))" by auto
   757   ultimately show ?thesis
   758    by (intro finite_mono_strict_prefix_implies_finite_fixpoint)
   759      (auto intro: finite_mono_remains_stable_implies_strict_prefix)
   760 qed
   761 
   762 lemma acc_subseteq_bacc:
   763   assumes "finite r"
   764   shows "Wellfounded.acc r \<subseteq> (\<Union>n. bacc r n)"
   765 proof
   766   fix x
   767   assume "x : Wellfounded.acc r"
   768   then have "\<exists> n. x : bacc r n"
   769   proof (induct x arbitrary: rule: acc.induct)
   770     case (accI x)
   771     then have "\<forall>y. \<exists> n. (y, x) \<in> r --> y : bacc r n" by simp
   772     from choice[OF this] obtain n where n: "\<forall>y. (y, x) \<in> r \<longrightarrow> y \<in> bacc r (n y)" ..
   773     obtain n where "\<And>y. (y, x) : r \<Longrightarrow> y : bacc r n"
   774     proof
   775       fix y assume y: "(y, x) : r"
   776       with n have "y : bacc r (n y)" by auto
   777       moreover have "n y <= Max ((%(y, x). n y) ` r)"
   778         using y `finite r` by (auto intro!: Max_ge)
   779       note bacc_mono[OF this, of r]
   780       ultimately show "y : bacc r (Max ((%(y, x). n y) ` r))" by auto
   781     qed
   782     then show ?case
   783       by (auto simp add: Let_def intro!: exI[of _ "Suc n"])
   784   qed
   785   then show "x : (UN n. bacc r n)" by auto
   786 qed
   787 
   788 lemma acc_bacc_eq:
   789   fixes A :: "('a :: finite \<times> 'a) set"
   790   assumes "finite A"
   791   shows "Wellfounded.acc A = bacc A (card (UNIV :: 'a set))"
   792   using assms by (metis acc_subseteq_bacc bacc_subseteq_acc bacc_upper_bound order_eq_iff)
   793 
   794 
   795 subsection {* Specification package -- Hilbertized version *}
   796 
   797 lemma exE_some: "[| Ex P ; c == Eps P |] ==> P c"
   798   by (simp only: someI_ex)
   799 
   800 ML_file "Tools/choice_specification.ML"
   801 
   802 end
   803