author | clasohm |
Wed, 02 Mar 1994 12:26:55 +0100 | |
changeset 48 | 21291189b51e |
parent 43 | 424c7b1489df |
child 57 | 194d088c1511 |
permissions | -rw-r--r-- |
1 | 1 |
(* Title: HOL/simpdata.ML |
2 |
ID: $Id$ |
|
3 |
Author: Tobias Nipkow |
|
4 |
Copyright 1991 University of Cambridge |
|
5 |
||
6 |
Instantiation of the generic simplifier |
|
7 |
*) |
|
8 |
||
0 | 9 |
open Simplifier; |
10 |
||
11 |
local |
|
12 |
||
13 |
fun prover s = prove_goal HOL.thy s (fn _ => [fast_tac HOL_cs 1]); |
|
14 |
||
15 |
val P_imp_P_iff_True = prover "P --> (P = True)" RS mp; |
|
16 |
val P_imp_P_eq_True = P_imp_P_iff_True RS eq_reflection; |
|
17 |
||
18 |
val not_P_imp_P_iff_F = prover "~P --> (P = False)" RS mp; |
|
19 |
val not_P_imp_P_eq_False = not_P_imp_P_iff_F RS eq_reflection; |
|
20 |
||
21 |
fun atomize r = |
|
22 |
case concl_of r of |
|
23 |
Const("Trueprop",_) $ p => |
|
24 |
(case p of |
|
25 |
Const("op -->",_)$_$_ => atomize(r RS mp) |
|
26 |
| Const("op &",_)$_$_ => atomize(r RS conjunct1) @ |
|
27 |
atomize(r RS conjunct2) |
|
28 |
| Const("All",_)$_ => atomize(r RS spec) |
|
29 |
| Const("True",_) => [] |
|
30 |
| Const("False",_) => [] |
|
31 |
| _ => [r]) |
|
32 |
| _ => [r]; |
|
33 |
||
34 |
fun mk_eq r = case concl_of r of |
|
35 |
Const("==",_)$_$_ => r |
|
36 |
| _$(Const("op =",_)$_$_) => r RS eq_reflection |
|
37 |
| _$(Const("not",_)$_) => r RS not_P_imp_P_eq_False |
|
38 |
| _ => r RS P_imp_P_eq_True; |
|
39 |
(* last 2 lines requires all formulae to be of the from Trueprop(.) *) |
|
40 |
||
41 |
fun gen_all th = forall_elim_vars (#maxidx(rep_thm th)+1) th; |
|
42 |
||
43 |
fun mk_rews thm = map mk_eq (atomize(gen_all thm)); |
|
44 |
||
1 | 45 |
val imp_cong = impI RSN |
0 | 46 |
(2, prove_goal HOL.thy "(P=P')--> (P'--> (Q=Q'))--> ((P-->Q) = (P'-->Q'))" |
47 |
(fn _=> [fast_tac HOL_cs 1]) RS mp RS mp); |
|
48 |
||
49 |
val o_apply = prove_goal HOL.thy "(f o g)(x) = f(g(x))" |
|
50 |
(fn _ => [ (stac o_def 1), (rtac refl 1) ]); |
|
51 |
||
52 |
val simp_thms = map prover |
|
53 |
[ "(x=x) = True", |
|
54 |
"(~True) = False", "(~False) = True", "(~ ~ P) = P", |
|
40 | 55 |
"(~P) ~= P", "P ~= (~P)", "(P ~= Q) = (P = (~Q))", |
0 | 56 |
"(True=P) = P", "(P=True) = P", |
57 |
"(True --> P) = P", "(False --> P) = True", |
|
58 |
"(P --> True) = True", "(P --> P) = True", |
|
59 |
"(P & True) = P", "(True & P) = P", |
|
60 |
"(P & False) = False", "(False & P) = False", "(P & P) = P", |
|
61 |
"(P | True) = True", "(True | P) = True", |
|
62 |
"(P | False) = P", "(False | P) = P", "(P | P) = P", |
|
30
2fdeeae553ac
modified solver of HOL_ss to take the new simplifier into account: the thm to
nipkow
parents:
3
diff
changeset
|
63 |
"(!x.P) = P", "(? x.P) = P", |
0 | 64 |
"(P|Q --> R) = ((P-->R)&(Q-->R))" ]; |
65 |
||
66 |
val meta_obj_reflection = prove_goal HOL.thy "x==y ==> x=y" |
|
67 |
(fn [prem] => [rewtac prem, rtac refl 1]); |
|
68 |
||
69 |
in |
|
70 |
||
34 | 71 |
val conj_assoc = prover "((P&Q)&R) = (P&(Q&R))"; |
0 | 72 |
|
73 |
val if_True = prove_goal HOL.thy "if(True,x,y) = x" |
|
74 |
(fn _=>[stac if_def 1, fast_tac (HOL_cs addIs [select_equality]) 1]); |
|
75 |
||
76 |
val if_False = prove_goal HOL.thy "if(False,x,y) = y" |
|
77 |
(fn _=>[stac if_def 1, fast_tac (HOL_cs addIs [select_equality]) 1]); |
|
78 |
||
79 |
val if_P = prove_goal HOL.thy "P ==> if(P,x,y) = x" |
|
80 |
(fn [prem] => [ stac (prem RS eqTrueI) 1, rtac if_True 1 ]); |
|
81 |
||
82 |
val if_not_P = prove_goal HOL.thy "~P ==> if(P,x,y) = y" |
|
83 |
(fn [prem] => [ stac (prem RS not_P_imp_P_iff_F) 1, rtac if_False 1 ]); |
|
84 |
||
85 |
val expand_if = prove_goal HOL.thy |
|
86 |
"P(if(Q,x,y)) = ((Q --> P(x)) & (~Q --> P(y)))" |
|
87 |
(fn _=> [ (res_inst_tac [("Q","Q")] (excluded_middle RS disjE) 1), |
|
88 |
rtac (if_P RS ssubst) 2, |
|
89 |
rtac (if_not_P RS ssubst) 1, |
|
90 |
REPEAT(fast_tac HOL_cs 1) ]); |
|
91 |
||
1 | 92 |
infix addcongs; |
93 |
fun ss addcongs congs = ss addeqcongs (congs RL [eq_reflection]); |
|
0 | 94 |
|
95 |
val HOL_ss = empty_ss |
|
96 |
setmksimps mk_rews |
|
43 | 97 |
setsolver (fn prems => resolve_tac (TrueI::refl::prems) ORELSE' atac |
98 |
ORELSE' etac FalseE) |
|
0 | 99 |
setsubgoaler asm_simp_tac |
34 | 100 |
addsimps ([if_True, if_False, o_apply, conj_assoc] @ simp_thms) |
1 | 101 |
addcongs [imp_cong]; |
0 | 102 |
|
103 |
fun split_tac splits = |
|
104 |
mk_case_split_tac (meta_obj_reflection RS iffD2) (map mk_eq splits); |
|
105 |
||
40 | 106 |
(* eliminiation of existential quantifiers in assumptions *) |
107 |
||
108 |
val ex_all_equiv = |
|
109 |
let val lemma1 = prove_goal HOL.thy |
|
110 |
"(? x. P(x) ==> PROP Q) ==> (!!x. P(x) ==> PROP Q)" |
|
111 |
(fn prems => [resolve_tac prems 1, etac exI 1]); |
|
112 |
val lemma2 = prove_goalw HOL.thy [Ex_def RS eq_reflection] |
|
113 |
"(!!x. P(x) ==> PROP Q) ==> (? x. P(x) ==> PROP Q)" |
|
114 |
(fn prems => [REPEAT(resolve_tac prems 1)]) |
|
115 |
in equal_intr lemma1 lemma2 end; |
|
116 |
||
30
2fdeeae553ac
modified solver of HOL_ss to take the new simplifier into account: the thm to
nipkow
parents:
3
diff
changeset
|
117 |
(** '&' congruence rule: not included by default! *) |
2fdeeae553ac
modified solver of HOL_ss to take the new simplifier into account: the thm to
nipkow
parents:
3
diff
changeset
|
118 |
|
2fdeeae553ac
modified solver of HOL_ss to take the new simplifier into account: the thm to
nipkow
parents:
3
diff
changeset
|
119 |
val conj_cong = impI RSN |
2fdeeae553ac
modified solver of HOL_ss to take the new simplifier into account: the thm to
nipkow
parents:
3
diff
changeset
|
120 |
(2, prove_goal HOL.thy "(P=P')--> (P'--> (Q=Q'))--> ((P&Q) = (P'&Q'))" |
2fdeeae553ac
modified solver of HOL_ss to take the new simplifier into account: the thm to
nipkow
parents:
3
diff
changeset
|
121 |
(fn _=> [fast_tac HOL_cs 1]) RS mp RS mp); |
2fdeeae553ac
modified solver of HOL_ss to take the new simplifier into account: the thm to
nipkow
parents:
3
diff
changeset
|
122 |
|
3
a910a65478be
This repeats a previous commit, which failed to update simpdata.ML because
lcp
parents:
1
diff
changeset
|
123 |
(** 'if' congruence rules: neither included by default! *) |
a910a65478be
This repeats a previous commit, which failed to update simpdata.ML because
lcp
parents:
1
diff
changeset
|
124 |
|
a910a65478be
This repeats a previous commit, which failed to update simpdata.ML because
lcp
parents:
1
diff
changeset
|
125 |
(*Simplifies x assuming c and y assuming ~c*) |
a910a65478be
This repeats a previous commit, which failed to update simpdata.ML because
lcp
parents:
1
diff
changeset
|
126 |
val if_cong = prove_goal HOL.thy |
a910a65478be
This repeats a previous commit, which failed to update simpdata.ML because
lcp
parents:
1
diff
changeset
|
127 |
"[| b=c; c ==> x=u; ~c ==> y=v |] ==> if(b,x,y) = if(c,u,v)" |
a910a65478be
This repeats a previous commit, which failed to update simpdata.ML because
lcp
parents:
1
diff
changeset
|
128 |
(fn rew::prems => |
a910a65478be
This repeats a previous commit, which failed to update simpdata.ML because
lcp
parents:
1
diff
changeset
|
129 |
[stac rew 1, stac expand_if 1, stac expand_if 1, |
a910a65478be
This repeats a previous commit, which failed to update simpdata.ML because
lcp
parents:
1
diff
changeset
|
130 |
fast_tac (HOL_cs addDs prems) 1]); |
a910a65478be
This repeats a previous commit, which failed to update simpdata.ML because
lcp
parents:
1
diff
changeset
|
131 |
|
a910a65478be
This repeats a previous commit, which failed to update simpdata.ML because
lcp
parents:
1
diff
changeset
|
132 |
(*Prevents simplification of x and y: much faster*) |
a910a65478be
This repeats a previous commit, which failed to update simpdata.ML because
lcp
parents:
1
diff
changeset
|
133 |
val if_weak_cong = prove_goal HOL.thy |
a910a65478be
This repeats a previous commit, which failed to update simpdata.ML because
lcp
parents:
1
diff
changeset
|
134 |
"b=c ==> if(b,x,y) = if(c,x,y)" |
a910a65478be
This repeats a previous commit, which failed to update simpdata.ML because
lcp
parents:
1
diff
changeset
|
135 |
(fn [prem] => [rtac (prem RS arg_cong) 1]); |
a910a65478be
This repeats a previous commit, which failed to update simpdata.ML because
lcp
parents:
1
diff
changeset
|
136 |
|
32 | 137 |
(*Prevents simplification of t: much faster*) |
138 |
val let_weak_cong = prove_goal HOL.thy |
|
139 |
"a = b ==> (let x=a in t(x)) = (let x=b in t(x))" |
|
140 |
(fn [prem] => [rtac (prem RS arg_cong) 1]); |
|
141 |
||
0 | 142 |
end; |