src/Pure/IsaPlanner/isa_fterm.ML
author skalberg
Sun, 13 Feb 2005 17:15:14 +0100
changeset 15531 08c8dad8e399
parent 15481 fc075ae929e4
child 15661 9ef583b08647
permissions -rw-r--r--
Deleted Library.option type.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     1
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     2
(*  Title:      libs/isa_fterm.ML
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     3
    Author:     Lucas Dixon, University of Edinburgh
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     4
                lucasd@dai.ed.ac.uk
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     5
    Date:       16 April 2003
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     6
*)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     7
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     8
(*  DESCRIPTION:
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     9
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    10
    Generic Foucs terms (like Zippers) instantiation for Isabelle terms. 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    11
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    12
*)   
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    13
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    14
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    15
signature ISA_ENCODE_TERM = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    16
  F_ENCODE_TERM_SIG
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    17
    where type term = Term.term type TypeT = Term.typ;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    18
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    19
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    20
(* signature BASIC_ISA_FTERM = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    21
FOCUS_TERM_SIG where type Term = ISA_ENCODE_TERM.Term *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    22
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    23
(* 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    24
signature ISA_FTERM =
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    25
sig
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    26
  include F_ENCODE_TERM_SIG
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    27
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    28
    val clean_match_ft :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    29
       Type.tsig ->
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    30
       FcTerm ->
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    31
       Term.term ->
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    32
       (
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    33
       ((Term.indexname * Term.typ) list * (Term.indexname * Term.term) list)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
    34
       * (string * Type) list * Term.term) option
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    35
    val clean_unify_ft :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    36
       Sign.sg ->
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    37
       int ->
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    38
       FcTerm ->
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    39
       Term.term ->
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    40
       (
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    41
       ((Term.indexname * Term.typ) list * (Term.indexname * Term.term) list)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    42
       * (string * Type) list * Term.term) Seq.seq
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    43
    val fakefree_badbounds :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    44
       (string * Term.typ) list -> int -> Term.term -> Term.term
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    45
    val find_fcterm_matches :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    46
       ((FcTerm -> 'a) -> FcTerm -> 'b) ->
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    47
       (FcTerm -> 'a) -> FcTerm -> 'b
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    48
    val find_sg_concl_matches :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    49
       ((FcTerm -> 'a) -> FcTerm -> 'b) ->
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    50
       (FcTerm -> 'a) -> int -> EncodeIsaFTerm.term -> 'b
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    51
    val find_sg_concl_thm_matches :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    52
       ((FcTerm -> 'a) -> FcTerm -> 'b) ->
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    53
       (FcTerm -> 'a) -> int -> Thm.thm -> 'b
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    54
    val find_sg_matches :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    55
       ((FcTerm -> 'a) -> FcTerm -> 'b) ->
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    56
       (FcTerm -> 'a) -> int -> EncodeIsaFTerm.term -> 'b
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    57
    val find_sg_thm_matches :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    58
       ((FcTerm -> 'a) -> FcTerm -> 'b) ->
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    59
       (FcTerm -> 'a) -> int -> Thm.thm -> 'b
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    60
    val focus_to_concl : FcTerm -> FcTerm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    61
    val focus_to_concl_of_term : EncodeIsaFTerm.term -> FcTerm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    62
    val focus_to_subgoal :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    63
       int -> FcTerm -> FcTerm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    64
    val focus_to_subgoal_of_term :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    65
       int -> EncodeIsaFTerm.term -> FcTerm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    66
    val focus_to_term_goal_prem :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    67
       int * int -> EncodeIsaFTerm.term -> FcTerm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    68
    val focuseq_to_subgoals :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    69
       FcTerm -> FcTerm Seq.seq
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    70
    exception isa_focus_term_exp of string
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    71
    val mk_foo_match :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    72
       (Term.term -> Term.term) ->
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    73
       ('a * Term.typ) list -> Term.term -> Term.term
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    74
    val prepmatch :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    75
       FcTerm ->
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    76
       Term.term * ((string * Type) list * Term.term)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    77
    val search_bl_ru_f :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    78
       (FcTerm -> 'a Seq.seq) ->
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    79
       FcTerm -> 'a Seq.seq
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    80
    val search_bl_ur_f :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    81
       (FcTerm -> 'a Seq.seq) ->
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    82
       FcTerm -> 'a Seq.seq
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    83
    val valid_match_start : FcTerm -> bool
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    84
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    85
end *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    86
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    87
structure BasicIsaFTerm = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    88
FocusTermFUN( structure EncodeTerm = EncodeIsaFTerm);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    89
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    90
(* HOL Dependent *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    91
structure IsaFTerm =
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    92
struct 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    93
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    94
(* include BasicIsaFTerm *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    95
(* open FocusTermFUN( structure EncodeTerm = EncodeIsaFTerm ); *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    96
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    97
open BasicIsaFTerm;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    98
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    99
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   100
(* SOME general search within a focus term... *)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   101
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   102
(* Note: only upterms with a free or constant are going to yeald a
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   103
match, thus if we get anything else (bound or var) skip it! This is
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   104
important if we switch to a unification net! in particular to avoid
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   105
vars. *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   106
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   107
fun valid_match_start ft =
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   108
    (case TermLib.bot_left_leaf_of (focus_of_fcterm ft) of 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   109
       Const _ => true
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   110
     | Free _ =>  true
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   111
     | Abs _ =>  true
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   112
     | _ => false);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   113
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   114
(* match starting at the bottom left, moving up to top of the term,
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   115
   then moving right to the next leaf and up again etc *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   116
(* FIXME: make properly lazy! *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   117
fun search_nonvar_bl_ur_f f ft = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   118
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   119
      val fts = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   120
          Seq.filter valid_match_start
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   121
                     (leaf_seq_of_fcterm ft)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   122
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   123
      fun mk_match_up_seq ft =
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   124
          case focus_up_abs_or_appr ft of 
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   125
            SOME ft' => Seq.append(f ft, mk_match_up_seq ft')
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   126
          | NONE => f ft
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   127
    in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   128
      Seq.flat (Seq.map mk_match_up_seq fts)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   129
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   130
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   131
fun search_bl_ur_f f ft = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   132
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   133
      val fts = (leaf_seq_of_fcterm ft)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   134
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   135
      fun mk_match_up_seq ft =
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   136
          case focus_up_abs_or_appr ft of 
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   137
            SOME ft' => Seq.append(f ft, mk_match_up_seq ft')
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   138
          | NONE => f ft
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   139
    in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   140
      Seq.flat (Seq.map mk_match_up_seq fts)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   141
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   142
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   143
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   144
(* FIXME: make properly lazy! *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   145
(* FIXME: make faking of bound vars local - for speeeeed *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   146
fun search_nonvar_bl_ru_f f ft = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   147
    let
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   148
      fun mauxtop ft = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   149
          if (valid_match_start ft)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   150
            then maux ft else Seq.empty
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   151
      and maux ft = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   152
          let val t' = (focus_of_fcterm ft) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   153
          (* val _ = writeln ("Examining: " ^ (TermLib.string_of_term t')) *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   154
          in 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   155
          (case t' of 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   156
            (_ $ _) => Seq.append (maux (focus_left ft), 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   157
                         Seq.append(mauxtop (focus_right ft), 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   158
                                    f ft))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   159
          | (Abs _) => Seq.append (maux (focus_abs ft), 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   160
                                   f ft)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   161
          | leaf => f ft) end
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   162
    in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   163
      mauxtop ft
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   164
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   165
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   166
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   167
fun search_bl_ru_f f ft = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   168
    let
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   169
      fun maux ft = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   170
          let val t' = (focus_of_fcterm ft) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   171
          (*   val _ = writeln ("Examining: " ^ (TermLib.string_of_term t')) *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   172
          in 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   173
          (case t' of 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   174
            (_ $ _) => Seq.append (maux (focus_left ft), 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   175
                         Seq.append(maux (focus_right ft), 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   176
                                    f ft))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   177
          | (Abs _) => Seq.append (maux (focus_abs ft), 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   178
                                   f ft)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   179
          | leaf => f ft) end
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   180
    in maux ft end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   181
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   182
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   183
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   184
exception isa_focus_term_exp of string;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   185
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   186
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   187
fun focus_to_dest_impl ft = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   188
    let val (lhs, rhs) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   189
            Logic.dest_implies (focus_of_fcterm ft) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   190
    in (focus_left ft, focus_right ft) end
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   191
    handle Term.TERM _ => 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   192
           raise isa_focus_term_exp 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   193
                   "focus_to_dest_impl: applied to non implication";
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   194
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   195
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   196
(* move focus to the conlusion *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   197
fun focus_to_concl ft = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   198
    let val (lhs, rhs) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   199
            Logic.dest_implies (focus_of_fcterm ft) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   200
    in focus_to_concl (focus_right ft) end
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   201
    handle Term.TERM _ =>  ft;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   202
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   203
val focus_to_concl_of_term = focus_to_concl o fcterm_of_term;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   204
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   205
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   206
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   207
(* give back sequence of focuses at different subgoals *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   208
(* FIXME: make truly lazy *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   209
fun focuseq_to_subgoals ft = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   210
    if (Logic.is_implies (focus_of_fcterm ft)) then 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   211
      Seq.cons (focus_right (focus_left ft), focuseq_to_subgoals (focus_right ft))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   212
    else
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   213
      Seq.empty;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   214
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   215
(* move focus to a specific subgoal, 0 is first  *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   216
fun focus_to_subgoal j ft = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   217
    let fun focus_to_subgoal' (ft, 0) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   218
            let val (lhs, rhs) = Logic.dest_implies (focus_of_fcterm ft) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   219
            in ft |> focus_left |> focus_right end
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   220
          | focus_to_subgoal' (ft, i) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   221
            let val (lhs, rhs) = Logic.dest_implies (focus_of_fcterm ft) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   222
            in focus_to_subgoal' (focus_right ft, i - 1) end
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   223
    in focus_to_subgoal' (ft, j - 1) end
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   224
    handle Term.TERM _ => 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   225
           raise isa_focus_term_exp 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   226
                   ("focus_to_subgoal: No such subgoal: " ^ 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   227
                    (string_of_int j));
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   228
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   229
fun focus_to_subgoal_of_term i t = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   230
    focus_to_subgoal i (fcterm_of_term t)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   231
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   232
(* move focus to a specific premise *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   233
(* fun focus_past_params i ft = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   234
    (focus_to_subgoal (focus_right ft, i))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   235
    handle isa_focus_term_exp _ => 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   236
           raise isa_focus_term_exp 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   237
             ("focus_to_prmise: No such premise: " ^ (string_of_int i)); *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   238
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   239
fun focus_to_term_goal_prem (premid,gaolid) t = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   240
    focus_to_subgoal premid (focus_to_subgoal_of_term gaolid t);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   241
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   242
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   243
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   244
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   245
(* T is outer bound vars, n is number of locally bound vars *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   246
fun fakefree_badbounds T n (a $ b) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   247
    (fakefree_badbounds T n a) $ (fakefree_badbounds T n b)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   248
  | fakefree_badbounds T n (Abs(s,ty,t)) =  
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   249
    Abs(s,ty, fakefree_badbounds T (n + 1) t)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   250
  | fakefree_badbounds T n (b as Bound i) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   251
    let fun mkfake_bound j [] = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   252
            raise ERROR_MESSAGE "fakefree_badbounds: bound is outside of the known types!"
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   253
          | mkfake_bound 0 ((s,ty)::Ts) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   254
            Free (RWTools.mk_fake_bound_name s,ty)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   255
          | mkfake_bound j (d::Ts) = mkfake_bound (j - 1) Ts
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   256
    in if n <= i then mkfake_bound (i - n) T else b end
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   257
  | fakefree_badbounds T n t = t;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   258
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   259
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   260
(* note: outerterm is the taget with the match replaced by a bound 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   261
         variable : ie: "P lhs" beocmes "%x. P x" 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   262
         insts is the types of instantiations of vars in lhs
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   263
         and typinsts is the type instantiations of types in the lhs
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   264
         Note: Final rule is the rule lifted into the ontext of the 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   265
         taget thm. *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   266
fun mk_foo_match mkuptermfunc Ts t = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   267
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   268
      val ty = Term.type_of t
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   269
      val bigtype = (rev (map snd Ts)) ---> ty
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   270
      fun mk_foo 0 t = t
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   271
        | mk_foo i t = mk_foo (i - 1) (t $ (Bound (i - 1)))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   272
      val num_of_bnds = (length Ts)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   273
      (* foo_term = "fooabs y0 ... yn" where y's are local bounds *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   274
      val foo_term = mk_foo num_of_bnds (Bound num_of_bnds)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   275
    in Abs("fooabs", bigtype, mkuptermfunc foo_term) end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   276
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   277
(* before matching we need to fake the bound vars that missing an
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   278
abstraction. In this function we additionally construct the
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   279
abstraction environment, and an outer context term (with the focus
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   280
abstracted out) for use in rewriting with RWInst.rw *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   281
fun prepmatch ft = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   282
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   283
      val t = focus_of_fcterm ft  
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   284
      val Ts = tyenv_of_focus ft
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   285
      val t' = fakefree_badbounds Ts 0 t
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   286
      fun mktermf t = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   287
          term_of_fcterm (set_focus_of_fcterm ft t)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   288
      val absterm = mk_foo_match mktermf Ts t'
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   289
    in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   290
      (t', (Ts, absterm))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   291
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   292
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   293
(* matching and unification for a focus term's focus *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   294
fun clean_match_ft tsig pat ft = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   295
    let val (t, (Ts,absterm)) = prepmatch ft in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   296
      case TermLib.clean_match tsig (pat, t) of 
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   297
        NONE => NONE 
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   298
      | SOME insts => SOME (insts, Ts, absterm) end;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   299
(* ix = max var index *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   300
fun clean_unify_ft sgn ix pat ft = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   301
    let val (t, (Ts,absterm)) = prepmatch ft in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   302
    Seq.map (fn insts => (insts, Ts, absterm)) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   303
            (TermLib.clean_unify sgn ix (t, pat)) end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   304
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   305
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   306
(* THINK: ? do we not need to incremement bound indices? *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   307
(* THINK: it should be the search which localisaes the search to the
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   308
current focus, not this hack in the matcher... ? *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   309
(* find matches below this particular focus term *) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   310
(* The search function is to find a match within a term... *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   311
(* the matcher is something that is applied to each node chosen by the
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   312
searchf and the results are flattened to form a lazy list. *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   313
fun find_fcterm_matches searchf matcher ft = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   314
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   315
      val ftupterm = upterm_of ft
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   316
      val focusft = focus_of_fcterm ft
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   317
      val add_uptermf = add_upterm ftupterm 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   318
    in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   319
    searchf
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   320
      (fn ft' => matcher (add_uptermf ft'))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   321
      (fcterm_of_term focusft)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   322
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   323
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   324
(* FIXME: move argument orders for efficiency... 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   325
i.e. wenzel style val foofunc = x o y;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   326
*)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   327
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   328
(* find the matches inside subgoal i of th *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   329
fun find_sg_matches searchf matcher i t = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   330
    let val subgoal_fcterm = focus_to_subgoal_of_term i t
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   331
    in find_fcterm_matches searchf matcher subgoal_fcterm end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   332
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   333
(* find the matches inside subgoal i of th *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   334
fun find_sg_thm_matches searchf matcher i th = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   335
    find_sg_matches searchf matcher i (Thm.prop_of th);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   336
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   337
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   338
(* find the matches inside subgoal i's conclusion of th *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   339
fun find_sg_concl_matches searchf matcher i t = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   340
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   341
      val subgoal_fcterm = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   342
          focus_to_concl (focus_to_subgoal_of_term i t)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   343
    in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   344
      find_fcterm_matches searchf matcher subgoal_fcterm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   345
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   346
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   347
(* find the matches inside subgoal i's conclusion of th *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   348
fun find_sg_concl_thm_matches searchf matcher i th = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   349
    find_sg_concl_matches searchf matcher i (Thm.prop_of th);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   350
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   351
end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   352
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   353
(* 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   354
test... 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   355
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   356
f_encode_isatermS.encode (read "P a");
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   357
isafocustermS.fcterm_of_term (read "f a");
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   358
isafocustermS.term_of_fcterm it;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   359
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   360
Goal "P b ==> P (suc b)";
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   361
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   362
TermLib.string_of_term ((focus_of_fcterm o focus_to_subgoal_of_term 1 o prop_of) (topthm()));
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   363
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   364
TermLib.string_of_term ((focus_of_fcterm o focus_to_concl o focus_to_subgoal_of_term 1 o prop_of) (topthm()));
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   365
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   366
TermLib.string_of_term ((focus_of_fcterm o focus_to_term_goal_prem (1,1) o prop_of) (topthm()));
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   367
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   368
*)