src/Provers/splitter.ML
author wenzelm
Sun, 29 Nov 1998 13:14:45 +0100
changeset 5986 6ebbc9e7cc20
parent 5553 ae42b36a50c2
child 6130 30b84ad2131d
permissions -rw-r--r--
added oct_char;
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.
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
     9
infix 4 addsplits delsplits;
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    10
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    11
signature SPLITTER_DATA =
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    12
sig
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    13
  structure Simplifier: SIMPLIFIER
5553
ae42b36a50c2 renamed mk_meta_eq to mk_eq
oheimb
parents: 5437
diff changeset
    14
  val mk_eq         : thm -> thm
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    15
  val meta_eq_to_iff: thm (* "x == y ==> x = y"                    *)
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    16
  val iffD          : thm (* "[| P = Q; Q |] ==> P"                *)
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    17
  val disjE         : thm (* "[| P | Q; P ==> R; Q ==> R |] ==> R" *)
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    18
  val conjE         : thm (* "[| P & Q; [| P; Q |] ==> R |] ==> R" *)
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    19
  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
    20
  val contrapos     : thm (* "[| ~ Q; P ==> Q |] ==> ~ P"          *)
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    21
  val contrapos2    : thm (* "[| Q; ~ P ==> ~ Q |] ==> P"          *)
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    22
  val notnotD       : thm (* "~ ~ P ==> P"                         *)
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    23
end
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    24
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    25
signature SPLITTER =
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    26
sig
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    27
  type simpset
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
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    35
end;
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    36
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    37
functor SplitterFun(Data: SPLITTER_DATA): SPLITTER =
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    38
struct 
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    39
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    40
type simpset = Data.Simplifier.simpset;
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    41
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    42
val Const ("==>", _) $ (Const ("Trueprop", _) $
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    43
         (Const (const_not, _) $ _    )) $ _ = #prop (rep_thm(Data.notnotD));
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    44
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    45
val Const ("==>", _) $ (Const ("Trueprop", _) $
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    46
         (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
    47
4668
131989b78417 Little reorganization. Loop tactics have names now.
nipkow
parents: 4519
diff changeset
    48
fun split_format_err() = error("Wrong format for split rule");
131989b78417 Little reorganization. Loop tactics have names now.
nipkow
parents: 4519
diff changeset
    49
5553
ae42b36a50c2 renamed mk_meta_eq to mk_eq
oheimb
parents: 5437
diff changeset
    50
fun split_thm_info thm = case concl_of (Data.mk_eq thm) of
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    51
     Const("==", _)$(Var _$t)$c =>
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    52
        (case strip_comb t of
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    53
           (Const(a,_),_) => (a,case c of (Const(s,_)$_)=>s=const_not|_=> false)
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    54
         | _              => split_format_err())
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    55
   | _ => split_format_err();
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    56
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    57
fun mk_case_split_tac order =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    58
let
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    59
1686
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
    60
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
    61
(************************************************************
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
    62
   Create lift-theorem "trlift" :
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
   [| !! x. Q(x)==R(x) ; P(R) == C |] ==> P(Q)==C
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
    65
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
    66
*************************************************************)
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    67
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    68
val meta_iffD = Data.meta_eq_to_iff RS Data.iffD;
943
8477483f663f Replaced read by read_cterm.
nipkow
parents: 927
diff changeset
    69
val lift =
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
    70
  let val ct = read_cterm (#sign(rep_thm Data.iffD))
943
8477483f663f Replaced read by read_cterm.
nipkow
parents: 927
diff changeset
    71
           ("[| !!x::'b::logic. Q(x) == R(x) |] ==> \
3835
9a5a4e123859 fixed dots;
wenzelm
parents: 3537
diff changeset
    72
            \P(%x. Q(x)) == P(%x. R(x))::'a::logic",propT)
943
8477483f663f Replaced read by read_cterm.
nipkow
parents: 927
diff changeset
    73
  in prove_goalw_cterm [] ct
8477483f663f Replaced read by read_cterm.
nipkow
parents: 927
diff changeset
    74
     (fn [prem] => [rewtac prem, rtac reflexive_thm 1])
8477483f663f Replaced read by read_cterm.
nipkow
parents: 927
diff changeset
    75
  end;
4
2695ba9b40f7 added header
nipkow
parents: 0
diff changeset
    76
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    77
val trlift = lift RS transitive_thm;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    78
val _ $ (Var(P,PT)$_) $ _ = concl_of trlift;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    79
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    80
1686
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
    81
(************************************************************************ 
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
    82
   Set up term for instantiation of P in the lift-theorem
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
    83
   
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
    84
   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
    85
   t     : lefthand side of meta-equality in subgoal
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
    86
           the lift theorem is applied to (see select)
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
    87
   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
    88
   T     : type of body of P(...)
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
    89
   maxi  : maximum index of Vars
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
    90
*************************************************************************)
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
    91
1030
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    92
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
    93
  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
    94
      fun down [] t i = Bound 0
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    95
        | down (p::ps) t i =
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    96
            let val (h,ts) = strip_comb t
2266
82aef6857c5b Replaced map...~~ by ListPair.map
paulson
parents: 2143
diff changeset
    97
                val v1 = ListPair.map var (take(p,ts), i upto (i+p-1))
1030
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    98
                val u::us = drop(p,ts)
2266
82aef6857c5b Replaced map...~~ by ListPair.map
paulson
parents: 2143
diff changeset
    99
                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
   100
      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
   101
  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
   102
1686
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   103
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   104
(************************************************************************ 
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   105
   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
   106
   P(...) == rhs
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   107
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   108
   t     : lefthand side of meta-equality in subgoal
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   109
           the split theorem is applied to (see select)
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   110
   T     : type of body of P(...)
4232
29f3875596ad Tuned function mk_cntxt_splitthm.
berghofe
parents: 4202
diff changeset
   111
   tt    : the term  Const(key,..) $ ...
1686
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   112
*************************************************************************)
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   113
4232
29f3875596ad Tuned function mk_cntxt_splitthm.
berghofe
parents: 4202
diff changeset
   114
fun mk_cntxt_splitthm t tt T =
29f3875596ad Tuned function mk_cntxt_splitthm.
berghofe
parents: 4202
diff changeset
   115
  let fun repl lev t =
29f3875596ad Tuned function mk_cntxt_splitthm.
berghofe
parents: 4202
diff changeset
   116
    if incr_boundvars lev tt = t then Bound lev
29f3875596ad Tuned function mk_cntxt_splitthm.
berghofe
parents: 4202
diff changeset
   117
    else case t of
29f3875596ad Tuned function mk_cntxt_splitthm.
berghofe
parents: 4202
diff changeset
   118
        (Abs (v, T2, t)) => Abs (v, T2, repl (lev+1) t)
29f3875596ad Tuned function mk_cntxt_splitthm.
berghofe
parents: 4202
diff changeset
   119
      | (Bound i) => Bound (if i>=lev then i+1 else i)
29f3875596ad Tuned function mk_cntxt_splitthm.
berghofe
parents: 4202
diff changeset
   120
      | (t1 $ t2) => (repl lev t1) $ (repl lev t2)
29f3875596ad Tuned function mk_cntxt_splitthm.
berghofe
parents: 4202
diff changeset
   121
      | t => t
29f3875596ad Tuned function mk_cntxt_splitthm.
berghofe
parents: 4202
diff changeset
   122
  in Abs("", T, repl 0 t) end;
1686
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   123
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   124
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   125
(* 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
   126
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
   127
1064
5d6fb2c938e0 Fixed bug.
nipkow
parents: 1030
diff changeset
   128
(* check if the innermost quantifier that needs to be removed
5d6fb2c938e0 Fixed bug.
nipkow
parents: 1030
diff changeset
   129
   has a body of type T; otherwise the expansion thm will fail later on
5d6fb2c938e0 Fixed bug.
nipkow
parents: 1030
diff changeset
   130
*)
5d6fb2c938e0 Fixed bug.
nipkow
parents: 1030
diff changeset
   131
fun type_test(T,lbnos,apsns) =
2143
093bbe6d333b Replaced min by Int.min
paulson
parents: 1721
diff changeset
   132
  let val (_,U,_) = nth_elem(foldl Int.min (hd lbnos, tl lbnos), apsns)
1064
5d6fb2c938e0 Fixed bug.
nipkow
parents: 1030
diff changeset
   133
  in T=U end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   134
1686
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   135
(*************************************************************************
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   136
   Create a "split_pack".
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   137
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   138
   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
   139
           is of the form
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   140
           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
   141
   T     : type of P(...)
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   142
   n     : number of arguments expected by Const(key,...)
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   143
   ts    : list of arguments actually found
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   144
   apsns : list of tuples of the form (T,U,pos), one tuple for each
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   145
           abstraction that is encountered on the way to the position where 
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   146
           Const(key, ...) $ ...  occurs, where
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   147
           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
   148
           U   : type of the abstraction's body
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   149
           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
   150
   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
   151
   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
   152
   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
   153
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   154
   A split pack is a tuple of the form
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   155
   (thm, apsns, pos, TB)
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   156
   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
   157
          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
   158
          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
   159
          lifting is required before applying the split-theorem.
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   160
******************************************************************************) 
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   161
1721
445654b6cb95 Rewrote mk_cntxt_splitthm. Added function mk_case_split_inside_tac.
berghofe
parents: 1686
diff changeset
   162
fun mk_split_pack(thm,T,n,ts,apsns,pos,TB,t) =
1064
5d6fb2c938e0 Fixed bug.
nipkow
parents: 1030
diff changeset
   163
  if n > length ts then []
5d6fb2c938e0 Fixed bug.
nipkow
parents: 1030
diff changeset
   164
  else let val lev = length apsns
1030
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   165
           val lbnos = foldl add_lbnos ([],take(n,ts))
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   166
           val flbnos = filter (fn i => i < lev) lbnos
4232
29f3875596ad Tuned function mk_cntxt_splitthm.
berghofe
parents: 4202
diff changeset
   167
           val tt = incr_boundvars (~lev) t
1721
445654b6cb95 Rewrote mk_cntxt_splitthm. Added function mk_case_split_inside_tac.
berghofe
parents: 1686
diff changeset
   168
       in if null flbnos then [(thm,[],pos,TB,tt)]
2143
093bbe6d333b Replaced min by Int.min
paulson
parents: 1721
diff changeset
   169
          else if type_test(T,flbnos,apsns) then [(thm, rev apsns,pos,TB,tt)] 
093bbe6d333b Replaced min by Int.min
paulson
parents: 1721
diff changeset
   170
               else []
1064
5d6fb2c938e0 Fixed bug.
nipkow
parents: 1030
diff changeset
   171
       end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   172
1686
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   173
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   174
(****************************************************************************
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   175
   Recursively scans term for occurences of Const(key,...) $ ...
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   176
   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
   177
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   178
   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
   179
          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
   180
          key : the theorem's key constant ( Const(key,...) $ ... )
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   181
          thm : the theorem itself
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   182
          T   : type of P( Const(key,...) $ ... )
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   183
          n   : number of arguments expected by Const(key,...)
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   184
   Ts   : types of parameters
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   185
   t    : the term to be scanned
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   186
******************************************************************************)
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   187
1030
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   188
fun split_posns cmap Ts t =
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   189
  let fun posns Ts pos apsns (Abs(_,T,t)) =
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   190
            let val U = fastype_of1(T::Ts,t)
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   191
            in posns (T::Ts) (0::pos) ((T,U,pos)::apsns) t end
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   192
        | posns Ts pos apsns t =
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   193
            let val (h,ts) = strip_comb t
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   194
                fun iter((i,a),t) = (i+1, (posns Ts (i::pos) apsns t) @ a);
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   195
                val a = case h of
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   196
                  Const(c,_) =>
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   197
                    (case assoc(cmap,c) of
4232
29f3875596ad Tuned function mk_cntxt_splitthm.
berghofe
parents: 4202
diff changeset
   198
                       Some(thm, T, n) =>
29f3875596ad Tuned function mk_cntxt_splitthm.
berghofe
parents: 4202
diff changeset
   199
                         let val t2 = list_comb (h, take (n, ts)) in
29f3875596ad Tuned function mk_cntxt_splitthm.
berghofe
parents: 4202
diff changeset
   200
                           mk_split_pack(thm,T,n,ts,apsns,pos,type_of1(Ts, t2),t2)
29f3875596ad Tuned function mk_cntxt_splitthm.
berghofe
parents: 4202
diff changeset
   201
                         end
1030
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   202
                     | None => [])
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   203
                | _ => []
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   204
             in snd(foldl iter ((0,a),ts)) end
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   205
  in posns Ts [] [] t end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   206
1686
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   207
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   208
fun nth_subgoal i thm = nth_elem(i-1,prems_of thm);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   209
1721
445654b6cb95 Rewrote mk_cntxt_splitthm. Added function mk_case_split_inside_tac.
berghofe
parents: 1686
diff changeset
   210
fun shorter((_,ps,pos,_,_),(_,qs,qos,_,_)) =
4519
055f2067d373 adapted to new sort function;
wenzelm
parents: 4453
diff changeset
   211
  prod_ord (int_ord o pairself length) (order o pairself length)
055f2067d373 adapted to new sort function;
wenzelm
parents: 4453
diff changeset
   212
    ((ps, pos), (qs, qos));
055f2067d373 adapted to new sort function;
wenzelm
parents: 4453
diff changeset
   213
1686
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   214
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   215
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   216
(************************************************************
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   217
   call split_posns with appropriate parameters
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   218
*************************************************************)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   219
1030
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   220
fun select cmap state i =
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   221
  let val goali = nth_subgoal i state
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   222
      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
   223
      val _ $ t $ _ = Logic.strip_assums_concl goali;
4519
055f2067d373 adapted to new sort function;
wenzelm
parents: 4453
diff changeset
   224
  in (Ts,t, sort shorter (split_posns cmap Ts t)) end;
1030
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   225
1686
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   226
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   227
(*************************************************************
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   228
   instantiate lift theorem
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   229
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   230
   if t is of the form
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   231
   ... ( Const(...,...) $ Abs( .... ) ) ...
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   232
   then
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   233
   P = %a.  ... ( Const(...,...) $ a ) ...
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   234
   where a has type T --> U
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   235
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   236
   Ts      : types of parameters
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   237
   t       : lefthand side of meta-equality in subgoal
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   238
             the split theorem is applied to (see cmap)
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   239
   T,U,pos : see mk_split_pack
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   240
   state   : current proof state
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   241
   lift    : the lift theorem
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   242
   i       : no. of subgoal
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   243
**************************************************************)
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   244
1030
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   245
fun inst_lift Ts t (T,U,pos) state lift i =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   246
  let val sg = #sign(rep_thm state)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   247
      val tsig = #tsig(Sign.rep_sg sg)
1030
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   248
      val cntxt = mk_cntxt Ts t pos (T-->U) (#maxidx(rep_thm lift))
231
cb6a24451544 Updated refs to old Sign functions
lcp
parents: 4
diff changeset
   249
      val cu = cterm_of sg cntxt
cb6a24451544 Updated refs to old Sign functions
lcp
parents: 4
diff changeset
   250
      val uT = #T(rep_cterm cu)
cb6a24451544 Updated refs to old Sign functions
lcp
parents: 4
diff changeset
   251
      val cP' = cterm_of sg (Var(P,uT))
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   252
      val ixnTs = Type.typ_match tsig ([],(PT,uT));
231
cb6a24451544 Updated refs to old Sign functions
lcp
parents: 4
diff changeset
   253
      val ixncTs = map (fn (x,y) => (x,ctyp_of sg y)) ixnTs;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   254
  in instantiate (ixncTs, [(cP',cu)]) lift end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   255
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   256
1686
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   257
(*************************************************************
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   258
   instantiate split theorem
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   259
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   260
   Ts    : types of parameters
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   261
   t     : lefthand side of meta-equality in subgoal
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   262
           the split theorem is applied to (see cmap)
4232
29f3875596ad Tuned function mk_cntxt_splitthm.
berghofe
parents: 4202
diff changeset
   263
   tt    : the term  Const(key,..) $ ...
1686
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   264
   thm   : the split theorem
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   265
   TB    : type of body of P(...)
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   266
   state : current proof state
4232
29f3875596ad Tuned function mk_cntxt_splitthm.
berghofe
parents: 4202
diff changeset
   267
   i     : number of subgoal
1686
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   268
**************************************************************)
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   269
4232
29f3875596ad Tuned function mk_cntxt_splitthm.
berghofe
parents: 4202
diff changeset
   270
fun inst_split Ts t tt thm TB state i =
29f3875596ad Tuned function mk_cntxt_splitthm.
berghofe
parents: 4202
diff changeset
   271
  let val _ $ ((Var (P2, PT2)) $ _) $ _ = concl_of thm;
1686
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   272
      val sg = #sign(rep_thm state)
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   273
      val tsig = #tsig(Sign.rep_sg sg)
4232
29f3875596ad Tuned function mk_cntxt_splitthm.
berghofe
parents: 4202
diff changeset
   274
      val cntxt = mk_cntxt_splitthm t tt TB;
4236
fc85fd718429 Fixed bug in inst_split.
berghofe
parents: 4232
diff changeset
   275
      val T = fastype_of1 (Ts, cntxt);
4232
29f3875596ad Tuned function mk_cntxt_splitthm.
berghofe
parents: 4202
diff changeset
   276
      val ixnTs = Type.typ_match tsig ([],(PT2, T))
29f3875596ad Tuned function mk_cntxt_splitthm.
berghofe
parents: 4202
diff changeset
   277
      val abss = foldl (fn (t, T) => Abs ("", T, t))
29f3875596ad Tuned function mk_cntxt_splitthm.
berghofe
parents: 4202
diff changeset
   278
  in
29f3875596ad Tuned function mk_cntxt_splitthm.
berghofe
parents: 4202
diff changeset
   279
    term_lift_inst_rule (state, i, ixnTs, [((P2, T), abss (cntxt, Ts))], thm)
29f3875596ad Tuned function mk_cntxt_splitthm.
berghofe
parents: 4202
diff changeset
   280
  end;
1686
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   281
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   282
(*****************************************************************************
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   283
   The split-tactic
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   284
   
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   285
   splits : list of split-theorems to be tried
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   286
   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
   287
*****************************************************************************)
c67d543bc395 Added functions mk_cntxt_splitthm and inst_split which instantiate
berghofe
parents: 1064
diff changeset
   288
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   289
fun split_tac [] i = no_tac
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   290
  | split_tac splits i =
5553
ae42b36a50c2 renamed mk_meta_eq to mk_eq
oheimb
parents: 5437
diff changeset
   291
  let val splits = map Data.mk_eq splits;
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   292
      fun const(thm) =
3918
94e0fdcb7b91 Added error messages.
nipkow
parents: 3835
diff changeset
   293
            (case concl_of thm of _$(t as _$lhs)$_ =>
94e0fdcb7b91 Added error messages.
nipkow
parents: 3835
diff changeset
   294
               (case strip_comb lhs of (Const(a,_),args) =>
94e0fdcb7b91 Added error messages.
nipkow
parents: 3835
diff changeset
   295
                  (a,(thm,fastype_of t,length args))
4668
131989b78417 Little reorganization. Loop tactics have names now.
nipkow
parents: 4519
diff changeset
   296
                | _ => split_format_err())
131989b78417 Little reorganization. Loop tactics have names now.
nipkow
parents: 4519
diff changeset
   297
             | _ => split_format_err())
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   298
      val cmap = map const splits;
3537
79ac9b475621 Removal of the tactical STATE
paulson
parents: 2266
diff changeset
   299
      fun lift_tac Ts t p st = (rtac (inst_lift Ts t p st trlift i) i) st
79ac9b475621 Removal of the tactical STATE
paulson
parents: 2266
diff changeset
   300
      fun lift_split_tac st = st |>
79ac9b475621 Removal of the tactical STATE
paulson
parents: 2266
diff changeset
   301
            let val (Ts,t,splits) = select cmap st i
1030
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   302
            in case splits of
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   303
                 [] => no_tac
1721
445654b6cb95 Rewrote mk_cntxt_splitthm. Added function mk_case_split_inside_tac.
berghofe
parents: 1686
diff changeset
   304
               | (thm,apsns,pos,TB,tt)::_ =>
1030
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   305
                   (case apsns of
3537
79ac9b475621 Removal of the tactical STATE
paulson
parents: 2266
diff changeset
   306
                      [] => (fn state => state |>
4232
29f3875596ad Tuned function mk_cntxt_splitthm.
berghofe
parents: 4202
diff changeset
   307
			           compose_tac (false, inst_split Ts t tt thm TB state i, 0) i)
3537
79ac9b475621 Removal of the tactical STATE
paulson
parents: 2266
diff changeset
   308
                    | p::_ => EVERY[lift_tac Ts t p,
1030
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   309
                                    rtac reflexive_thm (i+1),
3537
79ac9b475621 Removal of the tactical STATE
paulson
parents: 2266
diff changeset
   310
                                    lift_split_tac])
1030
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   311
            end
3537
79ac9b475621 Removal of the tactical STATE
paulson
parents: 2266
diff changeset
   312
  in COND (has_fewer_prems i) no_tac 
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   313
          (rtac meta_iffD i THEN lift_split_tac)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   314
  end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   315
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   316
in split_tac end;
1721
445654b6cb95 Rewrote mk_cntxt_splitthm. Added function mk_case_split_inside_tac.
berghofe
parents: 1686
diff changeset
   317
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   318
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   319
val split_tac        = mk_case_split_tac              int_ord;
4189
b8c7a6bc6c16 added split_prem_tac
oheimb
parents: 3918
diff changeset
   320
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   321
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
   322
4189
b8c7a6bc6c16 added split_prem_tac
oheimb
parents: 3918
diff changeset
   323
b8c7a6bc6c16 added split_prem_tac
oheimb
parents: 3918
diff changeset
   324
(*****************************************************************************
b8c7a6bc6c16 added split_prem_tac
oheimb
parents: 3918
diff changeset
   325
   The split-tactic for premises
b8c7a6bc6c16 added split_prem_tac
oheimb
parents: 3918
diff changeset
   326
   
b8c7a6bc6c16 added split_prem_tac
oheimb
parents: 3918
diff changeset
   327
   splits : list of split-theorems to be tried
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   328
****************************************************************************)
4202
96876d71eef5 renamed split_prem_tac to split_asm_tac
oheimb
parents: 4189
diff changeset
   329
fun split_asm_tac []     = K no_tac
96876d71eef5 renamed split_prem_tac to split_asm_tac
oheimb
parents: 4189
diff changeset
   330
  | split_asm_tac splits = 
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   331
4930
89271bc4e7ed extended addsplits and delsplits to handle also split rules for assumptions
oheimb
parents: 4668
diff changeset
   332
  let val cname_list = map (fst o split_thm_info) splits;
4189
b8c7a6bc6c16 added split_prem_tac
oheimb
parents: 3918
diff changeset
   333
      fun is_case (a,_) = a mem cname_list;
b8c7a6bc6c16 added split_prem_tac
oheimb
parents: 3918
diff changeset
   334
      fun tac (t,i) = 
b8c7a6bc6c16 added split_prem_tac
oheimb
parents: 3918
diff changeset
   335
	  let val n = find_index (exists_Const is_case) 
b8c7a6bc6c16 added split_prem_tac
oheimb
parents: 3918
diff changeset
   336
				 (Logic.strip_assums_hyp t);
b8c7a6bc6c16 added split_prem_tac
oheimb
parents: 3918
diff changeset
   337
	      fun first_prem_is_disj (Const ("==>", _) $ (Const ("Trueprop", _)
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   338
				 $ (Const (s, _) $ _ $ _ )) $ _ ) = (s=const_or)
4202
96876d71eef5 renamed split_prem_tac to split_asm_tac
oheimb
parents: 4189
diff changeset
   339
	      |   first_prem_is_disj (Const("all",_)$Abs(_,_,t)) = 
96876d71eef5 renamed split_prem_tac to split_asm_tac
oheimb
parents: 4189
diff changeset
   340
					first_prem_is_disj t
4189
b8c7a6bc6c16 added split_prem_tac
oheimb
parents: 3918
diff changeset
   341
	      |   first_prem_is_disj _ = false;
5437
f68b9d225942 added caveat; a real solution would be difficult
oheimb
parents: 5304
diff changeset
   342
      (* does not work properly if the split variable is bound by a quantfier *)
4202
96876d71eef5 renamed split_prem_tac to split_asm_tac
oheimb
parents: 4189
diff changeset
   343
	      fun flat_prems_tac i = SUBGOAL (fn (t,i) => 
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   344
			   (if first_prem_is_disj t
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   345
			    then EVERY[etac Data.disjE i,rotate_tac ~1 i,
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   346
				       rotate_tac ~1  (i+1),
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   347
				       flat_prems_tac (i+1)]
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   348
			    else all_tac) 
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   349
			   THEN REPEAT (eresolve_tac [Data.conjE,Data.exE] i)
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   350
			   THEN REPEAT (dresolve_tac [Data.notnotD]   i)) i;
4189
b8c7a6bc6c16 added split_prem_tac
oheimb
parents: 3918
diff changeset
   351
	  in if n<0 then no_tac else DETERM (EVERY'
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   352
		[rotate_tac n, etac Data.contrapos2,
4189
b8c7a6bc6c16 added split_prem_tac
oheimb
parents: 3918
diff changeset
   353
		 split_tac splits, 
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   354
		 rotate_tac ~1, etac Data.contrapos, rotate_tac ~1, 
4202
96876d71eef5 renamed split_prem_tac to split_asm_tac
oheimb
parents: 4189
diff changeset
   355
		 flat_prems_tac] i)
4189
b8c7a6bc6c16 added split_prem_tac
oheimb
parents: 3918
diff changeset
   356
	  end;
b8c7a6bc6c16 added split_prem_tac
oheimb
parents: 3918
diff changeset
   357
  in SUBGOAL tac
b8c7a6bc6c16 added split_prem_tac
oheimb
parents: 3918
diff changeset
   358
  end;
b8c7a6bc6c16 added split_prem_tac
oheimb
parents: 3918
diff changeset
   359
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   360
fun split_name name asm = "split " ^ name ^ (if asm then " asm" else "");
4189
b8c7a6bc6c16 added split_prem_tac
oheimb
parents: 3918
diff changeset
   361
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   362
fun ss addsplits splits =
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   363
  let fun addsplit (ss,split) =
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   364
        let val (name,asm) = split_thm_info split
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   365
        in Data.Simplifier.addloop(ss,(split_name name asm,
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   366
		       (if asm then split_asm_tac else split_tac) [split])) end
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   367
  in foldl addsplit (ss,splits) end;
1721
445654b6cb95 Rewrote mk_cntxt_splitthm. Added function mk_case_split_inside_tac.
berghofe
parents: 1686
diff changeset
   368
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   369
fun ss delsplits splits =
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   370
  let fun delsplit(ss,split) =
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   371
        let val (name,asm) = split_thm_info split
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   372
        in Data.Simplifier.delloop(ss,split_name name asm)
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   373
  end in foldl delsplit (ss,splits) end;
1721
445654b6cb95 Rewrote mk_cntxt_splitthm. Added function mk_case_split_inside_tac.
berghofe
parents: 1686
diff changeset
   374
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   375
fun Addsplits splits = (Data.Simplifier.simpset_ref() := 
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   376
			Data.Simplifier.simpset() addsplits splits);
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   377
fun Delsplits splits = (Data.Simplifier.simpset_ref() := 
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 4930
diff changeset
   378
			Data.Simplifier.simpset() delsplits splits);
4189
b8c7a6bc6c16 added split_prem_tac
oheimb
parents: 3918
diff changeset
   379
1721
445654b6cb95 Rewrote mk_cntxt_splitthm. Added function mk_case_split_inside_tac.
berghofe
parents: 1686
diff changeset
   380
end;