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