| author | wenzelm | 
| Tue, 07 Aug 2001 21:27:00 +0200 | |
| changeset 11472 | d08d4e17a5f6 | 
| parent 11451 | 8abfb4f7bd02 | 
| child 11687 | b0fe6e522559 | 
| permissions | -rw-r--r-- | 
| 923 | 1 | (* Title: HOL/HOL.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Tobias Nipkow | |
| 4 | Copyright 1993 University of Cambridge | |
| 5 | ||
| 2260 | 6 | Higher-Order Logic. | 
| 923 | 7 | *) | 
| 8 | ||
| 7357 | 9 | theory HOL = CPure | 
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11438diff
changeset | 10 | files ("HOL_lemmas.ML") ("cladata.ML") ("blastdata.ML") ("simpdata.ML"):
 | 
| 923 | 11 | |
| 2260 | 12 | |
| 13 | (** Core syntax **) | |
| 14 | ||
| 3947 | 15 | global | 
| 16 | ||
| 7357 | 17 | classes "term" < logic | 
| 18 | defaultsort "term" | |
| 923 | 19 | |
| 7357 | 20 | typedecl bool | 
| 923 | 21 | |
| 22 | arities | |
| 7357 | 23 | bool :: "term" | 
| 24 |   fun :: ("term", "term") "term"
 | |
| 923 | 25 | |
| 26 | consts | |
| 27 | ||
| 28 | (* Constants *) | |
| 29 | ||
| 7357 | 30 |   Trueprop      :: "bool => prop"                   ("(_)" 5)
 | 
| 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 | |
| 37 | (* Binders *) | |
| 38 | ||
| 11432 
8a203ae6efe3
added "The" (definite description operator) (by Larry);
 wenzelm parents: 
10489diff
changeset | 39 |   The           :: "('a => bool) => 'a"
 | 
| 7357 | 40 |   All           :: "('a => bool) => bool"           (binder "ALL " 10)
 | 
| 41 |   Ex            :: "('a => bool) => bool"           (binder "EX " 10)
 | |
| 42 |   Ex1           :: "('a => bool) => bool"           (binder "EX! " 10)
 | |
| 43 | Let :: "['a, 'a => 'b] => 'b" | |
| 923 | 44 | |
| 45 | (* Infixes *) | |
| 46 | ||
| 7357 | 47 | "=" :: "['a, 'a] => bool" (infixl 50) | 
| 48 | & :: "[bool, bool] => bool" (infixr 35) | |
| 49 | "|" :: "[bool, bool] => bool" (infixr 30) | |
| 50 | --> :: "[bool, bool] => bool" (infixr 25) | |
| 923 | 51 | |
| 10432 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 52 | local | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 53 | |
| 2260 | 54 | |
| 55 | (* Overloaded Constants *) | |
| 56 | ||
| 9869 | 57 | axclass zero < "term" | 
| 8940 | 58 | axclass plus < "term" | 
| 7357 | 59 | axclass minus < "term" | 
| 60 | axclass times < "term" | |
| 10432 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 61 | axclass inverse < "term" | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 62 | |
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 63 | global | 
| 3370 
5c5fdce3a4e4
Overloading of "^" requires new type class "power", with types "nat" and
 paulson parents: 
3320diff
changeset | 64 | |
| 2260 | 65 | consts | 
| 10432 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 66 |   "0"           :: "'a::zero"                       ("0")
 | 
| 7357 | 67 | "+" :: "['a::plus, 'a] => 'a" (infixl 65) | 
| 68 | - :: "['a::minus, 'a] => 'a" (infixl 65) | |
| 69 |   uminus        :: "['a::minus] => 'a"              ("- _" [81] 80)
 | |
| 7426 | 70 | * :: "['a::times, 'a] => 'a" (infixl 70) | 
| 10432 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 71 | |
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 72 | local | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 73 | |
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 74 | consts | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 75 | abs :: "'a::minus => 'a" | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 76 | inverse :: "'a::inverse => 'a" | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 77 | divide :: "['a::inverse, 'a] => 'a" (infixl "'/" 70) | 
| 2260 | 78 | |
| 10489 | 79 | syntax (xsymbols) | 
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11438diff
changeset | 80 |   abs :: "'a::minus => 'a"    ("\\<bar>_\\<bar>")
 | 
| 10489 | 81 | syntax (HTML output) | 
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11438diff
changeset | 82 |   abs :: "'a::minus => 'a"    ("\\<bar>_\\<bar>")
 | 
| 10489 | 83 | |
| 8959 | 84 | axclass plus_ac0 < plus, zero | 
| 10432 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 85 | commute: "x + y = y + x" | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 86 | assoc: "(x + y) + z = x + (y + z)" | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 87 | zero: "0 + x = x" | 
| 3820 | 88 | |
| 7238 
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
 wenzelm parents: 
7220diff
changeset | 89 | |
| 2260 | 90 | (** Additional concrete syntax **) | 
| 91 | ||
| 4868 | 92 | nonterminals | 
| 923 | 93 | letbinds letbind | 
| 94 | case_syn cases_syn | |
| 95 | ||
| 96 | syntax | |
| 7357 | 97 | ~= :: "['a, 'a] => bool" (infixl 50) | 
| 11432 
8a203ae6efe3
added "The" (definite description operator) (by Larry);
 wenzelm parents: 
10489diff
changeset | 98 |   "_The"        :: "[pttrn, bool] => 'a"                 ("(3THE _./ _)" [0, 10] 10)
 | 
| 923 | 99 | |
| 100 | (* Let expressions *) | |
| 101 | ||
| 7357 | 102 |   "_bind"       :: "[pttrn, 'a] => letbind"              ("(2_ =/ _)" 10)
 | 
| 103 |   ""            :: "letbind => letbinds"                 ("_")
 | |
| 104 |   "_binds"      :: "[letbind, letbinds] => letbinds"     ("_;/ _")
 | |
| 105 |   "_Let"        :: "[letbinds, 'a] => 'a"                ("(let (_)/ in (_))" 10)
 | |
| 923 | 106 | |
| 107 | (* Case expressions *) | |
| 108 | ||
| 9060 
b0dd884b1848
rename @case to _case_syntax (improves on low-level errors);
 wenzelm parents: 
8959diff
changeset | 109 |   "_case_syntax":: "['a, cases_syn] => 'b"               ("(case _ of/ _)" 10)
 | 
| 
b0dd884b1848
rename @case to _case_syntax (improves on low-level errors);
 wenzelm parents: 
8959diff
changeset | 110 |   "_case1"      :: "['a, 'b] => case_syn"                ("(2_ =>/ _)" 10)
 | 
| 7357 | 111 |   ""            :: "case_syn => cases_syn"               ("_")
 | 
| 9060 
b0dd884b1848
rename @case to _case_syntax (improves on low-level errors);
 wenzelm parents: 
8959diff
changeset | 112 |   "_case2"      :: "[case_syn, cases_syn] => cases_syn"  ("_/ | _")
 | 
| 923 | 113 | |
| 114 | translations | |
| 7238 
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
 wenzelm parents: 
7220diff
changeset | 115 | "x ~= y" == "~ (x = y)" | 
| 11432 
8a203ae6efe3
added "The" (definite description operator) (by Larry);
 wenzelm parents: 
10489diff
changeset | 116 | "THE x. P" == "The (%x. P)" | 
| 923 | 117 | "_Let (_binds b bs) e" == "_Let b (_Let bs e)" | 
| 1114 | 118 | "let x = a in e" == "Let a (%x. e)" | 
| 923 | 119 | |
| 3820 | 120 | syntax ("" output)
 | 
| 7357 | 121 |   "op ="        :: "['a, 'a] => bool"                    ("(_ =/ _)" [51, 51] 50)
 | 
| 122 |   "op ~="       :: "['a, 'a] => bool"                    ("(_ ~=/ _)" [51, 51] 50)
 | |
| 2260 | 123 | |
| 124 | syntax (symbols) | |
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11438diff
changeset | 125 |   Not           :: "bool => bool"                        ("\\<not> _" [40] 40)
 | 
| 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11438diff
changeset | 126 | "op &" :: "[bool, bool] => bool" (infixr "\\<and>" 35) | 
| 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11438diff
changeset | 127 | "op |" :: "[bool, bool] => bool" (infixr "\\<or>" 30) | 
| 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11438diff
changeset | 128 | "op -->" :: "[bool, bool] => bool" (infixr "\\<midarrow>\\<rightarrow>" 25) | 
| 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11438diff
changeset | 129 | "op ~=" :: "['a, 'a] => bool" (infixl "\\<noteq>" 50) | 
| 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11438diff
changeset | 130 |   "ALL "        :: "[idts, bool] => bool"                ("(3\\<forall>_./ _)" [0, 10] 10)
 | 
| 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11438diff
changeset | 131 |   "EX "         :: "[idts, bool] => bool"                ("(3\\<exists>_./ _)" [0, 10] 10)
 | 
| 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11438diff
changeset | 132 |   "EX! "        :: "[idts, bool] => bool"                ("(3\\<exists>!_./ _)" [0, 10] 10)
 | 
| 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11438diff
changeset | 133 |   "_case1"      :: "['a, 'b] => case_syn"                ("(2_ \\<Rightarrow>/ _)" 10)
 | 
| 9060 
b0dd884b1848
rename @case to _case_syntax (improves on low-level errors);
 wenzelm parents: 
8959diff
changeset | 134 | (*"_case2"      :: "[case_syn, cases_syn] => cases_syn"  ("_/ \\<orelse> _")*)
 | 
| 2372 | 135 | |
| 3820 | 136 | syntax (symbols output) | 
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11438diff
changeset | 137 |   "op ~="       :: "['a, 'a] => bool"                    ("(_ \\<noteq>/ _)" [51, 51] 50)
 | 
| 3820 | 138 | |
| 6027 
9dd06eeda95c
added new print_mode "xsymbols" for extended symbol support
 oheimb parents: 
5786diff
changeset | 139 | syntax (xsymbols) | 
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11438diff
changeset | 140 | "op -->" :: "[bool, bool] => bool" (infixr "\\<longrightarrow>" 25) | 
| 2260 | 141 | |
| 6340 | 142 | syntax (HTML output) | 
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11438diff
changeset | 143 |   Not           :: "bool => bool"                        ("\\<not> _" [40] 40)
 | 
| 6340 | 144 | |
| 7238 
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
 wenzelm parents: 
7220diff
changeset | 145 | syntax (HOL) | 
| 7357 | 146 |   "ALL "        :: "[idts, bool] => bool"                ("(3! _./ _)" [0, 10] 10)
 | 
| 147 |   "EX "         :: "[idts, bool] => bool"                ("(3? _./ _)" [0, 10] 10)
 | |
| 148 |   "EX! "        :: "[idts, bool] => bool"                ("(3?! _./ _)" [0, 10] 10)
 | |
| 7238 
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
 wenzelm parents: 
7220diff
changeset | 149 | |
| 
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
 wenzelm parents: 
7220diff
changeset | 150 | |
| 6340 | 151 | |
| 2260 | 152 | (** Rules and definitions **) | 
| 153 | ||
| 7357 | 154 | axioms | 
| 923 | 155 | |
| 7357 | 156 | eq_reflection: "(x=y) ==> (x==y)" | 
| 923 | 157 | |
| 158 | (* Basic Rules *) | |
| 159 | ||
| 7357 | 160 | refl: "t = (t::'a)" | 
| 161 | subst: "[| s = t; P(s) |] ==> P(t::'a)" | |
| 6289 | 162 | |
| 163 | (*Extensionality is built into the meta-logic, and this rule expresses | |
| 164 | a related property. It is an eta-expanded version of the traditional | |
| 165 | rule, and similar to the ABS rule of HOL.*) | |
| 7357 | 166 | ext: "(!!x::'a. (f x ::'b) = g x) ==> (%x. f x) = (%x. g x)" | 
| 6289 | 167 | |
| 11432 
8a203ae6efe3
added "The" (definite description operator) (by Larry);
 wenzelm parents: 
10489diff
changeset | 168 | the_eq_trivial: "(THE x. x = a) = (a::'a)" | 
| 923 | 169 | |
| 7357 | 170 | impI: "(P ==> Q) ==> P-->Q" | 
| 171 | mp: "[| P-->Q; P |] ==> Q" | |
| 923 | 172 | |
| 173 | defs | |
| 174 | ||
| 7357 | 175 | True_def: "True == ((%x::bool. x) = (%x. x))" | 
| 176 | All_def: "All(P) == (P = (%x. True))" | |
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11438diff
changeset | 177 | Ex_def: "Ex(P) == !Q. (!x. P x --> Q) --> Q" | 
| 7357 | 178 | False_def: "False == (!P. P)" | 
| 179 | not_def: "~ P == P-->False" | |
| 180 | and_def: "P & Q == !R. (P-->Q-->R) --> R" | |
| 181 | or_def: "P | Q == !R. (P-->R) --> (Q-->R) --> R" | |
| 182 | Ex1_def: "Ex1(P) == ? x. P(x) & (! y. P(y) --> y=x)" | |
| 923 | 183 | |
| 7357 | 184 | axioms | 
| 923 | 185 | (* Axioms *) | 
| 186 | ||
| 7357 | 187 | iff: "(P-->Q) --> (Q-->P) --> (P=Q)" | 
| 188 | True_or_False: "(P=True) | (P=False)" | |
| 923 | 189 | |
| 190 | defs | |
| 5069 | 191 | (*misc definitions*) | 
| 7357 | 192 | Let_def: "Let s f == f(s)" | 
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11438diff
changeset | 193 | if_def: "If P x y == THE z::'a. (P=True --> z=x) & (P=False --> z=y)" | 
| 5069 | 194 | |
| 195 | (*arbitrary is completely unspecified, but is made to appear as a | |
| 196 | definition syntactically*) | |
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11438diff
changeset | 197 | arbitrary_def: "False ==> arbitrary == (THE x. False)" | 
| 923 | 198 | |
| 3320 | 199 | |
| 4868 | 200 | |
| 7357 | 201 | (* theory and package setup *) | 
| 4868 | 202 | |
| 9736 | 203 | use "HOL_lemmas.ML" | 
| 9869 | 204 | |
| 11438 
3d9222b80989
declare trans [trans]  (*overridden in theory Calculation*);
 wenzelm parents: 
11432diff
changeset | 205 | declare trans [trans] (*overridden in theory Calculation*) | 
| 
3d9222b80989
declare trans [trans]  (*overridden in theory Calculation*);
 wenzelm parents: 
11432diff
changeset | 206 | |
| 10432 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 207 | lemma atomize_all: "(!!x. P x) == Trueprop (ALL x. P x)" | 
| 9488 | 208 | proof (rule equal_intr_rule) | 
| 209 | assume "!!x. P x" | |
| 10383 | 210 | show "ALL x. P x" by (rule allI) | 
| 9488 | 211 | next | 
| 212 | assume "ALL x. P x" | |
| 10383 | 213 | thus "!!x. P x" by (rule allE) | 
| 9488 | 214 | qed | 
| 215 | ||
| 10432 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 216 | lemma atomize_imp: "(A ==> B) == Trueprop (A --> B)" | 
| 9488 | 217 | proof (rule equal_intr_rule) | 
| 218 | assume r: "A ==> B" | |
| 10383 | 219 | show "A --> B" by (rule impI) (rule r) | 
| 9488 | 220 | next | 
| 221 | assume "A --> B" and A | |
| 10383 | 222 | thus B by (rule mp) | 
| 9488 | 223 | qed | 
| 224 | ||
| 10432 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 225 | lemma atomize_eq: "(x == y) == Trueprop (x = y)" | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 226 | proof (rule equal_intr_rule) | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 227 | assume "x == y" | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 228 | show "x = y" by (unfold prems) (rule refl) | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 229 | next | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 230 | assume "x = y" | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 231 | thus "x == y" by (rule eq_reflection) | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 232 | qed | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 233 | |
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 234 | lemmas atomize = atomize_all atomize_imp | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 235 | lemmas atomize' = atomize atomize_eq | 
| 9529 | 236 | |
| 10383 | 237 | use "cladata.ML" | 
| 238 | setup hypsubst_setup | |
| 239 | setup Classical.setup | |
| 240 | setup clasetup | |
| 241 | ||
| 9869 | 242 | use "blastdata.ML" | 
| 243 | setup Blast.setup | |
| 4868 | 244 | |
| 9869 | 245 | use "simpdata.ML" | 
| 246 | setup Simplifier.setup | |
| 247 | setup "Simplifier.method_setup Splitter.split_modifiers" setup simpsetup | |
| 248 | setup Splitter.setup setup Clasimp.setup | |
| 249 | ||
| 923 | 250 | end |