src/HOL/Hilbert_Choice.thy
changeset 14760 a08e916f4946
parent 14399 dc677b35e54f
child 14872 3f2144aebd76
equal deleted inserted replaced
14759:c90bed2d5bdf 14760:a08e916f4946
     1 (*  Title:      HOL/Hilbert_Choice.thy
     1 (*  Title:      HOL/Hilbert_Choice.thy
     2     ID:         $Id$
     2     ID: $Id$
     3     Author:     Lawrence C Paulson
     3     Author:     Lawrence C Paulson
     4     Copyright   2001  University of Cambridge
     4     Copyright   2001  University of Cambridge
     5 *)
     5 *)
     6 
     6 
     7 header {* Hilbert's epsilon-operator and everything to do with the Axiom of Choice *}
     7 header {* Hilbert's Epsilon-Operator and the Axiom of Choice *}
     8 
     8 
     9 theory Hilbert_Choice = NatArith
     9 theory Hilbert_Choice = NatArith
    10 files ("Hilbert_Choice_lemmas.ML") ("meson_lemmas.ML") ("Tools/meson.ML") ("Tools/specification_package.ML"):
    10 files ("Tools/meson.ML") ("Tools/specification_package.ML"):
    11 
    11 
    12 
    12 
    13 subsection {* Hilbert's epsilon *}
    13 subsection {* Hilbert's epsilon *}
    14 
    14 
    15 consts
    15 consts
    38 constdefs
    38 constdefs
    39   inv :: "('a => 'b) => ('b => 'a)"
    39   inv :: "('a => 'b) => ('b => 'a)"
    40   "inv(f :: 'a => 'b) == %y. SOME x. f x = y"
    40   "inv(f :: 'a => 'b) == %y. SOME x. f x = y"
    41 
    41 
    42   Inv :: "'a set => ('a => 'b) => ('b => 'a)"
    42   Inv :: "'a set => ('a => 'b) => ('b => 'a)"
    43   "Inv A f == %x. SOME y. y : A & f y = x"
    43   "Inv A f == %x. SOME y. y \<in> A & f y = x"
    44 
    44 
    45 
    45 
    46 use "Hilbert_Choice_lemmas.ML"
    46 subsection {*Hilbert's Epsilon-operator*}
    47 declare someI_ex [elim?];
    47 
       
    48 text{*Easier to apply than @{text someI} if the witness comes from an
       
    49 existential formula*}
       
    50 lemma someI_ex [elim?]: "\<exists>x. P x ==> P (SOME x. P x)"
       
    51 apply (erule exE)
       
    52 apply (erule someI)
       
    53 done
       
    54 
       
    55 text{*Easier to apply than @{text someI} because the conclusion has only one
       
    56 occurrence of @{term P}.*}
       
    57 lemma someI2: "[| P a;  !!x. P x ==> Q x |] ==> Q (SOME x. P x)"
       
    58 by (blast intro: someI)
       
    59 
       
    60 text{*Easier to apply than @{text someI2} if the witness comes from an
       
    61 existential formula*}
       
    62 lemma someI2_ex: "[| \<exists>a. P a; !!x. P x ==> Q x |] ==> Q (SOME x. P x)"
       
    63 by (blast intro: someI2)
       
    64 
       
    65 lemma some_equality [intro]:
       
    66      "[| P a;  !!x. P x ==> x=a |] ==> (SOME x. P x) = a"
       
    67 by (blast intro: someI2)
       
    68 
       
    69 lemma some1_equality: "[| EX!x. P x; P a |] ==> (SOME x. P x) = a"
       
    70 by (blast intro: some_equality)
       
    71 
       
    72 lemma some_eq_ex: "P (SOME x. P x) =  (\<exists>x. P x)"
       
    73 by (blast intro: someI)
       
    74 
       
    75 lemma some_eq_trivial [simp]: "(SOME y. y=x) = x"
       
    76 apply (rule some_equality)
       
    77 apply (rule refl, assumption)
       
    78 done
       
    79 
       
    80 lemma some_sym_eq_trivial [simp]: "(SOME y. x=y) = x"
       
    81 apply (rule some_equality)
       
    82 apply (rule refl)
       
    83 apply (erule sym)
       
    84 done
       
    85 
       
    86 
       
    87 subsection{*Axiom of Choice, Proved Using the Description Operator*}
       
    88 
       
    89 text{*Used in @{text "Tools/meson.ML"}*}
       
    90 lemma choice: "\<forall>x. \<exists>y. Q x y ==> \<exists>f. \<forall>x. Q x (f x)"
       
    91 by (fast elim: someI)
       
    92 
       
    93 lemma bchoice: "\<forall>x\<in>S. \<exists>y. Q x y ==> \<exists>f. \<forall>x\<in>S. Q x (f x)"
       
    94 by (fast elim: someI)
       
    95 
       
    96 
       
    97 subsection {*Function Inverse*}
       
    98 
       
    99 lemma inv_id [simp]: "inv id = id"
       
   100 by (simp add: inv_def id_def)
       
   101 
       
   102 text{*A one-to-one function has an inverse.*}
       
   103 lemma inv_f_f [simp]: "inj f ==> inv f (f x) = x"
       
   104 by (simp add: inv_def inj_eq)
       
   105 
       
   106 lemma inv_f_eq: "[| inj f;  f x = y |] ==> inv f y = x"
       
   107 apply (erule subst)
       
   108 apply (erule inv_f_f)
       
   109 done
       
   110 
       
   111 lemma inj_imp_inv_eq: "[| inj f; \<forall>x. f(g x) = x |] ==> inv f = g"
       
   112 by (blast intro: ext inv_f_eq)
       
   113 
       
   114 text{*But is it useful?*}
       
   115 lemma inj_transfer:
       
   116   assumes injf: "inj f" and minor: "!!y. y \<in> range(f) ==> P(inv f y)"
       
   117   shows "P x"
       
   118 proof -
       
   119   have "f x \<in> range f" by auto
       
   120   hence "P(inv f (f x))" by (rule minor)
       
   121   thus "P x" by (simp add: inv_f_f [OF injf])
       
   122 qed
       
   123 
       
   124 
       
   125 lemma inj_iff: "(inj f) = (inv f o f = id)"
       
   126 apply (simp add: o_def expand_fun_eq)
       
   127 apply (blast intro: inj_on_inverseI inv_f_f)
       
   128 done
       
   129 
       
   130 lemma inj_imp_surj_inv: "inj f ==> surj (inv f)"
       
   131 by (blast intro: surjI inv_f_f)
       
   132 
       
   133 lemma f_inv_f: "y \<in> range(f) ==> f(inv f y) = y"
       
   134 apply (simp add: inv_def)
       
   135 apply (fast intro: someI)
       
   136 done
       
   137 
       
   138 lemma surj_f_inv_f: "surj f ==> f(inv f y) = y"
       
   139 by (simp add: f_inv_f surj_range)
       
   140 
       
   141 lemma inv_injective:
       
   142   assumes eq: "inv f x = inv f y"
       
   143       and x: "x: range f"
       
   144       and y: "y: range f"
       
   145   shows "x=y"
       
   146 proof -
       
   147   have "f (inv f x) = f (inv f y)" using eq by simp
       
   148   thus ?thesis by (simp add: f_inv_f x y) 
       
   149 qed
       
   150 
       
   151 lemma inj_on_inv: "A <= range(f) ==> inj_on (inv f) A"
       
   152 by (fast intro: inj_onI elim: inv_injective injD)
       
   153 
       
   154 lemma surj_imp_inj_inv: "surj f ==> inj (inv f)"
       
   155 by (simp add: inj_on_inv surj_range)
       
   156 
       
   157 lemma surj_iff: "(surj f) = (f o inv f = id)"
       
   158 apply (simp add: o_def expand_fun_eq)
       
   159 apply (blast intro: surjI surj_f_inv_f)
       
   160 done
       
   161 
       
   162 lemma surj_imp_inv_eq: "[| surj f; \<forall>x. g(f x) = x |] ==> inv f = g"
       
   163 apply (rule ext)
       
   164 apply (drule_tac x = "inv f x" in spec)
       
   165 apply (simp add: surj_f_inv_f)
       
   166 done
       
   167 
       
   168 lemma bij_imp_bij_inv: "bij f ==> bij (inv f)"
       
   169 by (simp add: bij_def inj_imp_surj_inv surj_imp_inj_inv)
       
   170 
       
   171 lemma inv_equality: "[| !!x. g (f x) = x;  !!y. f (g y) = y |] ==> inv f = g"
       
   172 apply (rule ext)
       
   173 apply (auto simp add: inv_def)
       
   174 done
       
   175 
       
   176 lemma inv_inv_eq: "bij f ==> inv (inv f) = f"
       
   177 apply (rule inv_equality)
       
   178 apply (auto simp add: bij_def surj_f_inv_f)
       
   179 done
       
   180 
       
   181 (** bij(inv f) implies little about f.  Consider f::bool=>bool such that
       
   182     f(True)=f(False)=True.  Then it's consistent with axiom someI that
       
   183     inv f could be any function at all, including the identity function.
       
   184     If inv f=id then inv f is a bijection, but inj f, surj(f) and
       
   185     inv(inv f)=f all fail.
       
   186 **)
       
   187 
       
   188 lemma o_inv_distrib: "[| bij f; bij g |] ==> inv (f o g) = inv g o inv f"
       
   189 apply (rule inv_equality)
       
   190 apply (auto simp add: bij_def surj_f_inv_f)
       
   191 done
       
   192 
       
   193 
       
   194 lemma image_surj_f_inv_f: "surj f ==> f ` (inv f ` A) = A"
       
   195 by (simp add: image_eq_UN surj_f_inv_f)
       
   196 
       
   197 lemma image_inv_f_f: "inj f ==> (inv f) ` (f ` A) = A"
       
   198 by (simp add: image_eq_UN)
       
   199 
       
   200 lemma inv_image_comp: "inj f ==> inv f ` (f`X) = X"
       
   201 by (auto simp add: image_def)
       
   202 
       
   203 lemma bij_image_Collect_eq: "bij f ==> f ` Collect P = {y. P (inv f y)}"
       
   204 apply auto
       
   205 apply (force simp add: bij_is_inj)
       
   206 apply (blast intro: bij_is_surj [THEN surj_f_inv_f, symmetric])
       
   207 done
       
   208 
       
   209 lemma bij_vimage_eq_inv_image: "bij f ==> f -` A = inv f ` A" 
       
   210 apply (auto simp add: bij_is_surj [THEN surj_f_inv_f])
       
   211 apply (blast intro: bij_is_inj [THEN inv_f_f, symmetric])
       
   212 done
       
   213 
       
   214 
       
   215 subsection {*Inverse of a PI-function (restricted domain)*}
       
   216 
       
   217 lemma Inv_f_f: "[| inj_on f A;  x \<in> A |] ==> Inv A f (f x) = x"
       
   218 apply (simp add: Inv_def inj_on_def)
       
   219 apply (blast intro: someI2)
       
   220 done
       
   221 
       
   222 lemma f_Inv_f: "y \<in> f`A  ==> f (Inv A f y) = y"
       
   223 apply (simp add: Inv_def)
       
   224 apply (fast intro: someI2)
       
   225 done
       
   226 
       
   227 lemma Inv_injective:
       
   228   assumes eq: "Inv A f x = Inv A f y"
       
   229       and x: "x: f`A"
       
   230       and y: "y: f`A"
       
   231   shows "x=y"
       
   232 proof -
       
   233   have "f (Inv A f x) = f (Inv A f y)" using eq by simp
       
   234   thus ?thesis by (simp add: f_Inv_f x y) 
       
   235 qed
       
   236 
       
   237 lemma inj_on_Inv: "B <= f`A ==> inj_on (Inv A f) B"
       
   238 apply (rule inj_onI)
       
   239 apply (blast intro: inj_onI dest: Inv_injective injD)
       
   240 done
    48 
   241 
    49 lemma Inv_mem: "[| f ` A = B;  x \<in> B |] ==> Inv A f x \<in> A"
   242 lemma Inv_mem: "[| f ` A = B;  x \<in> B |] ==> Inv A f x \<in> A"
    50 apply (unfold Inv_def)
   243 apply (simp add: Inv_def)
    51 apply (fast intro: someI2)
   244 apply (fast intro: someI2)
    52 done
   245 done
    53 
   246 
    54 lemma Inv_f_eq:
   247 lemma Inv_f_eq: "[| inj_on f A; f x = y; x \<in> A |] ==> Inv A f y = x"
    55   "[| inj_on f A; f x = y; x : A |] ==> Inv A f y = x"
       
    56   apply (erule subst)
   248   apply (erule subst)
    57   apply (erule Inv_f_f)
   249   apply (erule Inv_f_f, assumption)
    58   apply assumption
       
    59   done
   250   done
    60 
   251 
    61 lemma Inv_comp:
   252 lemma Inv_comp:
    62   "[| inj_on f (g ` A); inj_on g A; x : f ` g ` A |] ==>
   253   "[| inj_on f (g ` A); inj_on g A; x \<in> f ` g ` A |] ==>
    63   Inv A (f o g) x = (Inv A g o Inv (g ` A) f) x"
   254   Inv A (f o g) x = (Inv A g o Inv (g ` A) f) x"
    64   apply simp
   255   apply simp
    65   apply (rule Inv_f_eq)
   256   apply (rule Inv_f_eq)
    66     apply (fast intro: comp_inj_on)
   257     apply (fast intro: comp_inj_on)
    67    apply (simp add: f_Inv_f Inv_mem)
   258    apply (simp add: f_Inv_f Inv_mem)
    68   apply (simp add: Inv_mem)
   259   apply (simp add: Inv_mem)
    69   done
   260   done
    70 
   261 
       
   262 
       
   263 subsection {*Other Consequences of Hilbert's Epsilon*}
       
   264 
       
   265 text {*Hilbert's Epsilon and the @{term split} Operator*}
       
   266 
       
   267 text{*Looping simprule*}
       
   268 lemma split_paired_Eps: "(SOME x. P x) = (SOME (a,b). P(a,b))"
       
   269 by (simp add: split_Pair_apply)
       
   270 
       
   271 lemma Eps_split: "Eps (split P) = (SOME xy. P (fst xy) (snd xy))"
       
   272 by (simp add: split_def)
       
   273 
       
   274 lemma Eps_split_eq [simp]: "(@(x',y'). x = x' & y = y') = (x,y)"
       
   275 by blast
       
   276 
       
   277 
       
   278 text{*A relation is wellfounded iff it has no infinite descending chain*}
       
   279 lemma wf_iff_no_infinite_down_chain:
       
   280   "wf r = (~(\<exists>f. \<forall>i. (f(Suc i),f i) \<in> r))"
       
   281 apply (simp only: wf_eq_minimal)
       
   282 apply (rule iffI)
       
   283  apply (rule notI)
       
   284  apply (erule exE)
       
   285  apply (erule_tac x = "{w. \<exists>i. w=f i}" in allE, blast)
       
   286 apply (erule contrapos_np, simp, clarify)
       
   287 apply (subgoal_tac "\<forall>n. nat_rec x (%i y. @z. z:Q & (z,y) :r) n \<in> Q")
       
   288  apply (rule_tac x = "nat_rec x (%i y. @z. z:Q & (z,y) :r)" in exI)
       
   289  apply (rule allI, simp)
       
   290  apply (rule someI2_ex, blast, blast)
       
   291 apply (rule allI)
       
   292 apply (induct_tac "n", simp_all)
       
   293 apply (rule someI2_ex, blast+)
       
   294 done
       
   295 
       
   296 text{*A dynamically-scoped fact for TFL *}
    71 lemma tfl_some: "\<forall>P x. P x --> P (Eps P)"
   297 lemma tfl_some: "\<forall>P x. P x --> P (Eps P)"
    72   -- {* dynamically-scoped fact for TFL *}
       
    73   by (blast intro: someI)
   298   by (blast intro: someI)
    74 
   299 
    75 
   300 
    76 subsection {* Least value operator *}
   301 subsection {* Least value operator *}
    77 
   302 
    78 constdefs
   303 constdefs
    79   LeastM :: "['a => 'b::ord, 'a => bool] => 'a"
   304   LeastM :: "['a => 'b::ord, 'a => bool] => 'a"
    80   "LeastM m P == SOME x. P x & (ALL y. P y --> m x <= m y)"
   305   "LeastM m P == SOME x. P x & (\<forall>y. P y --> m x <= m y)"
    81 
   306 
    82 syntax
   307 syntax
    83   "_LeastM" :: "[pttrn, 'a => 'b::ord, bool] => 'a"    ("LEAST _ WRT _. _" [0, 4, 10] 10)
   308   "_LeastM" :: "[pttrn, 'a => 'b::ord, bool] => 'a"    ("LEAST _ WRT _. _" [0, 4, 10] 10)
    84 translations
   309 translations
    85   "LEAST x WRT m. P" == "LeastM m (%x. P)"
   310   "LEAST x WRT m. P" == "LeastM m (%x. P)"
    86 
   311 
    87 lemma LeastMI2:
   312 lemma LeastMI2:
    88   "P x ==> (!!y. P y ==> m x <= m y)
   313   "P x ==> (!!y. P y ==> m x <= m y)
    89     ==> (!!x. P x ==> \<forall>y. P y --> m x \<le> m y ==> Q x)
   314     ==> (!!x. P x ==> \<forall>y. P y --> m x \<le> m y ==> Q x)
    90     ==> Q (LeastM m P)"
   315     ==> Q (LeastM m P)"
    91   apply (unfold LeastM_def)
   316   apply (simp add: LeastM_def)
    92   apply (rule someI2_ex, blast, blast)
   317   apply (rule someI2_ex, blast, blast)
    93   done
   318   done
    94 
   319 
    95 lemma LeastM_equality:
   320 lemma LeastM_equality:
    96   "P k ==> (!!x. P x ==> m k <= m x)
   321   "P k ==> (!!x. P x ==> m k <= m x)
    98   apply (rule LeastMI2, assumption, blast)
   323   apply (rule LeastMI2, assumption, blast)
    99   apply (blast intro!: order_antisym)
   324   apply (blast intro!: order_antisym)
   100   done
   325   done
   101 
   326 
   102 lemma wf_linord_ex_has_least:
   327 lemma wf_linord_ex_has_least:
   103   "wf r ==> ALL x y. ((x,y):r^+) = ((y,x)~:r^*) ==> P k
   328   "wf r ==> \<forall>x y. ((x,y):r^+) = ((y,x)~:r^*) ==> P k
   104     ==> EX x. P x & (!y. P y --> (m x,m y):r^*)"
   329     ==> \<exists>x. P x & (!y. P y --> (m x,m y):r^*)"
   105   apply (drule wf_trancl [THEN wf_eq_minimal [THEN iffD1]])
   330   apply (drule wf_trancl [THEN wf_eq_minimal [THEN iffD1]])
   106   apply (drule_tac x = "m`Collect P" in spec, force)
   331   apply (drule_tac x = "m`Collect P" in spec, force)
   107   done
   332   done
   108 
   333 
   109 lemma ex_has_least_nat:
   334 lemma ex_has_least_nat:
   110     "P k ==> EX x. P x & (ALL y. P y --> m x <= (m y::nat))"
   335     "P k ==> \<exists>x. P x & (\<forall>y. P y --> m x <= (m y::nat))"
   111   apply (simp only: pred_nat_trancl_eq_le [symmetric])
   336   apply (simp only: pred_nat_trancl_eq_le [symmetric])
   112   apply (rule wf_pred_nat [THEN wf_linord_ex_has_least])
   337   apply (rule wf_pred_nat [THEN wf_linord_ex_has_least])
   113    apply (simp add: less_eq not_le_iff_less pred_nat_trancl_eq_le, assumption)
   338    apply (simp add: less_eq not_le_iff_less pred_nat_trancl_eq_le, assumption)
   114   done
   339   done
   115 
   340 
   116 lemma LeastM_nat_lemma:
   341 lemma LeastM_nat_lemma:
   117     "P k ==> P (LeastM m P) & (ALL y. P y --> m (LeastM m P) <= (m y::nat))"
   342     "P k ==> P (LeastM m P) & (\<forall>y. P y --> m (LeastM m P) <= (m y::nat))"
   118   apply (unfold LeastM_def)
   343   apply (simp add: LeastM_def)
   119   apply (rule someI_ex)
   344   apply (rule someI_ex)
   120   apply (erule ex_has_least_nat)
   345   apply (erule ex_has_least_nat)
   121   done
   346   done
   122 
   347 
   123 lemmas LeastM_natI = LeastM_nat_lemma [THEN conjunct1, standard]
   348 lemmas LeastM_natI = LeastM_nat_lemma [THEN conjunct1, standard]
   128 
   353 
   129 subsection {* Greatest value operator *}
   354 subsection {* Greatest value operator *}
   130 
   355 
   131 constdefs
   356 constdefs
   132   GreatestM :: "['a => 'b::ord, 'a => bool] => 'a"
   357   GreatestM :: "['a => 'b::ord, 'a => bool] => 'a"
   133   "GreatestM m P == SOME x. P x & (ALL y. P y --> m y <= m x)"
   358   "GreatestM m P == SOME x. P x & (\<forall>y. P y --> m y <= m x)"
   134 
   359 
   135   Greatest :: "('a::ord => bool) => 'a"    (binder "GREATEST " 10)
   360   Greatest :: "('a::ord => bool) => 'a"    (binder "GREATEST " 10)
   136   "Greatest == GreatestM (%x. x)"
   361   "Greatest == GreatestM (%x. x)"
   137 
   362 
   138 syntax
   363 syntax
   144 
   369 
   145 lemma GreatestMI2:
   370 lemma GreatestMI2:
   146   "P x ==> (!!y. P y ==> m y <= m x)
   371   "P x ==> (!!y. P y ==> m y <= m x)
   147     ==> (!!x. P x ==> \<forall>y. P y --> m y \<le> m x ==> Q x)
   372     ==> (!!x. P x ==> \<forall>y. P y --> m y \<le> m x ==> Q x)
   148     ==> Q (GreatestM m P)"
   373     ==> Q (GreatestM m P)"
   149   apply (unfold GreatestM_def)
   374   apply (simp add: GreatestM_def)
   150   apply (rule someI2_ex, blast, blast)
   375   apply (rule someI2_ex, blast, blast)
   151   done
   376   done
   152 
   377 
   153 lemma GreatestM_equality:
   378 lemma GreatestM_equality:
   154  "P k ==> (!!x. P x ==> m x <= m k)
   379  "P k ==> (!!x. P x ==> m x <= m k)
   157   apply (blast intro!: order_antisym)
   382   apply (blast intro!: order_antisym)
   158   done
   383   done
   159 
   384 
   160 lemma Greatest_equality:
   385 lemma Greatest_equality:
   161   "P (k::'a::order) ==> (!!x. P x ==> x <= k) ==> (GREATEST x. P x) = k"
   386   "P (k::'a::order) ==> (!!x. P x ==> x <= k) ==> (GREATEST x. P x) = k"
   162   apply (unfold Greatest_def)
   387   apply (simp add: Greatest_def)
   163   apply (erule GreatestM_equality, blast)
   388   apply (erule GreatestM_equality, blast)
   164   done
   389   done
   165 
   390 
   166 lemma ex_has_greatest_nat_lemma:
   391 lemma ex_has_greatest_nat_lemma:
   167   "P k ==> ALL x. P x --> (EX y. P y & ~ ((m y::nat) <= m x))
   392   "P k ==> \<forall>x. P x --> (\<exists>y. P y & ~ ((m y::nat) <= m x))
   168     ==> EX y. P y & ~ (m y < m k + n)"
   393     ==> \<exists>y. P y & ~ (m y < m k + n)"
   169   apply (induct_tac n, force)
   394   apply (induct_tac n, force)
   170   apply (force simp add: le_Suc_eq)
   395   apply (force simp add: le_Suc_eq)
   171   done
   396   done
   172 
   397 
   173 lemma ex_has_greatest_nat:
   398 lemma ex_has_greatest_nat:
   174   "P k ==> ALL y. P y --> m y < b
   399   "P k ==> \<forall>y. P y --> m y < b
   175     ==> EX x. P x & (ALL y. P y --> (m y::nat) <= m x)"
   400     ==> \<exists>x. P x & (\<forall>y. P y --> (m y::nat) <= m x)"
   176   apply (rule ccontr)
   401   apply (rule ccontr)
   177   apply (cut_tac P = P and n = "b - m k" in ex_has_greatest_nat_lemma)
   402   apply (cut_tac P = P and n = "b - m k" in ex_has_greatest_nat_lemma)
   178     apply (subgoal_tac [3] "m k <= b", auto)
   403     apply (subgoal_tac [3] "m k <= b", auto)
   179   done
   404   done
   180 
   405 
   181 lemma GreatestM_nat_lemma:
   406 lemma GreatestM_nat_lemma:
   182   "P k ==> ALL y. P y --> m y < b
   407   "P k ==> \<forall>y. P y --> m y < b
   183     ==> P (GreatestM m P) & (ALL y. P y --> (m y::nat) <= m (GreatestM m P))"
   408     ==> P (GreatestM m P) & (\<forall>y. P y --> (m y::nat) <= m (GreatestM m P))"
   184   apply (unfold GreatestM_def)
   409   apply (simp add: GreatestM_def)
   185   apply (rule someI_ex)
   410   apply (rule someI_ex)
   186   apply (erule ex_has_greatest_nat, assumption)
   411   apply (erule ex_has_greatest_nat, assumption)
   187   done
   412   done
   188 
   413 
   189 lemmas GreatestM_natI = GreatestM_nat_lemma [THEN conjunct1, standard]
   414 lemmas GreatestM_natI = GreatestM_nat_lemma [THEN conjunct1, standard]
   190 
   415 
   191 lemma GreatestM_nat_le:
   416 lemma GreatestM_nat_le:
   192   "P x ==> ALL y. P y --> m y < b
   417   "P x ==> \<forall>y. P y --> m y < b
   193     ==> (m x::nat) <= m (GreatestM m P)"
   418     ==> (m x::nat) <= m (GreatestM m P)"
   194   apply (blast dest: GreatestM_nat_lemma [THEN conjunct2, THEN spec])
   419   apply (blast dest: GreatestM_nat_lemma [THEN conjunct2, THEN spec])
   195   done
   420   done
   196 
   421 
   197 
   422 
   198 text {* \medskip Specialization to @{text GREATEST}. *}
   423 text {* \medskip Specialization to @{text GREATEST}. *}
   199 
   424 
   200 lemma GreatestI: "P (k::nat) ==> ALL y. P y --> y < b ==> P (GREATEST x. P x)"
   425 lemma GreatestI: "P (k::nat) ==> \<forall>y. P y --> y < b ==> P (GREATEST x. P x)"
   201   apply (unfold Greatest_def)
   426   apply (simp add: Greatest_def)
   202   apply (rule GreatestM_natI, auto)
   427   apply (rule GreatestM_natI, auto)
   203   done
   428   done
   204 
   429 
   205 lemma Greatest_le:
   430 lemma Greatest_le:
   206     "P x ==> ALL y. P y --> y < b ==> (x::nat) <= (GREATEST x. P x)"
   431     "P x ==> \<forall>y. P y --> y < b ==> (x::nat) <= (GREATEST x. P x)"
   207   apply (unfold Greatest_def)
   432   apply (simp add: Greatest_def)
   208   apply (rule GreatestM_nat_le, auto)
   433   apply (rule GreatestM_nat_le, auto)
   209   done
   434   done
   210 
   435 
   211 
   436 
   212 subsection {* The Meson proof procedure *}
   437 subsection {* The Meson proof procedure *}
   216 text {* de Morgan laws *}
   441 text {* de Morgan laws *}
   217 
   442 
   218 lemma meson_not_conjD: "~(P&Q) ==> ~P | ~Q"
   443 lemma meson_not_conjD: "~(P&Q) ==> ~P | ~Q"
   219   and meson_not_disjD: "~(P|Q) ==> ~P & ~Q"
   444   and meson_not_disjD: "~(P|Q) ==> ~P & ~Q"
   220   and meson_not_notD: "~~P ==> P"
   445   and meson_not_notD: "~~P ==> P"
   221   and meson_not_allD: "!!P. ~(ALL x. P(x)) ==> EX x. ~P(x)"
   446   and meson_not_allD: "!!P. ~(\<forall>x. P(x)) ==> \<exists>x. ~P(x)"
   222   and meson_not_exD: "!!P. ~(EX x. P(x)) ==> ALL x. ~P(x)"
   447   and meson_not_exD: "!!P. ~(\<exists>x. P(x)) ==> \<forall>x. ~P(x)"
   223   by fast+
   448   by fast+
   224 
   449 
   225 text {* Removal of @{text "-->"} and @{text "<->"} (positive and
   450 text {* Removal of @{text "-->"} and @{text "<->"} (positive and
   226 negative occurrences) *}
   451 negative occurrences) *}
   227 
   452 
   235 
   460 
   236 subsubsection {* Pulling out the existential quantifiers *}
   461 subsubsection {* Pulling out the existential quantifiers *}
   237 
   462 
   238 text {* Conjunction *}
   463 text {* Conjunction *}
   239 
   464 
   240 lemma meson_conj_exD1: "!!P Q. (EX x. P(x)) & Q ==> EX x. P(x) & Q"
   465 lemma meson_conj_exD1: "!!P Q. (\<exists>x. P(x)) & Q ==> \<exists>x. P(x) & Q"
   241   and meson_conj_exD2: "!!P Q. P & (EX x. Q(x)) ==> EX x. P & Q(x)"
   466   and meson_conj_exD2: "!!P Q. P & (\<exists>x. Q(x)) ==> \<exists>x. P & Q(x)"
   242   by fast+
   467   by fast+
   243 
   468 
   244 
   469 
   245 text {* Disjunction *}
   470 text {* Disjunction *}
   246 
   471 
   247 lemma meson_disj_exD: "!!P Q. (EX x. P(x)) | (EX x. Q(x)) ==> EX x. P(x) | Q(x)"
   472 lemma meson_disj_exD: "!!P Q. (\<exists>x. P(x)) | (\<exists>x. Q(x)) ==> \<exists>x. P(x) | Q(x)"
   248   -- {* DO NOT USE with forall-Skolemization: makes fewer schematic variables!! *}
   473   -- {* DO NOT USE with forall-Skolemization: makes fewer schematic variables!! *}
   249   -- {* With ex-Skolemization, makes fewer Skolem constants *}
   474   -- {* With ex-Skolemization, makes fewer Skolem constants *}
   250   and meson_disj_exD1: "!!P Q. (EX x. P(x)) | Q ==> EX x. P(x) | Q"
   475   and meson_disj_exD1: "!!P Q. (\<exists>x. P(x)) | Q ==> \<exists>x. P(x) | Q"
   251   and meson_disj_exD2: "!!P Q. P | (EX x. Q(x)) ==> EX x. P | Q(x)"
   476   and meson_disj_exD2: "!!P Q. P | (\<exists>x. Q(x)) ==> \<exists>x. P | Q(x)"
   252   by fast+
   477   by fast+
   253 
   478 
   254 
   479 
   255 subsubsection {* Generating clauses for the Meson Proof Procedure *}
   480 subsubsection {* Generating clauses for the Meson Proof Procedure *}
   256 
   481 
   260   and meson_disj_comm: "P|Q ==> Q|P"
   485   and meson_disj_comm: "P|Q ==> Q|P"
   261   and meson_disj_FalseD1: "False|P ==> P"
   486   and meson_disj_FalseD1: "False|P ==> P"
   262   and meson_disj_FalseD2: "P|False ==> P"
   487   and meson_disj_FalseD2: "P|False ==> P"
   263   by fast+
   488   by fast+
   264 
   489 
   265 use "meson_lemmas.ML"
   490 
       
   491 subsection{*Lemmas for Meson, the Model Elimination Procedure*}
       
   492 
       
   493 
       
   494 text{* Generation of contrapositives *}
       
   495 
       
   496 text{*Inserts negated disjunct after removing the negation; P is a literal.
       
   497   Model elimination requires assuming the negation of every attempted subgoal,
       
   498   hence the negated disjuncts.*}
       
   499 lemma make_neg_rule: "~P|Q ==> ((~P==>P) ==> Q)"
       
   500 by blast
       
   501 
       
   502 text{*Version for Plaisted's "Postive refinement" of the Meson procedure*}
       
   503 lemma make_refined_neg_rule: "~P|Q ==> (P ==> Q)"
       
   504 by blast
       
   505 
       
   506 text{*@{term P} should be a literal*}
       
   507 lemma make_pos_rule: "P|Q ==> ((P==>~P) ==> Q)"
       
   508 by blast
       
   509 
       
   510 text{*Versions of @{text make_neg_rule} and @{text make_pos_rule} that don't
       
   511 insert new assumptions, for ordinary resolution.*}
       
   512 
       
   513 lemmas make_neg_rule' = make_refined_neg_rule
       
   514 
       
   515 lemma make_pos_rule': "[|P|Q; ~P|] ==> Q"
       
   516 by blast
       
   517 
       
   518 text{* Generation of a goal clause -- put away the final literal *}
       
   519 
       
   520 lemma make_neg_goal: "~P ==> ((~P==>P) ==> False)"
       
   521 by blast
       
   522 
       
   523 lemma make_pos_goal: "P ==> ((P==>~P) ==> False)"
       
   524 by blast
       
   525 
       
   526 
       
   527 subsubsection{* Lemmas for Forward Proof*}
       
   528 
       
   529 text{*There is a similarity to congruence rules*}
       
   530 
       
   531 (*NOTE: could handle conjunctions (faster?) by
       
   532     nf(th RS conjunct2) RS (nf(th RS conjunct1) RS conjI) *)
       
   533 lemma conj_forward: "[| P'&Q';  P' ==> P;  Q' ==> Q |] ==> P&Q"
       
   534 by blast
       
   535 
       
   536 lemma disj_forward: "[| P'|Q';  P' ==> P;  Q' ==> Q |] ==> P|Q"
       
   537 by blast
       
   538 
       
   539 (*Version of @{text disj_forward} for removal of duplicate literals*)
       
   540 lemma disj_forward2:
       
   541     "[| P'|Q';  P' ==> P;  [| Q'; P==>False |] ==> Q |] ==> P|Q"
       
   542 apply blast 
       
   543 done
       
   544 
       
   545 lemma all_forward: "[| \<forall>x. P'(x);  !!x. P'(x) ==> P(x) |] ==> \<forall>x. P(x)"
       
   546 by blast
       
   547 
       
   548 lemma ex_forward: "[| \<exists>x. P'(x);  !!x. P'(x) ==> P(x) |] ==> \<exists>x. P(x)"
       
   549 by blast
       
   550 
       
   551 ML
       
   552 {*
       
   553 val inv_def = thm "inv_def";
       
   554 val Inv_def = thm "Inv_def";
       
   555 
       
   556 val someI = thm "someI";
       
   557 val someI_ex = thm "someI_ex";
       
   558 val someI2 = thm "someI2";
       
   559 val someI2_ex = thm "someI2_ex";
       
   560 val some_equality = thm "some_equality";
       
   561 val some1_equality = thm "some1_equality";
       
   562 val some_eq_ex = thm "some_eq_ex";
       
   563 val some_eq_trivial = thm "some_eq_trivial";
       
   564 val some_sym_eq_trivial = thm "some_sym_eq_trivial";
       
   565 val choice = thm "choice";
       
   566 val bchoice = thm "bchoice";
       
   567 val inv_id = thm "inv_id";
       
   568 val inv_f_f = thm "inv_f_f";
       
   569 val inv_f_eq = thm "inv_f_eq";
       
   570 val inj_imp_inv_eq = thm "inj_imp_inv_eq";
       
   571 val inj_transfer = thm "inj_transfer";
       
   572 val inj_iff = thm "inj_iff";
       
   573 val inj_imp_surj_inv = thm "inj_imp_surj_inv";
       
   574 val f_inv_f = thm "f_inv_f";
       
   575 val surj_f_inv_f = thm "surj_f_inv_f";
       
   576 val inv_injective = thm "inv_injective";
       
   577 val inj_on_inv = thm "inj_on_inv";
       
   578 val surj_imp_inj_inv = thm "surj_imp_inj_inv";
       
   579 val surj_iff = thm "surj_iff";
       
   580 val surj_imp_inv_eq = thm "surj_imp_inv_eq";
       
   581 val bij_imp_bij_inv = thm "bij_imp_bij_inv";
       
   582 val inv_equality = thm "inv_equality";
       
   583 val inv_inv_eq = thm "inv_inv_eq";
       
   584 val o_inv_distrib = thm "o_inv_distrib";
       
   585 val image_surj_f_inv_f = thm "image_surj_f_inv_f";
       
   586 val image_inv_f_f = thm "image_inv_f_f";
       
   587 val inv_image_comp = thm "inv_image_comp";
       
   588 val bij_image_Collect_eq = thm "bij_image_Collect_eq";
       
   589 val bij_vimage_eq_inv_image = thm "bij_vimage_eq_inv_image";
       
   590 val Inv_f_f = thm "Inv_f_f";
       
   591 val f_Inv_f = thm "f_Inv_f";
       
   592 val Inv_injective = thm "Inv_injective";
       
   593 val inj_on_Inv = thm "inj_on_Inv";
       
   594 val split_paired_Eps = thm "split_paired_Eps";
       
   595 val Eps_split = thm "Eps_split";
       
   596 val Eps_split_eq = thm "Eps_split_eq";
       
   597 val wf_iff_no_infinite_down_chain = thm "wf_iff_no_infinite_down_chain";
       
   598 val Inv_mem = thm "Inv_mem";
       
   599 val Inv_f_eq = thm "Inv_f_eq";
       
   600 val Inv_comp = thm "Inv_comp";
       
   601 val tfl_some = thm "tfl_some";
       
   602 val make_neg_rule = thm "make_neg_rule";
       
   603 val make_refined_neg_rule = thm "make_refined_neg_rule";
       
   604 val make_pos_rule = thm "make_pos_rule";
       
   605 val make_neg_rule' = thm "make_neg_rule'";
       
   606 val make_pos_rule' = thm "make_pos_rule'";
       
   607 val make_neg_goal = thm "make_neg_goal";
       
   608 val make_pos_goal = thm "make_pos_goal";
       
   609 val conj_forward = thm "conj_forward";
       
   610 val disj_forward = thm "disj_forward";
       
   611 val disj_forward2 = thm "disj_forward2";
       
   612 val all_forward = thm "all_forward";
       
   613 val ex_forward = thm "ex_forward";
       
   614 *}
       
   615 
       
   616 
   266 use "Tools/meson.ML"
   617 use "Tools/meson.ML"
   267 setup meson_setup
   618 setup meson_setup
   268 
   619 
   269 use "Tools/specification_package.ML"
   620 use "Tools/specification_package.ML"
   270 
   621