author | wenzelm |
Wed, 09 Jun 2004 18:57:18 +0200 | |
changeset 14913 | e993eabc7197 |
parent 14854 | 61bdf2ae4dc5 |
child 15531 | 08c8dad8e399 |
permissions | -rw-r--r-- |
4 | 1 |
(* Title: Provers/splitter |
2 |
ID: $Id$ |
|
3 |
Author: Tobias Nipkow |
|
1030
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents:
943
diff
changeset
|
4 |
Copyright 1995 TU Munich |
4 | 5 |
|
6 |
Generic case-splitter, suitable for most logics. |
|
13157
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
7 |
Deals with equalities of the form ?P(f args) = ... |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
8 |
where "f args" must be a first-order term without duplicate variables. |
0 | 9 |
*) |
10 |
||
5304 | 11 |
infix 4 addsplits delsplits; |
12 |
||
13 |
signature SPLITTER_DATA = |
|
14 |
sig |
|
15 |
structure Simplifier: SIMPLIFIER |
|
5553 | 16 |
val mk_eq : thm -> thm |
5304 | 17 |
val meta_eq_to_iff: thm (* "x == y ==> x = y" *) |
18 |
val iffD : thm (* "[| P = Q; Q |] ==> P" *) |
|
19 |
val disjE : thm (* "[| P | Q; P ==> R; Q ==> R |] ==> R" *) |
|
20 |
val conjE : thm (* "[| P & Q; [| P; Q |] ==> R |] ==> R" *) |
|
21 |
val exE : thm (* "[| x. P x; !!x. P x ==> Q |] ==> Q" *) |
|
22 |
val contrapos : thm (* "[| ~ Q; P ==> Q |] ==> ~ P" *) |
|
23 |
val contrapos2 : thm (* "[| Q; ~ P ==> ~ Q |] ==> P" *) |
|
24 |
val notnotD : thm (* "~ ~ P ==> P" *) |
|
25 |
end |
|
26 |
||
27 |
signature SPLITTER = |
|
28 |
sig |
|
29 |
type simpset |
|
30 |
val split_tac : thm list -> int -> tactic |
|
31 |
val split_inside_tac: thm list -> int -> tactic |
|
32 |
val split_asm_tac : thm list -> int -> tactic |
|
33 |
val addsplits : simpset * thm list -> simpset |
|
34 |
val delsplits : simpset * thm list -> simpset |
|
35 |
val Addsplits : thm list -> unit |
|
36 |
val Delsplits : thm list -> unit |
|
8468 | 37 |
val split_add_global: theory attribute |
38 |
val split_del_global: theory attribute |
|
39 |
val split_add_local: Proof.context attribute |
|
40 |
val split_del_local: Proof.context attribute |
|
41 |
val split_modifiers : (Args.T list -> (Method.modifier * Args.T list)) list |
|
42 |
val setup: (theory -> theory) list |
|
5304 | 43 |
end; |
44 |
||
45 |
functor SplitterFun(Data: SPLITTER_DATA): SPLITTER = |
|
46 |
struct |
|
47 |
||
8468 | 48 |
structure Simplifier = Data.Simplifier; |
49 |
type simpset = Simplifier.simpset; |
|
5304 | 50 |
|
51 |
val Const ("==>", _) $ (Const ("Trueprop", _) $ |
|
52 |
(Const (const_not, _) $ _ )) $ _ = #prop (rep_thm(Data.notnotD)); |
|
53 |
||
54 |
val Const ("==>", _) $ (Const ("Trueprop", _) $ |
|
55 |
(Const (const_or , _) $ _ $ _)) $ _ = #prop (rep_thm(Data.disjE)); |
|
1721
445654b6cb95
Rewrote mk_cntxt_splitthm. Added function mk_case_split_inside_tac.
berghofe
parents:
1686
diff
changeset
|
56 |
|
4668
131989b78417
Little reorganization. Loop tactics have names now.
nipkow
parents:
4519
diff
changeset
|
57 |
fun split_format_err() = error("Wrong format for split rule"); |
131989b78417
Little reorganization. Loop tactics have names now.
nipkow
parents:
4519
diff
changeset
|
58 |
|
5553 | 59 |
fun split_thm_info thm = case concl_of (Data.mk_eq thm) of |
13855
644692eca537
addsplits / delsplits no longer ignore type of constant.
berghofe
parents:
13157
diff
changeset
|
60 |
Const("==", _) $ (Var _ $ t) $ c => (case strip_comb t of |
644692eca537
addsplits / delsplits no longer ignore type of constant.
berghofe
parents:
13157
diff
changeset
|
61 |
(Const p, _) => (p, case c of (Const (s, _) $ _) => s = const_not | _ => false) |
644692eca537
addsplits / delsplits no longer ignore type of constant.
berghofe
parents:
13157
diff
changeset
|
62 |
| _ => split_format_err ()) |
644692eca537
addsplits / delsplits no longer ignore type of constant.
berghofe
parents:
13157
diff
changeset
|
63 |
| _ => split_format_err (); |
5304 | 64 |
|
65 |
fun mk_case_split_tac order = |
|
0 | 66 |
let |
67 |
||
1686
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
68 |
|
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
69 |
(************************************************************ |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
70 |
Create lift-theorem "trlift" : |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
71 |
|
7672
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
72 |
[| !!x. Q x == R x; P(%x. R x) == C |] ==> P (%x. Q x) == C |
1686
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
73 |
|
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
74 |
*************************************************************) |
5304 | 75 |
|
76 |
val meta_iffD = Data.meta_eq_to_iff RS Data.iffD; |
|
943 | 77 |
val lift = |
5304 | 78 |
let val ct = read_cterm (#sign(rep_thm Data.iffD)) |
14854 | 79 |
("[| !!x. (Q::('b::{})=>('c::{}))(x) == R(x) |] ==> \ |
80 |
\P(%x. Q(x)) == P(%x. R(x))::'a::{}",propT) |
|
943 | 81 |
in prove_goalw_cterm [] ct |
82 |
(fn [prem] => [rewtac prem, rtac reflexive_thm 1]) |
|
83 |
end; |
|
4 | 84 |
|
0 | 85 |
val trlift = lift RS transitive_thm; |
7672
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
86 |
val _ $ (P $ _) $ _ = concl_of trlift; |
0 | 87 |
|
88 |
||
1686
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
89 |
(************************************************************************ |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
90 |
Set up term for instantiation of P in the lift-theorem |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
91 |
|
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
92 |
Ts : types of parameters (i.e. variables bound by meta-quantifiers) |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
93 |
t : lefthand side of meta-equality in subgoal |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
94 |
the lift theorem is applied to (see select) |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
95 |
pos : "path" leading to abstraction, coded as a list |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
96 |
T : type of body of P(...) |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
97 |
maxi : maximum index of Vars |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
98 |
*************************************************************************) |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
99 |
|
1030
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents:
943
diff
changeset
|
100 |
fun mk_cntxt Ts t pos T maxi = |
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents:
943
diff
changeset
|
101 |
let fun var (t,i) = Var(("X",i),type_of1(Ts,t)); |
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents:
943
diff
changeset
|
102 |
fun down [] t i = Bound 0 |
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents:
943
diff
changeset
|
103 |
| down (p::ps) t i = |
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents:
943
diff
changeset
|
104 |
let val (h,ts) = strip_comb t |
2266 | 105 |
val v1 = ListPair.map var (take(p,ts), i upto (i+p-1)) |
1030
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents:
943
diff
changeset
|
106 |
val u::us = drop(p,ts) |
2266 | 107 |
val v2 = ListPair.map var (us, (i+p) upto (i+length(ts)-2)) |
1030
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents:
943
diff
changeset
|
108 |
in list_comb(h,v1@[down ps u (i+length ts)]@v2) end; |
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents:
943
diff
changeset
|
109 |
in Abs("", T, down (rev pos) t maxi) end; |
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents:
943
diff
changeset
|
110 |
|
1686
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
111 |
|
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
112 |
(************************************************************************ |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
113 |
Set up term for instantiation of P in the split-theorem |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
114 |
P(...) == rhs |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
115 |
|
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
116 |
t : lefthand side of meta-equality in subgoal |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
117 |
the split theorem is applied to (see select) |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
118 |
T : type of body of P(...) |
4232 | 119 |
tt : the term Const(key,..) $ ... |
1686
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
120 |
*************************************************************************) |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
121 |
|
4232 | 122 |
fun mk_cntxt_splitthm t tt T = |
123 |
let fun repl lev t = |
|
7672
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
124 |
if incr_boundvars lev tt aconv t then Bound lev |
4232 | 125 |
else case t of |
126 |
(Abs (v, T2, t)) => Abs (v, T2, repl (lev+1) t) |
|
127 |
| (Bound i) => Bound (if i>=lev then i+1 else i) |
|
128 |
| (t1 $ t2) => (repl lev t1) $ (repl lev t2) |
|
129 |
| t => t |
|
130 |
in Abs("", T, repl 0 t) end; |
|
1686
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
131 |
|
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
132 |
|
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
133 |
(* add all loose bound variables in t to list is *) |
1030
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents:
943
diff
changeset
|
134 |
fun add_lbnos(is,t) = add_loose_bnos(t,0,is); |
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents:
943
diff
changeset
|
135 |
|
7672
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
136 |
(* check if the innermost abstraction that needs to be removed |
1064 | 137 |
has a body of type T; otherwise the expansion thm will fail later on |
138 |
*) |
|
139 |
fun type_test(T,lbnos,apsns) = |
|
2143 | 140 |
let val (_,U,_) = nth_elem(foldl Int.min (hd lbnos, tl lbnos), apsns) |
1064 | 141 |
in T=U end; |
0 | 142 |
|
1686
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
143 |
(************************************************************************* |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
144 |
Create a "split_pack". |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
145 |
|
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
146 |
thm : the relevant split-theorem, i.e. P(...) == rhs , where P(...) |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
147 |
is of the form |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
148 |
P( Const(key,...) $ t_1 $ ... $ t_n ) (e.g. key = "if") |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
149 |
T : type of P(...) |
7672
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
150 |
T' : type of term to be scanned |
1686
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
151 |
n : number of arguments expected by Const(key,...) |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
152 |
ts : list of arguments actually found |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
153 |
apsns : list of tuples of the form (T,U,pos), one tuple for each |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
154 |
abstraction that is encountered on the way to the position where |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
155 |
Const(key, ...) $ ... occurs, where |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
156 |
T : type of the variable bound by the abstraction |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
157 |
U : type of the abstraction's body |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
158 |
pos : "path" leading to the body of the abstraction |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
159 |
pos : "path" leading to the position where Const(key, ...) $ ... occurs. |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
160 |
TB : type of Const(key,...) $ t_1 $ ... $ t_n |
1721
445654b6cb95
Rewrote mk_cntxt_splitthm. Added function mk_case_split_inside_tac.
berghofe
parents:
1686
diff
changeset
|
161 |
t : the term Const(key,...) $ t_1 $ ... $ t_n |
1686
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
162 |
|
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
163 |
A split pack is a tuple of the form |
7672
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
164 |
(thm, apsns, pos, TB, tt) |
1686
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
165 |
Note : apsns is reversed, so that the outermost quantifier's position |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
166 |
comes first ! If the terms in ts don't contain variables bound |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
167 |
by other than meta-quantifiers, apsns is empty, because no further |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
168 |
lifting is required before applying the split-theorem. |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
169 |
******************************************************************************) |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
170 |
|
7672
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
171 |
fun mk_split_pack(thm, T, T', n, ts, apsns, pos, TB, t) = |
1064 | 172 |
if n > length ts then [] |
173 |
else let val lev = length apsns |
|
1030
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents:
943
diff
changeset
|
174 |
val lbnos = foldl add_lbnos ([],take(n,ts)) |
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents:
943
diff
changeset
|
175 |
val flbnos = filter (fn i => i < lev) lbnos |
4232 | 176 |
val tt = incr_boundvars (~lev) t |
7672
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
177 |
in if null flbnos then |
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
178 |
if T = T' then [(thm,[],pos,TB,tt)] else [] |
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
179 |
else if type_test(T,flbnos,apsns) then [(thm, rev apsns,pos,TB,tt)] |
2143 | 180 |
else [] |
1064 | 181 |
end; |
0 | 182 |
|
1686
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
183 |
|
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
184 |
(**************************************************************************** |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
185 |
Recursively scans term for occurences of Const(key,...) $ ... |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
186 |
Returns a list of "split-packs" (one for each occurence of Const(key,...) ) |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
187 |
|
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
188 |
cmap : association list of split-theorems that should be tried. |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
189 |
The elements have the format (key,(thm,T,n)) , where |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
190 |
key : the theorem's key constant ( Const(key,...) $ ... ) |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
191 |
thm : the theorem itself |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
192 |
T : type of P( Const(key,...) $ ... ) |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
193 |
n : number of arguments expected by Const(key,...) |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
194 |
Ts : types of parameters |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
195 |
t : the term to be scanned |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
196 |
******************************************************************************) |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
197 |
|
13157
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
198 |
(* Simplified first-order matching; |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
199 |
assumes that all Vars in the pattern are distinct; |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
200 |
see Pure/pattern.ML for the full version; |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
201 |
*) |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
202 |
local |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
203 |
exception MATCH |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
204 |
in |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
205 |
fun typ_match tsig args = (Type.typ_match tsig args) |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
206 |
handle Type.TYPE_MATCH => raise MATCH; |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
207 |
fun fomatch tsig args = |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
208 |
let |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
209 |
fun mtch tyinsts = fn |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
210 |
(Ts,Var(_,T), t) => typ_match tsig (tyinsts, (T, fastype_of1(Ts,t))) |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
211 |
| (_,Free (a,T), Free (b,U)) => |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
212 |
if a=b then typ_match tsig (tyinsts,(T,U)) else raise MATCH |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
213 |
| (_,Const (a,T), Const (b,U)) => |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
214 |
if a=b then typ_match tsig (tyinsts,(T,U)) else raise MATCH |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
215 |
| (_,Bound i, Bound j) => if i=j then tyinsts else raise MATCH |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
216 |
| (Ts,Abs(_,T,t), Abs(_,U,u)) => |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
217 |
mtch (typ_match tsig (tyinsts,(T,U))) (U::Ts,t,u) |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
218 |
| (Ts, f$t, g$u) => mtch (mtch tyinsts (Ts,f,g)) (Ts, t, u) |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
219 |
| _ => raise MATCH |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
220 |
in (mtch Vartab.empty args; true) handle MATCH => false end; |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
221 |
end |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
222 |
|
6130
30b84ad2131d
Fixed old bug: selection of constant to be split should depend not just on
nipkow
parents:
5553
diff
changeset
|
223 |
fun split_posns cmap sg Ts t = |
30b84ad2131d
Fixed old bug: selection of constant to be split should depend not just on
nipkow
parents:
5553
diff
changeset
|
224 |
let |
7672
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
225 |
val T' = fastype_of1 (Ts, t); |
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
226 |
fun posns Ts pos apsns (Abs (_, T, t)) = |
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
227 |
let val U = fastype_of1 (T::Ts,t) |
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
228 |
in posns (T::Ts) (0::pos) ((T, U, pos)::apsns) t end |
6130
30b84ad2131d
Fixed old bug: selection of constant to be split should depend not just on
nipkow
parents:
5553
diff
changeset
|
229 |
| posns Ts pos apsns t = |
30b84ad2131d
Fixed old bug: selection of constant to be split should depend not just on
nipkow
parents:
5553
diff
changeset
|
230 |
let |
7672
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
231 |
val (h, ts) = strip_comb t |
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
232 |
fun iter((i, a), t) = (i+1, (posns Ts (i::pos) apsns t) @ a); |
6130
30b84ad2131d
Fixed old bug: selection of constant to be split should depend not just on
nipkow
parents:
5553
diff
changeset
|
233 |
val a = case h of |
7672
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
234 |
Const(c, cT) => |
9267
dbf30a2d1b56
Now two split thms for same constant at different types is allowed.
nipkow
parents:
8815
diff
changeset
|
235 |
let fun find [] = [] |
13157
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
236 |
| find ((gcT, pat, thm, T, n)::tups) = |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
237 |
let val t2 = list_comb (h, take (n, ts)) |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
238 |
in if Sign.typ_instance sg (cT, gcT) |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
239 |
andalso fomatch (Sign.tsig_of sg) (Ts,pat,t2) |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
240 |
then mk_split_pack(thm,T,T',n,ts,apsns,pos,type_of1(Ts,t2),t2) |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
241 |
else find tups |
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
242 |
end |
9267
dbf30a2d1b56
Now two split thms for same constant at different types is allowed.
nipkow
parents:
8815
diff
changeset
|
243 |
in find (assocs cmap c) end |
6130
30b84ad2131d
Fixed old bug: selection of constant to be split should depend not just on
nipkow
parents:
5553
diff
changeset
|
244 |
| _ => [] |
7672
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
245 |
in snd(foldl iter ((0, a), ts)) end |
1030
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents:
943
diff
changeset
|
246 |
in posns Ts [] [] t end; |
0 | 247 |
|
1686
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
248 |
|
0 | 249 |
fun nth_subgoal i thm = nth_elem(i-1,prems_of thm); |
250 |
||
1721
445654b6cb95
Rewrote mk_cntxt_splitthm. Added function mk_case_split_inside_tac.
berghofe
parents:
1686
diff
changeset
|
251 |
fun shorter((_,ps,pos,_,_),(_,qs,qos,_,_)) = |
4519 | 252 |
prod_ord (int_ord o pairself length) (order o pairself length) |
253 |
((ps, pos), (qs, qos)); |
|
254 |
||
1686
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
255 |
|
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
256 |
|
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
257 |
(************************************************************ |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
258 |
call split_posns with appropriate parameters |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
259 |
*************************************************************) |
0 | 260 |
|
1030
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents:
943
diff
changeset
|
261 |
fun select cmap state i = |
6130
30b84ad2131d
Fixed old bug: selection of constant to be split should depend not just on
nipkow
parents:
5553
diff
changeset
|
262 |
let val sg = #sign(rep_thm state) |
30b84ad2131d
Fixed old bug: selection of constant to be split should depend not just on
nipkow
parents:
5553
diff
changeset
|
263 |
val goali = nth_subgoal i state |
1030
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents:
943
diff
changeset
|
264 |
val Ts = rev(map #2 (Logic.strip_params goali)) |
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents:
943
diff
changeset
|
265 |
val _ $ t $ _ = Logic.strip_assums_concl goali; |
6130
30b84ad2131d
Fixed old bug: selection of constant to be split should depend not just on
nipkow
parents:
5553
diff
changeset
|
266 |
in (Ts,t, sort shorter (split_posns cmap sg Ts t)) end; |
1030
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents:
943
diff
changeset
|
267 |
|
1686
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
268 |
|
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
269 |
(************************************************************* |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
270 |
instantiate lift theorem |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
271 |
|
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
272 |
if t is of the form |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
273 |
... ( Const(...,...) $ Abs( .... ) ) ... |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
274 |
then |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
275 |
P = %a. ... ( Const(...,...) $ a ) ... |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
276 |
where a has type T --> U |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
277 |
|
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
278 |
Ts : types of parameters |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
279 |
t : lefthand side of meta-equality in subgoal |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
280 |
the split theorem is applied to (see cmap) |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
281 |
T,U,pos : see mk_split_pack |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
282 |
state : current proof state |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
283 |
lift : the lift theorem |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
284 |
i : no. of subgoal |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
285 |
**************************************************************) |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
286 |
|
7672
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
287 |
fun inst_lift Ts t (T, U, pos) state i = |
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
288 |
let |
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
289 |
val cert = cterm_of (sign_of_thm state); |
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
290 |
val cntxt = mk_cntxt Ts t pos (T --> U) (#maxidx(rep_thm trlift)); |
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
291 |
in cterm_instantiate [(cert P, cert cntxt)] trlift |
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
292 |
end; |
0 | 293 |
|
294 |
||
1686
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
295 |
(************************************************************* |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
296 |
instantiate split theorem |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
297 |
|
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
298 |
Ts : types of parameters |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
299 |
t : lefthand side of meta-equality in subgoal |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
300 |
the split theorem is applied to (see cmap) |
4232 | 301 |
tt : the term Const(key,..) $ ... |
1686
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
302 |
thm : the split theorem |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
303 |
TB : type of body of P(...) |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
304 |
state : current proof state |
4232 | 305 |
i : number of subgoal |
1686
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
306 |
**************************************************************) |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
307 |
|
4232 | 308 |
fun inst_split Ts t tt thm TB state i = |
7672
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
309 |
let |
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
310 |
val thm' = Thm.lift_rule (state, i) thm; |
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
311 |
val (P, _) = strip_comb (fst (Logic.dest_equals |
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
312 |
(Logic.strip_assums_concl (#prop (rep_thm thm'))))); |
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
313 |
val cert = cterm_of (sign_of_thm state); |
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
314 |
val cntxt = mk_cntxt_splitthm t tt TB; |
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
315 |
val abss = foldl (fn (t, T) => Abs ("", T, t)); |
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
316 |
in cterm_instantiate [(cert P, cert (abss (cntxt, Ts)))] thm' |
4232 | 317 |
end; |
1686
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
318 |
|
7672
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
319 |
|
1686
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
320 |
(***************************************************************************** |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
321 |
The split-tactic |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
322 |
|
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
323 |
splits : list of split-theorems to be tried |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
324 |
i : number of subgoal the tactic should be applied to |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
325 |
*****************************************************************************) |
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents:
1064
diff
changeset
|
326 |
|
0 | 327 |
fun split_tac [] i = no_tac |
328 |
| split_tac splits i = |
|
5553 | 329 |
let val splits = map Data.mk_eq splits; |
9267
dbf30a2d1b56
Now two split thms for same constant at different types is allowed.
nipkow
parents:
8815
diff
changeset
|
330 |
fun add_thm(cmap,thm) = |
3918 | 331 |
(case concl_of thm of _$(t as _$lhs)$_ => |
6130
30b84ad2131d
Fixed old bug: selection of constant to be split should depend not just on
nipkow
parents:
5553
diff
changeset
|
332 |
(case strip_comb lhs of (Const(a,aT),args) => |
13157
4a4599f78f18
allowed more general split rules to cope with div/mod 2
nipkow
parents:
10821
diff
changeset
|
333 |
let val info = (aT,lhs,thm,fastype_of t,length args) |
9267
dbf30a2d1b56
Now two split thms for same constant at different types is allowed.
nipkow
parents:
8815
diff
changeset
|
334 |
in case assoc(cmap,a) of |
dbf30a2d1b56
Now two split thms for same constant at different types is allowed.
nipkow
parents:
8815
diff
changeset
|
335 |
Some infos => overwrite(cmap,(a,info::infos)) |
dbf30a2d1b56
Now two split thms for same constant at different types is allowed.
nipkow
parents:
8815
diff
changeset
|
336 |
| None => (a,[info])::cmap |
dbf30a2d1b56
Now two split thms for same constant at different types is allowed.
nipkow
parents:
8815
diff
changeset
|
337 |
end |
4668
131989b78417
Little reorganization. Loop tactics have names now.
nipkow
parents:
4519
diff
changeset
|
338 |
| _ => split_format_err()) |
131989b78417
Little reorganization. Loop tactics have names now.
nipkow
parents:
4519
diff
changeset
|
339 |
| _ => split_format_err()) |
9267
dbf30a2d1b56
Now two split thms for same constant at different types is allowed.
nipkow
parents:
8815
diff
changeset
|
340 |
val cmap = foldl add_thm ([],splits); |
7672
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
341 |
fun lift_tac Ts t p st = rtac (inst_lift Ts t p st i) i st |
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
342 |
fun lift_split_tac state = |
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
343 |
let val (Ts, t, splits) = select cmap state i |
1030
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents:
943
diff
changeset
|
344 |
in case splits of |
7672
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
345 |
[] => no_tac state |
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
346 |
| (thm, apsns, pos, TB, tt)::_ => |
1030
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents:
943
diff
changeset
|
347 |
(case apsns of |
7672
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
348 |
[] => compose_tac (false, inst_split Ts t tt thm TB state i, 0) i state |
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
349 |
| p::_ => EVERY [lift_tac Ts t p, |
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
350 |
rtac reflexive_thm (i+1), |
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
berghofe
parents:
6130
diff
changeset
|
351 |
lift_split_tac] state) |
1030
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents:
943
diff
changeset
|
352 |
end |
3537 | 353 |
in COND (has_fewer_prems i) no_tac |
5304 | 354 |
(rtac meta_iffD i THEN lift_split_tac) |
0 | 355 |
end; |
356 |
||
357 |
in split_tac end; |
|
1721
445654b6cb95
Rewrote mk_cntxt_splitthm. Added function mk_case_split_inside_tac.
berghofe
parents:
1686
diff
changeset
|
358 |
|
5304 | 359 |
|
360 |
val split_tac = mk_case_split_tac int_ord; |
|
4189 | 361 |
|
5304 | 362 |
val split_inside_tac = mk_case_split_tac (rev_order o int_ord); |
363 |
||
4189 | 364 |
|
365 |
(***************************************************************************** |
|
366 |
The split-tactic for premises |
|
367 |
||
368 |
splits : list of split-theorems to be tried |
|
5304 | 369 |
****************************************************************************) |
4202 | 370 |
fun split_asm_tac [] = K no_tac |
371 |
| split_asm_tac splits = |
|
5304 | 372 |
|
13855
644692eca537
addsplits / delsplits no longer ignore type of constant.
berghofe
parents:
13157
diff
changeset
|
373 |
let val cname_list = map (fst o fst o split_thm_info) splits; |
4189 | 374 |
fun is_case (a,_) = a mem cname_list; |
375 |
fun tac (t,i) = |
|
376 |
let val n = find_index (exists_Const is_case) |
|
377 |
(Logic.strip_assums_hyp t); |
|
378 |
fun first_prem_is_disj (Const ("==>", _) $ (Const ("Trueprop", _) |
|
5304 | 379 |
$ (Const (s, _) $ _ $ _ )) $ _ ) = (s=const_or) |
4202 | 380 |
| first_prem_is_disj (Const("all",_)$Abs(_,_,t)) = |
381 |
first_prem_is_disj t |
|
4189 | 382 |
| first_prem_is_disj _ = false; |
5437 | 383 |
(* does not work properly if the split variable is bound by a quantfier *) |
4202 | 384 |
fun flat_prems_tac i = SUBGOAL (fn (t,i) => |
5304 | 385 |
(if first_prem_is_disj t |
386 |
then EVERY[etac Data.disjE i,rotate_tac ~1 i, |
|
387 |
rotate_tac ~1 (i+1), |
|
388 |
flat_prems_tac (i+1)] |
|
389 |
else all_tac) |
|
390 |
THEN REPEAT (eresolve_tac [Data.conjE,Data.exE] i) |
|
391 |
THEN REPEAT (dresolve_tac [Data.notnotD] i)) i; |
|
4189 | 392 |
in if n<0 then no_tac else DETERM (EVERY' |
5304 | 393 |
[rotate_tac n, etac Data.contrapos2, |
4189 | 394 |
split_tac splits, |
5304 | 395 |
rotate_tac ~1, etac Data.contrapos, rotate_tac ~1, |
4202 | 396 |
flat_prems_tac] i) |
4189 | 397 |
end; |
398 |
in SUBGOAL tac |
|
399 |
end; |
|
400 |
||
10652 | 401 |
fun gen_split_tac [] = K no_tac |
402 |
| gen_split_tac (split::splits) = |
|
403 |
let val (_,asm) = split_thm_info split |
|
404 |
in (if asm then split_asm_tac else split_tac) [split] ORELSE' |
|
405 |
gen_split_tac splits |
|
406 |
end; |
|
8468 | 407 |
|
408 |
(** declare split rules **) |
|
409 |
||
410 |
(* addsplits / delsplits *) |
|
411 |
||
13859
adf68d9e5dec
split_name no longer uses Sign.string_of_typ to encode types, since
berghofe
parents:
13855
diff
changeset
|
412 |
fun string_of_typ (Type (s, Ts)) = (if null Ts then "" |
adf68d9e5dec
split_name no longer uses Sign.string_of_typ to encode types, since
berghofe
parents:
13855
diff
changeset
|
413 |
else enclose "(" ")" (commas (map string_of_typ Ts))) ^ s |
adf68d9e5dec
split_name no longer uses Sign.string_of_typ to encode types, since
berghofe
parents:
13855
diff
changeset
|
414 |
| string_of_typ _ = "_"; |
adf68d9e5dec
split_name no longer uses Sign.string_of_typ to encode types, since
berghofe
parents:
13855
diff
changeset
|
415 |
|
adf68d9e5dec
split_name no longer uses Sign.string_of_typ to encode types, since
berghofe
parents:
13855
diff
changeset
|
416 |
fun split_name (name, T) asm = "split " ^ |
adf68d9e5dec
split_name no longer uses Sign.string_of_typ to encode types, since
berghofe
parents:
13855
diff
changeset
|
417 |
(if asm then "asm " else "") ^ name ^ " :: " ^ string_of_typ T; |
4189 | 418 |
|
5304 | 419 |
fun ss addsplits splits = |
420 |
let fun addsplit (ss,split) = |
|
421 |
let val (name,asm) = split_thm_info split |
|
13859
adf68d9e5dec
split_name no longer uses Sign.string_of_typ to encode types, since
berghofe
parents:
13855
diff
changeset
|
422 |
in Simplifier.addloop (ss, (split_name name asm, |
5304 | 423 |
(if asm then split_asm_tac else split_tac) [split])) end |
424 |
in foldl addsplit (ss,splits) end; |
|
1721
445654b6cb95
Rewrote mk_cntxt_splitthm. Added function mk_case_split_inside_tac.
berghofe
parents:
1686
diff
changeset
|
425 |
|
5304 | 426 |
fun ss delsplits splits = |
427 |
let fun delsplit(ss,split) = |
|
428 |
let val (name,asm) = split_thm_info split |
|
13859
adf68d9e5dec
split_name no longer uses Sign.string_of_typ to encode types, since
berghofe
parents:
13855
diff
changeset
|
429 |
in Simplifier.delloop (ss, split_name name asm) |
5304 | 430 |
end in foldl delsplit (ss,splits) end; |
1721
445654b6cb95
Rewrote mk_cntxt_splitthm. Added function mk_case_split_inside_tac.
berghofe
parents:
1686
diff
changeset
|
431 |
|
8468 | 432 |
fun Addsplits splits = (Simplifier.simpset_ref() := |
433 |
Simplifier.simpset() addsplits splits); |
|
434 |
fun Delsplits splits = (Simplifier.simpset_ref() := |
|
435 |
Simplifier.simpset() delsplits splits); |
|
436 |
||
437 |
||
438 |
(* attributes *) |
|
439 |
||
440 |
val splitN = "split"; |
|
441 |
||
442 |
val split_add_global = Simplifier.change_global_ss (op addsplits); |
|
443 |
val split_del_global = Simplifier.change_global_ss (op delsplits); |
|
444 |
val split_add_local = Simplifier.change_local_ss (op addsplits); |
|
445 |
val split_del_local = Simplifier.change_local_ss (op delsplits); |
|
446 |
||
8634 | 447 |
val split_attr = |
448 |
(Attrib.add_del_args split_add_global split_del_global, |
|
449 |
Attrib.add_del_args split_add_local split_del_local); |
|
450 |
||
451 |
||
9703 | 452 |
(* methods *) |
8468 | 453 |
|
454 |
val split_modifiers = |
|
8815 | 455 |
[Args.$$$ splitN -- Args.colon >> K ((I, split_add_local): Method.modifier), |
10034 | 456 |
Args.$$$ splitN -- Args.add -- Args.colon >> K (I, split_add_local), |
457 |
Args.$$$ splitN -- Args.del -- Args.colon >> K (I, split_del_local)]; |
|
8468 | 458 |
|
10652 | 459 |
val split_args = #2 oo Method.syntax Attrib.local_thms; |
9807 | 460 |
|
10821 | 461 |
fun split_meth ths = Method.SIMPLE_METHOD' HEADGOAL (CHANGED_PROP o gen_split_tac ths); |
9703 | 462 |
|
8468 | 463 |
|
464 |
||
465 |
(** theory setup **) |
|
466 |
||
9703 | 467 |
val setup = |
9900 | 468 |
[Attrib.add_attributes [(splitN, split_attr, "declaration of case split rule")], |
469 |
Method.add_methods [(splitN, split_meth oo split_args, "apply case split rule")]]; |
|
4189 | 470 |
|
1721
445654b6cb95
Rewrote mk_cntxt_splitthm. Added function mk_case_split_inside_tac.
berghofe
parents:
1686
diff
changeset
|
471 |
end; |