src/Provers/splitter.ML
author nipkow
Sun, 16 Apr 1995 11:56:11 +0200
changeset 1064 5d6fb2c938e0
parent 1030 1d8fa2fc4b9c
child 1686 c67d543bc395
permissions -rw-r--r--
Fixed bug.
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.
2695ba9b40f7 added header
nipkow
parents: 0
diff changeset
     7
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
Use:
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
val split_tac = mk_case_split_tac iffD;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    11
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    12
by(case_split_tac splits i);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    13
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    14
where splits = [P(elim(...)) == rhs, ...]
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    15
      iffD  = [| P <-> Q; Q |] ==> P (* is called iffD2 in HOL *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    16
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    17
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    18
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
fun mk_case_split_tac iffD =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    20
let
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    21
943
8477483f663f Replaced read by read_cterm.
nipkow
parents: 927
diff changeset
    22
val lift =
8477483f663f Replaced read by read_cterm.
nipkow
parents: 927
diff changeset
    23
  let val ct = read_cterm (#sign(rep_thm iffD))
8477483f663f Replaced read by read_cterm.
nipkow
parents: 927
diff changeset
    24
           ("[| !!x::'b::logic. Q(x) == R(x) |] ==> \
8477483f663f Replaced read by read_cterm.
nipkow
parents: 927
diff changeset
    25
            \P(%x.Q(x)) == P(%x.R(x))::'a::logic",propT)
8477483f663f Replaced read by read_cterm.
nipkow
parents: 927
diff changeset
    26
  in prove_goalw_cterm [] ct
8477483f663f Replaced read by read_cterm.
nipkow
parents: 927
diff changeset
    27
     (fn [prem] => [rewtac prem, rtac reflexive_thm 1])
8477483f663f Replaced read by read_cterm.
nipkow
parents: 927
diff changeset
    28
  end;
4
2695ba9b40f7 added header
nipkow
parents: 0
diff changeset
    29
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    30
val trlift = lift RS transitive_thm;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    31
val _ $ (Var(P,PT)$_) $ _ = concl_of trlift;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    32
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    33
1030
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    34
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
    35
  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
    36
      fun down [] t i = Bound 0
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    37
        | down (p::ps) t i =
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    38
            let val (h,ts) = strip_comb t
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    39
                val v1 = map var (take(p,ts) ~~ (i upto (i+p-1)))
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    40
                val u::us = drop(p,ts)
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    41
                val v2 = map var (us ~~ ((i+p) upto (i+length(ts)-2)))
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    42
      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
    43
  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
    44
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    45
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
    46
1064
5d6fb2c938e0 Fixed bug.
nipkow
parents: 1030
diff changeset
    47
(* check if the innermost quantifier that needs to be removed
5d6fb2c938e0 Fixed bug.
nipkow
parents: 1030
diff changeset
    48
   has a body of type T; otherwise the expansion thm will fail later on
5d6fb2c938e0 Fixed bug.
nipkow
parents: 1030
diff changeset
    49
*)
5d6fb2c938e0 Fixed bug.
nipkow
parents: 1030
diff changeset
    50
fun type_test(T,lbnos,apsns) =
5d6fb2c938e0 Fixed bug.
nipkow
parents: 1030
diff changeset
    51
  let val (_,U,_) = nth_elem(min lbnos,apsns)
5d6fb2c938e0 Fixed bug.
nipkow
parents: 1030
diff changeset
    52
  in T=U end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    53
1030
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    54
fun mk_split_pack(thm,T,n,ts,apsns) =
1064
5d6fb2c938e0 Fixed bug.
nipkow
parents: 1030
diff changeset
    55
  if n > length ts then []
5d6fb2c938e0 Fixed bug.
nipkow
parents: 1030
diff changeset
    56
  else let val lev = length apsns
1030
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    57
           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
    58
           val flbnos = filter (fn i => i < lev) lbnos
1064
5d6fb2c938e0 Fixed bug.
nipkow
parents: 1030
diff changeset
    59
       in if null flbnos then [(thm,[])]
5d6fb2c938e0 Fixed bug.
nipkow
parents: 1030
diff changeset
    60
          else if type_test(T,flbnos,apsns) then [(thm, rev apsns)] else []
5d6fb2c938e0 Fixed bug.
nipkow
parents: 1030
diff changeset
    61
       end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    62
1030
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    63
fun split_posns cmap Ts t =
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    64
  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
    65
            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
    66
            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
    67
        | posns Ts pos apsns t =
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    68
            let val (h,ts) = strip_comb t
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    69
                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
    70
                val a = case h of
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    71
                  Const(c,_) =>
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    72
                    (case assoc(cmap,c) of
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    73
                       Some(thm,T,n) => mk_split_pack(thm,T,n,ts,apsns)
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    74
                     | None => [])
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    75
                | _ => []
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    76
             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
    77
  in posns Ts [] [] t end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    78
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    79
fun nth_subgoal i thm = nth_elem(i-1,prems_of thm);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    80
1030
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    81
fun shorter((_,ps),(_,qs)) = length ps <= length qs;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    82
1030
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    83
fun select cmap state i =
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    84
  let val goali = nth_subgoal i state
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    85
      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
    86
      val _ $ t $ _ = Logic.strip_assums_concl goali;
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    87
  in (Ts,t,sort shorter (split_posns cmap Ts t)) end;
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    88
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
    89
fun inst_lift Ts t (T,U,pos) state lift i =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    90
  let val sg = #sign(rep_thm state)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    91
      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
    92
      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
    93
      val cu = cterm_of sg cntxt
cb6a24451544 Updated refs to old Sign functions
lcp
parents: 4
diff changeset
    94
      val uT = #T(rep_cterm cu)
cb6a24451544 Updated refs to old Sign functions
lcp
parents: 4
diff changeset
    95
      val cP' = cterm_of sg (Var(P,uT))
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    96
      val ixnTs = Type.typ_match tsig ([],(PT,uT));
231
cb6a24451544 Updated refs to old Sign functions
lcp
parents: 4
diff changeset
    97
      val ixncTs = map (fn (x,y) => (x,ctyp_of sg y)) ixnTs;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    98
  in instantiate (ixncTs, [(cP',cu)]) lift end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    99
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   100
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   101
fun split_tac [] i = no_tac
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   102
  | split_tac splits i =
1030
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   103
  let fun const(thm) = let val _$(t as _$lhs)$_ = concl_of thm
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   104
                           val (Const(a,_),args) = strip_comb lhs
1030
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   105
                       in (a,(thm,fastype_of t,length args)) end
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   106
      val cmap = map const splits;
1030
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   107
      fun lift Ts t p state = rtac (inst_lift Ts t p state trlift i) i
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   108
      fun lift_split state =
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   109
            let val (Ts,t,splits) = select cmap state i
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   110
            in case splits of
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   111
                 [] => no_tac
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   112
               | (thm,apsns)::_ =>
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   113
                   (case apsns of
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   114
                      [] => rtac thm i
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   115
                    | p::_ => EVERY[STATE(lift Ts t p),
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   116
                                    rtac reflexive_thm (i+1),
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   117
                                    STATE lift_split])
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   118
            end
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   119
  in STATE(fn thm =>
1030
1d8fa2fc4b9c Completely rewrote split_tac. The old one failed in strange circumstances.
nipkow
parents: 943
diff changeset
   120
       if i <= nprems_of thm then rtac iffD i THEN STATE lift_split
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   121
       else no_tac)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   122
  end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   123
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   124
in split_tac end;