| author | blanchet | 
| Thu, 06 Jun 2013 21:12:08 +0200 | |
| changeset 52323 | a11bbb5fef56 | 
| parent 52131 | 366fa32ee2a3 | 
| child 54216 | c0c453ce70a7 | 
| permissions | -rw-r--r-- | 
| 32174 | 1  | 
(* Title: Provers/splitter.ML  | 
| 4 | 2  | 
Author: Tobias Nipkow  | 
| 
1030
 
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
 
nipkow 
parents: 
943 
diff
changeset
 | 
3  | 
Copyright 1995 TU Munich  | 
| 4 | 4  | 
|
5  | 
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
 | 
6  | 
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
 | 
7  | 
where "f args" must be a first-order term without duplicate variables.  | 
| 0 | 8  | 
*)  | 
9  | 
||
| 5304 | 10  | 
signature SPLITTER_DATA =  | 
11  | 
sig  | 
|
| 
32177
 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 
wenzelm 
parents: 
32174 
diff
changeset
 | 
12  | 
val thy : theory  | 
| 5553 | 13  | 
val mk_eq : thm -> thm  | 
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
14  | 
val meta_eq_to_iff: thm (* "x == y ==> x = y" *)  | 
| 
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
15  | 
val iffD : thm (* "[| P = Q; Q |] ==> P" *)  | 
| 
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
16  | 
val disjE : thm (* "[| P | Q; P ==> R; Q ==> R |] ==> R" *)  | 
| 
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
17  | 
val conjE : thm (* "[| P & Q; [| P; Q |] ==> R |] ==> R" *)  | 
| 
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
18  | 
val exE : thm (* "[| EX x. P x; !!x. P x ==> Q |] ==> Q" *)  | 
| 
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
19  | 
val contrapos : thm (* "[| ~ Q; P ==> Q |] ==> ~ P" *)  | 
| 
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
20  | 
val contrapos2 : thm (* "[| Q; ~ P ==> ~ Q |] ==> P" *)  | 
| 
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
21  | 
val notnotD : thm (* "~ ~ P ==> P" *)  | 
| 5304 | 22  | 
end  | 
23  | 
||
24  | 
signature SPLITTER =  | 
|
25  | 
sig  | 
|
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
26  | 
(* somewhat more internal functions *)  | 
| 33242 | 27  | 
val cmap_of_split_thms: thm list -> (string * (typ * term * thm * typ * int) list) list  | 
28  | 
val split_posns: (string * (typ * term * thm * typ * int) list) list ->  | 
|
29  | 
theory -> typ list -> term -> (thm * (typ * typ * int list) list * int list * typ * term) list  | 
|
30  | 
(* first argument is a "cmap", returns a list of "split packs" *)  | 
|
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
31  | 
(* the "real" interface, providing a number of tactics *)  | 
| 5304 | 32  | 
val split_tac : thm list -> int -> tactic  | 
33  | 
val split_inside_tac: thm list -> int -> tactic  | 
|
34  | 
val split_asm_tac : thm list -> int -> tactic  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
45620 
diff
changeset
 | 
35  | 
val add_split: thm -> Proof.context -> Proof.context  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
45620 
diff
changeset
 | 
36  | 
val del_split: thm -> Proof.context -> Proof.context  | 
| 18728 | 37  | 
val split_add: attribute  | 
38  | 
val split_del: attribute  | 
|
| 30513 | 39  | 
val split_modifiers : Method.modifier parser list  | 
| 18708 | 40  | 
val setup: theory -> theory  | 
| 5304 | 41  | 
end;  | 
42  | 
||
| 
32177
 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 
wenzelm 
parents: 
32174 
diff
changeset
 | 
43  | 
functor Splitter(Data: SPLITTER_DATA): SPLITTER =  | 
| 17881 | 44  | 
struct  | 
| 5304 | 45  | 
|
| 18545 | 46  | 
val Const (const_not, _) $ _ =  | 
| 35625 | 47  | 
Object_Logic.drop_judgment Data.thy  | 
| 18545 | 48  | 
(#1 (Logic.dest_implies (Thm.prop_of Data.notnotD)));  | 
| 5304 | 49  | 
|
| 18545 | 50  | 
val Const (const_or , _) $ _ $ _ =  | 
| 35625 | 51  | 
Object_Logic.drop_judgment Data.thy  | 
| 18545 | 52  | 
(#1 (Logic.dest_implies (Thm.prop_of Data.disjE)));  | 
53  | 
||
| 35625 | 54  | 
val const_Trueprop = Object_Logic.judgment_name Data.thy;  | 
| 18545 | 55  | 
|
| 
1721
 
445654b6cb95
Rewrote mk_cntxt_splitthm. Added function mk_case_split_inside_tac.
 
berghofe 
parents: 
1686 
diff
changeset
 | 
56  | 
|
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
57  | 
fun split_format_err () = error "Wrong format for split rule";  | 
| 
4668
 
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  | 
|
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
65  | 
fun cmap_of_split_thms thms =  | 
| 
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
66  | 
let  | 
| 
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
67  | 
val splits = map Data.mk_eq thms  | 
| 33242 | 68  | 
fun add_thm thm cmap =  | 
69  | 
(case concl_of thm of _ $ (t as _ $ lhs) $ _ =>  | 
|
70  | 
(case strip_comb lhs of (Const(a,aT),args) =>  | 
|
71  | 
let val info = (aT,lhs,thm,fastype_of t,length args)  | 
|
72  | 
in case AList.lookup (op =) cmap a of  | 
|
73  | 
SOME infos => AList.update (op =) (a, info::infos) cmap  | 
|
74  | 
| NONE => (a,[info])::cmap  | 
|
75  | 
end  | 
|
76  | 
| _ => split_format_err())  | 
|
77  | 
| _ => split_format_err())  | 
|
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
78  | 
in  | 
| 33242 | 79  | 
fold add_thm splits []  | 
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
80  | 
end;  | 
| 
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
81  | 
|
| 
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
82  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
83  | 
(* mk_case_split_tac *)  | 
| 
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
84  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
85  | 
|
| 5304 | 86  | 
fun mk_case_split_tac order =  | 
| 0 | 87  | 
let  | 
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  | 
Create lift-theorem "trlift" :  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
91  | 
|
| 
7672
 
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
 
berghofe 
parents: 
6130 
diff
changeset
 | 
92  | 
[| !!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
 | 
93  | 
|
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
94  | 
*************************************************************)  | 
| 5304 | 95  | 
|
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
96  | 
val meta_iffD = Data.meta_eq_to_iff RS Data.iffD; (* (P == Q) ==> Q ==> P *)  | 
| 
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
97  | 
|
| 22838 | 98  | 
val lift = Goal.prove_global Pure.thy ["P", "Q", "R"]  | 
| 24707 | 99  | 
[Syntax.read_prop_global Pure.thy "!!x :: 'b. Q(x) == R(x) :: 'c"]  | 
100  | 
(Syntax.read_prop_global Pure.thy "P(%x. Q(x)) == P(%x. R(x))")  | 
|
| 
28839
 
32d498cf7595
eliminated rewrite_tac/fold_tac, which are not well-formed tactics due to change of main conclusion;
 
wenzelm 
parents: 
26711 
diff
changeset
 | 
101  | 
  (fn {prems, ...} => rewrite_goals_tac prems THEN rtac reflexive_thm 1)
 | 
| 4 | 102  | 
|
| 0 | 103  | 
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
 | 
104  | 
val _ $ (P $ _) $ _ = concl_of trlift;  | 
| 0 | 105  | 
|
106  | 
||
| 17881 | 107  | 
(************************************************************************  | 
| 
1686
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
108  | 
Set up term for instantiation of P in the lift-theorem  | 
| 17881 | 109  | 
|
| 
1686
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
110  | 
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
 | 
111  | 
t : lefthand side of meta-equality in subgoal  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
112  | 
the lift theorem is applied to (see select)  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
113  | 
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
 | 
114  | 
T : type of body of P(...)  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
115  | 
maxi : maximum index of Vars  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
116  | 
*************************************************************************)  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
117  | 
|
| 
1030
 
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
 
nipkow 
parents: 
943 
diff
changeset
 | 
118  | 
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
 | 
119  | 
  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
 | 
120  | 
fun down [] t i = Bound 0  | 
| 
 
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
 
nipkow 
parents: 
943 
diff
changeset
 | 
121  | 
| down (p::ps) t i =  | 
| 
 
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
 
nipkow 
parents: 
943 
diff
changeset
 | 
122  | 
let val (h,ts) = strip_comb t  | 
| 33955 | 123  | 
val v1 = ListPair.map var (take p ts, i upto (i+p-1))  | 
124  | 
val u::us = drop p ts  | 
|
| 2266 | 125  | 
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
 | 
126  | 
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
 | 
127  | 
  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
 | 
128  | 
|
| 
1686
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
129  | 
|
| 17881 | 130  | 
(************************************************************************  | 
| 
1686
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
131  | 
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
 | 
132  | 
P(...) == rhs  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
133  | 
|
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
134  | 
t : lefthand side of meta-equality in subgoal  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
135  | 
the split theorem is applied to (see select)  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
136  | 
T : type of body of P(...)  | 
| 4232 | 137  | 
tt : the term Const(key,..) $ ...  | 
| 
1686
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
138  | 
*************************************************************************)  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
139  | 
|
| 4232 | 140  | 
fun mk_cntxt_splitthm t tt T =  | 
141  | 
let fun repl lev t =  | 
|
| 52131 | 142  | 
if Envir.aeconv(incr_boundvars lev tt, t) then Bound lev  | 
| 4232 | 143  | 
else case t of  | 
144  | 
(Abs (v, T2, t)) => Abs (v, T2, repl (lev+1) t)  | 
|
145  | 
| (Bound i) => Bound (if i>=lev then i+1 else i)  | 
|
146  | 
| (t1 $ t2) => (repl lev t1) $ (repl lev t2)  | 
|
147  | 
| t => t  | 
|
148  | 
  in Abs("", T, repl 0 t) end;
 | 
|
| 
1686
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
149  | 
|
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
150  | 
|
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
151  | 
(* add all loose bound variables in t to list is *)  | 
| 33245 | 152  | 
fun add_lbnos t is = add_loose_bnos (t, 0, is);  | 
| 
1030
 
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
 
nipkow 
parents: 
943 
diff
changeset
 | 
153  | 
|
| 
7672
 
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
 
berghofe 
parents: 
6130 
diff
changeset
 | 
154  | 
(* check if the innermost abstraction that needs to be removed  | 
| 1064 | 155  | 
has a body of type T; otherwise the expansion thm will fail later on  | 
156  | 
*)  | 
|
| 33029 | 157  | 
fun type_test (T, lbnos, apsns) =  | 
| 42364 | 158  | 
let val (_, U: typ, _) = nth apsns (foldl1 Int.min lbnos)  | 
| 33029 | 159  | 
in T = U end;  | 
| 0 | 160  | 
|
| 
1686
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
161  | 
(*************************************************************************  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
162  | 
Create a "split_pack".  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
163  | 
|
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
164  | 
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
 | 
165  | 
is of the form  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
166  | 
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
 | 
167  | 
T : type of P(...)  | 
| 
7672
 
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
 
berghofe 
parents: 
6130 
diff
changeset
 | 
168  | 
T' : type of term to be scanned  | 
| 
1686
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
169  | 
n : number of arguments expected by Const(key,...)  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
170  | 
ts : list of arguments actually found  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
171  | 
apsns : list of tuples of the form (T,U,pos), one tuple for each  | 
| 17881 | 172  | 
abstraction that is encountered on the way to the position where  | 
| 
1686
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
173  | 
Const(key, ...) $ ... occurs, where  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
174  | 
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
 | 
175  | 
U : type of the abstraction's body  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
176  | 
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
 | 
177  | 
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
 | 
178  | 
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
 | 
179  | 
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
 | 
180  | 
|
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
181  | 
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
 | 
182  | 
(thm, apsns, pos, TB, tt)  | 
| 
1686
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
183  | 
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
 | 
184  | 
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
 | 
185  | 
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
 | 
186  | 
lifting is required before applying the split-theorem.  | 
| 17881 | 187  | 
******************************************************************************)  | 
| 
1686
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
188  | 
|
| 20664 | 189  | 
fun mk_split_pack (thm, T: typ, T', n, ts, apsns, pos, TB, t) =  | 
| 1064 | 190  | 
if n > length ts then []  | 
191  | 
else let val lev = length apsns  | 
|
| 33955 | 192  | 
val lbnos = fold add_lbnos (take n ts) []  | 
| 33317 | 193  | 
val flbnos = filter (fn i => i < lev) lbnos  | 
| 4232 | 194  | 
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
 | 
195  | 
in if null flbnos then  | 
| 
 
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
 
berghofe 
parents: 
6130 
diff
changeset
 | 
196  | 
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
 | 
197  | 
else if type_test(T,flbnos,apsns) then [(thm, rev apsns,pos,TB,tt)]  | 
| 2143 | 198  | 
else []  | 
| 1064 | 199  | 
end;  | 
| 0 | 200  | 
|
| 
1686
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
201  | 
|
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
202  | 
(****************************************************************************  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
203  | 
Recursively scans term for occurences of Const(key,...) $ ...  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
204  | 
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
 | 
205  | 
|
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
206  | 
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
 | 
207  | 
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
 | 
208  | 
key : the theorem's key constant ( Const(key,...) $ ... )  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
209  | 
thm : the theorem itself  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
210  | 
T : type of P( Const(key,...) $ ... )  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
211  | 
n : number of arguments expected by Const(key,...)  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
212  | 
Ts : types of parameters  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
213  | 
t : the term to be scanned  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
214  | 
******************************************************************************)  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
215  | 
|
| 
13157
 
4a4599f78f18
allowed more general split rules to cope with div/mod 2
 
nipkow 
parents: 
10821 
diff
changeset
 | 
216  | 
(* Simplified first-order matching;  | 
| 
 
4a4599f78f18
allowed more general split rules to cope with div/mod 2
 
nipkow 
parents: 
10821 
diff
changeset
 | 
217  | 
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
 | 
218  | 
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
 | 
219  | 
*)  | 
| 
 
4a4599f78f18
allowed more general split rules to cope with div/mod 2
 
nipkow 
parents: 
10821 
diff
changeset
 | 
220  | 
local  | 
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
221  | 
exception MATCH  | 
| 
13157
 
4a4599f78f18
allowed more general split rules to cope with div/mod 2
 
nipkow 
parents: 
10821 
diff
changeset
 | 
222  | 
in  | 
| 
42367
 
577d85fb5862
more direct Thm.cprem_of (with exception THM instead of Subscript);
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
223  | 
fun typ_match thy (tyenv, TU) = Sign.typ_match thy TU tyenv  | 
| 
 
577d85fb5862
more direct Thm.cprem_of (with exception THM instead of Subscript);
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
224  | 
handle Type.TYPE_MATCH => raise MATCH;  | 
| 33242 | 225  | 
|
| 
42367
 
577d85fb5862
more direct Thm.cprem_of (with exception THM instead of Subscript);
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
226  | 
fun fomatch thy args =  | 
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
227  | 
let  | 
| 
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
228  | 
fun mtch tyinsts = fn  | 
| 
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
229  | 
(Ts, Var(_,T), t) =>  | 
| 
42367
 
577d85fb5862
more direct Thm.cprem_of (with exception THM instead of Subscript);
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
230  | 
typ_match thy (tyinsts, (T, fastype_of1(Ts,t)))  | 
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
231  | 
| (_, Free (a,T), Free (b,U)) =>  | 
| 
42367
 
577d85fb5862
more direct Thm.cprem_of (with exception THM instead of Subscript);
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
232  | 
if a=b then typ_match thy (tyinsts,(T,U)) else raise MATCH  | 
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
233  | 
| (_, Const (a,T), Const (b,U)) =>  | 
| 
42367
 
577d85fb5862
more direct Thm.cprem_of (with exception THM instead of Subscript);
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
234  | 
if a=b then typ_match thy (tyinsts,(T,U)) else raise MATCH  | 
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
235  | 
| (_, Bound i, Bound j) =>  | 
| 
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
236  | 
if i=j then tyinsts else raise MATCH  | 
| 
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
237  | 
| (Ts, Abs(_,T,t), Abs(_,U,u)) =>  | 
| 
42367
 
577d85fb5862
more direct Thm.cprem_of (with exception THM instead of Subscript);
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
238  | 
mtch (typ_match thy (tyinsts,(T,U))) (U::Ts,t,u)  | 
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
239  | 
| (Ts, f$t, g$u) =>  | 
| 
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
240  | 
mtch (mtch tyinsts (Ts,f,g)) (Ts, t, u)  | 
| 
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
241  | 
| _ => raise MATCH  | 
| 
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
242  | 
in (mtch Vartab.empty args; true) handle MATCH => false end;  | 
| 33242 | 243  | 
end;  | 
| 
13157
 
4a4599f78f18
allowed more general split rules to cope with div/mod 2
 
nipkow 
parents: 
10821 
diff
changeset
 | 
244  | 
|
| 
42367
 
577d85fb5862
more direct Thm.cprem_of (with exception THM instead of Subscript);
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
245  | 
fun split_posns (cmap : (string * (typ * term * thm * typ * int) list) list) thy Ts t =  | 
| 
6130
 
30b84ad2131d
Fixed old bug: selection of constant to be split should depend not just on
 
nipkow 
parents: 
5553 
diff
changeset
 | 
246  | 
let  | 
| 
7672
 
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
 
berghofe 
parents: 
6130 
diff
changeset
 | 
247  | 
val T' = fastype_of1 (Ts, t);  | 
| 
 
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
 
berghofe 
parents: 
6130 
diff
changeset
 | 
248  | 
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
 | 
249  | 
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
 | 
250  | 
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
 | 
251  | 
| 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
 | 
252  | 
let  | 
| 
7672
 
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
 
berghofe 
parents: 
6130 
diff
changeset
 | 
253  | 
val (h, ts) = strip_comb t  | 
| 33245 | 254  | 
fun iter t (i, a) = (i+1, (posns Ts (i::pos) apsns t) @ a);  | 
255  | 
val a =  | 
|
256  | 
case h of  | 
|
257  | 
Const(c, cT) =>  | 
|
258  | 
let fun find [] = []  | 
|
259  | 
| find ((gcT, pat, thm, T, n)::tups) =  | 
|
| 
42367
 
577d85fb5862
more direct Thm.cprem_of (with exception THM instead of Subscript);
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
260  | 
let val t2 = list_comb (h, take n ts) in  | 
| 
 
577d85fb5862
more direct Thm.cprem_of (with exception THM instead of Subscript);
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
261  | 
if Sign.typ_instance thy (cT, gcT) andalso fomatch thy (Ts, pat, t2)  | 
| 
 
577d85fb5862
more direct Thm.cprem_of (with exception THM instead of Subscript);
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
262  | 
then mk_split_pack(thm,T,T',n,ts,apsns,pos,type_of1(Ts,t2),t2)  | 
| 
 
577d85fb5862
more direct Thm.cprem_of (with exception THM instead of Subscript);
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
263  | 
else find tups  | 
| 33245 | 264  | 
end  | 
265  | 
in find (these (AList.lookup (op =) cmap c)) end  | 
|
266  | 
| _ => []  | 
|
267  | 
in snd (fold iter ts (0, a)) end  | 
|
| 
1030
 
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
 
nipkow 
parents: 
943 
diff
changeset
 | 
268  | 
in posns Ts [] [] t end;  | 
| 0 | 269  | 
|
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
270  | 
fun shorter ((_,ps,pos,_,_), (_,qs,qos,_,_)) =  | 
| 4519 | 271  | 
prod_ord (int_ord o pairself length) (order o pairself length)  | 
272  | 
((ps, pos), (qs, qos));  | 
|
273  | 
||
| 
1686
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
274  | 
|
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
275  | 
(************************************************************  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
276  | 
call split_posns with appropriate parameters  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
277  | 
*************************************************************)  | 
| 0 | 278  | 
|
| 
1030
 
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
 
nipkow 
parents: 
943 
diff
changeset
 | 
279  | 
fun select cmap state i =  | 
| 
42367
 
577d85fb5862
more direct Thm.cprem_of (with exception THM instead of Subscript);
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
280  | 
let  | 
| 
 
577d85fb5862
more direct Thm.cprem_of (with exception THM instead of Subscript);
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
281  | 
val thy = Thm.theory_of_thm state  | 
| 
 
577d85fb5862
more direct Thm.cprem_of (with exception THM instead of Subscript);
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
282  | 
val goal = term_of (Thm.cprem_of state i);  | 
| 
 
577d85fb5862
more direct Thm.cprem_of (with exception THM instead of Subscript);
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
283  | 
val Ts = rev (map #2 (Logic.strip_params goal));  | 
| 
 
577d85fb5862
more direct Thm.cprem_of (with exception THM instead of Subscript);
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
284  | 
val _ $ t $ _ = Logic.strip_assums_concl goal;  | 
| 
 
577d85fb5862
more direct Thm.cprem_of (with exception THM instead of Subscript);
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
285  | 
in (Ts, t, sort shorter (split_posns cmap thy Ts t)) end;  | 
| 
1030
 
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
 
nipkow 
parents: 
943 
diff
changeset
 | 
286  | 
|
| 
42367
 
577d85fb5862
more direct Thm.cprem_of (with exception THM instead of Subscript);
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
287  | 
fun exported_split_posns cmap thy Ts t =  | 
| 
 
577d85fb5862
more direct Thm.cprem_of (with exception THM instead of Subscript);
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
288  | 
sort shorter (split_posns cmap thy Ts t);  | 
| 
1686
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
289  | 
|
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
290  | 
(*************************************************************  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
291  | 
instantiate lift theorem  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
292  | 
|
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
293  | 
if t is of the form  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
294  | 
... ( Const(...,...) $ Abs( .... ) ) ...  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
295  | 
then  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
296  | 
P = %a. ... ( Const(...,...) $ a ) ...  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
297  | 
where a has type T --> U  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
298  | 
|
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
299  | 
Ts : types of parameters  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
300  | 
t : lefthand side of meta-equality in subgoal  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
301  | 
the split theorem is applied to (see cmap)  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
302  | 
T,U,pos : see mk_split_pack  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
303  | 
state : current proof state  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
304  | 
lift : the lift theorem  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
305  | 
i : no. of subgoal  | 
| 
 
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  | 
|
| 
7672
 
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
 
berghofe 
parents: 
6130 
diff
changeset
 | 
308  | 
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
 | 
309  | 
let  | 
| 22578 | 310  | 
val cert = cterm_of (Thm.theory_of_thm state);  | 
| 22596 | 311  | 
val cntxt = mk_cntxt Ts t pos (T --> U) (Thm.maxidx_of trlift);  | 
| 
7672
 
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
 
berghofe 
parents: 
6130 
diff
changeset
 | 
312  | 
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
 | 
313  | 
end;  | 
| 0 | 314  | 
|
315  | 
||
| 
1686
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
316  | 
(*************************************************************  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
317  | 
instantiate split theorem  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
318  | 
|
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
319  | 
Ts : types of parameters  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
320  | 
t : lefthand side of meta-equality in subgoal  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
321  | 
the split theorem is applied to (see cmap)  | 
| 4232 | 322  | 
tt : the term Const(key,..) $ ...  | 
| 
1686
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
323  | 
thm : the split theorem  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
324  | 
TB : type of body of P(...)  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
325  | 
state : current proof state  | 
| 4232 | 326  | 
i : number of subgoal  | 
| 
1686
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
327  | 
**************************************************************)  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
328  | 
|
| 4232 | 329  | 
fun inst_split Ts t tt thm TB state i =  | 
| 17881 | 330  | 
let  | 
| 18145 | 331  | 
val thm' = Thm.lift_rule (Thm.cprem_of state i) thm;  | 
| 
7672
 
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
 
berghofe 
parents: 
6130 
diff
changeset
 | 
332  | 
val (P, _) = strip_comb (fst (Logic.dest_equals  | 
| 22596 | 333  | 
(Logic.strip_assums_concl (Thm.prop_of thm'))));  | 
| 22578 | 334  | 
val cert = cterm_of (Thm.theory_of_thm state);  | 
| 
7672
 
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
 
berghofe 
parents: 
6130 
diff
changeset
 | 
335  | 
val cntxt = mk_cntxt_splitthm t tt TB;  | 
| 33245 | 336  | 
    val abss = fold (fn T => fn t => Abs ("", T, t));
 | 
337  | 
in cterm_instantiate [(cert P, cert (abss Ts cntxt))] thm'  | 
|
| 4232 | 338  | 
end;  | 
| 
1686
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
339  | 
|
| 
7672
 
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
 
berghofe 
parents: 
6130 
diff
changeset
 | 
340  | 
|
| 
1686
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
341  | 
(*****************************************************************************  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
342  | 
The split-tactic  | 
| 17881 | 343  | 
|
| 
1686
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
344  | 
splits : list of split-theorems to be tried  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
345  | 
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
 | 
346  | 
*****************************************************************************)  | 
| 
 
c67d543bc395
Added functions mk_cntxt_splitthm and inst_split which instantiate
 
berghofe 
parents: 
1064 
diff
changeset
 | 
347  | 
|
| 0 | 348  | 
fun split_tac [] i = no_tac  | 
349  | 
| split_tac splits i =  | 
|
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
350  | 
let val cmap = cmap_of_split_thms splits  | 
| 
7672
 
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
 
berghofe 
parents: 
6130 
diff
changeset
 | 
351  | 
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
 | 
352  | 
fun lift_split_tac state =  | 
| 
 
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
 
berghofe 
parents: 
6130 
diff
changeset
 | 
353  | 
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
 | 
354  | 
in case splits of  | 
| 
7672
 
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
 
berghofe 
parents: 
6130 
diff
changeset
 | 
355  | 
[] => no_tac state  | 
| 
 
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
 
berghofe 
parents: 
6130 
diff
changeset
 | 
356  | 
| (thm, apsns, pos, TB, tt)::_ =>  | 
| 
1030
 
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
 
nipkow 
parents: 
943 
diff
changeset
 | 
357  | 
(case apsns of  | 
| 
7672
 
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
 
berghofe 
parents: 
6130 
diff
changeset
 | 
358  | 
[] => 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
 | 
359  | 
| 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
 | 
360  | 
rtac reflexive_thm (i+1),  | 
| 
 
c092e67d12f8
- Fixed bug in mk_split_pack which caused application of expansion theorem
 
berghofe 
parents: 
6130 
diff
changeset
 | 
361  | 
lift_split_tac] state)  | 
| 
1030
 
1d8fa2fc4b9c
Completely rewrote split_tac. The old one failed in strange circumstances.
 
nipkow 
parents: 
943 
diff
changeset
 | 
362  | 
end  | 
| 17881 | 363  | 
in COND (has_fewer_prems i) no_tac  | 
| 5304 | 364  | 
(rtac meta_iffD i THEN lift_split_tac)  | 
| 0 | 365  | 
end;  | 
366  | 
||
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
367  | 
in (split_tac, exported_split_posns) end; (* mk_case_split_tac *)  | 
| 
1721
 
445654b6cb95
Rewrote mk_cntxt_splitthm. Added function mk_case_split_inside_tac.
 
berghofe 
parents: 
1686 
diff
changeset
 | 
368  | 
|
| 5304 | 369  | 
|
| 33242 | 370  | 
val (split_tac, split_posns) = mk_case_split_tac int_ord;  | 
| 4189 | 371  | 
|
| 33242 | 372  | 
val (split_inside_tac, _) = mk_case_split_tac (rev_order o int_ord);  | 
| 5304 | 373  | 
|
| 4189 | 374  | 
|
375  | 
(*****************************************************************************  | 
|
376  | 
The split-tactic for premises  | 
|
| 17881 | 377  | 
|
| 4189 | 378  | 
splits : list of split-theorems to be tried  | 
| 5304 | 379  | 
****************************************************************************)  | 
| 33242 | 380  | 
fun split_asm_tac [] = K no_tac  | 
| 17881 | 381  | 
| split_asm_tac splits =  | 
| 5304 | 382  | 
|
| 
13855
 
644692eca537
addsplits / delsplits no longer ignore type of constant.
 
berghofe 
parents: 
13157 
diff
changeset
 | 
383  | 
let val cname_list = map (fst o fst o split_thm_info) splits;  | 
| 17881 | 384  | 
fun tac (t,i) =  | 
| 20664 | 385  | 
let val n = find_index (exists_Const (member (op =) cname_list o #1))  | 
| 17881 | 386  | 
(Logic.strip_assums_hyp t);  | 
| 18545 | 387  | 
              fun first_prem_is_disj (Const ("==>", _) $ (Const (c, _)
 | 
388  | 
$ (Const (s, _) $ _ $ _ )) $ _ ) = c = const_Trueprop andalso s = const_or  | 
|
| 17881 | 389  | 
              |   first_prem_is_disj (Const("all",_)$Abs(_,_,t)) =
 | 
390  | 
first_prem_is_disj t  | 
|
391  | 
| first_prem_is_disj _ = false;  | 
|
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
392  | 
(* does not work properly if the split variable is bound by a quantifier *)  | 
| 17881 | 393  | 
fun flat_prems_tac i = SUBGOAL (fn (t,i) =>  | 
394  | 
(if first_prem_is_disj t  | 
|
395  | 
then EVERY[etac Data.disjE i,rotate_tac ~1 i,  | 
|
396  | 
rotate_tac ~1 (i+1),  | 
|
397  | 
flat_prems_tac (i+1)]  | 
|
398  | 
else all_tac)  | 
|
399  | 
THEN REPEAT (eresolve_tac [Data.conjE,Data.exE] i)  | 
|
400  | 
THEN REPEAT (dresolve_tac [Data.notnotD] i)) i;  | 
|
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
401  | 
in if n<0 then no_tac else (DETERM (EVERY'  | 
| 17881 | 402  | 
[rotate_tac n, etac Data.contrapos2,  | 
403  | 
split_tac splits,  | 
|
404  | 
rotate_tac ~1, etac Data.contrapos, rotate_tac ~1,  | 
|
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18988 
diff
changeset
 | 
405  | 
flat_prems_tac] i))  | 
| 17881 | 406  | 
end;  | 
| 4189 | 407  | 
in SUBGOAL tac  | 
408  | 
end;  | 
|
409  | 
||
| 10652 | 410  | 
fun gen_split_tac [] = K no_tac  | 
411  | 
| gen_split_tac (split::splits) =  | 
|
412  | 
let val (_,asm) = split_thm_info split  | 
|
413  | 
in (if asm then split_asm_tac else split_tac) [split] ORELSE'  | 
|
414  | 
gen_split_tac splits  | 
|
415  | 
end;  | 
|
| 8468 | 416  | 
|
| 18688 | 417  | 
|
| 8468 | 418  | 
(** declare split rules **)  | 
419  | 
||
| 
45620
 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 
wenzelm 
parents: 
42367 
diff
changeset
 | 
420  | 
(* add_split / del_split *)  | 
| 8468 | 421  | 
|
| 33242 | 422  | 
fun string_of_typ (Type (s, Ts)) =  | 
423  | 
      (if null Ts then "" else enclose "(" ")" (commas (map string_of_typ Ts))) ^ s
 | 
|
| 
13859
 
adf68d9e5dec
split_name no longer uses Sign.string_of_typ to encode types, since
 
berghofe 
parents: 
13855 
diff
changeset
 | 
424  | 
| string_of_typ _ = "_";  | 
| 
 
adf68d9e5dec
split_name no longer uses Sign.string_of_typ to encode types, since
 
berghofe 
parents: 
13855 
diff
changeset
 | 
425  | 
|
| 17881 | 426  | 
fun split_name (name, T) asm = "split " ^  | 
| 
13859
 
adf68d9e5dec
split_name no longer uses Sign.string_of_typ to encode types, since
 
berghofe 
parents: 
13855 
diff
changeset
 | 
427  | 
(if asm then "asm " else "") ^ name ^ " :: " ^ string_of_typ T;  | 
| 4189 | 428  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
45620 
diff
changeset
 | 
429  | 
fun add_split split ctxt =  | 
| 33242 | 430  | 
let  | 
| 
45620
 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 
wenzelm 
parents: 
42367 
diff
changeset
 | 
431  | 
val (name, asm) = split_thm_info split  | 
| 
 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 
wenzelm 
parents: 
42367 
diff
changeset
 | 
432  | 
val tac = (if asm then split_asm_tac else split_tac) [split]  | 
| 52037 | 433  | 
in Simplifier.addloop (ctxt, (split_name name asm, K tac)) end;  | 
| 
1721
 
445654b6cb95
Rewrote mk_cntxt_splitthm. Added function mk_case_split_inside_tac.
 
berghofe 
parents: 
1686 
diff
changeset
 | 
434  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
45620 
diff
changeset
 | 
435  | 
fun del_split split ctxt =  | 
| 
45620
 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 
wenzelm 
parents: 
42367 
diff
changeset
 | 
436  | 
let val (name, asm) = split_thm_info split  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
45620 
diff
changeset
 | 
437  | 
in Simplifier.delloop (ctxt, split_name name asm) end;  | 
| 
1721
 
445654b6cb95
Rewrote mk_cntxt_splitthm. Added function mk_case_split_inside_tac.
 
berghofe 
parents: 
1686 
diff
changeset
 | 
438  | 
|
| 8468 | 439  | 
|
440  | 
(* attributes *)  | 
|
441  | 
||
442  | 
val splitN = "split";  | 
|
443  | 
||
| 
45620
 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 
wenzelm 
parents: 
42367 
diff
changeset
 | 
444  | 
val split_add = Simplifier.attrib add_split;  | 
| 
 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 
wenzelm 
parents: 
42367 
diff
changeset
 | 
445  | 
val split_del = Simplifier.attrib del_split;  | 
| 8634 | 446  | 
|
447  | 
||
| 9703 | 448  | 
(* methods *)  | 
| 8468 | 449  | 
|
450  | 
val split_modifiers =  | 
|
| 18728 | 451  | 
[Args.$$$ splitN -- Args.colon >> K ((I, split_add): Method.modifier),  | 
452  | 
Args.$$$ splitN -- Args.add -- Args.colon >> K (I, split_add),  | 
|
453  | 
Args.$$$ splitN -- Args.del -- Args.colon >> K (I, split_del)];  | 
|
| 8468 | 454  | 
|
455  | 
||
| 18688 | 456  | 
(* theory setup *)  | 
| 8468 | 457  | 
|
| 9703 | 458  | 
val setup =  | 
| 33242 | 459  | 
  Attrib.setup @{binding split}
 | 
460  | 
(Attrib.add_del split_add split_del) "declare case split rule" #>  | 
|
| 
30722
 
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
 
wenzelm 
parents: 
30609 
diff
changeset
 | 
461  | 
  Method.setup @{binding split}
 | 
| 
 
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
 
wenzelm 
parents: 
30609 
diff
changeset
 | 
462  | 
(Attrib.thms >> (fn ths => K (SIMPLE_METHOD' (CHANGED_PROP o gen_split_tac ths))))  | 
| 
 
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
 
wenzelm 
parents: 
30609 
diff
changeset
 | 
463  | 
"apply case split rule";  | 
| 4189 | 464  | 
|
| 
1721
 
445654b6cb95
Rewrote mk_cntxt_splitthm. Added function mk_case_split_inside_tac.
 
berghofe 
parents: 
1686 
diff
changeset
 | 
465  | 
end;  |