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