| author | ballarin | 
| Fri, 14 Mar 2003 18:00:16 +0100 | |
| changeset 13864 | f44f121dd275 | 
| parent 13764 | 3e180bf68496 | 
| child 14201 | 7ad7ab89c402 | 
| permissions | -rw-r--r-- | 
| 923 | 1 | (* Title: HOL/HOL.thy | 
| 2 | ID: $Id$ | |
| 11750 | 3 | Author: Tobias Nipkow, Markus Wenzel, and Larry Paulson | 
| 12386 | 4 | License: GPL (GNU GENERAL PUBLIC LICENSE) | 
| 11750 | 5 | *) | 
| 923 | 6 | |
| 11750 | 7 | header {* The basis of Higher-Order Logic *}
 | 
| 923 | 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 | |
| 11750 | 13 | subsection {* Primitive logic *}
 | 
| 14 | ||
| 15 | subsubsection {* Core syntax *}
 | |
| 2260 | 16 | |
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12281diff
changeset | 17 | classes type < logic | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12281diff
changeset | 18 | defaultsort type | 
| 3947 | 19 | |
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12281diff
changeset | 20 | global | 
| 923 | 21 | |
| 7357 | 22 | typedecl bool | 
| 923 | 23 | |
| 24 | arities | |
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12281diff
changeset | 25 | bool :: type | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12281diff
changeset | 26 | fun :: (type, type) type | 
| 923 | 27 | |
| 11750 | 28 | judgment | 
| 29 |   Trueprop      :: "bool => prop"                   ("(_)" 5)
 | |
| 923 | 30 | |
| 11750 | 31 | consts | 
| 7357 | 32 |   Not           :: "bool => bool"                   ("~ _" [40] 40)
 | 
| 33 | True :: bool | |
| 34 | False :: bool | |
| 35 |   If            :: "[bool, 'a, 'a] => 'a"           ("(if (_)/ then (_)/ else (_))" 10)
 | |
| 3947 | 36 | arbitrary :: 'a | 
| 923 | 37 | |
| 11432 
8a203ae6efe3
added "The" (definite description operator) (by Larry);
 wenzelm parents: 
10489diff
changeset | 38 |   The           :: "('a => bool) => 'a"
 | 
| 7357 | 39 |   All           :: "('a => bool) => bool"           (binder "ALL " 10)
 | 
| 40 |   Ex            :: "('a => bool) => bool"           (binder "EX " 10)
 | |
| 41 |   Ex1           :: "('a => bool) => bool"           (binder "EX! " 10)
 | |
| 42 | Let :: "['a, 'a => 'b] => 'b" | |
| 923 | 43 | |
| 7357 | 44 | "=" :: "['a, 'a] => bool" (infixl 50) | 
| 45 | & :: "[bool, bool] => bool" (infixr 35) | |
| 46 | "|" :: "[bool, bool] => bool" (infixr 30) | |
| 47 | --> :: "[bool, bool] => bool" (infixr 25) | |
| 923 | 48 | |
| 10432 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 49 | local | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 50 | |
| 2260 | 51 | |
| 11750 | 52 | subsubsection {* Additional concrete syntax *}
 | 
| 2260 | 53 | |
| 4868 | 54 | nonterminals | 
| 923 | 55 | letbinds letbind | 
| 56 | case_syn cases_syn | |
| 57 | ||
| 58 | syntax | |
| 12650 | 59 | "_not_equal" :: "['a, 'a] => bool" (infixl "~=" 50) | 
| 11432 
8a203ae6efe3
added "The" (definite description operator) (by Larry);
 wenzelm parents: 
10489diff
changeset | 60 |   "_The"        :: "[pttrn, bool] => 'a"                 ("(3THE _./ _)" [0, 10] 10)
 | 
| 923 | 61 | |
| 7357 | 62 |   "_bind"       :: "[pttrn, 'a] => letbind"              ("(2_ =/ _)" 10)
 | 
| 63 |   ""            :: "letbind => letbinds"                 ("_")
 | |
| 64 |   "_binds"      :: "[letbind, letbinds] => letbinds"     ("_;/ _")
 | |
| 65 |   "_Let"        :: "[letbinds, 'a] => 'a"                ("(let (_)/ in (_))" 10)
 | |
| 923 | 66 | |
| 9060 
b0dd884b1848
rename @case to _case_syntax (improves on low-level errors);
 wenzelm parents: 
8959diff
changeset | 67 |   "_case_syntax":: "['a, cases_syn] => 'b"               ("(case _ of/ _)" 10)
 | 
| 
b0dd884b1848
rename @case to _case_syntax (improves on low-level errors);
 wenzelm parents: 
8959diff
changeset | 68 |   "_case1"      :: "['a, 'b] => case_syn"                ("(2_ =>/ _)" 10)
 | 
| 7357 | 69 |   ""            :: "case_syn => cases_syn"               ("_")
 | 
| 9060 
b0dd884b1848
rename @case to _case_syntax (improves on low-level errors);
 wenzelm parents: 
8959diff
changeset | 70 |   "_case2"      :: "[case_syn, cases_syn] => cases_syn"  ("_/ | _")
 | 
| 923 | 71 | |
| 72 | translations | |
| 7238 
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
 wenzelm parents: 
7220diff
changeset | 73 | "x ~= y" == "~ (x = y)" | 
| 13764 | 74 | "THE x. P" == "The (%x. P)" | 
| 923 | 75 | "_Let (_binds b bs) e" == "_Let b (_Let bs e)" | 
| 1114 | 76 | "let x = a in e" == "Let a (%x. e)" | 
| 923 | 77 | |
| 13763 
f94b569cd610
added print translations tha avoid eta contraction for important binders.
 nipkow parents: 
13723diff
changeset | 78 | print_translation {*
 | 
| 
f94b569cd610
added print translations tha avoid eta contraction for important binders.
 nipkow parents: 
13723diff
changeset | 79 | (* To avoid eta-contraction of body: *) | 
| 
f94b569cd610
added print translations tha avoid eta contraction for important binders.
 nipkow parents: 
13723diff
changeset | 80 | [("The", fn [Abs abs] =>
 | 
| 
f94b569cd610
added print translations tha avoid eta contraction for important binders.
 nipkow parents: 
13723diff
changeset | 81 | let val (x,t) = atomic_abs_tr' abs | 
| 
f94b569cd610
added print translations tha avoid eta contraction for important binders.
 nipkow parents: 
13723diff
changeset | 82 | in Syntax.const "_The" $ x $ t end)] | 
| 
f94b569cd610
added print translations tha avoid eta contraction for important binders.
 nipkow parents: 
13723diff
changeset | 83 | *} | 
| 
f94b569cd610
added print translations tha avoid eta contraction for important binders.
 nipkow parents: 
13723diff
changeset | 84 | |
| 12633 | 85 | syntax (output) | 
| 11687 | 86 | "=" :: "['a, 'a] => bool" (infix 50) | 
| 12650 | 87 | "_not_equal" :: "['a, 'a] => bool" (infix "~=" 50) | 
| 2260 | 88 | |
| 12114 
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
 wenzelm parents: 
12023diff
changeset | 89 | syntax (xsymbols) | 
| 11687 | 90 |   Not           :: "bool => bool"                        ("\<not> _" [40] 40)
 | 
| 91 | "op &" :: "[bool, bool] => bool" (infixr "\<and>" 35) | |
| 92 | "op |" :: "[bool, bool] => bool" (infixr "\<or>" 30) | |
| 12114 
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
 wenzelm parents: 
12023diff
changeset | 93 | "op -->" :: "[bool, bool] => bool" (infixr "\<longrightarrow>" 25) | 
| 12650 | 94 | "_not_equal" :: "['a, 'a] => bool" (infix "\<noteq>" 50) | 
| 11687 | 95 |   "ALL "        :: "[idts, bool] => bool"                ("(3\<forall>_./ _)" [0, 10] 10)
 | 
| 96 |   "EX "         :: "[idts, bool] => bool"                ("(3\<exists>_./ _)" [0, 10] 10)
 | |
| 97 |   "EX! "        :: "[idts, bool] => bool"                ("(3\<exists>!_./ _)" [0, 10] 10)
 | |
| 98 |   "_case1"      :: "['a, 'b] => case_syn"                ("(2_ \<Rightarrow>/ _)" 10)
 | |
| 9060 
b0dd884b1848
rename @case to _case_syntax (improves on low-level errors);
 wenzelm parents: 
8959diff
changeset | 99 | (*"_case2"      :: "[case_syn, cases_syn] => cases_syn"  ("_/ \\<orelse> _")*)
 | 
| 2372 | 100 | |
| 12114 
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
 wenzelm parents: 
12023diff
changeset | 101 | syntax (xsymbols output) | 
| 12650 | 102 | "_not_equal" :: "['a, 'a] => bool" (infix "\<noteq>" 50) | 
| 3820 | 103 | |
| 6340 | 104 | syntax (HTML output) | 
| 11687 | 105 |   Not           :: "bool => bool"                        ("\<not> _" [40] 40)
 | 
| 6340 | 106 | |
| 7238 
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
 wenzelm parents: 
7220diff
changeset | 107 | syntax (HOL) | 
| 7357 | 108 |   "ALL "        :: "[idts, bool] => bool"                ("(3! _./ _)" [0, 10] 10)
 | 
| 109 |   "EX "         :: "[idts, bool] => bool"                ("(3? _./ _)" [0, 10] 10)
 | |
| 110 |   "EX! "        :: "[idts, bool] => bool"                ("(3?! _./ _)" [0, 10] 10)
 | |
| 7238 
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
 wenzelm parents: 
7220diff
changeset | 111 | |
| 
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
 wenzelm parents: 
7220diff
changeset | 112 | |
| 11750 | 113 | subsubsection {* Axioms and basic definitions *}
 | 
| 2260 | 114 | |
| 7357 | 115 | axioms | 
| 116 | eq_reflection: "(x=y) ==> (x==y)" | |
| 923 | 117 | |
| 7357 | 118 | refl: "t = (t::'a)" | 
| 119 | subst: "[| s = t; P(s) |] ==> P(t::'a)" | |
| 6289 | 120 | |
| 7357 | 121 | ext: "(!!x::'a. (f x ::'b) = g x) ==> (%x. f x) = (%x. g x)" | 
| 11750 | 122 |     -- {* Extensionality is built into the meta-logic, and this rule expresses *}
 | 
| 123 |     -- {* a related property.  It is an eta-expanded version of the traditional *}
 | |
| 124 |     -- {* rule, and similar to the ABS rule of HOL *}
 | |
| 6289 | 125 | |
| 11432 
8a203ae6efe3
added "The" (definite description operator) (by Larry);
 wenzelm parents: 
10489diff
changeset | 126 | the_eq_trivial: "(THE x. x = a) = (a::'a)" | 
| 923 | 127 | |
| 7357 | 128 | impI: "(P ==> Q) ==> P-->Q" | 
| 129 | mp: "[| P-->Q; P |] ==> Q" | |
| 923 | 130 | |
| 131 | defs | |
| 7357 | 132 | True_def: "True == ((%x::bool. x) = (%x. x))" | 
| 133 | All_def: "All(P) == (P = (%x. True))" | |
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11438diff
changeset | 134 | Ex_def: "Ex(P) == !Q. (!x. P x --> Q) --> Q" | 
| 7357 | 135 | False_def: "False == (!P. P)" | 
| 136 | not_def: "~ P == P-->False" | |
| 137 | and_def: "P & Q == !R. (P-->Q-->R) --> R" | |
| 138 | or_def: "P | Q == !R. (P-->R) --> (Q-->R) --> R" | |
| 139 | Ex1_def: "Ex1(P) == ? x. P(x) & (! y. P(y) --> y=x)" | |
| 923 | 140 | |
| 7357 | 141 | axioms | 
| 142 | iff: "(P-->Q) --> (Q-->P) --> (P=Q)" | |
| 143 | True_or_False: "(P=True) | (P=False)" | |
| 923 | 144 | |
| 145 | defs | |
| 7357 | 146 | Let_def: "Let s f == f(s)" | 
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11438diff
changeset | 147 | if_def: "If P x y == THE z::'a. (P=True --> z=x) & (P=False --> z=y)" | 
| 5069 | 148 | |
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11438diff
changeset | 149 | arbitrary_def: "False ==> arbitrary == (THE x. False)" | 
| 11750 | 150 |     -- {* @{term arbitrary} is completely unspecified, but is made to appear as a
 | 
| 151 | definition syntactically *} | |
| 923 | 152 | |
| 3320 | 153 | |
| 11750 | 154 | subsubsection {* Generic algebraic operations *}
 | 
| 4868 | 155 | |
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12281diff
changeset | 156 | axclass zero < type | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12281diff
changeset | 157 | axclass one < type | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12281diff
changeset | 158 | axclass plus < type | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12281diff
changeset | 159 | axclass minus < type | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12281diff
changeset | 160 | axclass times < type | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12281diff
changeset | 161 | axclass inverse < type | 
| 11750 | 162 | |
| 163 | global | |
| 164 | ||
| 165 | consts | |
| 166 |   "0"           :: "'a::zero"                       ("0")
 | |
| 167 |   "1"           :: "'a::one"                        ("1")
 | |
| 168 | "+" :: "['a::plus, 'a] => 'a" (infixl 65) | |
| 169 | - :: "['a::minus, 'a] => 'a" (infixl 65) | |
| 170 |   uminus        :: "['a::minus] => 'a"              ("- _" [81] 80)
 | |
| 171 | * :: "['a::times, 'a] => 'a" (infixl 70) | |
| 172 | ||
| 13456 
42601eb7553f
special syntax for index "1" (plain numeral hidden by "1" symbol in HOL);
 wenzelm parents: 
13438diff
changeset | 173 | syntax | 
| 
42601eb7553f
special syntax for index "1" (plain numeral hidden by "1" symbol in HOL);
 wenzelm parents: 
13438diff
changeset | 174 |   "_index1"  :: index    ("\<^sub>1")
 | 
| 
42601eb7553f
special syntax for index "1" (plain numeral hidden by "1" symbol in HOL);
 wenzelm parents: 
13438diff
changeset | 175 | translations | 
| 
42601eb7553f
special syntax for index "1" (plain numeral hidden by "1" symbol in HOL);
 wenzelm parents: 
13438diff
changeset | 176 | (index) "\<^sub>1" == "_index 1" | 
| 
42601eb7553f
special syntax for index "1" (plain numeral hidden by "1" symbol in HOL);
 wenzelm parents: 
13438diff
changeset | 177 | |
| 11750 | 178 | local | 
| 179 | ||
| 180 | typed_print_translation {*
 | |
| 181 | let | |
| 182 | fun tr' c = (c, fn show_sorts => fn T => fn ts => | |
| 183 | if T = dummyT orelse not (! show_types) andalso can Term.dest_Type T then raise Match | |
| 184 | else Syntax.const Syntax.constrainC $ Syntax.const c $ Syntax.term_of_typ show_sorts T); | |
| 185 | in [tr' "0", tr' "1"] end; | |
| 186 | *} -- {* show types that are presumably too general *}
 | |
| 187 | ||
| 188 | ||
| 189 | consts | |
| 190 | abs :: "'a::minus => 'a" | |
| 191 | inverse :: "'a::inverse => 'a" | |
| 192 | divide :: "['a::inverse, 'a] => 'a" (infixl "'/" 70) | |
| 193 | ||
| 194 | syntax (xsymbols) | |
| 195 |   abs :: "'a::minus => 'a"    ("\<bar>_\<bar>")
 | |
| 196 | syntax (HTML output) | |
| 197 |   abs :: "'a::minus => 'a"    ("\<bar>_\<bar>")
 | |
| 198 | ||
| 199 | axclass plus_ac0 < plus, zero | |
| 200 | commute: "x + y = y + x" | |
| 201 | assoc: "(x + y) + z = x + (y + z)" | |
| 202 | zero: "0 + x = x" | |
| 203 | ||
| 204 | ||
| 205 | subsection {* Theory and package setup *}
 | |
| 206 | ||
| 207 | subsubsection {* Basic lemmas *}
 | |
| 4868 | 208 | |
| 9736 | 209 | use "HOL_lemmas.ML" | 
| 11687 | 210 | theorems case_split = case_split_thm [case_names True False] | 
| 9869 | 211 | |
| 12386 | 212 | |
| 213 | subsubsection {* Intuitionistic Reasoning *}
 | |
| 214 | ||
| 215 | lemma impE': | |
| 12937 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 216 | assumes 1: "P --> Q" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 217 | and 2: "Q ==> R" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 218 | and 3: "P --> Q ==> P" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 219 | shows R | 
| 12386 | 220 | proof - | 
| 221 | from 3 and 1 have P . | |
| 222 | with 1 have Q by (rule impE) | |
| 223 | with 2 show R . | |
| 224 | qed | |
| 225 | ||
| 226 | lemma allE': | |
| 12937 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 227 | assumes 1: "ALL x. P x" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 228 | and 2: "P x ==> ALL x. P x ==> Q" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 229 | shows Q | 
| 12386 | 230 | proof - | 
| 231 | from 1 have "P x" by (rule spec) | |
| 232 | from this and 1 show Q by (rule 2) | |
| 233 | qed | |
| 234 | ||
| 12937 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 235 | lemma notE': | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 236 | assumes 1: "~ P" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 237 | and 2: "~ P ==> P" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 238 | shows R | 
| 12386 | 239 | proof - | 
| 240 | from 2 and 1 have P . | |
| 241 | with 1 show R by (rule notE) | |
| 242 | qed | |
| 243 | ||
| 244 | lemmas [CPure.elim!] = disjE iffE FalseE conjE exE | |
| 245 | and [CPure.intro!] = iffI conjI impI TrueI notI allI refl | |
| 246 | and [CPure.elim 2] = allE notE' impE' | |
| 247 | and [CPure.intro] = exI disjI2 disjI1 | |
| 248 | ||
| 249 | lemmas [trans] = trans | |
| 250 | and [sym] = sym not_sym | |
| 251 | and [CPure.elim?] = iffD1 iffD2 impE | |
| 11750 | 252 | |
| 11438 
3d9222b80989
declare trans [trans]  (*overridden in theory Calculation*);
 wenzelm parents: 
11432diff
changeset | 253 | |
| 11750 | 254 | subsubsection {* Atomizing meta-level connectives *}
 | 
| 255 | ||
| 256 | lemma atomize_all [atomize]: "(!!x. P x) == Trueprop (ALL x. P x)" | |
| 12003 | 257 | proof | 
| 9488 | 258 | assume "!!x. P x" | 
| 10383 | 259 | show "ALL x. P x" by (rule allI) | 
| 9488 | 260 | next | 
| 261 | assume "ALL x. P x" | |
| 10383 | 262 | thus "!!x. P x" by (rule allE) | 
| 9488 | 263 | qed | 
| 264 | ||
| 11750 | 265 | lemma atomize_imp [atomize]: "(A ==> B) == Trueprop (A --> B)" | 
| 12003 | 266 | proof | 
| 9488 | 267 | assume r: "A ==> B" | 
| 10383 | 268 | show "A --> B" by (rule impI) (rule r) | 
| 9488 | 269 | next | 
| 270 | assume "A --> B" and A | |
| 10383 | 271 | thus B by (rule mp) | 
| 9488 | 272 | qed | 
| 273 | ||
| 11750 | 274 | lemma atomize_eq [atomize]: "(x == y) == Trueprop (x = y)" | 
| 12003 | 275 | proof | 
| 10432 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 276 | assume "x == y" | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 277 | show "x = y" by (unfold prems) (rule refl) | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 278 | next | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 279 | assume "x = y" | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 280 | thus "x == y" by (rule eq_reflection) | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 281 | qed | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 282 | |
| 12023 | 283 | lemma atomize_conj [atomize]: | 
| 284 | "(!!C. (A ==> B ==> PROP C) ==> PROP C) == Trueprop (A & B)" | |
| 12003 | 285 | proof | 
| 11953 | 286 | assume "!!C. (A ==> B ==> PROP C) ==> PROP C" | 
| 287 | show "A & B" by (rule conjI) | |
| 288 | next | |
| 289 | fix C | |
| 290 | assume "A & B" | |
| 291 | assume "A ==> B ==> PROP C" | |
| 292 | thus "PROP C" | |
| 293 | proof this | |
| 294 | show A by (rule conjunct1) | |
| 295 | show B by (rule conjunct2) | |
| 296 | qed | |
| 297 | qed | |
| 298 | ||
| 12386 | 299 | lemmas [symmetric, rulify] = atomize_all atomize_imp | 
| 300 | ||
| 11750 | 301 | |
| 302 | subsubsection {* Classical Reasoner setup *}
 | |
| 9529 | 303 | |
| 10383 | 304 | use "cladata.ML" | 
| 305 | setup hypsubst_setup | |
| 11977 | 306 | |
| 12386 | 307 | ML_setup {*
 | 
| 308 | Context.>> (ContextRules.addSWrapper (fn tac => hyp_subst_tac' ORELSE' tac)); | |
| 309 | *} | |
| 11977 | 310 | |
| 10383 | 311 | setup Classical.setup | 
| 312 | setup clasetup | |
| 313 | ||
| 12386 | 314 | lemmas [intro?] = ext | 
| 315 | and [elim?] = ex1_implies_ex | |
| 11977 | 316 | |
| 9869 | 317 | use "blastdata.ML" | 
| 318 | setup Blast.setup | |
| 4868 | 319 | |
| 11750 | 320 | |
| 321 | subsubsection {* Simplifier setup *}
 | |
| 322 | ||
| 12281 | 323 | lemma meta_eq_to_obj_eq: "x == y ==> x = y" | 
| 324 | proof - | |
| 325 | assume r: "x == y" | |
| 326 | show "x = y" by (unfold r) (rule refl) | |
| 327 | qed | |
| 328 | ||
| 329 | lemma eta_contract_eq: "(%s. f s) = f" .. | |
| 330 | ||
| 331 | lemma simp_thms: | |
| 12937 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 332 | shows not_not: "(~ ~ P) = P" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 333 | and | 
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 334 | "(P ~= Q) = (P = (~Q))" | 
| 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 335 | "(P | ~P) = True" "(~P | P) = True" | 
| 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 336 | "((~P) = (~Q)) = (P=Q)" | 
| 12281 | 337 | "(x = x) = True" | 
| 338 | "(~True) = False" "(~False) = True" | |
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 339 | "(~P) ~= P" "P ~= (~P)" | 
| 12281 | 340 | "(True=P) = P" "(P=True) = P" "(False=P) = (~P)" "(P=False) = (~P)" | 
| 341 | "(True --> P) = P" "(False --> P) = True" | |
| 342 | "(P --> True) = True" "(P --> P) = True" | |
| 343 | "(P --> False) = (~P)" "(P --> ~P) = (~P)" | |
| 344 | "(P & True) = P" "(True & P) = P" | |
| 345 | "(P & False) = False" "(False & P) = False" | |
| 346 | "(P & P) = P" "(P & (P & Q)) = (P & Q)" | |
| 347 | "(P & ~P) = False" "(~P & P) = False" | |
| 348 | "(P | True) = True" "(True | P) = True" | |
| 349 | "(P | False) = P" "(False | P) = P" | |
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 350 | "(P | P) = P" "(P | (P | Q)) = (P | Q)" and | 
| 12281 | 351 | "(ALL x. P) = P" "(EX x. P) = P" "EX x. x=t" "EX x. t=x" | 
| 352 |     -- {* needed for the one-point-rule quantifier simplification procs *}
 | |
| 353 |     -- {* essential for termination!! *} and
 | |
| 354 | "!!P. (EX x. x=t & P(x)) = P(t)" | |
| 355 | "!!P. (EX x. t=x & P(x)) = P(t)" | |
| 356 | "!!P. (ALL x. x=t --> P(x)) = P(t)" | |
| 12937 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 357 | "!!P. (ALL x. t=x --> P(x)) = P(t)" | 
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 358 | by (blast, blast, blast, blast, blast, rules+) | 
| 13421 | 359 | |
| 12281 | 360 | lemma imp_cong: "(P = P') ==> (P' ==> (Q = Q')) ==> ((P --> Q) = (P' --> Q'))" | 
| 12354 | 361 | by rules | 
| 12281 | 362 | |
| 363 | lemma ex_simps: | |
| 364 | "!!P Q. (EX x. P x & Q) = ((EX x. P x) & Q)" | |
| 365 | "!!P Q. (EX x. P & Q x) = (P & (EX x. Q x))" | |
| 366 | "!!P Q. (EX x. P x | Q) = ((EX x. P x) | Q)" | |
| 367 | "!!P Q. (EX x. P | Q x) = (P | (EX x. Q x))" | |
| 368 | "!!P Q. (EX x. P x --> Q) = ((ALL x. P x) --> Q)" | |
| 369 | "!!P Q. (EX x. P --> Q x) = (P --> (EX x. Q x))" | |
| 370 |   -- {* Miniscoping: pushing in existential quantifiers. *}
 | |
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 371 | by (rules | blast)+ | 
| 12281 | 372 | |
| 373 | lemma all_simps: | |
| 374 | "!!P Q. (ALL x. P x & Q) = ((ALL x. P x) & Q)" | |
| 375 | "!!P Q. (ALL x. P & Q x) = (P & (ALL x. Q x))" | |
| 376 | "!!P Q. (ALL x. P x | Q) = ((ALL x. P x) | Q)" | |
| 377 | "!!P Q. (ALL x. P | Q x) = (P | (ALL x. Q x))" | |
| 378 | "!!P Q. (ALL x. P x --> Q) = ((EX x. P x) --> Q)" | |
| 379 | "!!P Q. (ALL x. P --> Q x) = (P --> (ALL x. Q x))" | |
| 380 |   -- {* Miniscoping: pushing in universal quantifiers. *}
 | |
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 381 | by (rules | blast)+ | 
| 12281 | 382 | |
| 383 | lemma eq_ac: | |
| 12937 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 384 | shows eq_commute: "(a=b) = (b=a)" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 385 | and eq_left_commute: "(P=(Q=R)) = (Q=(P=R))" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 386 | and eq_assoc: "((P=Q)=R) = (P=(Q=R))" by (rules, blast+) | 
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 387 | lemma neq_commute: "(a~=b) = (b~=a)" by rules | 
| 12281 | 388 | |
| 389 | lemma conj_comms: | |
| 12937 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 390 | shows conj_commute: "(P&Q) = (Q&P)" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 391 | and conj_left_commute: "(P&(Q&R)) = (Q&(P&R))" by rules+ | 
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 392 | lemma conj_assoc: "((P&Q)&R) = (P&(Q&R))" by rules | 
| 12281 | 393 | |
| 394 | lemma disj_comms: | |
| 12937 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 395 | shows disj_commute: "(P|Q) = (Q|P)" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 396 | and disj_left_commute: "(P|(Q|R)) = (Q|(P|R))" by rules+ | 
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 397 | lemma disj_assoc: "((P|Q)|R) = (P|(Q|R))" by rules | 
| 12281 | 398 | |
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 399 | lemma conj_disj_distribL: "(P&(Q|R)) = (P&Q | P&R)" by rules | 
| 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 400 | lemma conj_disj_distribR: "((P|Q)&R) = (P&R | Q&R)" by rules | 
| 12281 | 401 | |
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 402 | lemma disj_conj_distribL: "(P|(Q&R)) = ((P|Q) & (P|R))" by rules | 
| 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 403 | lemma disj_conj_distribR: "((P&Q)|R) = ((P|R) & (Q|R))" by rules | 
| 12281 | 404 | |
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 405 | lemma imp_conjR: "(P --> (Q&R)) = ((P-->Q) & (P-->R))" by rules | 
| 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 406 | lemma imp_conjL: "((P&Q) -->R) = (P --> (Q --> R))" by rules | 
| 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 407 | lemma imp_disjL: "((P|Q) --> R) = ((P-->R)&(Q-->R))" by rules | 
| 12281 | 408 | |
| 409 | text {* These two are specialized, but @{text imp_disj_not1} is useful in @{text "Auth/Yahalom"}. *}
 | |
| 410 | lemma imp_disj_not1: "(P --> Q | R) = (~Q --> P --> R)" by blast | |
| 411 | lemma imp_disj_not2: "(P --> Q | R) = (~R --> P --> Q)" by blast | |
| 412 | ||
| 413 | lemma imp_disj1: "((P-->Q)|R) = (P--> Q|R)" by blast | |
| 414 | lemma imp_disj2: "(Q|(P-->R)) = (P--> Q|R)" by blast | |
| 415 | ||
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 416 | lemma de_Morgan_disj: "(~(P | Q)) = (~P & ~Q)" by rules | 
| 12281 | 417 | lemma de_Morgan_conj: "(~(P & Q)) = (~P | ~Q)" by blast | 
| 418 | lemma not_imp: "(~(P --> Q)) = (P & ~Q)" by blast | |
| 419 | lemma not_iff: "(P~=Q) = (P = (~Q))" by blast | |
| 420 | lemma disj_not1: "(~P | Q) = (P --> Q)" by blast | |
| 421 | lemma disj_not2: "(P | ~Q) = (Q --> P)"  -- {* changes orientation :-( *}
 | |
| 422 | by blast | |
| 423 | lemma imp_conv_disj: "(P --> Q) = ((~P) | Q)" by blast | |
| 424 | ||
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 425 | lemma iff_conv_conj_imp: "(P = Q) = ((P --> Q) & (Q --> P))" by rules | 
| 12281 | 426 | |
| 427 | ||
| 428 | lemma cases_simp: "((P --> Q) & (~P --> Q)) = Q" | |
| 429 |   -- {* Avoids duplication of subgoals after @{text split_if}, when the true and false *}
 | |
| 430 |   -- {* cases boil down to the same thing. *}
 | |
| 431 | by blast | |
| 432 | ||
| 433 | lemma not_all: "(~ (! x. P(x))) = (? x.~P(x))" by blast | |
| 434 | lemma imp_all: "((! x. P x) --> Q) = (? x. P x --> Q)" by blast | |
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 435 | lemma not_ex: "(~ (? x. P(x))) = (! x.~P(x))" by rules | 
| 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 436 | lemma imp_ex: "((? x. P x) --> Q) = (! x. P x --> Q)" by rules | 
| 12281 | 437 | |
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 438 | lemma ex_disj_distrib: "(? x. P(x) | Q(x)) = ((? x. P(x)) | (? x. Q(x)))" by rules | 
| 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 439 | lemma all_conj_distrib: "(!x. P(x) & Q(x)) = ((! x. P(x)) & (! x. Q(x)))" by rules | 
| 12281 | 440 | |
| 441 | text {*
 | |
| 442 |   \medskip The @{text "&"} congruence rule: not included by default!
 | |
| 443 | May slow rewrite proofs down by as much as 50\% *} | |
| 444 | ||
| 445 | lemma conj_cong: | |
| 446 | "(P = P') ==> (P' ==> (Q = Q')) ==> ((P & Q) = (P' & Q'))" | |
| 12354 | 447 | by rules | 
| 12281 | 448 | |
| 449 | lemma rev_conj_cong: | |
| 450 | "(Q = Q') ==> (Q' ==> (P = P')) ==> ((P & Q) = (P' & Q'))" | |
| 12354 | 451 | by rules | 
| 12281 | 452 | |
| 453 | text {* The @{text "|"} congruence rule: not included by default! *}
 | |
| 454 | ||
| 455 | lemma disj_cong: | |
| 456 | "(P = P') ==> (~P' ==> (Q = Q')) ==> ((P | Q) = (P' | Q'))" | |
| 457 | by blast | |
| 458 | ||
| 459 | lemma eq_sym_conv: "(x = y) = (y = x)" | |
| 12354 | 460 | by rules | 
| 12281 | 461 | |
| 462 | ||
| 463 | text {* \medskip if-then-else rules *}
 | |
| 464 | ||
| 465 | lemma if_True: "(if True then x else y) = x" | |
| 466 | by (unfold if_def) blast | |
| 467 | ||
| 468 | lemma if_False: "(if False then x else y) = y" | |
| 469 | by (unfold if_def) blast | |
| 470 | ||
| 471 | lemma if_P: "P ==> (if P then x else y) = x" | |
| 472 | by (unfold if_def) blast | |
| 473 | ||
| 474 | lemma if_not_P: "~P ==> (if P then x else y) = y" | |
| 475 | by (unfold if_def) blast | |
| 476 | ||
| 477 | lemma split_if: "P (if Q then x else y) = ((Q --> P(x)) & (~Q --> P(y)))" | |
| 478 | apply (rule case_split [of Q]) | |
| 479 | apply (subst if_P) | |
| 480 | prefer 3 apply (subst if_not_P) | |
| 481 | apply blast+ | |
| 482 | done | |
| 483 | ||
| 484 | lemma split_if_asm: "P (if Q then x else y) = (~((Q & ~P x) | (~Q & ~P y)))" | |
| 485 | apply (subst split_if) | |
| 486 | apply blast | |
| 487 | done | |
| 488 | ||
| 489 | lemmas if_splits = split_if split_if_asm | |
| 490 | ||
| 491 | lemma if_def2: "(if Q then x else y) = ((Q --> x) & (~ Q --> y))" | |
| 492 | by (rule split_if) | |
| 493 | ||
| 494 | lemma if_cancel: "(if c then x else x) = x" | |
| 495 | apply (subst split_if) | |
| 496 | apply blast | |
| 497 | done | |
| 498 | ||
| 499 | lemma if_eq_cancel: "(if x = y then y else x) = x" | |
| 500 | apply (subst split_if) | |
| 501 | apply blast | |
| 502 | done | |
| 503 | ||
| 504 | lemma if_bool_eq_conj: "(if P then Q else R) = ((P-->Q) & (~P-->R))" | |
| 505 |   -- {* This form is useful for expanding @{text if}s on the RIGHT of the @{text "==>"} symbol. *}
 | |
| 506 | by (rule split_if) | |
| 507 | ||
| 508 | lemma if_bool_eq_disj: "(if P then Q else R) = ((P&Q) | (~P&R))" | |
| 509 |   -- {* And this form is useful for expanding @{text if}s on the LEFT. *}
 | |
| 510 | apply (subst split_if) | |
| 511 | apply blast | |
| 512 | done | |
| 513 | ||
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 514 | lemma Eq_TrueI: "P ==> P == True" by (unfold atomize_eq) rules | 
| 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 515 | lemma Eq_FalseI: "~P ==> P == False" by (unfold atomize_eq) rules | 
| 12281 | 516 | |
| 9869 | 517 | use "simpdata.ML" | 
| 518 | setup Simplifier.setup | |
| 519 | setup "Simplifier.method_setup Splitter.split_modifiers" setup simpsetup | |
| 520 | setup Splitter.setup setup Clasimp.setup | |
| 521 | ||
| 13723 | 522 | lemma ex1_eq[iff]: "EX! x. x = t" "EX! x. t = x" | 
| 523 | by blast+ | |
| 524 | ||
| 13638 | 525 | theorem choice_eq: "(ALL x. EX! y. P x y) = (EX! f. ALL x. P x (f x))" | 
| 526 | apply (rule iffI) | |
| 527 | apply (rule_tac a = "%x. THE y. P x y" in ex1I) | |
| 528 | apply (fast dest!: theI') | |
| 529 | apply (fast intro: ext the1_equality [symmetric]) | |
| 530 | apply (erule ex1E) | |
| 531 | apply (rule allI) | |
| 532 | apply (rule ex1I) | |
| 533 | apply (erule spec) | |
| 534 | apply (rule ccontr) | |
| 535 | apply (erule_tac x = "%z. if z = x then y else f z" in allE) | |
| 536 | apply (erule impE) | |
| 537 | apply (rule allI) | |
| 538 | apply (rule_tac P = "xa = x" in case_split_thm) | |
| 539 | apply (drule_tac [3] x = x in fun_cong) | |
| 540 | apply simp_all | |
| 541 | done | |
| 542 | ||
| 13438 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 543 | text{*Needs only HOL-lemmas:*}
 | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 544 | lemma mk_left_commute: | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 545 | assumes a: "\<And>x y z. f (f x y) z = f x (f y z)" and | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 546 | c: "\<And>x y. f x y = f y x" | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 547 | shows "f x (f y z) = f y (f x z)" | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 548 | by(rule trans[OF trans[OF c a] arg_cong[OF c, of "f y"]]) | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 549 | |
| 11750 | 550 | |
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 551 | subsubsection {* Generic cases and induction *}
 | 
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 552 | |
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 553 | constdefs | 
| 11989 | 554 |   induct_forall :: "('a => bool) => bool"
 | 
| 555 | "induct_forall P == \<forall>x. P x" | |
| 556 | induct_implies :: "bool => bool => bool" | |
| 557 | "induct_implies A B == A --> B" | |
| 558 | induct_equal :: "'a => 'a => bool" | |
| 559 | "induct_equal x y == x = y" | |
| 560 | induct_conj :: "bool => bool => bool" | |
| 561 | "induct_conj A B == A & B" | |
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 562 | |
| 11989 | 563 | lemma induct_forall_eq: "(!!x. P x) == Trueprop (induct_forall (\<lambda>x. P x))" | 
| 564 | by (simp only: atomize_all induct_forall_def) | |
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 565 | |
| 11989 | 566 | lemma induct_implies_eq: "(A ==> B) == Trueprop (induct_implies A B)" | 
| 567 | by (simp only: atomize_imp induct_implies_def) | |
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 568 | |
| 11989 | 569 | lemma induct_equal_eq: "(x == y) == Trueprop (induct_equal x y)" | 
| 570 | by (simp only: atomize_eq induct_equal_def) | |
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 571 | |
| 11989 | 572 | lemma induct_forall_conj: "induct_forall (\<lambda>x. induct_conj (A x) (B x)) = | 
| 573 | induct_conj (induct_forall A) (induct_forall B)" | |
| 12354 | 574 | by (unfold induct_forall_def induct_conj_def) rules | 
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 575 | |
| 11989 | 576 | lemma induct_implies_conj: "induct_implies C (induct_conj A B) = | 
| 577 | induct_conj (induct_implies C A) (induct_implies C B)" | |
| 12354 | 578 | by (unfold induct_implies_def induct_conj_def) rules | 
| 11989 | 579 | |
| 13598 
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
 berghofe parents: 
13596diff
changeset | 580 | lemma induct_conj_curry: "(induct_conj A B ==> PROP C) == (A ==> B ==> PROP C)" | 
| 
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
 berghofe parents: 
13596diff
changeset | 581 | proof | 
| 
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
 berghofe parents: 
13596diff
changeset | 582 | assume r: "induct_conj A B ==> PROP C" and A B | 
| 
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
 berghofe parents: 
13596diff
changeset | 583 | show "PROP C" by (rule r) (simp! add: induct_conj_def) | 
| 
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
 berghofe parents: 
13596diff
changeset | 584 | next | 
| 
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
 berghofe parents: 
13596diff
changeset | 585 | assume r: "A ==> B ==> PROP C" and "induct_conj A B" | 
| 
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
 berghofe parents: 
13596diff
changeset | 586 | show "PROP C" by (rule r) (simp! add: induct_conj_def)+ | 
| 
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
 berghofe parents: 
13596diff
changeset | 587 | qed | 
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 588 | |
| 11989 | 589 | lemma induct_impliesI: "(A ==> B) ==> induct_implies A B" | 
| 590 | by (simp add: induct_implies_def) | |
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 591 | |
| 12161 | 592 | lemmas induct_atomize = atomize_conj induct_forall_eq induct_implies_eq induct_equal_eq | 
| 593 | lemmas induct_rulify1 [symmetric, standard] = induct_forall_eq induct_implies_eq induct_equal_eq | |
| 594 | lemmas induct_rulify2 = induct_forall_def induct_implies_def induct_equal_def induct_conj_def | |
| 11989 | 595 | 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 | 596 | |
| 11989 | 597 | hide const induct_forall induct_implies induct_equal induct_conj | 
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 598 | |
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 599 | |
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 600 | text {* Method setup. *}
 | 
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 601 | |
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 602 | ML {*
 | 
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 603 | structure InductMethod = InductMethodFun | 
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 604 | (struct | 
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 605 | val dest_concls = HOLogic.dest_concls; | 
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 606 | val cases_default = thm "case_split"; | 
| 11989 | 607 | val local_impI = thm "induct_impliesI"; | 
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 608 | val conjI = thm "conjI"; | 
| 11989 | 609 | val atomize = thms "induct_atomize"; | 
| 610 | val rulify1 = thms "induct_rulify1"; | |
| 611 | val rulify2 = thms "induct_rulify2"; | |
| 12240 | 612 | val localize = [Thm.symmetric (thm "induct_implies_def")]; | 
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 613 | end); | 
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 614 | *} | 
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 615 | |
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 616 | setup InductMethod.setup | 
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 617 | |
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 618 | |
| 11750 | 619 | subsection {* Order signatures and orders *}
 | 
| 620 | ||
| 621 | axclass | |
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12281diff
changeset | 622 | ord < type | 
| 11750 | 623 | |
| 624 | syntax | |
| 625 |   "op <"        :: "['a::ord, 'a] => bool"             ("op <")
 | |
| 626 |   "op <="       :: "['a::ord, 'a] => bool"             ("op <=")
 | |
| 627 | ||
| 628 | global | |
| 629 | ||
| 630 | consts | |
| 631 |   "op <"        :: "['a::ord, 'a] => bool"             ("(_/ < _)"  [50, 51] 50)
 | |
| 632 |   "op <="       :: "['a::ord, 'a] => bool"             ("(_/ <= _)" [50, 51] 50)
 | |
| 633 | ||
| 634 | local | |
| 635 | ||
| 12114 
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
 wenzelm parents: 
12023diff
changeset | 636 | syntax (xsymbols) | 
| 11750 | 637 |   "op <="       :: "['a::ord, 'a] => bool"             ("op \<le>")
 | 
| 638 |   "op <="       :: "['a::ord, 'a] => bool"             ("(_/ \<le> _)"  [50, 51] 50)
 | |
| 639 | ||
| 640 | ||
| 641 | subsubsection {* Monotonicity *}
 | |
| 642 | ||
| 13412 | 643 | locale mono = | 
| 644 | fixes f | |
| 645 | assumes mono: "A <= B ==> f A <= f B" | |
| 11750 | 646 | |
| 13421 | 647 | lemmas monoI [intro?] = mono.intro | 
| 13412 | 648 | and monoD [dest?] = mono.mono | 
| 11750 | 649 | |
| 650 | constdefs | |
| 651 | min :: "['a::ord, 'a] => 'a" | |
| 652 | "min a b == (if a <= b then a else b)" | |
| 653 | max :: "['a::ord, 'a] => 'a" | |
| 654 | "max a b == (if a <= b then b else a)" | |
| 655 | ||
| 656 | lemma min_leastL: "(!!x. least <= x) ==> min least x = least" | |
| 657 | by (simp add: min_def) | |
| 658 | ||
| 659 | lemma min_of_mono: | |
| 660 | "ALL x y. (f x <= f y) = (x <= y) ==> min (f m) (f n) = f (min m n)" | |
| 661 | by (simp add: min_def) | |
| 662 | ||
| 663 | lemma max_leastL: "(!!x. least <= x) ==> max least x = x" | |
| 664 | by (simp add: max_def) | |
| 665 | ||
| 666 | lemma max_of_mono: | |
| 667 | "ALL x y. (f x <= f y) = (x <= y) ==> max (f m) (f n) = f (max m n)" | |
| 668 | by (simp add: max_def) | |
| 669 | ||
| 670 | ||
| 671 | subsubsection "Orders" | |
| 672 | ||
| 673 | axclass order < ord | |
| 674 | order_refl [iff]: "x <= x" | |
| 675 | order_trans: "x <= y ==> y <= z ==> x <= z" | |
| 676 | order_antisym: "x <= y ==> y <= x ==> x = y" | |
| 677 | order_less_le: "(x < y) = (x <= y & x ~= y)" | |
| 678 | ||
| 679 | ||
| 680 | text {* Reflexivity. *}
 | |
| 681 | ||
| 682 | lemma order_eq_refl: "!!x::'a::order. x = y ==> x <= y" | |
| 683 |     -- {* This form is useful with the classical reasoner. *}
 | |
| 684 | apply (erule ssubst) | |
| 685 | apply (rule order_refl) | |
| 686 | done | |
| 687 | ||
| 13553 | 688 | lemma order_less_irrefl [iff]: "~ x < (x::'a::order)" | 
| 11750 | 689 | by (simp add: order_less_le) | 
| 690 | ||
| 691 | lemma order_le_less: "((x::'a::order) <= y) = (x < y | x = y)" | |
| 692 |     -- {* NOT suitable for iff, since it can cause PROOF FAILED. *}
 | |
| 693 | apply (simp add: order_less_le) | |
| 12256 | 694 | apply blast | 
| 11750 | 695 | done | 
| 696 | ||
| 697 | lemmas order_le_imp_less_or_eq = order_le_less [THEN iffD1, standard] | |
| 698 | ||
| 699 | lemma order_less_imp_le: "!!x::'a::order. x < y ==> x <= y" | |
| 700 | by (simp add: order_less_le) | |
| 701 | ||
| 702 | ||
| 703 | text {* Asymmetry. *}
 | |
| 704 | ||
| 705 | lemma order_less_not_sym: "(x::'a::order) < y ==> ~ (y < x)" | |
| 706 | by (simp add: order_less_le order_antisym) | |
| 707 | ||
| 708 | lemma order_less_asym: "x < (y::'a::order) ==> (~P ==> y < x) ==> P" | |
| 709 | apply (drule order_less_not_sym) | |
| 710 | apply (erule contrapos_np) | |
| 711 | apply simp | |
| 712 | done | |
| 713 | ||
| 714 | ||
| 715 | text {* Transitivity. *}
 | |
| 716 | ||
| 717 | lemma order_less_trans: "!!x::'a::order. [| x < y; y < z |] ==> x < z" | |
| 718 | apply (simp add: order_less_le) | |
| 719 | apply (blast intro: order_trans order_antisym) | |
| 720 | done | |
| 721 | ||
| 722 | lemma order_le_less_trans: "!!x::'a::order. [| x <= y; y < z |] ==> x < z" | |
| 723 | apply (simp add: order_less_le) | |
| 724 | apply (blast intro: order_trans order_antisym) | |
| 725 | done | |
| 726 | ||
| 727 | lemma order_less_le_trans: "!!x::'a::order. [| x < y; y <= z |] ==> x < z" | |
| 728 | apply (simp add: order_less_le) | |
| 729 | apply (blast intro: order_trans order_antisym) | |
| 730 | done | |
| 731 | ||
| 732 | ||
| 733 | text {* Useful for simplification, but too risky to include by default. *}
 | |
| 734 | ||
| 735 | lemma order_less_imp_not_less: "(x::'a::order) < y ==> (~ y < x) = True" | |
| 736 | by (blast elim: order_less_asym) | |
| 737 | ||
| 738 | lemma order_less_imp_triv: "(x::'a::order) < y ==> (y < x --> P) = True" | |
| 739 | by (blast elim: order_less_asym) | |
| 740 | ||
| 741 | lemma order_less_imp_not_eq: "(x::'a::order) < y ==> (x = y) = False" | |
| 742 | by auto | |
| 743 | ||
| 744 | lemma order_less_imp_not_eq2: "(x::'a::order) < y ==> (y = x) = False" | |
| 745 | by auto | |
| 746 | ||
| 747 | ||
| 748 | text {* Other operators. *}
 | |
| 749 | ||
| 750 | lemma min_leastR: "(!!x::'a::order. least <= x) ==> min x least = least" | |
| 751 | apply (simp add: min_def) | |
| 752 | apply (blast intro: order_antisym) | |
| 753 | done | |
| 754 | ||
| 755 | lemma max_leastR: "(!!x::'a::order. least <= x) ==> max x least = x" | |
| 756 | apply (simp add: max_def) | |
| 757 | apply (blast intro: order_antisym) | |
| 758 | done | |
| 759 | ||
| 760 | ||
| 761 | subsubsection {* Least value operator *}
 | |
| 762 | ||
| 763 | constdefs | |
| 764 |   Least :: "('a::ord => bool) => 'a"               (binder "LEAST " 10)
 | |
| 765 | "Least P == THE x. P x & (ALL y. P y --> x <= y)" | |
| 766 |     -- {* We can no longer use LeastM because the latter requires Hilbert-AC. *}
 | |
| 767 | ||
| 768 | lemma LeastI2: | |
| 769 | "[| P (x::'a::order); | |
| 770 | !!y. P y ==> x <= y; | |
| 771 | !!x. [| P x; ALL y. P y --> x \<le> y |] ==> Q x |] | |
| 12281 | 772 | ==> Q (Least P)" | 
| 11750 | 773 | apply (unfold Least_def) | 
| 774 | apply (rule theI2) | |
| 775 | apply (blast intro: order_antisym)+ | |
| 776 | done | |
| 777 | ||
| 778 | lemma Least_equality: | |
| 12281 | 779 | "[| P (k::'a::order); !!x. P x ==> k <= x |] ==> (LEAST x. P x) = k" | 
| 11750 | 780 | apply (simp add: Least_def) | 
| 781 | apply (rule the_equality) | |
| 782 | apply (auto intro!: order_antisym) | |
| 783 | done | |
| 784 | ||
| 785 | ||
| 786 | subsubsection "Linear / total orders" | |
| 787 | ||
| 788 | axclass linorder < order | |
| 789 | linorder_linear: "x <= y | y <= x" | |
| 790 | ||
| 791 | lemma linorder_less_linear: "!!x::'a::linorder. x<y | x=y | y<x" | |
| 792 | apply (simp add: order_less_le) | |
| 793 | apply (insert linorder_linear) | |
| 794 | apply blast | |
| 795 | done | |
| 796 | ||
| 797 | lemma linorder_cases [case_names less equal greater]: | |
| 798 | "((x::'a::linorder) < y ==> P) ==> (x = y ==> P) ==> (y < x ==> P) ==> P" | |
| 799 | apply (insert linorder_less_linear) | |
| 800 | apply blast | |
| 801 | done | |
| 802 | ||
| 803 | lemma linorder_not_less: "!!x::'a::linorder. (~ x < y) = (y <= x)" | |
| 804 | apply (simp add: order_less_le) | |
| 805 | apply (insert linorder_linear) | |
| 806 | apply (blast intro: order_antisym) | |
| 807 | done | |
| 808 | ||
| 809 | lemma linorder_not_le: "!!x::'a::linorder. (~ x <= y) = (y < x)" | |
| 810 | apply (simp add: order_less_le) | |
| 811 | apply (insert linorder_linear) | |
| 812 | apply (blast intro: order_antisym) | |
| 813 | done | |
| 814 | ||
| 815 | lemma linorder_neq_iff: "!!x::'a::linorder. (x ~= y) = (x<y | y<x)" | |
| 816 | apply (cut_tac x = x and y = y in linorder_less_linear) | |
| 817 | apply auto | |
| 818 | done | |
| 819 | ||
| 820 | lemma linorder_neqE: "x ~= (y::'a::linorder) ==> (x < y ==> R) ==> (y < x ==> R) ==> R" | |
| 821 | apply (simp add: linorder_neq_iff) | |
| 822 | apply blast | |
| 823 | done | |
| 824 | ||
| 825 | ||
| 826 | subsubsection "Min and max on (linear) orders" | |
| 827 | ||
| 828 | lemma min_same [simp]: "min (x::'a::order) x = x" | |
| 829 | by (simp add: min_def) | |
| 830 | ||
| 831 | lemma max_same [simp]: "max (x::'a::order) x = x" | |
| 832 | by (simp add: max_def) | |
| 833 | ||
| 834 | lemma le_max_iff_disj: "!!z::'a::linorder. (z <= max x y) = (z <= x | z <= y)" | |
| 835 | apply (simp add: max_def) | |
| 836 | apply (insert linorder_linear) | |
| 837 | apply (blast intro: order_trans) | |
| 838 | done | |
| 839 | ||
| 840 | lemma le_maxI1: "(x::'a::linorder) <= max x y" | |
| 841 | by (simp add: le_max_iff_disj) | |
| 842 | ||
| 843 | lemma le_maxI2: "(y::'a::linorder) <= max x y" | |
| 844 |     -- {* CANNOT use with @{text "[intro!]"} because blast will give PROOF FAILED. *}
 | |
| 845 | by (simp add: le_max_iff_disj) | |
| 846 | ||
| 847 | lemma less_max_iff_disj: "!!z::'a::linorder. (z < max x y) = (z < x | z < y)" | |
| 848 | apply (simp add: max_def order_le_less) | |
| 849 | apply (insert linorder_less_linear) | |
| 850 | apply (blast intro: order_less_trans) | |
| 851 | done | |
| 852 | ||
| 853 | lemma max_le_iff_conj [simp]: | |
| 854 | "!!z::'a::linorder. (max x y <= z) = (x <= z & y <= z)" | |
| 855 | apply (simp add: max_def) | |
| 856 | apply (insert linorder_linear) | |
| 857 | apply (blast intro: order_trans) | |
| 858 | done | |
| 859 | ||
| 860 | lemma max_less_iff_conj [simp]: | |
| 861 | "!!z::'a::linorder. (max x y < z) = (x < z & y < z)" | |
| 862 | apply (simp add: order_le_less max_def) | |
| 863 | apply (insert linorder_less_linear) | |
| 864 | apply (blast intro: order_less_trans) | |
| 865 | done | |
| 866 | ||
| 867 | lemma le_min_iff_conj [simp]: | |
| 868 | "!!z::'a::linorder. (z <= min x y) = (z <= x & z <= y)" | |
| 12892 | 869 |     -- {* @{text "[iff]"} screws up a @{text blast} in MiniML *}
 | 
| 11750 | 870 | apply (simp add: min_def) | 
| 871 | apply (insert linorder_linear) | |
| 872 | apply (blast intro: order_trans) | |
| 873 | done | |
| 874 | ||
| 875 | lemma min_less_iff_conj [simp]: | |
| 876 | "!!z::'a::linorder. (z < min x y) = (z < x & z < y)" | |
| 877 | apply (simp add: order_le_less min_def) | |
| 878 | apply (insert linorder_less_linear) | |
| 879 | apply (blast intro: order_less_trans) | |
| 880 | done | |
| 881 | ||
| 882 | lemma min_le_iff_disj: "!!z::'a::linorder. (min x y <= z) = (x <= z | y <= z)" | |
| 883 | apply (simp add: min_def) | |
| 884 | apply (insert linorder_linear) | |
| 885 | apply (blast intro: order_trans) | |
| 886 | done | |
| 887 | ||
| 888 | lemma min_less_iff_disj: "!!z::'a::linorder. (min x y < z) = (x < z | y < z)" | |
| 889 | apply (simp add: min_def order_le_less) | |
| 890 | apply (insert linorder_less_linear) | |
| 891 | apply (blast intro: order_less_trans) | |
| 892 | done | |
| 893 | ||
| 13438 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 894 | lemma max_assoc: "!!x::'a::linorder. max (max x y) z = max x (max y z)" | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 895 | apply(simp add:max_def) | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 896 | apply(rule conjI) | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 897 | apply(blast intro:order_trans) | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 898 | apply(simp add:linorder_not_le) | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 899 | apply(blast dest: order_less_trans order_le_less_trans) | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 900 | done | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 901 | |
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 902 | lemma max_commute: "!!x::'a::linorder. max x y = max y x" | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 903 | apply(simp add:max_def) | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 904 | apply(rule conjI) | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 905 | apply(blast intro:order_antisym) | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 906 | apply(simp add:linorder_not_le) | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 907 | apply(blast dest: order_less_trans) | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 908 | done | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 909 | |
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 910 | lemmas max_ac = max_assoc max_commute | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 911 | mk_left_commute[of max,OF max_assoc max_commute] | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 912 | |
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 913 | lemma min_assoc: "!!x::'a::linorder. min (min x y) z = min x (min y z)" | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 914 | apply(simp add:min_def) | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 915 | apply(rule conjI) | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 916 | apply(blast intro:order_trans) | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 917 | apply(simp add:linorder_not_le) | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 918 | apply(blast dest: order_less_trans order_le_less_trans) | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 919 | done | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 920 | |
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 921 | lemma min_commute: "!!x::'a::linorder. min x y = min y x" | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 922 | apply(simp add:min_def) | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 923 | apply(rule conjI) | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 924 | apply(blast intro:order_antisym) | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 925 | apply(simp add:linorder_not_le) | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 926 | apply(blast dest: order_less_trans) | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 927 | done | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 928 | |
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 929 | lemmas min_ac = min_assoc min_commute | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 930 | mk_left_commute[of min,OF min_assoc min_commute] | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 931 | |
| 11750 | 932 | lemma split_min: | 
| 933 | "P (min (i::'a::linorder) j) = ((i <= j --> P(i)) & (~ i <= j --> P(j)))" | |
| 934 | by (simp add: min_def) | |
| 935 | ||
| 936 | lemma split_max: | |
| 937 | "P (max (i::'a::linorder) j) = ((i <= j --> P(j)) & (~ i <= j --> P(i)))" | |
| 938 | by (simp add: max_def) | |
| 939 | ||
| 940 | ||
| 941 | subsubsection "Bounded quantifiers" | |
| 942 | ||
| 943 | syntax | |
| 944 |   "_lessAll" :: "[idt, 'a, bool] => bool"   ("(3ALL _<_./ _)"  [0, 0, 10] 10)
 | |
| 945 |   "_lessEx"  :: "[idt, 'a, bool] => bool"   ("(3EX _<_./ _)"  [0, 0, 10] 10)
 | |
| 946 |   "_leAll"   :: "[idt, 'a, bool] => bool"   ("(3ALL _<=_./ _)" [0, 0, 10] 10)
 | |
| 947 |   "_leEx"    :: "[idt, 'a, bool] => bool"   ("(3EX _<=_./ _)" [0, 0, 10] 10)
 | |
| 948 | ||
| 12114 
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
 wenzelm parents: 
12023diff
changeset | 949 | syntax (xsymbols) | 
| 11750 | 950 |   "_lessAll" :: "[idt, 'a, bool] => bool"   ("(3\<forall>_<_./ _)"  [0, 0, 10] 10)
 | 
| 951 |   "_lessEx"  :: "[idt, 'a, bool] => bool"   ("(3\<exists>_<_./ _)"  [0, 0, 10] 10)
 | |
| 952 |   "_leAll"   :: "[idt, 'a, bool] => bool"   ("(3\<forall>_\<le>_./ _)" [0, 0, 10] 10)
 | |
| 953 |   "_leEx"    :: "[idt, 'a, bool] => bool"   ("(3\<exists>_\<le>_./ _)" [0, 0, 10] 10)
 | |
| 954 | ||
| 955 | syntax (HOL) | |
| 956 |   "_lessAll" :: "[idt, 'a, bool] => bool"   ("(3! _<_./ _)"  [0, 0, 10] 10)
 | |
| 957 |   "_lessEx"  :: "[idt, 'a, bool] => bool"   ("(3? _<_./ _)"  [0, 0, 10] 10)
 | |
| 958 |   "_leAll"   :: "[idt, 'a, bool] => bool"   ("(3! _<=_./ _)" [0, 0, 10] 10)
 | |
| 959 |   "_leEx"    :: "[idt, 'a, bool] => bool"   ("(3? _<=_./ _)" [0, 0, 10] 10)
 | |
| 960 | ||
| 961 | translations | |
| 962 | "ALL x<y. P" => "ALL x. x < y --> P" | |
| 963 | "EX x<y. P" => "EX x. x < y & P" | |
| 964 | "ALL x<=y. P" => "ALL x. x <= y --> P" | |
| 965 | "EX x<=y. P" => "EX x. x <= y & P" | |
| 966 | ||
| 923 | 967 | end |