| author | haftmann | 
| Sun, 05 Jun 2005 14:41:23 +0200 | |
| changeset 16276 | 3a50bf1f04d0 | 
| parent 16121 | a80aa66d2271 | 
| child 16417 | 9bc16273c2d4 | 
| 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 | |
| 15131 | 8 | theory HOL | 
| 15140 | 9 | imports CPure | 
| 15524 
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
 nipkow parents: 
15481diff
changeset | 10 | files ("cladata.ML") ("blastdata.ML") ("simpdata.ML") ("eqrule_HOL_data.ML")
 | 
| 15481 | 11 |       ("~~/src/Provers/eqsubst.ML")
 | 
| 15131 | 12 | begin | 
| 2260 | 13 | |
| 11750 | 14 | subsection {* Primitive logic *}
 | 
| 15 | ||
| 16 | subsubsection {* Core syntax *}
 | |
| 2260 | 17 | |
| 14854 | 18 | classes type | 
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12281diff
changeset | 19 | defaultsort type | 
| 3947 | 20 | |
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12281diff
changeset | 21 | global | 
| 923 | 22 | |
| 7357 | 23 | typedecl bool | 
| 923 | 24 | |
| 25 | arities | |
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12281diff
changeset | 26 | bool :: type | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12281diff
changeset | 27 | fun :: (type, type) type | 
| 923 | 28 | |
| 11750 | 29 | judgment | 
| 30 |   Trueprop      :: "bool => prop"                   ("(_)" 5)
 | |
| 923 | 31 | |
| 11750 | 32 | consts | 
| 7357 | 33 |   Not           :: "bool => bool"                   ("~ _" [40] 40)
 | 
| 34 | True :: bool | |
| 35 | False :: bool | |
| 36 |   If            :: "[bool, 'a, 'a] => 'a"           ("(if (_)/ then (_)/ else (_))" 10)
 | |
| 3947 | 37 | arbitrary :: 'a | 
| 923 | 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 | |
| 7357 | 45 | "=" :: "['a, 'a] => bool" (infixl 50) | 
| 46 | & :: "[bool, bool] => bool" (infixr 35) | |
| 47 | "|" :: "[bool, bool] => bool" (infixr 30) | |
| 48 | --> :: "[bool, bool] => bool" (infixr 25) | |
| 923 | 49 | |
| 10432 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 50 | local | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 51 | |
| 2260 | 52 | |
| 11750 | 53 | subsubsection {* Additional concrete syntax *}
 | 
| 2260 | 54 | |
| 4868 | 55 | nonterminals | 
| 923 | 56 | letbinds letbind | 
| 57 | case_syn cases_syn | |
| 58 | ||
| 59 | syntax | |
| 12650 | 60 | "_not_equal" :: "['a, 'a] => bool" (infixl "~=" 50) | 
| 11432 
8a203ae6efe3
added "The" (definite description operator) (by Larry);
 wenzelm parents: 
10489diff
changeset | 61 |   "_The"        :: "[pttrn, bool] => 'a"                 ("(3THE _./ _)" [0, 10] 10)
 | 
| 923 | 62 | |
| 7357 | 63 |   "_bind"       :: "[pttrn, 'a] => letbind"              ("(2_ =/ _)" 10)
 | 
| 64 |   ""            :: "letbind => letbinds"                 ("_")
 | |
| 65 |   "_binds"      :: "[letbind, letbinds] => letbinds"     ("_;/ _")
 | |
| 66 |   "_Let"        :: "[letbinds, 'a] => 'a"                ("(let (_)/ in (_))" 10)
 | |
| 923 | 67 | |
| 9060 
b0dd884b1848
rename @case to _case_syntax (improves on low-level errors);
 wenzelm parents: 
8959diff
changeset | 68 |   "_case_syntax":: "['a, cases_syn] => 'b"               ("(case _ of/ _)" 10)
 | 
| 
b0dd884b1848
rename @case to _case_syntax (improves on low-level errors);
 wenzelm parents: 
8959diff
changeset | 69 |   "_case1"      :: "['a, 'b] => case_syn"                ("(2_ =>/ _)" 10)
 | 
| 7357 | 70 |   ""            :: "case_syn => cases_syn"               ("_")
 | 
| 9060 
b0dd884b1848
rename @case to _case_syntax (improves on low-level errors);
 wenzelm parents: 
8959diff
changeset | 71 |   "_case2"      :: "[case_syn, cases_syn] => cases_syn"  ("_/ | _")
 | 
| 923 | 72 | |
| 73 | translations | |
| 7238 
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
 wenzelm parents: 
7220diff
changeset | 74 | "x ~= y" == "~ (x = y)" | 
| 13764 | 75 | "THE x. P" == "The (%x. P)" | 
| 923 | 76 | "_Let (_binds b bs) e" == "_Let b (_Let bs e)" | 
| 1114 | 77 | "let x = a in e" == "Let a (%x. e)" | 
| 923 | 78 | |
| 13763 
f94b569cd610
added print translations tha avoid eta contraction for important binders.
 nipkow parents: 
13723diff
changeset | 79 | print_translation {*
 | 
| 
f94b569cd610
added print translations tha avoid eta contraction for important binders.
 nipkow parents: 
13723diff
changeset | 80 | (* To avoid eta-contraction of body: *) | 
| 
f94b569cd610
added print translations tha avoid eta contraction for important binders.
 nipkow parents: 
13723diff
changeset | 81 | [("The", fn [Abs abs] =>
 | 
| 
f94b569cd610
added print translations tha avoid eta contraction for important binders.
 nipkow parents: 
13723diff
changeset | 82 | let val (x,t) = atomic_abs_tr' abs | 
| 
f94b569cd610
added print translations tha avoid eta contraction for important binders.
 nipkow parents: 
13723diff
changeset | 83 | in Syntax.const "_The" $ x $ t end)] | 
| 
f94b569cd610
added print translations tha avoid eta contraction for important binders.
 nipkow parents: 
13723diff
changeset | 84 | *} | 
| 
f94b569cd610
added print translations tha avoid eta contraction for important binders.
 nipkow parents: 
13723diff
changeset | 85 | |
| 12633 | 86 | syntax (output) | 
| 11687 | 87 | "=" :: "['a, 'a] => bool" (infix 50) | 
| 12650 | 88 | "_not_equal" :: "['a, 'a] => bool" (infix "~=" 50) | 
| 2260 | 89 | |
| 12114 
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
 wenzelm parents: 
12023diff
changeset | 90 | syntax (xsymbols) | 
| 11687 | 91 |   Not           :: "bool => bool"                        ("\<not> _" [40] 40)
 | 
| 92 | "op &" :: "[bool, bool] => bool" (infixr "\<and>" 35) | |
| 93 | "op |" :: "[bool, bool] => bool" (infixr "\<or>" 30) | |
| 12114 
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
 wenzelm parents: 
12023diff
changeset | 94 | "op -->" :: "[bool, bool] => bool" (infixr "\<longrightarrow>" 25) | 
| 12650 | 95 | "_not_equal" :: "['a, 'a] => bool" (infix "\<noteq>" 50) | 
| 11687 | 96 |   "ALL "        :: "[idts, bool] => bool"                ("(3\<forall>_./ _)" [0, 10] 10)
 | 
| 97 |   "EX "         :: "[idts, bool] => bool"                ("(3\<exists>_./ _)" [0, 10] 10)
 | |
| 98 |   "EX! "        :: "[idts, bool] => bool"                ("(3\<exists>!_./ _)" [0, 10] 10)
 | |
| 99 |   "_case1"      :: "['a, 'b] => case_syn"                ("(2_ \<Rightarrow>/ _)" 10)
 | |
| 14361 
ad2f5da643b4
* Support for raw latex output in control symbols: \<^raw...>
 schirmer parents: 
14357diff
changeset | 100 | (*"_case2"      :: "[case_syn, cases_syn] => cases_syn"  ("_/ \<orelse> _")*)
 | 
| 2372 | 101 | |
| 12114 
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
 wenzelm parents: 
12023diff
changeset | 102 | syntax (xsymbols output) | 
| 12650 | 103 | "_not_equal" :: "['a, 'a] => bool" (infix "\<noteq>" 50) | 
| 3820 | 104 | |
| 6340 | 105 | syntax (HTML output) | 
| 14565 | 106 | "_not_equal" :: "['a, 'a] => bool" (infix "\<noteq>" 50) | 
| 11687 | 107 |   Not           :: "bool => bool"                        ("\<not> _" [40] 40)
 | 
| 14565 | 108 | "op &" :: "[bool, bool] => bool" (infixr "\<and>" 35) | 
| 109 | "op |" :: "[bool, bool] => bool" (infixr "\<or>" 30) | |
| 110 | "_not_equal" :: "['a, 'a] => bool" (infix "\<noteq>" 50) | |
| 111 |   "ALL "        :: "[idts, bool] => bool"                ("(3\<forall>_./ _)" [0, 10] 10)
 | |
| 112 |   "EX "         :: "[idts, bool] => bool"                ("(3\<exists>_./ _)" [0, 10] 10)
 | |
| 113 |   "EX! "        :: "[idts, bool] => bool"                ("(3\<exists>!_./ _)" [0, 10] 10)
 | |
| 6340 | 114 | |
| 7238 
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
 wenzelm parents: 
7220diff
changeset | 115 | syntax (HOL) | 
| 7357 | 116 |   "ALL "        :: "[idts, bool] => bool"                ("(3! _./ _)" [0, 10] 10)
 | 
| 117 |   "EX "         :: "[idts, bool] => bool"                ("(3? _./ _)" [0, 10] 10)
 | |
| 118 |   "EX! "        :: "[idts, bool] => bool"                ("(3?! _./ _)" [0, 10] 10)
 | |
| 7238 
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
 wenzelm parents: 
7220diff
changeset | 119 | |
| 
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
 wenzelm parents: 
7220diff
changeset | 120 | |
| 11750 | 121 | subsubsection {* Axioms and basic definitions *}
 | 
| 2260 | 122 | |
| 7357 | 123 | axioms | 
| 15380 | 124 | eq_reflection: "(x=y) ==> (x==y)" | 
| 923 | 125 | |
| 15380 | 126 | refl: "t = (t::'a)" | 
| 6289 | 127 | |
| 15380 | 128 | ext: "(!!x::'a. (f x ::'b) = g x) ==> (%x. f x) = (%x. g x)" | 
| 129 |     -- {*Extensionality is built into the meta-logic, and this rule expresses
 | |
| 130 | a related property. It is an eta-expanded version of the traditional | |
| 131 | rule, and similar to the ABS rule of HOL*} | |
| 6289 | 132 | |
| 11432 
8a203ae6efe3
added "The" (definite description operator) (by Larry);
 wenzelm parents: 
10489diff
changeset | 133 | the_eq_trivial: "(THE x. x = a) = (a::'a)" | 
| 923 | 134 | |
| 15380 | 135 | impI: "(P ==> Q) ==> P-->Q" | 
| 136 | mp: "[| P-->Q; P |] ==> Q" | |
| 137 | ||
| 138 | ||
| 139 | text{*Thanks to Stephan Merz*}
 | |
| 140 | theorem subst: | |
| 141 | assumes eq: "s = t" and p: "P(s)" | |
| 142 | shows "P(t::'a)" | |
| 143 | proof - | |
| 144 | from eq have meta: "s \<equiv> t" | |
| 145 | by (rule eq_reflection) | |
| 146 | from p show ?thesis | |
| 147 | by (unfold meta) | |
| 148 | qed | |
| 923 | 149 | |
| 150 | defs | |
| 7357 | 151 | True_def: "True == ((%x::bool. x) = (%x. x))" | 
| 152 | All_def: "All(P) == (P = (%x. True))" | |
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11438diff
changeset | 153 | Ex_def: "Ex(P) == !Q. (!x. P x --> Q) --> Q" | 
| 7357 | 154 | False_def: "False == (!P. P)" | 
| 155 | not_def: "~ P == P-->False" | |
| 156 | and_def: "P & Q == !R. (P-->Q-->R) --> R" | |
| 157 | or_def: "P | Q == !R. (P-->R) --> (Q-->R) --> R" | |
| 158 | Ex1_def: "Ex1(P) == ? x. P(x) & (! y. P(y) --> y=x)" | |
| 923 | 159 | |
| 7357 | 160 | axioms | 
| 161 | iff: "(P-->Q) --> (Q-->P) --> (P=Q)" | |
| 162 | True_or_False: "(P=True) | (P=False)" | |
| 923 | 163 | |
| 164 | defs | |
| 7357 | 165 | Let_def: "Let s f == f(s)" | 
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11438diff
changeset | 166 | if_def: "If P x y == THE z::'a. (P=True --> z=x) & (P=False --> z=y)" | 
| 5069 | 167 | |
| 14223 
0ee05eef881b
Added support for making constants final, that is, ensuring that no
 skalberg parents: 
14208diff
changeset | 168 | finalconsts | 
| 
0ee05eef881b
Added support for making constants final, that is, ensuring that no
 skalberg parents: 
14208diff
changeset | 169 | "op =" | 
| 
0ee05eef881b
Added support for making constants final, that is, ensuring that no
 skalberg parents: 
14208diff
changeset | 170 | "op -->" | 
| 
0ee05eef881b
Added support for making constants final, that is, ensuring that no
 skalberg parents: 
14208diff
changeset | 171 | The | 
| 
0ee05eef881b
Added support for making constants final, that is, ensuring that no
 skalberg parents: 
14208diff
changeset | 172 | arbitrary | 
| 3320 | 173 | |
| 11750 | 174 | subsubsection {* Generic algebraic operations *}
 | 
| 4868 | 175 | |
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12281diff
changeset | 176 | axclass zero < type | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12281diff
changeset | 177 | axclass one < type | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12281diff
changeset | 178 | axclass plus < type | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12281diff
changeset | 179 | axclass minus < type | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12281diff
changeset | 180 | axclass times < type | 
| 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12281diff
changeset | 181 | axclass inverse < type | 
| 11750 | 182 | |
| 183 | global | |
| 184 | ||
| 185 | consts | |
| 186 |   "0"           :: "'a::zero"                       ("0")
 | |
| 187 |   "1"           :: "'a::one"                        ("1")
 | |
| 188 | "+" :: "['a::plus, 'a] => 'a" (infixl 65) | |
| 189 | - :: "['a::minus, 'a] => 'a" (infixl 65) | |
| 190 |   uminus        :: "['a::minus] => 'a"              ("- _" [81] 80)
 | |
| 191 | * :: "['a::times, 'a] => 'a" (infixl 70) | |
| 192 | ||
| 13456 
42601eb7553f
special syntax for index "1" (plain numeral hidden by "1" symbol in HOL);
 wenzelm parents: 
13438diff
changeset | 193 | syntax | 
| 
42601eb7553f
special syntax for index "1" (plain numeral hidden by "1" symbol in HOL);
 wenzelm parents: 
13438diff
changeset | 194 |   "_index1"  :: index    ("\<^sub>1")
 | 
| 
42601eb7553f
special syntax for index "1" (plain numeral hidden by "1" symbol in HOL);
 wenzelm parents: 
13438diff
changeset | 195 | translations | 
| 14690 | 196 | (index) "\<^sub>1" => (index) "\<^bsub>\<struct>\<^esub>" | 
| 13456 
42601eb7553f
special syntax for index "1" (plain numeral hidden by "1" symbol in HOL);
 wenzelm parents: 
13438diff
changeset | 197 | |
| 11750 | 198 | local | 
| 199 | ||
| 200 | typed_print_translation {*
 | |
| 201 | let | |
| 202 | fun tr' c = (c, fn show_sorts => fn T => fn ts => | |
| 203 | if T = dummyT orelse not (! show_types) andalso can Term.dest_Type T then raise Match | |
| 204 | else Syntax.const Syntax.constrainC $ Syntax.const c $ Syntax.term_of_typ show_sorts T); | |
| 205 | in [tr' "0", tr' "1"] end; | |
| 206 | *} -- {* show types that are presumably too general *}
 | |
| 207 | ||
| 208 | ||
| 209 | consts | |
| 210 | abs :: "'a::minus => 'a" | |
| 211 | inverse :: "'a::inverse => 'a" | |
| 212 | divide :: "['a::inverse, 'a] => 'a" (infixl "'/" 70) | |
| 213 | ||
| 214 | syntax (xsymbols) | |
| 215 |   abs :: "'a::minus => 'a"    ("\<bar>_\<bar>")
 | |
| 216 | syntax (HTML output) | |
| 217 |   abs :: "'a::minus => 'a"    ("\<bar>_\<bar>")
 | |
| 218 | ||
| 219 | ||
| 15411 | 220 | subsection {*Equality*}
 | 
| 221 | ||
| 222 | lemma sym: "s=t ==> t=s" | |
| 223 | apply (erule subst) | |
| 224 | apply (rule refl) | |
| 225 | done | |
| 226 | ||
| 227 | (*calling "standard" reduces maxidx to 0*) | |
| 228 | lemmas ssubst = sym [THEN subst, standard] | |
| 229 | ||
| 230 | lemma trans: "[| r=s; s=t |] ==> r=t" | |
| 231 | apply (erule subst , assumption) | |
| 232 | done | |
| 233 | ||
| 234 | lemma def_imp_eq: assumes meq: "A == B" shows "A = B" | |
| 235 | apply (unfold meq) | |
| 236 | apply (rule refl) | |
| 237 | done | |
| 238 | ||
| 239 | (*Useful with eresolve_tac for proving equalties from known equalities. | |
| 240 | a = b | |
| 241 | | | | |
| 242 | c = d *) | |
| 243 | lemma box_equals: "[| a=b; a=c; b=d |] ==> c=d" | |
| 244 | apply (rule trans) | |
| 245 | apply (rule trans) | |
| 246 | apply (rule sym) | |
| 247 | apply assumption+ | |
| 248 | done | |
| 249 | ||
| 15524 
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
 nipkow parents: 
15481diff
changeset | 250 | text {* For calculational reasoning: *}
 | 
| 
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
 nipkow parents: 
15481diff
changeset | 251 | |
| 
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
 nipkow parents: 
15481diff
changeset | 252 | lemma forw_subst: "a = b ==> P b ==> P a" | 
| 
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
 nipkow parents: 
15481diff
changeset | 253 | by (rule ssubst) | 
| 
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
 nipkow parents: 
15481diff
changeset | 254 | |
| 
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
 nipkow parents: 
15481diff
changeset | 255 | lemma back_subst: "P a ==> a = b ==> P b" | 
| 
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
 nipkow parents: 
15481diff
changeset | 256 | by (rule subst) | 
| 
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
 nipkow parents: 
15481diff
changeset | 257 | |
| 15411 | 258 | |
| 259 | subsection {*Congruence rules for application*}
 | |
| 260 | ||
| 261 | (*similar to AP_THM in Gordon's HOL*) | |
| 262 | lemma fun_cong: "(f::'a=>'b) = g ==> f(x)=g(x)" | |
| 263 | apply (erule subst) | |
| 264 | apply (rule refl) | |
| 265 | done | |
| 266 | ||
| 267 | (*similar to AP_TERM in Gordon's HOL and FOL's subst_context*) | |
| 268 | lemma arg_cong: "x=y ==> f(x)=f(y)" | |
| 269 | apply (erule subst) | |
| 270 | apply (rule refl) | |
| 271 | done | |
| 272 | ||
| 15655 | 273 | lemma arg_cong2: "\<lbrakk> a = b; c = d \<rbrakk> \<Longrightarrow> f a c = f b d" | 
| 274 | apply (erule ssubst)+ | |
| 275 | apply (rule refl) | |
| 276 | done | |
| 277 | ||
| 278 | ||
| 15411 | 279 | lemma cong: "[| f = g; (x::'a) = y |] ==> f(x) = g(y)" | 
| 280 | apply (erule subst)+ | |
| 281 | apply (rule refl) | |
| 282 | done | |
| 283 | ||
| 284 | ||
| 285 | subsection {*Equality of booleans -- iff*}
 | |
| 286 | ||
| 287 | lemma iffI: assumes prems: "P ==> Q" "Q ==> P" shows "P=Q" | |
| 288 | apply (rules intro: iff [THEN mp, THEN mp] impI prems) | |
| 289 | done | |
| 290 | ||
| 291 | lemma iffD2: "[| P=Q; Q |] ==> P" | |
| 292 | apply (erule ssubst) | |
| 293 | apply assumption | |
| 294 | done | |
| 295 | ||
| 296 | lemma rev_iffD2: "[| Q; P=Q |] ==> P" | |
| 297 | apply (erule iffD2) | |
| 298 | apply assumption | |
| 299 | done | |
| 300 | ||
| 301 | lemmas iffD1 = sym [THEN iffD2, standard] | |
| 302 | lemmas rev_iffD1 = sym [THEN [2] rev_iffD2, standard] | |
| 303 | ||
| 304 | lemma iffE: | |
| 305 | assumes major: "P=Q" | |
| 306 | and minor: "[| P --> Q; Q --> P |] ==> R" | |
| 307 | shows "R" | |
| 308 | by (rules intro: minor impI major [THEN iffD2] major [THEN iffD1]) | |
| 309 | ||
| 310 | ||
| 311 | subsection {*True*}
 | |
| 312 | ||
| 313 | lemma TrueI: "True" | |
| 314 | apply (unfold True_def) | |
| 315 | apply (rule refl) | |
| 316 | done | |
| 317 | ||
| 318 | lemma eqTrueI: "P ==> P=True" | |
| 319 | by (rules intro: iffI TrueI) | |
| 320 | ||
| 321 | lemma eqTrueE: "P=True ==> P" | |
| 322 | apply (erule iffD2) | |
| 323 | apply (rule TrueI) | |
| 324 | done | |
| 325 | ||
| 326 | ||
| 327 | subsection {*Universal quantifier*}
 | |
| 328 | ||
| 329 | lemma allI: assumes p: "!!x::'a. P(x)" shows "ALL x. P(x)" | |
| 330 | apply (unfold All_def) | |
| 331 | apply (rules intro: ext eqTrueI p) | |
| 332 | done | |
| 333 | ||
| 334 | lemma spec: "ALL x::'a. P(x) ==> P(x)" | |
| 335 | apply (unfold All_def) | |
| 336 | apply (rule eqTrueE) | |
| 337 | apply (erule fun_cong) | |
| 338 | done | |
| 339 | ||
| 340 | lemma allE: | |
| 341 | assumes major: "ALL x. P(x)" | |
| 342 | and minor: "P(x) ==> R" | |
| 343 | shows "R" | |
| 344 | by (rules intro: minor major [THEN spec]) | |
| 345 | ||
| 346 | lemma all_dupE: | |
| 347 | assumes major: "ALL x. P(x)" | |
| 348 | and minor: "[| P(x); ALL x. P(x) |] ==> R" | |
| 349 | shows "R" | |
| 350 | by (rules intro: minor major major [THEN spec]) | |
| 351 | ||
| 352 | ||
| 353 | subsection {*False*}
 | |
| 354 | (*Depends upon spec; it is impossible to do propositional logic before quantifiers!*) | |
| 355 | ||
| 356 | lemma FalseE: "False ==> P" | |
| 357 | apply (unfold False_def) | |
| 358 | apply (erule spec) | |
| 359 | done | |
| 360 | ||
| 361 | lemma False_neq_True: "False=True ==> P" | |
| 362 | by (erule eqTrueE [THEN FalseE]) | |
| 363 | ||
| 364 | ||
| 365 | subsection {*Negation*}
 | |
| 366 | ||
| 367 | lemma notI: | |
| 368 | assumes p: "P ==> False" | |
| 369 | shows "~P" | |
| 370 | apply (unfold not_def) | |
| 371 | apply (rules intro: impI p) | |
| 372 | done | |
| 373 | ||
| 374 | lemma False_not_True: "False ~= True" | |
| 375 | apply (rule notI) | |
| 376 | apply (erule False_neq_True) | |
| 377 | done | |
| 378 | ||
| 379 | lemma True_not_False: "True ~= False" | |
| 380 | apply (rule notI) | |
| 381 | apply (drule sym) | |
| 382 | apply (erule False_neq_True) | |
| 383 | done | |
| 384 | ||
| 385 | lemma notE: "[| ~P; P |] ==> R" | |
| 386 | apply (unfold not_def) | |
| 387 | apply (erule mp [THEN FalseE]) | |
| 388 | apply assumption | |
| 389 | done | |
| 390 | ||
| 391 | (* Alternative ~ introduction rule: [| P ==> ~ Pa; P ==> Pa |] ==> ~ P *) | |
| 392 | lemmas notI2 = notE [THEN notI, standard] | |
| 393 | ||
| 394 | ||
| 395 | subsection {*Implication*}
 | |
| 396 | ||
| 397 | lemma impE: | |
| 398 | assumes "P-->Q" "P" "Q ==> R" | |
| 399 | shows "R" | |
| 400 | by (rules intro: prems mp) | |
| 401 | ||
| 402 | (* Reduces Q to P-->Q, allowing substitution in P. *) | |
| 403 | lemma rev_mp: "[| P; P --> Q |] ==> Q" | |
| 404 | by (rules intro: mp) | |
| 405 | ||
| 406 | lemma contrapos_nn: | |
| 407 | assumes major: "~Q" | |
| 408 | and minor: "P==>Q" | |
| 409 | shows "~P" | |
| 410 | by (rules intro: notI minor major [THEN notE]) | |
| 411 | ||
| 412 | (*not used at all, but we already have the other 3 combinations *) | |
| 413 | lemma contrapos_pn: | |
| 414 | assumes major: "Q" | |
| 415 | and minor: "P ==> ~Q" | |
| 416 | shows "~P" | |
| 417 | by (rules intro: notI minor major notE) | |
| 418 | ||
| 419 | lemma not_sym: "t ~= s ==> s ~= t" | |
| 420 | apply (erule contrapos_nn) | |
| 421 | apply (erule sym) | |
| 422 | done | |
| 423 | ||
| 424 | (*still used in HOLCF*) | |
| 425 | lemma rev_contrapos: | |
| 426 | assumes pq: "P ==> Q" | |
| 427 | and nq: "~Q" | |
| 428 | shows "~P" | |
| 429 | apply (rule nq [THEN contrapos_nn]) | |
| 430 | apply (erule pq) | |
| 431 | done | |
| 432 | ||
| 433 | subsection {*Existential quantifier*}
 | |
| 434 | ||
| 435 | lemma exI: "P x ==> EX x::'a. P x" | |
| 436 | apply (unfold Ex_def) | |
| 437 | apply (rules intro: allI allE impI mp) | |
| 438 | done | |
| 439 | ||
| 440 | lemma exE: | |
| 441 | assumes major: "EX x::'a. P(x)" | |
| 442 | and minor: "!!x. P(x) ==> Q" | |
| 443 | shows "Q" | |
| 444 | apply (rule major [unfolded Ex_def, THEN spec, THEN mp]) | |
| 445 | apply (rules intro: impI [THEN allI] minor) | |
| 446 | done | |
| 447 | ||
| 448 | ||
| 449 | subsection {*Conjunction*}
 | |
| 450 | ||
| 451 | lemma conjI: "[| P; Q |] ==> P&Q" | |
| 452 | apply (unfold and_def) | |
| 453 | apply (rules intro: impI [THEN allI] mp) | |
| 454 | done | |
| 455 | ||
| 456 | lemma conjunct1: "[| P & Q |] ==> P" | |
| 457 | apply (unfold and_def) | |
| 458 | apply (rules intro: impI dest: spec mp) | |
| 459 | done | |
| 460 | ||
| 461 | lemma conjunct2: "[| P & Q |] ==> Q" | |
| 462 | apply (unfold and_def) | |
| 463 | apply (rules intro: impI dest: spec mp) | |
| 464 | done | |
| 465 | ||
| 466 | lemma conjE: | |
| 467 | assumes major: "P&Q" | |
| 468 | and minor: "[| P; Q |] ==> R" | |
| 469 | shows "R" | |
| 470 | apply (rule minor) | |
| 471 | apply (rule major [THEN conjunct1]) | |
| 472 | apply (rule major [THEN conjunct2]) | |
| 473 | done | |
| 474 | ||
| 475 | lemma context_conjI: | |
| 476 | assumes prems: "P" "P ==> Q" shows "P & Q" | |
| 477 | by (rules intro: conjI prems) | |
| 478 | ||
| 479 | ||
| 480 | subsection {*Disjunction*}
 | |
| 481 | ||
| 482 | lemma disjI1: "P ==> P|Q" | |
| 483 | apply (unfold or_def) | |
| 484 | apply (rules intro: allI impI mp) | |
| 485 | done | |
| 486 | ||
| 487 | lemma disjI2: "Q ==> P|Q" | |
| 488 | apply (unfold or_def) | |
| 489 | apply (rules intro: allI impI mp) | |
| 490 | done | |
| 491 | ||
| 492 | lemma disjE: | |
| 493 | assumes major: "P|Q" | |
| 494 | and minorP: "P ==> R" | |
| 495 | and minorQ: "Q ==> R" | |
| 496 | shows "R" | |
| 497 | by (rules intro: minorP minorQ impI | |
| 498 | major [unfolded or_def, THEN spec, THEN mp, THEN mp]) | |
| 499 | ||
| 500 | ||
| 501 | subsection {*Classical logic*}
 | |
| 502 | ||
| 503 | ||
| 504 | lemma classical: | |
| 505 | assumes prem: "~P ==> P" | |
| 506 | shows "P" | |
| 507 | apply (rule True_or_False [THEN disjE, THEN eqTrueE]) | |
| 508 | apply assumption | |
| 509 | apply (rule notI [THEN prem, THEN eqTrueI]) | |
| 510 | apply (erule subst) | |
| 511 | apply assumption | |
| 512 | done | |
| 513 | ||
| 514 | lemmas ccontr = FalseE [THEN classical, standard] | |
| 515 | ||
| 516 | (*notE with premises exchanged; it discharges ~R so that it can be used to | |
| 517 | make elimination rules*) | |
| 518 | lemma rev_notE: | |
| 519 | assumes premp: "P" | |
| 520 | and premnot: "~R ==> ~P" | |
| 521 | shows "R" | |
| 522 | apply (rule ccontr) | |
| 523 | apply (erule notE [OF premnot premp]) | |
| 524 | done | |
| 525 | ||
| 526 | (*Double negation law*) | |
| 527 | lemma notnotD: "~~P ==> P" | |
| 528 | apply (rule classical) | |
| 529 | apply (erule notE) | |
| 530 | apply assumption | |
| 531 | done | |
| 532 | ||
| 533 | lemma contrapos_pp: | |
| 534 | assumes p1: "Q" | |
| 535 | and p2: "~P ==> ~Q" | |
| 536 | shows "P" | |
| 537 | by (rules intro: classical p1 p2 notE) | |
| 538 | ||
| 539 | ||
| 540 | subsection {*Unique existence*}
 | |
| 541 | ||
| 542 | lemma ex1I: | |
| 543 | assumes prems: "P a" "!!x. P(x) ==> x=a" | |
| 544 | shows "EX! x. P(x)" | |
| 545 | by (unfold Ex1_def, rules intro: prems exI conjI allI impI) | |
| 546 | ||
| 547 | text{*Sometimes easier to use: the premises have no shared variables.  Safe!*}
 | |
| 548 | lemma ex_ex1I: | |
| 549 | assumes ex_prem: "EX x. P(x)" | |
| 550 | and eq: "!!x y. [| P(x); P(y) |] ==> x=y" | |
| 551 | shows "EX! x. P(x)" | |
| 552 | by (rules intro: ex_prem [THEN exE] ex1I eq) | |
| 553 | ||
| 554 | lemma ex1E: | |
| 555 | assumes major: "EX! x. P(x)" | |
| 556 | and minor: "!!x. [| P(x); ALL y. P(y) --> y=x |] ==> R" | |
| 557 | shows "R" | |
| 558 | apply (rule major [unfolded Ex1_def, THEN exE]) | |
| 559 | apply (erule conjE) | |
| 560 | apply (rules intro: minor) | |
| 561 | done | |
| 562 | ||
| 563 | lemma ex1_implies_ex: "EX! x. P x ==> EX x. P x" | |
| 564 | apply (erule ex1E) | |
| 565 | apply (rule exI) | |
| 566 | apply assumption | |
| 567 | done | |
| 568 | ||
| 569 | ||
| 570 | subsection {*THE: definite description operator*}
 | |
| 571 | ||
| 572 | lemma the_equality: | |
| 573 | assumes prema: "P a" | |
| 574 | and premx: "!!x. P x ==> x=a" | |
| 575 | shows "(THE x. P x) = a" | |
| 576 | apply (rule trans [OF _ the_eq_trivial]) | |
| 577 | apply (rule_tac f = "The" in arg_cong) | |
| 578 | apply (rule ext) | |
| 579 | apply (rule iffI) | |
| 580 | apply (erule premx) | |
| 581 | apply (erule ssubst, rule prema) | |
| 582 | done | |
| 583 | ||
| 584 | lemma theI: | |
| 585 | assumes "P a" and "!!x. P x ==> x=a" | |
| 586 | shows "P (THE x. P x)" | |
| 587 | by (rules intro: prems the_equality [THEN ssubst]) | |
| 588 | ||
| 589 | lemma theI': "EX! x. P x ==> P (THE x. P x)" | |
| 590 | apply (erule ex1E) | |
| 591 | apply (erule theI) | |
| 592 | apply (erule allE) | |
| 593 | apply (erule mp) | |
| 594 | apply assumption | |
| 595 | done | |
| 596 | ||
| 597 | (*Easier to apply than theI: only one occurrence of P*) | |
| 598 | lemma theI2: | |
| 599 | assumes "P a" "!!x. P x ==> x=a" "!!x. P x ==> Q x" | |
| 600 | shows "Q (THE x. P x)" | |
| 601 | by (rules intro: prems theI) | |
| 602 | ||
| 603 | lemma the1_equality: "[| EX!x. P x; P a |] ==> (THE x. P x) = a" | |
| 604 | apply (rule the_equality) | |
| 605 | apply assumption | |
| 606 | apply (erule ex1E) | |
| 607 | apply (erule all_dupE) | |
| 608 | apply (drule mp) | |
| 609 | apply assumption | |
| 610 | apply (erule ssubst) | |
| 611 | apply (erule allE) | |
| 612 | apply (erule mp) | |
| 613 | apply assumption | |
| 614 | done | |
| 615 | ||
| 616 | lemma the_sym_eq_trivial: "(THE y. x=y) = x" | |
| 617 | apply (rule the_equality) | |
| 618 | apply (rule refl) | |
| 619 | apply (erule sym) | |
| 620 | done | |
| 621 | ||
| 622 | ||
| 623 | subsection {*Classical intro rules for disjunction and existential quantifiers*}
 | |
| 624 | ||
| 625 | lemma disjCI: | |
| 626 | assumes "~Q ==> P" shows "P|Q" | |
| 627 | apply (rule classical) | |
| 628 | apply (rules intro: prems disjI1 disjI2 notI elim: notE) | |
| 629 | done | |
| 630 | ||
| 631 | lemma excluded_middle: "~P | P" | |
| 632 | by (rules intro: disjCI) | |
| 633 | ||
| 634 | text{*case distinction as a natural deduction rule. Note that @{term "~P"}
 | |
| 635 | is the second case, not the first.*} | |
| 636 | lemma case_split_thm: | |
| 637 | assumes prem1: "P ==> Q" | |
| 638 | and prem2: "~P ==> Q" | |
| 639 | shows "Q" | |
| 640 | apply (rule excluded_middle [THEN disjE]) | |
| 641 | apply (erule prem2) | |
| 642 | apply (erule prem1) | |
| 643 | done | |
| 644 | ||
| 645 | (*Classical implies (-->) elimination. *) | |
| 646 | lemma impCE: | |
| 647 | assumes major: "P-->Q" | |
| 648 | and minor: "~P ==> R" "Q ==> R" | |
| 649 | shows "R" | |
| 650 | apply (rule excluded_middle [of P, THEN disjE]) | |
| 651 | apply (rules intro: minor major [THEN mp])+ | |
| 652 | done | |
| 653 | ||
| 654 | (*This version of --> elimination works on Q before P. It works best for | |
| 655 | those cases in which P holds "almost everywhere". Can't install as | |
| 656 | default: would break old proofs.*) | |
| 657 | lemma impCE': | |
| 658 | assumes major: "P-->Q" | |
| 659 | and minor: "Q ==> R" "~P ==> R" | |
| 660 | shows "R" | |
| 661 | apply (rule excluded_middle [of P, THEN disjE]) | |
| 662 | apply (rules intro: minor major [THEN mp])+ | |
| 663 | done | |
| 664 | ||
| 665 | (*Classical <-> elimination. *) | |
| 666 | lemma iffCE: | |
| 667 | assumes major: "P=Q" | |
| 668 | and minor: "[| P; Q |] ==> R" "[| ~P; ~Q |] ==> R" | |
| 669 | shows "R" | |
| 670 | apply (rule major [THEN iffE]) | |
| 671 | apply (rules intro: minor elim: impCE notE) | |
| 672 | done | |
| 673 | ||
| 674 | lemma exCI: | |
| 675 | assumes "ALL x. ~P(x) ==> P(a)" | |
| 676 | shows "EX x. P(x)" | |
| 677 | apply (rule ccontr) | |
| 678 | apply (rules intro: prems exI allI notI notE [of "\<exists>x. P x"]) | |
| 679 | done | |
| 680 | ||
| 681 | ||
| 682 | ||
| 11750 | 683 | subsection {* Theory and package setup *}
 | 
| 684 | ||
| 15411 | 685 | ML | 
| 686 | {*
 | |
| 687 | val plusI = thm "plusI" | |
| 688 | val minusI = thm "minusI" | |
| 689 | val timesI = thm "timesI" | |
| 690 | val eq_reflection = thm "eq_reflection" | |
| 691 | val refl = thm "refl" | |
| 692 | val subst = thm "subst" | |
| 693 | val ext = thm "ext" | |
| 694 | val impI = thm "impI" | |
| 695 | val mp = thm "mp" | |
| 696 | val True_def = thm "True_def" | |
| 697 | val All_def = thm "All_def" | |
| 698 | val Ex_def = thm "Ex_def" | |
| 699 | val False_def = thm "False_def" | |
| 700 | val not_def = thm "not_def" | |
| 701 | val and_def = thm "and_def" | |
| 702 | val or_def = thm "or_def" | |
| 703 | val Ex1_def = thm "Ex1_def" | |
| 704 | val iff = thm "iff" | |
| 705 | val True_or_False = thm "True_or_False" | |
| 706 | val Let_def = thm "Let_def" | |
| 707 | val if_def = thm "if_def" | |
| 708 | val sym = thm "sym" | |
| 709 | val ssubst = thm "ssubst" | |
| 710 | val trans = thm "trans" | |
| 711 | val def_imp_eq = thm "def_imp_eq" | |
| 712 | val box_equals = thm "box_equals" | |
| 713 | val fun_cong = thm "fun_cong" | |
| 714 | val arg_cong = thm "arg_cong" | |
| 715 | val cong = thm "cong" | |
| 716 | val iffI = thm "iffI" | |
| 717 | val iffD2 = thm "iffD2" | |
| 718 | val rev_iffD2 = thm "rev_iffD2" | |
| 719 | val iffD1 = thm "iffD1" | |
| 720 | val rev_iffD1 = thm "rev_iffD1" | |
| 721 | val iffE = thm "iffE" | |
| 722 | val TrueI = thm "TrueI" | |
| 723 | val eqTrueI = thm "eqTrueI" | |
| 724 | val eqTrueE = thm "eqTrueE" | |
| 725 | val allI = thm "allI" | |
| 726 | val spec = thm "spec" | |
| 727 | val allE = thm "allE" | |
| 728 | val all_dupE = thm "all_dupE" | |
| 729 | val FalseE = thm "FalseE" | |
| 730 | val False_neq_True = thm "False_neq_True" | |
| 731 | val notI = thm "notI" | |
| 732 | val False_not_True = thm "False_not_True" | |
| 733 | val True_not_False = thm "True_not_False" | |
| 734 | val notE = thm "notE" | |
| 735 | val notI2 = thm "notI2" | |
| 736 | val impE = thm "impE" | |
| 737 | val rev_mp = thm "rev_mp" | |
| 738 | val contrapos_nn = thm "contrapos_nn" | |
| 739 | val contrapos_pn = thm "contrapos_pn" | |
| 740 | val not_sym = thm "not_sym" | |
| 741 | val rev_contrapos = thm "rev_contrapos" | |
| 742 | val exI = thm "exI" | |
| 743 | val exE = thm "exE" | |
| 744 | val conjI = thm "conjI" | |
| 745 | val conjunct1 = thm "conjunct1" | |
| 746 | val conjunct2 = thm "conjunct2" | |
| 747 | val conjE = thm "conjE" | |
| 748 | val context_conjI = thm "context_conjI" | |
| 749 | val disjI1 = thm "disjI1" | |
| 750 | val disjI2 = thm "disjI2" | |
| 751 | val disjE = thm "disjE" | |
| 752 | val classical = thm "classical" | |
| 753 | val ccontr = thm "ccontr" | |
| 754 | val rev_notE = thm "rev_notE" | |
| 755 | val notnotD = thm "notnotD" | |
| 756 | val contrapos_pp = thm "contrapos_pp" | |
| 757 | val ex1I = thm "ex1I" | |
| 758 | val ex_ex1I = thm "ex_ex1I" | |
| 759 | val ex1E = thm "ex1E" | |
| 760 | val ex1_implies_ex = thm "ex1_implies_ex" | |
| 761 | val the_equality = thm "the_equality" | |
| 762 | val theI = thm "theI" | |
| 763 | val theI' = thm "theI'" | |
| 764 | val theI2 = thm "theI2" | |
| 765 | val the1_equality = thm "the1_equality" | |
| 766 | val the_sym_eq_trivial = thm "the_sym_eq_trivial" | |
| 767 | val disjCI = thm "disjCI" | |
| 768 | val excluded_middle = thm "excluded_middle" | |
| 769 | val case_split_thm = thm "case_split_thm" | |
| 770 | val impCE = thm "impCE" | |
| 771 | val impCE = thm "impCE" | |
| 772 | val iffCE = thm "iffCE" | |
| 773 | val exCI = thm "exCI" | |
| 4868 | 774 | |
| 15411 | 775 | (* combination of (spec RS spec RS ...(j times) ... spec RS mp) *) | 
| 776 | local | |
| 777 |   fun wrong_prem (Const ("All", _) $ (Abs (_, _, t))) = wrong_prem t
 | |
| 778 | | wrong_prem (Bound _) = true | |
| 779 | | wrong_prem _ = false | |
| 15570 | 780 | val filter_right = List.filter (fn t => not (wrong_prem (HOLogic.dest_Trueprop (hd (Thm.prems_of t))))) | 
| 15411 | 781 | in | 
| 782 | fun smp i = funpow i (fn m => filter_right ([spec] RL m)) ([mp]) | |
| 783 | fun smp_tac j = EVERY'[dresolve_tac (smp j), atac] | |
| 784 | end | |
| 785 | ||
| 786 | ||
| 787 | fun strip_tac i = REPEAT(resolve_tac [impI,allI] i) | |
| 788 | ||
| 789 | (*Obsolete form of disjunctive case analysis*) | |
| 790 | fun excluded_middle_tac sP = | |
| 791 |     res_inst_tac [("Q",sP)] (excluded_middle RS disjE)
 | |
| 792 | ||
| 793 | fun case_tac a = res_inst_tac [("P",a)] case_split_thm
 | |
| 794 | *} | |
| 795 | ||
| 11687 | 796 | theorems case_split = case_split_thm [case_names True False] | 
| 9869 | 797 | |
| 12386 | 798 | |
| 799 | subsubsection {* Intuitionistic Reasoning *}
 | |
| 800 | ||
| 801 | lemma impE': | |
| 12937 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 802 | assumes 1: "P --> Q" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 803 | and 2: "Q ==> R" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 804 | and 3: "P --> Q ==> P" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 805 | shows R | 
| 12386 | 806 | proof - | 
| 807 | from 3 and 1 have P . | |
| 808 | with 1 have Q by (rule impE) | |
| 809 | with 2 show R . | |
| 810 | qed | |
| 811 | ||
| 812 | lemma allE': | |
| 12937 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 813 | assumes 1: "ALL x. P x" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 814 | and 2: "P x ==> ALL x. P x ==> Q" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 815 | shows Q | 
| 12386 | 816 | proof - | 
| 817 | from 1 have "P x" by (rule spec) | |
| 818 | from this and 1 show Q by (rule 2) | |
| 819 | qed | |
| 820 | ||
| 12937 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 821 | lemma notE': | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 822 | assumes 1: "~ P" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 823 | and 2: "~ P ==> P" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 824 | shows R | 
| 12386 | 825 | proof - | 
| 826 | from 2 and 1 have P . | |
| 827 | with 1 show R by (rule notE) | |
| 828 | qed | |
| 829 | ||
| 15801 | 830 | lemmas [Pure.elim!] = disjE iffE FalseE conjE exE | 
| 831 | and [Pure.intro!] = iffI conjI impI TrueI notI allI refl | |
| 832 | and [Pure.elim 2] = allE notE' impE' | |
| 833 | and [Pure.intro] = exI disjI2 disjI1 | |
| 12386 | 834 | |
| 835 | lemmas [trans] = trans | |
| 836 | and [sym] = sym not_sym | |
| 15801 | 837 | and [Pure.elim?] = iffD1 iffD2 impE | 
| 11750 | 838 | |
| 11438 
3d9222b80989
declare trans [trans]  (*overridden in theory Calculation*);
 wenzelm parents: 
11432diff
changeset | 839 | |
| 11750 | 840 | subsubsection {* Atomizing meta-level connectives *}
 | 
| 841 | ||
| 842 | lemma atomize_all [atomize]: "(!!x. P x) == Trueprop (ALL x. P x)" | |
| 12003 | 843 | proof | 
| 9488 | 844 | assume "!!x. P x" | 
| 10383 | 845 | show "ALL x. P x" by (rule allI) | 
| 9488 | 846 | next | 
| 847 | assume "ALL x. P x" | |
| 10383 | 848 | thus "!!x. P x" by (rule allE) | 
| 9488 | 849 | qed | 
| 850 | ||
| 11750 | 851 | lemma atomize_imp [atomize]: "(A ==> B) == Trueprop (A --> B)" | 
| 12003 | 852 | proof | 
| 9488 | 853 | assume r: "A ==> B" | 
| 10383 | 854 | show "A --> B" by (rule impI) (rule r) | 
| 9488 | 855 | next | 
| 856 | assume "A --> B" and A | |
| 10383 | 857 | thus B by (rule mp) | 
| 9488 | 858 | qed | 
| 859 | ||
| 14749 | 860 | lemma atomize_not: "(A ==> False) == Trueprop (~A)" | 
| 861 | proof | |
| 862 | assume r: "A ==> False" | |
| 863 | show "~A" by (rule notI) (rule r) | |
| 864 | next | |
| 865 | assume "~A" and A | |
| 866 | thus False by (rule notE) | |
| 867 | qed | |
| 868 | ||
| 11750 | 869 | lemma atomize_eq [atomize]: "(x == y) == Trueprop (x = y)" | 
| 12003 | 870 | proof | 
| 10432 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 871 | assume "x == y" | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 872 | show "x = y" by (unfold prems) (rule refl) | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 873 | next | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 874 | assume "x = y" | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 875 | thus "x == y" by (rule eq_reflection) | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 876 | qed | 
| 
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
 wenzelm parents: 
10383diff
changeset | 877 | |
| 12023 | 878 | lemma atomize_conj [atomize]: | 
| 879 | "(!!C. (A ==> B ==> PROP C) ==> PROP C) == Trueprop (A & B)" | |
| 12003 | 880 | proof | 
| 11953 | 881 | assume "!!C. (A ==> B ==> PROP C) ==> PROP C" | 
| 882 | show "A & B" by (rule conjI) | |
| 883 | next | |
| 884 | fix C | |
| 885 | assume "A & B" | |
| 886 | assume "A ==> B ==> PROP C" | |
| 887 | thus "PROP C" | |
| 888 | proof this | |
| 889 | show A by (rule conjunct1) | |
| 890 | show B by (rule conjunct2) | |
| 891 | qed | |
| 892 | qed | |
| 893 | ||
| 12386 | 894 | lemmas [symmetric, rulify] = atomize_all atomize_imp | 
| 895 | ||
| 11750 | 896 | |
| 897 | subsubsection {* Classical Reasoner setup *}
 | |
| 9529 | 898 | |
| 10383 | 899 | use "cladata.ML" | 
| 900 | setup hypsubst_setup | |
| 11977 | 901 | |
| 16121 | 902 | setup {*
 | 
| 903 | [ContextRules.addSWrapper (fn tac => hyp_subst_tac' ORELSE' tac)] | |
| 12386 | 904 | *} | 
| 11977 | 905 | |
| 10383 | 906 | setup Classical.setup | 
| 907 | setup clasetup | |
| 908 | ||
| 12386 | 909 | lemmas [intro?] = ext | 
| 910 | and [elim?] = ex1_implies_ex | |
| 11977 | 911 | |
| 9869 | 912 | use "blastdata.ML" | 
| 913 | setup Blast.setup | |
| 4868 | 914 | |
| 11750 | 915 | |
| 15481 | 916 | subsection {* Simplifier setup *}
 | 
| 11750 | 917 | |
| 12281 | 918 | lemma meta_eq_to_obj_eq: "x == y ==> x = y" | 
| 919 | proof - | |
| 920 | assume r: "x == y" | |
| 921 | show "x = y" by (unfold r) (rule refl) | |
| 922 | qed | |
| 923 | ||
| 924 | lemma eta_contract_eq: "(%s. f s) = f" .. | |
| 925 | ||
| 926 | lemma simp_thms: | |
| 12937 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 927 | shows not_not: "(~ ~ P) = P" | 
| 15354 | 928 | and Not_eq_iff: "((~P) = (~Q)) = (P = Q)" | 
| 12937 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 929 | and | 
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 930 | "(P ~= Q) = (P = (~Q))" | 
| 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 931 | "(P | ~P) = True" "(~P | P) = True" | 
| 12281 | 932 | "(x = x) = True" | 
| 933 | "(~True) = False" "(~False) = True" | |
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 934 | "(~P) ~= P" "P ~= (~P)" | 
| 12281 | 935 | "(True=P) = P" "(P=True) = P" "(False=P) = (~P)" "(P=False) = (~P)" | 
| 936 | "(True --> P) = P" "(False --> P) = True" | |
| 937 | "(P --> True) = True" "(P --> P) = True" | |
| 938 | "(P --> False) = (~P)" "(P --> ~P) = (~P)" | |
| 939 | "(P & True) = P" "(True & P) = P" | |
| 940 | "(P & False) = False" "(False & P) = False" | |
| 941 | "(P & P) = P" "(P & (P & Q)) = (P & Q)" | |
| 942 | "(P & ~P) = False" "(~P & P) = False" | |
| 943 | "(P | True) = True" "(True | P) = True" | |
| 944 | "(P | False) = P" "(False | P) = P" | |
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 945 | "(P | P) = P" "(P | (P | Q)) = (P | Q)" and | 
| 12281 | 946 | "(ALL x. P) = P" "(EX x. P) = P" "EX x. x=t" "EX x. t=x" | 
| 947 |     -- {* needed for the one-point-rule quantifier simplification procs *}
 | |
| 948 |     -- {* essential for termination!! *} and
 | |
| 949 | "!!P. (EX x. x=t & P(x)) = P(t)" | |
| 950 | "!!P. (EX x. t=x & P(x)) = P(t)" | |
| 951 | "!!P. (ALL x. x=t --> P(x)) = P(t)" | |
| 12937 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 952 | "!!P. (ALL x. t=x --> P(x)) = P(t)" | 
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 953 | by (blast, blast, blast, blast, blast, rules+) | 
| 13421 | 954 | |
| 12281 | 955 | lemma imp_cong: "(P = P') ==> (P' ==> (Q = Q')) ==> ((P --> Q) = (P' --> Q'))" | 
| 12354 | 956 | by rules | 
| 12281 | 957 | |
| 958 | lemma ex_simps: | |
| 959 | "!!P Q. (EX x. P x & Q) = ((EX x. P x) & Q)" | |
| 960 | "!!P Q. (EX x. P & Q x) = (P & (EX x. Q x))" | |
| 961 | "!!P Q. (EX x. P x | Q) = ((EX x. P x) | Q)" | |
| 962 | "!!P Q. (EX x. P | Q x) = (P | (EX x. Q x))" | |
| 963 | "!!P Q. (EX x. P x --> Q) = ((ALL x. P x) --> Q)" | |
| 964 | "!!P Q. (EX x. P --> Q x) = (P --> (EX x. Q x))" | |
| 965 |   -- {* Miniscoping: pushing in existential quantifiers. *}
 | |
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 966 | by (rules | blast)+ | 
| 12281 | 967 | |
| 968 | lemma all_simps: | |
| 969 | "!!P Q. (ALL x. P x & Q) = ((ALL x. P x) & Q)" | |
| 970 | "!!P Q. (ALL x. P & Q x) = (P & (ALL x. Q x))" | |
| 971 | "!!P Q. (ALL x. P x | Q) = ((ALL x. P x) | Q)" | |
| 972 | "!!P Q. (ALL x. P | Q x) = (P | (ALL x. Q x))" | |
| 973 | "!!P Q. (ALL x. P x --> Q) = ((EX x. P x) --> Q)" | |
| 974 | "!!P Q. (ALL x. P --> Q x) = (P --> (ALL x. Q x))" | |
| 975 |   -- {* Miniscoping: pushing in universal quantifiers. *}
 | |
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 976 | by (rules | blast)+ | 
| 12281 | 977 | |
| 14201 | 978 | lemma disj_absorb: "(A | A) = A" | 
| 979 | by blast | |
| 980 | ||
| 981 | lemma disj_left_absorb: "(A | (A | B)) = (A | B)" | |
| 982 | by blast | |
| 983 | ||
| 984 | lemma conj_absorb: "(A & A) = A" | |
| 985 | by blast | |
| 986 | ||
| 987 | lemma conj_left_absorb: "(A & (A & B)) = (A & B)" | |
| 988 | by blast | |
| 989 | ||
| 12281 | 990 | lemma eq_ac: | 
| 12937 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 991 | shows eq_commute: "(a=b) = (b=a)" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 992 | and eq_left_commute: "(P=(Q=R)) = (Q=(P=R))" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 993 | 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 | 994 | lemma neq_commute: "(a~=b) = (b~=a)" by rules | 
| 12281 | 995 | |
| 996 | lemma conj_comms: | |
| 12937 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 997 | shows conj_commute: "(P&Q) = (Q&P)" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 998 | 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 | 999 | lemma conj_assoc: "((P&Q)&R) = (P&(Q&R))" by rules | 
| 12281 | 1000 | |
| 1001 | lemma disj_comms: | |
| 12937 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 1002 | shows disj_commute: "(P|Q) = (Q|P)" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12892diff
changeset | 1003 | 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 | 1004 | lemma disj_assoc: "((P|Q)|R) = (P|(Q|R))" by rules | 
| 12281 | 1005 | |
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 1006 | 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 | 1007 | lemma conj_disj_distribR: "((P|Q)&R) = (P&R | Q&R)" by rules | 
| 12281 | 1008 | |
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 1009 | 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 | 1010 | lemma disj_conj_distribR: "((P&Q)|R) = ((P|R) & (Q|R))" by rules | 
| 12281 | 1011 | |
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 1012 | lemma imp_conjR: "(P --> (Q&R)) = ((P-->Q) & (P-->R))" by rules | 
| 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 1013 | lemma imp_conjL: "((P&Q) -->R) = (P --> (Q --> R))" by rules | 
| 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 1014 | lemma imp_disjL: "((P|Q) --> R) = ((P-->R)&(Q-->R))" by rules | 
| 12281 | 1015 | |
| 1016 | text {* These two are specialized, but @{text imp_disj_not1} is useful in @{text "Auth/Yahalom"}. *}
 | |
| 1017 | lemma imp_disj_not1: "(P --> Q | R) = (~Q --> P --> R)" by blast | |
| 1018 | lemma imp_disj_not2: "(P --> Q | R) = (~R --> P --> Q)" by blast | |
| 1019 | ||
| 1020 | lemma imp_disj1: "((P-->Q)|R) = (P--> Q|R)" by blast | |
| 1021 | lemma imp_disj2: "(Q|(P-->R)) = (P--> Q|R)" by blast | |
| 1022 | ||
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 1023 | lemma de_Morgan_disj: "(~(P | Q)) = (~P & ~Q)" by rules | 
| 12281 | 1024 | lemma de_Morgan_conj: "(~(P & Q)) = (~P | ~Q)" by blast | 
| 1025 | lemma not_imp: "(~(P --> Q)) = (P & ~Q)" by blast | |
| 1026 | lemma not_iff: "(P~=Q) = (P = (~Q))" by blast | |
| 1027 | lemma disj_not1: "(~P | Q) = (P --> Q)" by blast | |
| 1028 | lemma disj_not2: "(P | ~Q) = (Q --> P)"  -- {* changes orientation :-( *}
 | |
| 1029 | by blast | |
| 1030 | lemma imp_conv_disj: "(P --> Q) = ((~P) | Q)" by blast | |
| 1031 | ||
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 1032 | lemma iff_conv_conj_imp: "(P = Q) = ((P --> Q) & (Q --> P))" by rules | 
| 12281 | 1033 | |
| 1034 | ||
| 1035 | lemma cases_simp: "((P --> Q) & (~P --> Q)) = Q" | |
| 1036 |   -- {* Avoids duplication of subgoals after @{text split_if}, when the true and false *}
 | |
| 1037 |   -- {* cases boil down to the same thing. *}
 | |
| 1038 | by blast | |
| 1039 | ||
| 1040 | lemma not_all: "(~ (! x. P(x))) = (? x.~P(x))" by blast | |
| 1041 | 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 | 1042 | lemma not_ex: "(~ (? x. P(x))) = (! x.~P(x))" by rules | 
| 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 1043 | lemma imp_ex: "((? x. P x) --> Q) = (! x. P x --> Q)" by rules | 
| 12281 | 1044 | |
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 1045 | 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 | 1046 | lemma all_conj_distrib: "(!x. P(x) & Q(x)) = ((! x. P(x)) & (! x. Q(x)))" by rules | 
| 12281 | 1047 | |
| 1048 | text {*
 | |
| 1049 |   \medskip The @{text "&"} congruence rule: not included by default!
 | |
| 1050 | May slow rewrite proofs down by as much as 50\% *} | |
| 1051 | ||
| 1052 | lemma conj_cong: | |
| 1053 | "(P = P') ==> (P' ==> (Q = Q')) ==> ((P & Q) = (P' & Q'))" | |
| 12354 | 1054 | by rules | 
| 12281 | 1055 | |
| 1056 | lemma rev_conj_cong: | |
| 1057 | "(Q = Q') ==> (Q' ==> (P = P')) ==> ((P & Q) = (P' & Q'))" | |
| 12354 | 1058 | by rules | 
| 12281 | 1059 | |
| 1060 | text {* The @{text "|"} congruence rule: not included by default! *}
 | |
| 1061 | ||
| 1062 | lemma disj_cong: | |
| 1063 | "(P = P') ==> (~P' ==> (Q = Q')) ==> ((P | Q) = (P' | Q'))" | |
| 1064 | by blast | |
| 1065 | ||
| 1066 | lemma eq_sym_conv: "(x = y) = (y = x)" | |
| 12354 | 1067 | by rules | 
| 12281 | 1068 | |
| 1069 | ||
| 1070 | text {* \medskip if-then-else rules *}
 | |
| 1071 | ||
| 1072 | lemma if_True: "(if True then x else y) = x" | |
| 1073 | by (unfold if_def) blast | |
| 1074 | ||
| 1075 | lemma if_False: "(if False then x else y) = y" | |
| 1076 | by (unfold if_def) blast | |
| 1077 | ||
| 1078 | lemma if_P: "P ==> (if P then x else y) = x" | |
| 1079 | by (unfold if_def) blast | |
| 1080 | ||
| 1081 | lemma if_not_P: "~P ==> (if P then x else y) = y" | |
| 1082 | by (unfold if_def) blast | |
| 1083 | ||
| 1084 | lemma split_if: "P (if Q then x else y) = ((Q --> P(x)) & (~Q --> P(y)))" | |
| 1085 | apply (rule case_split [of Q]) | |
| 15481 | 1086 | apply (simplesubst if_P) | 
| 1087 | prefer 3 apply (simplesubst if_not_P, blast+) | |
| 12281 | 1088 | done | 
| 1089 | ||
| 1090 | lemma split_if_asm: "P (if Q then x else y) = (~((Q & ~P x) | (~Q & ~P y)))" | |
| 15481 | 1091 | by (simplesubst split_if, blast) | 
| 12281 | 1092 | |
| 1093 | lemmas if_splits = split_if split_if_asm | |
| 1094 | ||
| 1095 | lemma if_def2: "(if Q then x else y) = ((Q --> x) & (~ Q --> y))" | |
| 1096 | by (rule split_if) | |
| 1097 | ||
| 1098 | lemma if_cancel: "(if c then x else x) = x" | |
| 15481 | 1099 | by (simplesubst split_if, blast) | 
| 12281 | 1100 | |
| 1101 | lemma if_eq_cancel: "(if x = y then y else x) = x" | |
| 15481 | 1102 | by (simplesubst split_if, blast) | 
| 12281 | 1103 | |
| 1104 | lemma if_bool_eq_conj: "(if P then Q else R) = ((P-->Q) & (~P-->R))" | |
| 1105 |   -- {* This form is useful for expanding @{text if}s on the RIGHT of the @{text "==>"} symbol. *}
 | |
| 1106 | by (rule split_if) | |
| 1107 | ||
| 1108 | lemma if_bool_eq_disj: "(if P then Q else R) = ((P&Q) | (~P&R))" | |
| 1109 |   -- {* And this form is useful for expanding @{text if}s on the LEFT. *}
 | |
| 15481 | 1110 | apply (simplesubst split_if, blast) | 
| 12281 | 1111 | done | 
| 1112 | ||
| 12436 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 1113 | lemma Eq_TrueI: "P ==> P == True" by (unfold atomize_eq) rules | 
| 
a2df07fefed7
Replaced several occurrences of "blast" by "rules".
 berghofe parents: 
12386diff
changeset | 1114 | lemma Eq_FalseI: "~P ==> P == False" by (unfold atomize_eq) rules | 
| 12281 | 1115 | |
| 15423 | 1116 | text {* \medskip let rules for simproc *}
 | 
| 1117 | ||
| 1118 | lemma Let_folded: "f x \<equiv> g x \<Longrightarrow> Let x f \<equiv> Let x g" | |
| 1119 | by (unfold Let_def) | |
| 1120 | ||
| 1121 | lemma Let_unfold: "f x \<equiv> g \<Longrightarrow> Let x f \<equiv> g" | |
| 1122 | by (unfold Let_def) | |
| 1123 | ||
| 14201 | 1124 | subsubsection {* Actual Installation of the Simplifier *}
 | 
| 1125 | ||
| 9869 | 1126 | use "simpdata.ML" | 
| 1127 | setup "Simplifier.method_setup Splitter.split_modifiers" setup simpsetup | |
| 1128 | setup Splitter.setup setup Clasimp.setup | |
| 1129 | ||
| 15481 | 1130 | |
| 1131 | subsubsection {* Lucas Dixon's eqstep tactic *}
 | |
| 1132 | ||
| 1133 | use "~~/src/Provers/eqsubst.ML"; | |
| 1134 | use "eqrule_HOL_data.ML"; | |
| 1135 | ||
| 1136 | setup EQSubstTac.setup | |
| 1137 | ||
| 1138 | ||
| 1139 | subsection {* Other simple lemmas *}
 | |
| 1140 | ||
| 15411 | 1141 | declare disj_absorb [simp] conj_absorb [simp] | 
| 14201 | 1142 | |
| 13723 | 1143 | lemma ex1_eq[iff]: "EX! x. x = t" "EX! x. t = x" | 
| 1144 | by blast+ | |
| 1145 | ||
| 15481 | 1146 | |
| 13638 | 1147 | theorem choice_eq: "(ALL x. EX! y. P x y) = (EX! f. ALL x. P x (f x))" | 
| 1148 | apply (rule iffI) | |
| 1149 | apply (rule_tac a = "%x. THE y. P x y" in ex1I) | |
| 1150 | apply (fast dest!: theI') | |
| 1151 | apply (fast intro: ext the1_equality [symmetric]) | |
| 1152 | apply (erule ex1E) | |
| 1153 | apply (rule allI) | |
| 1154 | apply (rule ex1I) | |
| 1155 | apply (erule spec) | |
| 1156 | apply (erule_tac x = "%z. if z = x then y else f z" in allE) | |
| 1157 | apply (erule impE) | |
| 1158 | apply (rule allI) | |
| 1159 | apply (rule_tac P = "xa = x" in case_split_thm) | |
| 14208 | 1160 | apply (drule_tac [3] x = x in fun_cong, simp_all) | 
| 13638 | 1161 | done | 
| 1162 | ||
| 13438 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 1163 | text{*Needs only HOL-lemmas:*}
 | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 1164 | lemma mk_left_commute: | 
| 
527811f00c56
added mk_left_commute to HOL.thy and used it "everywhere"
 nipkow parents: 
13421diff
changeset | 1165 | 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 | 1166 | 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 | 1167 | 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 | 1168 | 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 | 1169 | |
| 11750 | 1170 | |
| 15481 | 1171 | subsection {* Generic cases and induction *}
 | 
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 1172 | |
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 1173 | constdefs | 
| 11989 | 1174 |   induct_forall :: "('a => bool) => bool"
 | 
| 1175 | "induct_forall P == \<forall>x. P x" | |
| 1176 | induct_implies :: "bool => bool => bool" | |
| 1177 | "induct_implies A B == A --> B" | |
| 1178 | induct_equal :: "'a => 'a => bool" | |
| 1179 | "induct_equal x y == x = y" | |
| 1180 | induct_conj :: "bool => bool => bool" | |
| 1181 | "induct_conj A B == A & B" | |
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 1182 | |
| 11989 | 1183 | lemma induct_forall_eq: "(!!x. P x) == Trueprop (induct_forall (\<lambda>x. P x))" | 
| 1184 | by (simp only: atomize_all induct_forall_def) | |
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 1185 | |
| 11989 | 1186 | lemma induct_implies_eq: "(A ==> B) == Trueprop (induct_implies A B)" | 
| 1187 | by (simp only: atomize_imp induct_implies_def) | |
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 1188 | |
| 11989 | 1189 | lemma induct_equal_eq: "(x == y) == Trueprop (induct_equal x y)" | 
| 1190 | by (simp only: atomize_eq induct_equal_def) | |
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 1191 | |
| 11989 | 1192 | lemma induct_forall_conj: "induct_forall (\<lambda>x. induct_conj (A x) (B x)) = | 
| 1193 | induct_conj (induct_forall A) (induct_forall B)" | |
| 12354 | 1194 | by (unfold induct_forall_def induct_conj_def) rules | 
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 1195 | |
| 11989 | 1196 | lemma induct_implies_conj: "induct_implies C (induct_conj A B) = | 
| 1197 | induct_conj (induct_implies C A) (induct_implies C B)" | |
| 12354 | 1198 | by (unfold induct_implies_def induct_conj_def) rules | 
| 11989 | 1199 | |
| 13598 
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
 berghofe parents: 
13596diff
changeset | 1200 | 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 | 1201 | proof | 
| 
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
 berghofe parents: 
13596diff
changeset | 1202 | 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 | 1203 | 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 | 1204 | next | 
| 
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
 berghofe parents: 
13596diff
changeset | 1205 | 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 | 1206 | 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 | 1207 | qed | 
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 1208 | |
| 11989 | 1209 | lemma induct_impliesI: "(A ==> B) ==> induct_implies A B" | 
| 1210 | by (simp add: induct_implies_def) | |
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 1211 | |
| 12161 | 1212 | lemmas induct_atomize = atomize_conj induct_forall_eq induct_implies_eq induct_equal_eq | 
| 1213 | lemmas induct_rulify1 [symmetric, standard] = induct_forall_eq induct_implies_eq induct_equal_eq | |
| 1214 | lemmas induct_rulify2 = induct_forall_def induct_implies_def induct_equal_def induct_conj_def | |
| 11989 | 1215 | 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 | 1216 | |
| 11989 | 1217 | hide const induct_forall induct_implies induct_equal induct_conj | 
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 1218 | |
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 1219 | |
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 1220 | text {* Method setup. *}
 | 
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 1221 | |
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 1222 | ML {*
 | 
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 1223 | structure InductMethod = InductMethodFun | 
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 1224 | (struct | 
| 15411 | 1225 | val dest_concls = HOLogic.dest_concls | 
| 1226 | val cases_default = thm "case_split" | |
| 1227 | val local_impI = thm "induct_impliesI" | |
| 1228 | val conjI = thm "conjI" | |
| 1229 | val atomize = thms "induct_atomize" | |
| 1230 | val rulify1 = thms "induct_rulify1" | |
| 1231 | val rulify2 = thms "induct_rulify2" | |
| 1232 | val localize = [Thm.symmetric (thm "induct_implies_def")] | |
| 11824 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 1233 | end); | 
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 1234 | *} | 
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 1235 | |
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 1236 | setup InductMethod.setup | 
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 1237 | |
| 
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
 wenzelm parents: 
11770diff
changeset | 1238 | |
| 14357 | 1239 | end | 
| 15411 | 1240 |