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