src/HOL/Hilbert_Choice.thy
author nipkow
Wed, 03 Sep 2003 18:20:57 +0200
changeset 14180 d2e550609c40
parent 14115 65ec3f73d00b
child 14208 144f45277d5a
permissions -rw-r--r--
Introduced new syntax for maplets x |-> y
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Hilbert_Choice.thy
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
     2
    ID:         $Id$
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
     4
    Copyright   2001  University of Cambridge
12023
wenzelm
parents: 11506
diff changeset
     5
*)
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
     6
12023
wenzelm
parents: 11506
diff changeset
     7
header {* Hilbert's epsilon-operator and everything to do with the Axiom of Choice *}
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
     8
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
     9
theory Hilbert_Choice = NatArith
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents: 13764
diff changeset
    10
files ("Hilbert_Choice_lemmas.ML") ("meson_lemmas.ML") ("Tools/meson.ML") ("Tools/specification_package.ML"):
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
    11
12298
wenzelm
parents: 12023
diff changeset
    12
wenzelm
parents: 12023
diff changeset
    13
subsection {* Hilbert's epsilon *}
wenzelm
parents: 12023
diff changeset
    14
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
    15
consts
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
    16
  Eps           :: "('a => bool) => 'a"
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
    17
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
    18
syntax (input)
12298
wenzelm
parents: 12023
diff changeset
    19
  "_Eps"        :: "[pttrn, bool] => 'a"    ("(3\<epsilon>_./ _)" [0, 10] 10)
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
    20
syntax (HOL)
12298
wenzelm
parents: 12023
diff changeset
    21
  "_Eps"        :: "[pttrn, bool] => 'a"    ("(3@ _./ _)" [0, 10] 10)
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
    22
syntax
12298
wenzelm
parents: 12023
diff changeset
    23
  "_Eps"        :: "[pttrn, bool] => 'a"    ("(3SOME _./ _)" [0, 10] 10)
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
    24
translations
13764
3e180bf68496 removed some problems with print translations
nipkow
parents: 13763
diff changeset
    25
  "SOME x. P" == "Eps (%x. P)"
13763
f94b569cd610 added print translations tha avoid eta contraction for important binders.
nipkow
parents: 13585
diff changeset
    26
f94b569cd610 added print translations tha avoid eta contraction for important binders.
nipkow
parents: 13585
diff changeset
    27
print_translation {*
f94b569cd610 added print translations tha avoid eta contraction for important binders.
nipkow
parents: 13585
diff changeset
    28
(* to avoid eta-contraction of body *)
f94b569cd610 added print translations tha avoid eta contraction for important binders.
nipkow
parents: 13585
diff changeset
    29
[("Eps", fn [Abs abs] =>
f94b569cd610 added print translations tha avoid eta contraction for important binders.
nipkow
parents: 13585
diff changeset
    30
     let val (x,t) = atomic_abs_tr' abs
f94b569cd610 added print translations tha avoid eta contraction for important binders.
nipkow
parents: 13585
diff changeset
    31
     in Syntax.const "_Eps" $ x $ t end)]
f94b569cd610 added print translations tha avoid eta contraction for important binders.
nipkow
parents: 13585
diff changeset
    32
*}
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
    33
12298
wenzelm
parents: 12023
diff changeset
    34
axioms
wenzelm
parents: 12023
diff changeset
    35
  someI: "P (x::'a) ==> P (SOME x. P x)"
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
    36
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
    37
12298
wenzelm
parents: 12023
diff changeset
    38
constdefs
wenzelm
parents: 12023
diff changeset
    39
  inv :: "('a => 'b) => ('b => 'a)"
wenzelm
parents: 12023
diff changeset
    40
  "inv(f :: 'a => 'b) == %y. SOME x. f x = y"
11454
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11451
diff changeset
    41
12298
wenzelm
parents: 12023
diff changeset
    42
  Inv :: "'a set => ('a => 'b) => ('b => 'a)"
wenzelm
parents: 12023
diff changeset
    43
  "Inv A f == %x. SOME y. y : A & f y = x"
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
    44
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
    45
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
    46
use "Hilbert_Choice_lemmas.ML"
12372
cd3a09c7dac9 tuned declarations;
wenzelm
parents: 12298
diff changeset
    47
declare someI_ex [elim?];
cd3a09c7dac9 tuned declarations;
wenzelm
parents: 12298
diff changeset
    48
13585
db4005b40cc6 Converted Fun to Isar style.
paulson
parents: 12372
diff changeset
    49
lemma Inv_mem: "[| f ` A = B;  x \<in> B |] ==> Inv A f x \<in> A"
db4005b40cc6 Converted Fun to Isar style.
paulson
parents: 12372
diff changeset
    50
apply (unfold Inv_def)
db4005b40cc6 Converted Fun to Isar style.
paulson
parents: 12372
diff changeset
    51
apply (fast intro: someI2)
db4005b40cc6 Converted Fun to Isar style.
paulson
parents: 12372
diff changeset
    52
done
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
    53
12298
wenzelm
parents: 12023
diff changeset
    54
lemma tfl_some: "\<forall>P x. P x --> P (Eps P)"
wenzelm
parents: 12023
diff changeset
    55
  -- {* dynamically-scoped fact for TFL *}
wenzelm
parents: 12023
diff changeset
    56
  by (blast intro: someI)
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
    57
12298
wenzelm
parents: 12023
diff changeset
    58
wenzelm
parents: 12023
diff changeset
    59
subsection {* Least value operator *}
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
    60
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
    61
constdefs
12298
wenzelm
parents: 12023
diff changeset
    62
  LeastM :: "['a => 'b::ord, 'a => bool] => 'a"
wenzelm
parents: 12023
diff changeset
    63
  "LeastM m P == SOME x. P x & (ALL y. P y --> m x <= m y)"
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
    64
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
    65
syntax
12298
wenzelm
parents: 12023
diff changeset
    66
  "_LeastM" :: "[pttrn, 'a => 'b::ord, bool] => 'a"    ("LEAST _ WRT _. _" [0, 4, 10] 10)
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
    67
translations
12298
wenzelm
parents: 12023
diff changeset
    68
  "LEAST x WRT m. P" == "LeastM m (%x. P)"
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
    69
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
    70
lemma LeastMI2:
12298
wenzelm
parents: 12023
diff changeset
    71
  "P x ==> (!!y. P y ==> m x <= m y)
wenzelm
parents: 12023
diff changeset
    72
    ==> (!!x. P x ==> \<forall>y. P y --> m x \<le> m y ==> Q x)
wenzelm
parents: 12023
diff changeset
    73
    ==> Q (LeastM m P)"
wenzelm
parents: 12023
diff changeset
    74
  apply (unfold LeastM_def)
wenzelm
parents: 12023
diff changeset
    75
  apply (rule someI2_ex)
wenzelm
parents: 12023
diff changeset
    76
   apply blast
wenzelm
parents: 12023
diff changeset
    77
  apply blast
wenzelm
parents: 12023
diff changeset
    78
  done
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
    79
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
    80
lemma LeastM_equality:
12298
wenzelm
parents: 12023
diff changeset
    81
  "P k ==> (!!x. P x ==> m k <= m x)
wenzelm
parents: 12023
diff changeset
    82
    ==> m (LEAST x WRT m. P x) = (m k::'a::order)"
wenzelm
parents: 12023
diff changeset
    83
  apply (rule LeastMI2)
wenzelm
parents: 12023
diff changeset
    84
    apply assumption
wenzelm
parents: 12023
diff changeset
    85
   apply blast
wenzelm
parents: 12023
diff changeset
    86
  apply (blast intro!: order_antisym)
wenzelm
parents: 12023
diff changeset
    87
  done
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
    88
11454
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11451
diff changeset
    89
lemma wf_linord_ex_has_least:
12298
wenzelm
parents: 12023
diff changeset
    90
  "wf r ==> ALL x y. ((x,y):r^+) = ((y,x)~:r^*) ==> P k
wenzelm
parents: 12023
diff changeset
    91
    ==> EX x. P x & (!y. P y --> (m x,m y):r^*)"
wenzelm
parents: 12023
diff changeset
    92
  apply (drule wf_trancl [THEN wf_eq_minimal [THEN iffD1]])
wenzelm
parents: 12023
diff changeset
    93
  apply (drule_tac x = "m`Collect P" in spec)
wenzelm
parents: 12023
diff changeset
    94
  apply force
wenzelm
parents: 12023
diff changeset
    95
  done
11454
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11451
diff changeset
    96
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11451
diff changeset
    97
lemma ex_has_least_nat:
12298
wenzelm
parents: 12023
diff changeset
    98
    "P k ==> EX x. P x & (ALL y. P y --> m x <= (m y::nat))"
wenzelm
parents: 12023
diff changeset
    99
  apply (simp only: pred_nat_trancl_eq_le [symmetric])
wenzelm
parents: 12023
diff changeset
   100
  apply (rule wf_pred_nat [THEN wf_linord_ex_has_least])
wenzelm
parents: 12023
diff changeset
   101
   apply (simp add: less_eq not_le_iff_less pred_nat_trancl_eq_le)
wenzelm
parents: 12023
diff changeset
   102
  apply assumption
wenzelm
parents: 12023
diff changeset
   103
  done
11454
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11451
diff changeset
   104
12298
wenzelm
parents: 12023
diff changeset
   105
lemma LeastM_nat_lemma:
wenzelm
parents: 12023
diff changeset
   106
    "P k ==> P (LeastM m P) & (ALL y. P y --> m (LeastM m P) <= (m y::nat))"
wenzelm
parents: 12023
diff changeset
   107
  apply (unfold LeastM_def)
wenzelm
parents: 12023
diff changeset
   108
  apply (rule someI_ex)
wenzelm
parents: 12023
diff changeset
   109
  apply (erule ex_has_least_nat)
wenzelm
parents: 12023
diff changeset
   110
  done
11454
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11451
diff changeset
   111
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11451
diff changeset
   112
lemmas LeastM_natI = LeastM_nat_lemma [THEN conjunct1, standard]
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11451
diff changeset
   113
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11451
diff changeset
   114
lemma LeastM_nat_le: "P x ==> m (LeastM m P) <= (m x::nat)"
12298
wenzelm
parents: 12023
diff changeset
   115
  apply (rule LeastM_nat_lemma [THEN conjunct2, THEN spec, THEN mp])
wenzelm
parents: 12023
diff changeset
   116
   apply assumption
wenzelm
parents: 12023
diff changeset
   117
  apply assumption
wenzelm
parents: 12023
diff changeset
   118
  done
11454
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11451
diff changeset
   119
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   120
12298
wenzelm
parents: 12023
diff changeset
   121
subsection {* Greatest value operator *}
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   122
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   123
constdefs
12298
wenzelm
parents: 12023
diff changeset
   124
  GreatestM :: "['a => 'b::ord, 'a => bool] => 'a"
wenzelm
parents: 12023
diff changeset
   125
  "GreatestM m P == SOME x. P x & (ALL y. P y --> m y <= m x)"
wenzelm
parents: 12023
diff changeset
   126
wenzelm
parents: 12023
diff changeset
   127
  Greatest :: "('a::ord => bool) => 'a"    (binder "GREATEST " 10)
wenzelm
parents: 12023
diff changeset
   128
  "Greatest == GreatestM (%x. x)"
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   129
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   130
syntax
12298
wenzelm
parents: 12023
diff changeset
   131
  "_GreatestM" :: "[pttrn, 'a=>'b::ord, bool] => 'a"
wenzelm
parents: 12023
diff changeset
   132
      ("GREATEST _ WRT _. _" [0, 4, 10] 10)
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   133
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   134
translations
12298
wenzelm
parents: 12023
diff changeset
   135
  "GREATEST x WRT m. P" == "GreatestM m (%x. P)"
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   136
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   137
lemma GreatestMI2:
12298
wenzelm
parents: 12023
diff changeset
   138
  "P x ==> (!!y. P y ==> m y <= m x)
wenzelm
parents: 12023
diff changeset
   139
    ==> (!!x. P x ==> \<forall>y. P y --> m y \<le> m x ==> Q x)
wenzelm
parents: 12023
diff changeset
   140
    ==> Q (GreatestM m P)"
wenzelm
parents: 12023
diff changeset
   141
  apply (unfold GreatestM_def)
wenzelm
parents: 12023
diff changeset
   142
  apply (rule someI2_ex)
wenzelm
parents: 12023
diff changeset
   143
   apply blast
wenzelm
parents: 12023
diff changeset
   144
  apply blast
wenzelm
parents: 12023
diff changeset
   145
  done
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   146
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   147
lemma GreatestM_equality:
12298
wenzelm
parents: 12023
diff changeset
   148
 "P k ==> (!!x. P x ==> m x <= m k)
wenzelm
parents: 12023
diff changeset
   149
    ==> m (GREATEST x WRT m. P x) = (m k::'a::order)"
wenzelm
parents: 12023
diff changeset
   150
  apply (rule_tac m = m in GreatestMI2)
wenzelm
parents: 12023
diff changeset
   151
    apply assumption
wenzelm
parents: 12023
diff changeset
   152
   apply blast
wenzelm
parents: 12023
diff changeset
   153
  apply (blast intro!: order_antisym)
wenzelm
parents: 12023
diff changeset
   154
  done
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   155
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   156
lemma Greatest_equality:
12298
wenzelm
parents: 12023
diff changeset
   157
  "P (k::'a::order) ==> (!!x. P x ==> x <= k) ==> (GREATEST x. P x) = k"
wenzelm
parents: 12023
diff changeset
   158
  apply (unfold Greatest_def)
wenzelm
parents: 12023
diff changeset
   159
  apply (erule GreatestM_equality)
wenzelm
parents: 12023
diff changeset
   160
  apply blast
wenzelm
parents: 12023
diff changeset
   161
  done
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   162
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   163
lemma ex_has_greatest_nat_lemma:
12298
wenzelm
parents: 12023
diff changeset
   164
  "P k ==> ALL x. P x --> (EX y. P y & ~ ((m y::nat) <= m x))
wenzelm
parents: 12023
diff changeset
   165
    ==> EX y. P y & ~ (m y < m k + n)"
wenzelm
parents: 12023
diff changeset
   166
  apply (induct_tac n)
wenzelm
parents: 12023
diff changeset
   167
   apply force
wenzelm
parents: 12023
diff changeset
   168
  apply (force simp add: le_Suc_eq)
wenzelm
parents: 12023
diff changeset
   169
  done
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   170
12298
wenzelm
parents: 12023
diff changeset
   171
lemma ex_has_greatest_nat:
wenzelm
parents: 12023
diff changeset
   172
  "P k ==> ALL y. P y --> m y < b
wenzelm
parents: 12023
diff changeset
   173
    ==> EX x. P x & (ALL y. P y --> (m y::nat) <= m x)"
wenzelm
parents: 12023
diff changeset
   174
  apply (rule ccontr)
wenzelm
parents: 12023
diff changeset
   175
  apply (cut_tac P = P and n = "b - m k" in ex_has_greatest_nat_lemma)
wenzelm
parents: 12023
diff changeset
   176
    apply (subgoal_tac [3] "m k <= b")
wenzelm
parents: 12023
diff changeset
   177
     apply auto
wenzelm
parents: 12023
diff changeset
   178
  done
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   179
12298
wenzelm
parents: 12023
diff changeset
   180
lemma GreatestM_nat_lemma:
wenzelm
parents: 12023
diff changeset
   181
  "P k ==> ALL y. P y --> m y < b
wenzelm
parents: 12023
diff changeset
   182
    ==> P (GreatestM m P) & (ALL y. P y --> (m y::nat) <= m (GreatestM m P))"
wenzelm
parents: 12023
diff changeset
   183
  apply (unfold GreatestM_def)
wenzelm
parents: 12023
diff changeset
   184
  apply (rule someI_ex)
wenzelm
parents: 12023
diff changeset
   185
  apply (erule ex_has_greatest_nat)
wenzelm
parents: 12023
diff changeset
   186
  apply assumption
wenzelm
parents: 12023
diff changeset
   187
  done
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   188
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   189
lemmas GreatestM_natI = GreatestM_nat_lemma [THEN conjunct1, standard]
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   190
12298
wenzelm
parents: 12023
diff changeset
   191
lemma GreatestM_nat_le:
wenzelm
parents: 12023
diff changeset
   192
  "P x ==> ALL y. P y --> m y < b
wenzelm
parents: 12023
diff changeset
   193
    ==> (m x::nat) <= m (GreatestM m P)"
wenzelm
parents: 12023
diff changeset
   194
  apply (blast dest: GreatestM_nat_lemma [THEN conjunct2, THEN spec])
wenzelm
parents: 12023
diff changeset
   195
  done
wenzelm
parents: 12023
diff changeset
   196
wenzelm
parents: 12023
diff changeset
   197
wenzelm
parents: 12023
diff changeset
   198
text {* \medskip Specialization to @{text GREATEST}. *}
wenzelm
parents: 12023
diff changeset
   199
wenzelm
parents: 12023
diff changeset
   200
lemma GreatestI: "P (k::nat) ==> ALL y. P y --> y < b ==> P (GREATEST x. P x)"
wenzelm
parents: 12023
diff changeset
   201
  apply (unfold Greatest_def)
wenzelm
parents: 12023
diff changeset
   202
  apply (rule GreatestM_natI)
wenzelm
parents: 12023
diff changeset
   203
   apply auto
wenzelm
parents: 12023
diff changeset
   204
  done
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   205
12298
wenzelm
parents: 12023
diff changeset
   206
lemma Greatest_le:
wenzelm
parents: 12023
diff changeset
   207
    "P x ==> ALL y. P y --> y < b ==> (x::nat) <= (GREATEST x. P x)"
wenzelm
parents: 12023
diff changeset
   208
  apply (unfold Greatest_def)
wenzelm
parents: 12023
diff changeset
   209
  apply (rule GreatestM_nat_le)
wenzelm
parents: 12023
diff changeset
   210
   apply auto
wenzelm
parents: 12023
diff changeset
   211
  done
wenzelm
parents: 12023
diff changeset
   212
wenzelm
parents: 12023
diff changeset
   213
wenzelm
parents: 12023
diff changeset
   214
subsection {* The Meson proof procedure *}
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   215
12298
wenzelm
parents: 12023
diff changeset
   216
subsubsection {* Negation Normal Form *}
wenzelm
parents: 12023
diff changeset
   217
wenzelm
parents: 12023
diff changeset
   218
text {* de Morgan laws *}
wenzelm
parents: 12023
diff changeset
   219
wenzelm
parents: 12023
diff changeset
   220
lemma meson_not_conjD: "~(P&Q) ==> ~P | ~Q"
wenzelm
parents: 12023
diff changeset
   221
  and meson_not_disjD: "~(P|Q) ==> ~P & ~Q"
wenzelm
parents: 12023
diff changeset
   222
  and meson_not_notD: "~~P ==> P"
wenzelm
parents: 12023
diff changeset
   223
  and meson_not_allD: "!!P. ~(ALL x. P(x)) ==> EX x. ~P(x)"
wenzelm
parents: 12023
diff changeset
   224
  and meson_not_exD: "!!P. ~(EX x. P(x)) ==> ALL x. ~P(x)"
wenzelm
parents: 12023
diff changeset
   225
  by fast+
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   226
12298
wenzelm
parents: 12023
diff changeset
   227
text {* Removal of @{text "-->"} and @{text "<->"} (positive and
wenzelm
parents: 12023
diff changeset
   228
negative occurrences) *}
wenzelm
parents: 12023
diff changeset
   229
wenzelm
parents: 12023
diff changeset
   230
lemma meson_imp_to_disjD: "P-->Q ==> ~P | Q"
wenzelm
parents: 12023
diff changeset
   231
  and meson_not_impD: "~(P-->Q) ==> P & ~Q"
wenzelm
parents: 12023
diff changeset
   232
  and meson_iff_to_disjD: "P=Q ==> (~P | Q) & (~Q | P)"
wenzelm
parents: 12023
diff changeset
   233
  and meson_not_iffD: "~(P=Q) ==> (P | Q) & (~P | ~Q)"
wenzelm
parents: 12023
diff changeset
   234
    -- {* Much more efficient than @{prop "(P & ~Q) | (Q & ~P)"} for computing CNF *}
wenzelm
parents: 12023
diff changeset
   235
  by fast+
wenzelm
parents: 12023
diff changeset
   236
wenzelm
parents: 12023
diff changeset
   237
wenzelm
parents: 12023
diff changeset
   238
subsubsection {* Pulling out the existential quantifiers *}
wenzelm
parents: 12023
diff changeset
   239
wenzelm
parents: 12023
diff changeset
   240
text {* Conjunction *}
wenzelm
parents: 12023
diff changeset
   241
wenzelm
parents: 12023
diff changeset
   242
lemma meson_conj_exD1: "!!P Q. (EX x. P(x)) & Q ==> EX x. P(x) & Q"
wenzelm
parents: 12023
diff changeset
   243
  and meson_conj_exD2: "!!P Q. P & (EX x. Q(x)) ==> EX x. P & Q(x)"
wenzelm
parents: 12023
diff changeset
   244
  by fast+
wenzelm
parents: 12023
diff changeset
   245
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   246
12298
wenzelm
parents: 12023
diff changeset
   247
text {* Disjunction *}
wenzelm
parents: 12023
diff changeset
   248
wenzelm
parents: 12023
diff changeset
   249
lemma meson_disj_exD: "!!P Q. (EX x. P(x)) | (EX x. Q(x)) ==> EX x. P(x) | Q(x)"
wenzelm
parents: 12023
diff changeset
   250
  -- {* DO NOT USE with forall-Skolemization: makes fewer schematic variables!! *}
wenzelm
parents: 12023
diff changeset
   251
  -- {* With ex-Skolemization, makes fewer Skolem constants *}
wenzelm
parents: 12023
diff changeset
   252
  and meson_disj_exD1: "!!P Q. (EX x. P(x)) | Q ==> EX x. P(x) | Q"
wenzelm
parents: 12023
diff changeset
   253
  and meson_disj_exD2: "!!P Q. P | (EX x. Q(x)) ==> EX x. P | Q(x)"
wenzelm
parents: 12023
diff changeset
   254
  by fast+
wenzelm
parents: 12023
diff changeset
   255
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   256
12298
wenzelm
parents: 12023
diff changeset
   257
subsubsection {* Generating clauses for the Meson Proof Procedure *}
wenzelm
parents: 12023
diff changeset
   258
wenzelm
parents: 12023
diff changeset
   259
text {* Disjunctions *}
wenzelm
parents: 12023
diff changeset
   260
wenzelm
parents: 12023
diff changeset
   261
lemma meson_disj_assoc: "(P|Q)|R ==> P|(Q|R)"
wenzelm
parents: 12023
diff changeset
   262
  and meson_disj_comm: "P|Q ==> Q|P"
wenzelm
parents: 12023
diff changeset
   263
  and meson_disj_FalseD1: "False|P ==> P"
wenzelm
parents: 12023
diff changeset
   264
  and meson_disj_FalseD2: "P|False ==> P"
wenzelm
parents: 12023
diff changeset
   265
  by fast+
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   266
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   267
use "meson_lemmas.ML"
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   268
use "Tools/meson.ML"
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   269
setup meson_setup
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   270
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents: 13764
diff changeset
   271
use "Tools/specification_package.ML"
65ec3f73d00b Added package for definition by specification.
skalberg
parents: 13764
diff changeset
   272
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
diff changeset
   273
end