| author | berghofe | 
| Mon, 19 Nov 2001 17:40:45 +0100 | |
| changeset 12238 | 09966ccbc84c | 
| parent 12161 | ea4fbf26a945 | 
| child 12240 | 0760eda193c4 | 
| permissions | -rw-r--r-- | 
| 923 | 1 | (* Title: HOL/HOL.thy | 
| 2 | ID: $Id$ | |
| 11750 | 3 | Author: Tobias Nipkow, Markus Wenzel, and Larry Paulson | 
| 4 | *) | |
| 923 | 5 | |
| 11750 | 6 | header {* The basis of Higher-Order Logic *}
 | 
| 923 | 7 | |
| 7357 | 8 | theory HOL = CPure | 
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11438diff
changeset | 9 | files ("HOL_lemmas.ML") ("cladata.ML") ("blastdata.ML") ("simpdata.ML"):
 | 
| 923 | 10 | |
| 2260 | 11 | |
| 11750 | 12 | subsection {* Primitive logic *}
 | 
| 13 | ||
| 14 | subsubsection {* Core syntax *}
 | |
| 2260 | 15 | |
| 3947 | 16 | global | 
| 17 | ||
| 7357 | 18 | classes "term" < logic | 
| 19 | defaultsort "term" | |
| 923 | 20 | |
| 7357 | 21 | typedecl bool | 
| 923 | 22 | |
| 23 | arities | |
| 7357 | 24 | bool :: "term" | 
| 25 |   fun :: ("term", "term") "term"
 | |
| 923 | 26 | |
| 11750 | 27 | judgment | 
| 28 |   Trueprop      :: "bool => prop"                   ("(_)" 5)
 | |
| 923 | 29 | |
| 11750 | 30 | consts | 
| 7357 | 31 |   Not           :: "bool => bool"                   ("~ _" [40] 40)
 | 
| 32 | True :: bool | |
| 33 | False :: bool | |
| 34 |   If            :: "[bool, 'a, 'a] => 'a"           ("(if (_)/ then (_)/ else (_))" 10)
 | |
| 3947 | 35 | arbitrary :: 'a | 
| 923 | 36 | |
| 11432 
8a203ae6efe3
added "The" (definite description operator) (by Larry);
 wenzelm parents: 
10489diff
changeset | 37 |   The           :: "('a => bool) => 'a"
 | 
| 7357 | 38 |   All           :: "('a => bool) => bool"           (binder "ALL " 10)
 | 
| 39 |   Ex            :: "('a => bool) => bool"           (binder "EX " 10)
 | |
| 40 |   Ex1           :: "('a => bool) => bool"           (binder "EX! " 10)
 | |
| 41 | Let :: "['a, 'a => 'b] => 'b" | |
| 923 | 42 | |
| 7357 | 43 | "=" :: "['a, 'a] => bool" (infixl 50) | 
| 44 | & :: "[bool, bool] => bool" (infixr 35) | |
| 45 | "|" :: "[bool, bool] => bool" (infixr 30) | |
| 46 | --> :: "[bool, bool] => bool" (infixr 25) | |
| 923 | 47 | |
| 10432 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 48 | local | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 49 | |
| 2260 | 50 | |
| 11750 | 51 | subsubsection {* Additional concrete syntax *}
 | 
| 2260 | 52 | |
| 4868 | 53 | nonterminals | 
| 923 | 54 | letbinds letbind | 
| 55 | case_syn cases_syn | |
| 56 | ||
| 57 | syntax | |
| 7357 | 58 | ~= :: "['a, 'a] => bool" (infixl 50) | 
| 11432 
8a203ae6efe3
added "The" (definite description operator) (by Larry);
 wenzelm parents: 
10489diff
changeset | 59 |   "_The"        :: "[pttrn, bool] => 'a"                 ("(3THE _./ _)" [0, 10] 10)
 | 
| 923 | 60 | |
| 7357 | 61 |   "_bind"       :: "[pttrn, 'a] => letbind"              ("(2_ =/ _)" 10)
 | 
| 62 |   ""            :: "letbind => letbinds"                 ("_")
 | |
| 63 |   "_binds"      :: "[letbind, letbinds] => letbinds"     ("_;/ _")
 | |
| 64 |   "_Let"        :: "[letbinds, 'a] => 'a"                ("(let (_)/ in (_))" 10)
 | |
| 923 | 65 | |
| 9060 
b0dd884b1848
rename @case to _case_syntax (improves on low-level errors);
 wenzelm parents: 
8959diff
changeset | 66 |   "_case_syntax":: "['a, cases_syn] => 'b"               ("(case _ of/ _)" 10)
 | 
| 
b0dd884b1848
rename @case to _case_syntax (improves on low-level errors);
 wenzelm parents: 
8959diff
changeset | 67 |   "_case1"      :: "['a, 'b] => case_syn"                ("(2_ =>/ _)" 10)
 | 
| 7357 | 68 |   ""            :: "case_syn => cases_syn"               ("_")
 | 
| 9060 
b0dd884b1848
rename @case to _case_syntax (improves on low-level errors);
 wenzelm parents: 
8959diff
changeset | 69 |   "_case2"      :: "[case_syn, cases_syn] => cases_syn"  ("_/ | _")
 | 
| 923 | 70 | |
| 71 | translations | |
| 7238 
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
 wenzelm parents: 
7220diff
changeset | 72 | "x ~= y" == "~ (x = y)" | 
| 11432 
8a203ae6efe3
added "The" (definite description operator) (by Larry);
 wenzelm parents: 
10489diff
changeset | 73 | "THE x. P" == "The (%x. P)" | 
| 923 | 74 | "_Let (_binds b bs) e" == "_Let b (_Let bs e)" | 
| 1114 | 75 | "let x = a in e" == "Let a (%x. e)" | 
| 923 | 76 | |
| 3820 | 77 | syntax ("" output)
 | 
| 11687 | 78 | "=" :: "['a, 'a] => bool" (infix 50) | 
| 79 | "~=" :: "['a, 'a] => bool" (infix 50) | |
| 2260 | 80 | |
| 12114 
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
 wenzelm parents: 
12023diff
changeset | 81 | syntax (xsymbols) | 
| 11687 | 82 |   Not           :: "bool => bool"                        ("\<not> _" [40] 40)
 | 
| 83 | "op &" :: "[bool, bool] => bool" (infixr "\<and>" 35) | |
| 84 | "op |" :: "[bool, bool] => bool" (infixr "\<or>" 30) | |
| 12114 
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
 wenzelm parents: 
12023diff
changeset | 85 | "op -->" :: "[bool, bool] => bool" (infixr "\<longrightarrow>" 25) | 
| 11687 | 86 | "op ~=" :: "['a, 'a] => bool" (infix "\<noteq>" 50) | 
| 87 |   "ALL "        :: "[idts, bool] => bool"                ("(3\<forall>_./ _)" [0, 10] 10)
 | |
| 88 |   "EX "         :: "[idts, bool] => bool"                ("(3\<exists>_./ _)" [0, 10] 10)
 | |
| 89 |   "EX! "        :: "[idts, bool] => bool"                ("(3\<exists>!_./ _)" [0, 10] 10)
 | |
| 90 |   "_case1"      :: "['a, 'b] => case_syn"                ("(2_ \<Rightarrow>/ _)" 10)
 | |
| 9060 
b0dd884b1848
rename @case to _case_syntax (improves on low-level errors);
 wenzelm parents: 
8959diff
changeset | 91 | (*"_case2"      :: "[case_syn, cases_syn] => cases_syn"  ("_/ \\<orelse> _")*)
 | 
| 2372 | 92 | |
| 12114 
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
 wenzelm parents: 
12023diff
changeset | 93 | syntax (xsymbols output) | 
| 11687 | 94 | "op ~=" :: "['a, 'a] => bool" (infix "\<noteq>" 50) | 
| 3820 | 95 | |
| 6340 | 96 | syntax (HTML output) | 
| 11687 | 97 |   Not           :: "bool => bool"                        ("\<not> _" [40] 40)
 | 
| 6340 | 98 | |
| 7238 
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
 wenzelm parents: 
7220diff
changeset | 99 | syntax (HOL) | 
| 7357 | 100 |   "ALL "        :: "[idts, bool] => bool"                ("(3! _./ _)" [0, 10] 10)
 | 
| 101 |   "EX "         :: "[idts, bool] => bool"                ("(3? _./ _)" [0, 10] 10)
 | |
| 102 |   "EX! "        :: "[idts, bool] => bool"                ("(3?! _./ _)" [0, 10] 10)
 | |
| 7238 
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
 wenzelm parents: 
7220diff
changeset | 103 | |
| 
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
 wenzelm parents: 
7220diff
changeset | 104 | |
| 11750 | 105 | subsubsection {* Axioms and basic definitions *}
 | 
| 2260 | 106 | |
| 7357 | 107 | axioms | 
| 108 | eq_reflection: "(x=y) ==> (x==y)" | |
| 923 | 109 | |
| 7357 | 110 | refl: "t = (t::'a)" | 
| 111 | subst: "[| s = t; P(s) |] ==> P(t::'a)" | |
| 6289 | 112 | |
| 7357 | 113 | ext: "(!!x::'a. (f x ::'b) = g x) ==> (%x. f x) = (%x. g x)" | 
| 11750 | 114 |     -- {* Extensionality is built into the meta-logic, and this rule expresses *}
 | 
| 115 |     -- {* a related property.  It is an eta-expanded version of the traditional *}
 | |
| 116 |     -- {* rule, and similar to the ABS rule of HOL *}
 | |
| 6289 | 117 | |
| 11432 
8a203ae6efe3
added "The" (definite description operator) (by Larry);
 wenzelm parents: 
10489diff
changeset | 118 | the_eq_trivial: "(THE x. x = a) = (a::'a)" | 
| 923 | 119 | |
| 7357 | 120 | impI: "(P ==> Q) ==> P-->Q" | 
| 121 | mp: "[| P-->Q; P |] ==> Q" | |
| 923 | 122 | |
| 123 | defs | |
| 7357 | 124 | True_def: "True == ((%x::bool. x) = (%x. x))" | 
| 125 | All_def: "All(P) == (P = (%x. True))" | |
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11438diff
changeset | 126 | Ex_def: "Ex(P) == !Q. (!x. P x --> Q) --> Q" | 
| 7357 | 127 | False_def: "False == (!P. P)" | 
| 128 | not_def: "~ P == P-->False" | |
| 129 | and_def: "P & Q == !R. (P-->Q-->R) --> R" | |
| 130 | or_def: "P | Q == !R. (P-->R) --> (Q-->R) --> R" | |
| 131 | Ex1_def: "Ex1(P) == ? x. P(x) & (! y. P(y) --> y=x)" | |
| 923 | 132 | |
| 7357 | 133 | axioms | 
| 134 | iff: "(P-->Q) --> (Q-->P) --> (P=Q)" | |
| 135 | True_or_False: "(P=True) | (P=False)" | |
| 923 | 136 | |
| 137 | defs | |
| 7357 | 138 | Let_def: "Let s f == f(s)" | 
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11438diff
changeset | 139 | if_def: "If P x y == THE z::'a. (P=True --> z=x) & (P=False --> z=y)" | 
| 5069 | 140 | |
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11438diff
changeset | 141 | arbitrary_def: "False ==> arbitrary == (THE x. False)" | 
| 11750 | 142 |     -- {* @{term arbitrary} is completely unspecified, but is made to appear as a
 | 
| 143 | definition syntactically *} | |
| 923 | 144 | |
| 3320 | 145 | |
| 11750 | 146 | subsubsection {* Generic algebraic operations *}
 | 
| 4868 | 147 | |
| 11750 | 148 | axclass zero < "term" | 
| 149 | axclass one < "term" | |
| 150 | axclass plus < "term" | |
| 151 | axclass minus < "term" | |
| 152 | axclass times < "term" | |
| 153 | axclass inverse < "term" | |
| 154 | ||
| 155 | global | |
| 156 | ||
| 157 | consts | |
| 158 |   "0"           :: "'a::zero"                       ("0")
 | |
| 159 |   "1"           :: "'a::one"                        ("1")
 | |
| 160 | "+" :: "['a::plus, 'a] => 'a" (infixl 65) | |
| 161 | - :: "['a::minus, 'a] => 'a" (infixl 65) | |
| 162 |   uminus        :: "['a::minus] => 'a"              ("- _" [81] 80)
 | |
| 163 | * :: "['a::times, 'a] => 'a" (infixl 70) | |
| 164 | ||
| 165 | local | |
| 166 | ||
| 167 | typed_print_translation {*
 | |
| 168 | let | |
| 169 | fun tr' c = (c, fn show_sorts => fn T => fn ts => | |
| 170 | if T = dummyT orelse not (! show_types) andalso can Term.dest_Type T then raise Match | |
| 171 | else Syntax.const Syntax.constrainC $ Syntax.const c $ Syntax.term_of_typ show_sorts T); | |
| 172 | in [tr' "0", tr' "1"] end; | |
| 173 | *} -- {* show types that are presumably too general *}
 | |
| 174 | ||
| 175 | ||
| 176 | consts | |
| 177 | abs :: "'a::minus => 'a" | |
| 178 | inverse :: "'a::inverse => 'a" | |
| 179 | divide :: "['a::inverse, 'a] => 'a" (infixl "'/" 70) | |
| 180 | ||
| 181 | syntax (xsymbols) | |
| 182 |   abs :: "'a::minus => 'a"    ("\<bar>_\<bar>")
 | |
| 183 | syntax (HTML output) | |
| 184 |   abs :: "'a::minus => 'a"    ("\<bar>_\<bar>")
 | |
| 185 | ||
| 186 | axclass plus_ac0 < plus, zero | |
| 187 | commute: "x + y = y + x" | |
| 188 | assoc: "(x + y) + z = x + (y + z)" | |
| 189 | zero: "0 + x = x" | |
| 190 | ||
| 191 | ||
| 192 | subsection {* Theory and package setup *}
 | |
| 193 | ||
| 194 | subsubsection {* Basic lemmas *}
 | |
| 4868 | 195 | |
| 9736 | 196 | use "HOL_lemmas.ML" | 
| 11687 | 197 | theorems case_split = case_split_thm [case_names True False] | 
| 9869 | 198 | |
| 11750 | 199 | declare trans [trans] | 
| 200 | declare impE [CPure.elim] iffD1 [CPure.elim] iffD2 [CPure.elim] | |
| 201 | ||
| 11438 
3d9222b80989
declare trans [trans]  (*overridden in theory Calculation*);
 wenzelm parents: 
11432diff
changeset | 202 | |
| 11750 | 203 | subsubsection {* Atomizing meta-level connectives *}
 | 
| 204 | ||
| 205 | lemma atomize_all [atomize]: "(!!x. P x) == Trueprop (ALL x. P x)" | |
| 12003 | 206 | proof | 
| 9488 | 207 | assume "!!x. P x" | 
| 10383 | 208 | show "ALL x. P x" by (rule allI) | 
| 9488 | 209 | next | 
| 210 | assume "ALL x. P x" | |
| 10383 | 211 | thus "!!x. P x" by (rule allE) | 
| 9488 | 212 | qed | 
| 213 | ||
| 11750 | 214 | lemma atomize_imp [atomize]: "(A ==> B) == Trueprop (A --> B)" | 
| 12003 | 215 | proof | 
| 9488 | 216 | assume r: "A ==> B" | 
| 10383 | 217 | show "A --> B" by (rule impI) (rule r) | 
| 9488 | 218 | next | 
| 219 | assume "A --> B" and A | |
| 10383 | 220 | thus B by (rule mp) | 
| 9488 | 221 | qed | 
| 222 | ||
| 11750 | 223 | lemma atomize_eq [atomize]: "(x == y) == Trueprop (x = y)" | 
| 12003 | 224 | proof | 
| 10432 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 225 | assume "x == y" | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 226 | show "x = y" by (unfold prems) (rule refl) | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 227 | next | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 228 | assume "x = y" | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 229 | thus "x == y" by (rule eq_reflection) | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 230 | qed | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 231 | |
| 12023 | 232 | lemma atomize_conj [atomize]: | 
| 233 | "(!!C. (A ==> B ==> PROP C) ==> PROP C) == Trueprop (A & B)" | |
| 12003 | 234 | proof | 
| 11953 | 235 | assume "!!C. (A ==> B ==> PROP C) ==> PROP C" | 
| 236 | show "A & B" by (rule conjI) | |
| 237 | next | |
| 238 | fix C | |
| 239 | assume "A & B" | |
| 240 | assume "A ==> B ==> PROP C" | |
| 241 | thus "PROP C" | |
| 242 | proof this | |
| 243 | show A by (rule conjunct1) | |
| 244 | show B by (rule conjunct2) | |
| 245 | qed | |
| 246 | qed | |
| 247 | ||
| 11750 | 248 | |
| 249 | subsubsection {* Classical Reasoner setup *}
 | |
| 9529 | 250 | |
| 10383 | 251 | use "cladata.ML" | 
| 252 | setup hypsubst_setup | |
| 11977 | 253 | |
| 11770 | 254 | declare atomize_all [symmetric, rulify] atomize_imp [symmetric, rulify] | 
| 11977 | 255 | |
| 10383 | 256 | setup Classical.setup | 
| 257 | setup clasetup | |
| 258 | ||
| 11977 | 259 | declare ext [intro?] | 
| 260 | declare disjI1 [elim?] disjI2 [elim?] ex1_implies_ex [elim?] sym [elim?] | |
| 261 | ||
| 9869 | 262 | use "blastdata.ML" | 
| 263 | setup Blast.setup | |
| 4868 | 264 | |
| 11750 | 265 | |
| 266 | subsubsection {* Simplifier setup *}
 | |
| 267 | ||
| 9869 | 268 | use "simpdata.ML" | 
| 269 | setup Simplifier.setup | |
| 270 | setup "Simplifier.method_setup Splitter.split_modifiers" setup simpsetup | |
| 271 | setup Splitter.setup setup Clasimp.setup | |
| 272 | ||
| 11750 | 273 | |
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 274 | subsubsection {* Generic cases and induction *}
 | 
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 275 | |
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 276 | constdefs | 
| 11989 | 277 |   induct_forall :: "('a => bool) => bool"
 | 
| 278 | "induct_forall P == \<forall>x. P x" | |
| 279 | induct_implies :: "bool => bool => bool" | |
| 280 | "induct_implies A B == A --> B" | |
| 281 | induct_equal :: "'a => 'a => bool" | |
| 282 | "induct_equal x y == x = y" | |
| 283 | induct_conj :: "bool => bool => bool" | |
| 284 | "induct_conj A B == A & B" | |
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 285 | |
| 11989 | 286 | lemma induct_forall_eq: "(!!x. P x) == Trueprop (induct_forall (\<lambda>x. P x))" | 
| 287 | by (simp only: atomize_all induct_forall_def) | |
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 288 | |
| 11989 | 289 | lemma induct_implies_eq: "(A ==> B) == Trueprop (induct_implies A B)" | 
| 290 | by (simp only: atomize_imp induct_implies_def) | |
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 291 | |
| 11989 | 292 | lemma induct_equal_eq: "(x == y) == Trueprop (induct_equal x y)" | 
| 293 | by (simp only: atomize_eq induct_equal_def) | |
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 294 | |
| 11989 | 295 | lemma induct_forall_conj: "induct_forall (\<lambda>x. induct_conj (A x) (B x)) = | 
| 296 | induct_conj (induct_forall A) (induct_forall B)" | |
| 297 | by (unfold induct_forall_def induct_conj_def) blast | |
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 298 | |
| 11989 | 299 | lemma induct_implies_conj: "induct_implies C (induct_conj A B) = | 
| 300 | induct_conj (induct_implies C A) (induct_implies C B)" | |
| 301 | by (unfold induct_implies_def induct_conj_def) blast | |
| 302 | ||
| 303 | lemma induct_conj_curry: "(induct_conj A B ==> C) == (A ==> B ==> C)" | |
| 304 | by (simp only: atomize_imp atomize_eq induct_conj_def) (rule equal_intr_rule, blast+) | |
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 305 | |
| 11989 | 306 | lemma induct_impliesI: "(A ==> B) ==> induct_implies A B" | 
| 307 | by (simp add: induct_implies_def) | |
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 308 | |
| 12161 | 309 | lemmas induct_atomize = atomize_conj induct_forall_eq induct_implies_eq induct_equal_eq | 
| 310 | lemmas induct_rulify1 [symmetric, standard] = induct_forall_eq induct_implies_eq induct_equal_eq | |
| 311 | lemmas induct_rulify2 = induct_forall_def induct_implies_def induct_equal_def induct_conj_def | |
| 11989 | 312 | lemmas induct_conj = induct_forall_conj induct_implies_conj induct_conj_curry | 
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 313 | |
| 11989 | 314 | hide const induct_forall induct_implies induct_equal induct_conj | 
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 315 | |
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 316 | |
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 317 | text {* Method setup. *}
 | 
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 318 | |
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 319 | ML {*
 | 
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 320 | structure InductMethod = InductMethodFun | 
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 321 | (struct | 
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 322 | val dest_concls = HOLogic.dest_concls; | 
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 323 | val cases_default = thm "case_split"; | 
| 12161 | 324 | val local_imp_def = thm "induct_implies_def"; | 
| 11989 | 325 | val local_impI = thm "induct_impliesI"; | 
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 326 | val conjI = thm "conjI"; | 
| 11989 | 327 | val atomize = thms "induct_atomize"; | 
| 328 | val rulify1 = thms "induct_rulify1"; | |
| 329 | val rulify2 = thms "induct_rulify2"; | |
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 330 | end); | 
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 331 | *} | 
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 332 | |
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 333 | setup InductMethod.setup | 
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 334 | |
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 335 | |
| 11750 | 336 | subsection {* Order signatures and orders *}
 | 
| 337 | ||
| 338 | axclass | |
| 339 | ord < "term" | |
| 340 | ||
| 341 | syntax | |
| 342 |   "op <"        :: "['a::ord, 'a] => bool"             ("op <")
 | |
| 343 |   "op <="       :: "['a::ord, 'a] => bool"             ("op <=")
 | |
| 344 | ||
| 345 | global | |
| 346 | ||
| 347 | consts | |
| 348 |   "op <"        :: "['a::ord, 'a] => bool"             ("(_/ < _)"  [50, 51] 50)
 | |
| 349 |   "op <="       :: "['a::ord, 'a] => bool"             ("(_/ <= _)" [50, 51] 50)
 | |
| 350 | ||
| 351 | local | |
| 352 | ||
| 12114 
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
 wenzelm parents: 
12023diff
changeset | 353 | syntax (xsymbols) | 
| 11750 | 354 |   "op <="       :: "['a::ord, 'a] => bool"             ("op \<le>")
 | 
| 355 |   "op <="       :: "['a::ord, 'a] => bool"             ("(_/ \<le> _)"  [50, 51] 50)
 | |
| 356 | ||
| 357 | (*Tell blast about overloading of < and <= to reduce the risk of | |
| 358 | its applying a rule for the wrong type*) | |
| 359 | ML {*
 | |
| 360 | Blast.overloaded ("op <" , domain_type);
 | |
| 361 | Blast.overloaded ("op <=", domain_type);
 | |
| 362 | *} | |
| 363 | ||
| 364 | ||
| 365 | subsubsection {* Monotonicity *}
 | |
| 366 | ||
| 367 | constdefs | |
| 368 | mono :: "['a::ord => 'b::ord] => bool" | |
| 369 | "mono f == ALL A B. A <= B --> f A <= f B" | |
| 370 | ||
| 371 | lemma monoI [intro?]: "(!!A B. A <= B ==> f A <= f B) ==> mono f" | |
| 372 | by (unfold mono_def) blast | |
| 373 | ||
| 374 | lemma monoD [dest?]: "mono f ==> A <= B ==> f A <= f B" | |
| 375 | by (unfold mono_def) blast | |
| 376 | ||
| 377 | constdefs | |
| 378 | min :: "['a::ord, 'a] => 'a" | |
| 379 | "min a b == (if a <= b then a else b)" | |
| 380 | max :: "['a::ord, 'a] => 'a" | |
| 381 | "max a b == (if a <= b then b else a)" | |
| 382 | ||
| 383 | lemma min_leastL: "(!!x. least <= x) ==> min least x = least" | |
| 384 | by (simp add: min_def) | |
| 385 | ||
| 386 | lemma min_of_mono: | |
| 387 | "ALL x y. (f x <= f y) = (x <= y) ==> min (f m) (f n) = f (min m n)" | |
| 388 | by (simp add: min_def) | |
| 389 | ||
| 390 | lemma max_leastL: "(!!x. least <= x) ==> max least x = x" | |
| 391 | by (simp add: max_def) | |
| 392 | ||
| 393 | lemma max_of_mono: | |
| 394 | "ALL x y. (f x <= f y) = (x <= y) ==> max (f m) (f n) = f (max m n)" | |
| 395 | by (simp add: max_def) | |
| 396 | ||
| 397 | ||
| 398 | subsubsection "Orders" | |
| 399 | ||
| 400 | axclass order < ord | |
| 401 | order_refl [iff]: "x <= x" | |
| 402 | order_trans: "x <= y ==> y <= z ==> x <= z" | |
| 403 | order_antisym: "x <= y ==> y <= x ==> x = y" | |
| 404 | order_less_le: "(x < y) = (x <= y & x ~= y)" | |
| 405 | ||
| 406 | ||
| 407 | text {* Reflexivity. *}
 | |
| 408 | ||
| 409 | lemma order_eq_refl: "!!x::'a::order. x = y ==> x <= y" | |
| 410 |     -- {* This form is useful with the classical reasoner. *}
 | |
| 411 | apply (erule ssubst) | |
| 412 | apply (rule order_refl) | |
| 413 | done | |
| 414 | ||
| 415 | lemma order_less_irrefl [simp]: "~ x < (x::'a::order)" | |
| 416 | by (simp add: order_less_le) | |
| 417 | ||
| 418 | lemma order_le_less: "((x::'a::order) <= y) = (x < y | x = y)" | |
| 419 |     -- {* NOT suitable for iff, since it can cause PROOF FAILED. *}
 | |
| 420 | apply (simp add: order_less_le) | |
| 421 | apply (blast intro!: order_refl) | |
| 422 | done | |
| 423 | ||
| 424 | lemmas order_le_imp_less_or_eq = order_le_less [THEN iffD1, standard] | |
| 425 | ||
| 426 | lemma order_less_imp_le: "!!x::'a::order. x < y ==> x <= y" | |
| 427 | by (simp add: order_less_le) | |
| 428 | ||
| 429 | ||
| 430 | text {* Asymmetry. *}
 | |
| 431 | ||
| 432 | lemma order_less_not_sym: "(x::'a::order) < y ==> ~ (y < x)" | |
| 433 | by (simp add: order_less_le order_antisym) | |
| 434 | ||
| 435 | lemma order_less_asym: "x < (y::'a::order) ==> (~P ==> y < x) ==> P" | |
| 436 | apply (drule order_less_not_sym) | |
| 437 | apply (erule contrapos_np) | |
| 438 | apply simp | |
| 439 | done | |
| 440 | ||
| 441 | ||
| 442 | text {* Transitivity. *}
 | |
| 443 | ||
| 444 | lemma order_less_trans: "!!x::'a::order. [| x < y; y < z |] ==> x < z" | |
| 445 | apply (simp add: order_less_le) | |
| 446 | apply (blast intro: order_trans order_antisym) | |
| 447 | done | |
| 448 | ||
| 449 | lemma order_le_less_trans: "!!x::'a::order. [| x <= y; y < z |] ==> x < z" | |
| 450 | apply (simp add: order_less_le) | |
| 451 | apply (blast intro: order_trans order_antisym) | |
| 452 | done | |
| 453 | ||
| 454 | lemma order_less_le_trans: "!!x::'a::order. [| x < y; y <= z |] ==> x < z" | |
| 455 | apply (simp add: order_less_le) | |
| 456 | apply (blast intro: order_trans order_antisym) | |
| 457 | done | |
| 458 | ||
| 459 | ||
| 460 | text {* Useful for simplification, but too risky to include by default. *}
 | |
| 461 | ||
| 462 | lemma order_less_imp_not_less: "(x::'a::order) < y ==> (~ y < x) = True" | |
| 463 | by (blast elim: order_less_asym) | |
| 464 | ||
| 465 | lemma order_less_imp_triv: "(x::'a::order) < y ==> (y < x --> P) = True" | |
| 466 | by (blast elim: order_less_asym) | |
| 467 | ||
| 468 | lemma order_less_imp_not_eq: "(x::'a::order) < y ==> (x = y) = False" | |
| 469 | by auto | |
| 470 | ||
| 471 | lemma order_less_imp_not_eq2: "(x::'a::order) < y ==> (y = x) = False" | |
| 472 | by auto | |
| 473 | ||
| 474 | ||
| 475 | text {* Other operators. *}
 | |
| 476 | ||
| 477 | lemma min_leastR: "(!!x::'a::order. least <= x) ==> min x least = least" | |
| 478 | apply (simp add: min_def) | |
| 479 | apply (blast intro: order_antisym) | |
| 480 | done | |
| 481 | ||
| 482 | lemma max_leastR: "(!!x::'a::order. least <= x) ==> max x least = x" | |
| 483 | apply (simp add: max_def) | |
| 484 | apply (blast intro: order_antisym) | |
| 485 | done | |
| 486 | ||
| 487 | ||
| 488 | subsubsection {* Least value operator *}
 | |
| 489 | ||
| 490 | constdefs | |
| 491 |   Least :: "('a::ord => bool) => 'a"               (binder "LEAST " 10)
 | |
| 492 | "Least P == THE x. P x & (ALL y. P y --> x <= y)" | |
| 493 |     -- {* We can no longer use LeastM because the latter requires Hilbert-AC. *}
 | |
| 494 | ||
| 495 | lemma LeastI2: | |
| 496 | "[| P (x::'a::order); | |
| 497 | !!y. P y ==> x <= y; | |
| 498 | !!x. [| P x; ALL y. P y --> x \<le> y |] ==> Q x |] | |
| 499 | ==> Q (Least P)"; | |
| 500 | apply (unfold Least_def) | |
| 501 | apply (rule theI2) | |
| 502 | apply (blast intro: order_antisym)+ | |
| 503 | done | |
| 504 | ||
| 505 | lemma Least_equality: | |
| 506 | "[| P (k::'a::order); !!x. P x ==> k <= x |] ==> (LEAST x. P x) = k"; | |
| 507 | apply (simp add: Least_def) | |
| 508 | apply (rule the_equality) | |
| 509 | apply (auto intro!: order_antisym) | |
| 510 | done | |
| 511 | ||
| 512 | ||
| 513 | subsubsection "Linear / total orders" | |
| 514 | ||
| 515 | axclass linorder < order | |
| 516 | linorder_linear: "x <= y | y <= x" | |
| 517 | ||
| 518 | lemma linorder_less_linear: "!!x::'a::linorder. x<y | x=y | y<x" | |
| 519 | apply (simp add: order_less_le) | |
| 520 | apply (insert linorder_linear) | |
| 521 | apply blast | |
| 522 | done | |
| 523 | ||
| 524 | lemma linorder_cases [case_names less equal greater]: | |
| 525 | "((x::'a::linorder) < y ==> P) ==> (x = y ==> P) ==> (y < x ==> P) ==> P" | |
| 526 | apply (insert linorder_less_linear) | |
| 527 | apply blast | |
| 528 | done | |
| 529 | ||
| 530 | lemma linorder_not_less: "!!x::'a::linorder. (~ x < y) = (y <= x)" | |
| 531 | apply (simp add: order_less_le) | |
| 532 | apply (insert linorder_linear) | |
| 533 | apply (blast intro: order_antisym) | |
| 534 | done | |
| 535 | ||
| 536 | lemma linorder_not_le: "!!x::'a::linorder. (~ x <= y) = (y < x)" | |
| 537 | apply (simp add: order_less_le) | |
| 538 | apply (insert linorder_linear) | |
| 539 | apply (blast intro: order_antisym) | |
| 540 | done | |
| 541 | ||
| 542 | lemma linorder_neq_iff: "!!x::'a::linorder. (x ~= y) = (x<y | y<x)" | |
| 543 | apply (cut_tac x = x and y = y in linorder_less_linear) | |
| 544 | apply auto | |
| 545 | done | |
| 546 | ||
| 547 | lemma linorder_neqE: "x ~= (y::'a::linorder) ==> (x < y ==> R) ==> (y < x ==> R) ==> R" | |
| 548 | apply (simp add: linorder_neq_iff) | |
| 549 | apply blast | |
| 550 | done | |
| 551 | ||
| 552 | ||
| 553 | subsubsection "Min and max on (linear) orders" | |
| 554 | ||
| 555 | lemma min_same [simp]: "min (x::'a::order) x = x" | |
| 556 | by (simp add: min_def) | |
| 557 | ||
| 558 | lemma max_same [simp]: "max (x::'a::order) x = x" | |
| 559 | by (simp add: max_def) | |
| 560 | ||
| 561 | lemma le_max_iff_disj: "!!z::'a::linorder. (z <= max x y) = (z <= x | z <= y)" | |
| 562 | apply (simp add: max_def) | |
| 563 | apply (insert linorder_linear) | |
| 564 | apply (blast intro: order_trans) | |
| 565 | done | |
| 566 | ||
| 567 | lemma le_maxI1: "(x::'a::linorder) <= max x y" | |
| 568 | by (simp add: le_max_iff_disj) | |
| 569 | ||
| 570 | lemma le_maxI2: "(y::'a::linorder) <= max x y" | |
| 571 |     -- {* CANNOT use with @{text "[intro!]"} because blast will give PROOF FAILED. *}
 | |
| 572 | by (simp add: le_max_iff_disj) | |
| 573 | ||
| 574 | lemma less_max_iff_disj: "!!z::'a::linorder. (z < max x y) = (z < x | z < y)" | |
| 575 | apply (simp add: max_def order_le_less) | |
| 576 | apply (insert linorder_less_linear) | |
| 577 | apply (blast intro: order_less_trans) | |
| 578 | done | |
| 579 | ||
| 580 | lemma max_le_iff_conj [simp]: | |
| 581 | "!!z::'a::linorder. (max x y <= z) = (x <= z & y <= z)" | |
| 582 | apply (simp add: max_def) | |
| 583 | apply (insert linorder_linear) | |
| 584 | apply (blast intro: order_trans) | |
| 585 | done | |
| 586 | ||
| 587 | lemma max_less_iff_conj [simp]: | |
| 588 | "!!z::'a::linorder. (max x y < z) = (x < z & y < z)" | |
| 589 | apply (simp add: order_le_less max_def) | |
| 590 | apply (insert linorder_less_linear) | |
| 591 | apply (blast intro: order_less_trans) | |
| 592 | done | |
| 593 | ||
| 594 | lemma le_min_iff_conj [simp]: | |
| 595 | "!!z::'a::linorder. (z <= min x y) = (z <= x & z <= y)" | |
| 596 |     -- {* @{text "[iff]"} screws up a Q{text blast} in MiniML *}
 | |
| 597 | apply (simp add: min_def) | |
| 598 | apply (insert linorder_linear) | |
| 599 | apply (blast intro: order_trans) | |
| 600 | done | |
| 601 | ||
| 602 | lemma min_less_iff_conj [simp]: | |
| 603 | "!!z::'a::linorder. (z < min x y) = (z < x & z < y)" | |
| 604 | apply (simp add: order_le_less min_def) | |
| 605 | apply (insert linorder_less_linear) | |
| 606 | apply (blast intro: order_less_trans) | |
| 607 | done | |
| 608 | ||
| 609 | lemma min_le_iff_disj: "!!z::'a::linorder. (min x y <= z) = (x <= z | y <= z)" | |
| 610 | apply (simp add: min_def) | |
| 611 | apply (insert linorder_linear) | |
| 612 | apply (blast intro: order_trans) | |
| 613 | done | |
| 614 | ||
| 615 | lemma min_less_iff_disj: "!!z::'a::linorder. (min x y < z) = (x < z | y < z)" | |
| 616 | apply (simp add: min_def order_le_less) | |
| 617 | apply (insert linorder_less_linear) | |
| 618 | apply (blast intro: order_less_trans) | |
| 619 | done | |
| 620 | ||
| 621 | lemma split_min: | |
| 622 | "P (min (i::'a::linorder) j) = ((i <= j --> P(i)) & (~ i <= j --> P(j)))" | |
| 623 | by (simp add: min_def) | |
| 624 | ||
| 625 | lemma split_max: | |
| 626 | "P (max (i::'a::linorder) j) = ((i <= j --> P(j)) & (~ i <= j --> P(i)))" | |
| 627 | by (simp add: max_def) | |
| 628 | ||
| 629 | ||
| 630 | subsubsection "Bounded quantifiers" | |
| 631 | ||
| 632 | syntax | |
| 633 |   "_lessAll" :: "[idt, 'a, bool] => bool"   ("(3ALL _<_./ _)"  [0, 0, 10] 10)
 | |
| 634 |   "_lessEx"  :: "[idt, 'a, bool] => bool"   ("(3EX _<_./ _)"  [0, 0, 10] 10)
 | |
| 635 |   "_leAll"   :: "[idt, 'a, bool] => bool"   ("(3ALL _<=_./ _)" [0, 0, 10] 10)
 | |
| 636 |   "_leEx"    :: "[idt, 'a, bool] => bool"   ("(3EX _<=_./ _)" [0, 0, 10] 10)
 | |
| 637 | ||
| 12114 
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
 wenzelm parents: 
12023diff
changeset | 638 | syntax (xsymbols) | 
| 11750 | 639 |   "_lessAll" :: "[idt, 'a, bool] => bool"   ("(3\<forall>_<_./ _)"  [0, 0, 10] 10)
 | 
| 640 |   "_lessEx"  :: "[idt, 'a, bool] => bool"   ("(3\<exists>_<_./ _)"  [0, 0, 10] 10)
 | |
| 641 |   "_leAll"   :: "[idt, 'a, bool] => bool"   ("(3\<forall>_\<le>_./ _)" [0, 0, 10] 10)
 | |
| 642 |   "_leEx"    :: "[idt, 'a, bool] => bool"   ("(3\<exists>_\<le>_./ _)" [0, 0, 10] 10)
 | |
| 643 | ||
| 644 | syntax (HOL) | |
| 645 |   "_lessAll" :: "[idt, 'a, bool] => bool"   ("(3! _<_./ _)"  [0, 0, 10] 10)
 | |
| 646 |   "_lessEx"  :: "[idt, 'a, bool] => bool"   ("(3? _<_./ _)"  [0, 0, 10] 10)
 | |
| 647 |   "_leAll"   :: "[idt, 'a, bool] => bool"   ("(3! _<=_./ _)" [0, 0, 10] 10)
 | |
| 648 |   "_leEx"    :: "[idt, 'a, bool] => bool"   ("(3? _<=_./ _)" [0, 0, 10] 10)
 | |
| 649 | ||
| 650 | translations | |
| 651 | "ALL x<y. P" => "ALL x. x < y --> P" | |
| 652 | "EX x<y. P" => "EX x. x < y & P" | |
| 653 | "ALL x<=y. P" => "ALL x. x <= y --> P" | |
| 654 | "EX x<=y. P" => "EX x. x <= y & P" | |
| 655 | ||
| 923 | 656 | end |