src/Pure/IsaPlanner/focus_term_lib.ML
author paulson
Thu, 28 Jul 2005 17:54:39 +0200
changeset 16954 82d0a25c5a1d
parent 16179 fa7e70be26b0
child 17044 94d38d9fac40
permissions -rw-r--r--
new function trim_ends
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     1
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) 
16179
fa7e70be26b0 header;
wenzelm
parents: 15814
diff changeset
     2
(*  Title:      Pure/IsaPlanner/focus_term_lib.ML
fa7e70be26b0 header;
wenzelm
parents: 15814
diff changeset
     3
    ID:		$Id$
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     4
    Author:     Lucas Dixon, University of Edinburgh
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     5
                lucasd@dai.ed.ac.uk
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     6
    Date:       16 April 2003
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
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     9
(*  DESCRIPTION:
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    10
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    11
    Generic Foucs terms (like Zippers), which allows modification and
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    12
    manipulation of term keeping track of how we got to the position
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    13
    in the term. We provide a signature for terms, ie any kind of
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    14
    lamda calculus with abs and application, and some notion of types
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    15
    and naming of abstracted vars. 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    16
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    17
    FIXME: at some point in the future make a library to work simply
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    18
    with polymorphic upterms - that way if I want to use them without
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    19
    the focus part, then I don't need to include all the focus stuff.
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    20
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    21
*)   
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
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    25
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    26
(*  to endoce and decode general terms into the type needed by focus
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    27
term, this is an easy thing to write, but needs to be done for each
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    28
type that is going to be used with focus terms. *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    29
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    30
signature F_ENCODE_TERM_SIG = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    31
sig
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
  type term (* type of term to be encoded into TermT for focus term manip *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    34
  type TypeT (* type annotation for abstractions *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    35
  type LeafT (* type for other leaf types in term sturcture *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    36
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    37
  (* the type to be encoded into *)
15632
bb178a7a69c1 op vor infix-Konstruktoren im datatype binding zum besseren Parsen
gagern
parents: 15531
diff changeset
    38
  datatype TermT = op $ of TermT * TermT
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    39
    | Abs of string * TypeT * TermT
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    40
    | lf of LeafT
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    41
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    42
  (* the encode and decode functions *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    43
  val fakebounds : string * TypeT -> term -> term
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    44
  val encode : term -> TermT
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    45
  val decode : TermT -> term
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    46
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    47
end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    48
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    49
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    50
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    51
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    52
signature FOCUS_TERM_SIG = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    53
sig
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    54
     structure MinTermS : F_ENCODE_TERM_SIG
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    55
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    56
     exception focus_term_exp of string;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    57
     exception out_of_term_exception of string;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    58
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    59
     type Term = MinTermS.TermT;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    60
     type Type = MinTermS.TypeT;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    61
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    62
     type UpTerm = (Term,Type) UpTermLib.T;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    63
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    64
(*   datatype
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    65
       UpTerm =
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    66
           abs of string * Type * UpTerm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    67
         | appl of Term * UpTerm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    68
         | appr of Term * UpTerm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    69
         | root *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    70
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    71
     datatype FcTerm = focus of Term * UpTerm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    72
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    73
     (* translating *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    74
     val fcterm_of_term : MinTermS.term -> FcTerm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    75
     val term_of_fcterm : FcTerm -> MinTermS.term
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    76
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    77
     (* editing/constrution *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    78
     val enc_appl : (MinTermS.term * UpTerm) -> UpTerm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    79
     val enc_appr : (MinTermS.term * UpTerm) -> UpTerm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    80
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    81
     (* upterm creatioin *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    82
     val upterm_of : FcTerm -> UpTerm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    83
     val add_upterm : UpTerm -> FcTerm -> FcTerm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    84
     val mk_term_of_upterm : (MinTermS.term * UpTerm) -> MinTermS.term
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    85
     val mk_termf_of_upterm : UpTerm -> 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    86
            (((string * Type) list) * 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    87
             (MinTermS.term -> MinTermS.term))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    88
     val pure_mk_termf_of_upterm : (MinTermS.term, Type) UpTermLib.T -> 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    89
                                   (((string * Type) list) * 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    90
                                    (MinTermS.term -> MinTermS.term))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    91
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    92
     (* focusing, throws exceptions *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    93
     val focus_bot_left_leaf : FcTerm -> FcTerm
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15632
diff changeset
    94
     val focus_bot_left_nonabs_leaf : FcTerm -> FcTerm
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    95
     val focus_left : FcTerm -> FcTerm
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15632
diff changeset
    96
     val focus_strict_left : FcTerm -> FcTerm
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    97
     val focus_right : FcTerm -> FcTerm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    98
     val focus_abs : FcTerm -> FcTerm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    99
     val focus_fake_abs : FcTerm -> FcTerm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   100
     val focus_to_top : FcTerm -> FcTerm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   101
     val focus_up : FcTerm -> FcTerm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   102
     val focus_up_right : FcTerm -> FcTerm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   103
     val focus_up_right1 : FcTerm -> FcTerm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   104
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   105
     (* optional focus changes *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   106
     val focus_up_abs : FcTerm -> FcTerm option
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   107
     val focus_up_appr : FcTerm -> FcTerm option
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   108
     val focus_up_appl : FcTerm -> FcTerm option
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   109
     val focus_up_abs_or_appr : FcTerm -> FcTerm option
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   110
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   111
     val tyenv_of_focus : FcTerm -> (string * Type) list
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   112
     val tyenv_of_focus' : FcTerm -> Type list
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
     (* set/get the focus term *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   115
     val set_focus_of_fcterm : FcTerm -> MinTermS.term -> FcTerm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   116
     val focus_of_fcterm : FcTerm -> MinTermS.term
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   117
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   118
     (* leaf navigation *) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   119
     val leaf_seq_of_fcterm : FcTerm -> FcTerm Seq.seq
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   120
     val next_leaf_fcterm : FcTerm -> FcTerm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   121
     val next_leaf_of_fcterm_seq : FcTerm -> FcTerm Seq.seq
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
     (* analysis *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   124
     val upsize_of_fcterm : FcTerm -> int
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   125
     val pretty_fcterm : FcTerm -> Pretty.T
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   126
end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   127
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   128
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   129
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
(* 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   132
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   133
   NOTE!!! I think this can be done in a purely functional way with a
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   134
   pair of a term (the focus) and a function, that when applied
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   135
   unfocuses one level... maybe think about this as it would be a very
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   136
   nice generic consrtuction, without the need for encoding/decoding
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   137
   strutcures.
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   138
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   139
*)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   140
functor FocusTermFUN(structure EncodeTerm : F_ENCODE_TERM_SIG)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   141
  : FOCUS_TERM_SIG = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   142
struct
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
structure MinTermS = EncodeTerm;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   145
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   146
local open MinTermS open UpTermLib in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   147
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   148
  type Term = MinTermS.TermT;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   149
  type Type = MinTermS.TypeT;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   150
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   151
  exception focus_term_exp of string;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   152
  exception out_of_term_exception of string;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   153
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   154
  infix 9 $
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   155
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   156
  (* datatype to hold a term tree and the path to where you are in the term *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   157
(*   datatype UpTerm = root
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   158
                  | appl of (Term * UpTerm)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   159
                  | appr of (Term * UpTerm)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   160
                  | abs of (string * Type * UpTerm); *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   161
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   162
  type UpTerm = (Term,Type) UpTermLib.T;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   163
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   164
  fun enc_appl (t,u) = appl((MinTermS.encode t),u);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   165
  fun enc_appr (t,u) = appr((MinTermS.encode t),u);
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
  datatype FcTerm = focus of (Term * UpTerm);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   168
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   169
  (* the the term of the upterm *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   170
  fun mk_term_of_upt_aux (t, root) = MinTermS.decode t
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   171
    | mk_term_of_upt_aux (t, appl (l,m)) =  mk_term_of_upt_aux(l$t, m)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   172
    | mk_term_of_upt_aux (t, appr (r,m)) =  mk_term_of_upt_aux(t$r, m)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   173
    | mk_term_of_upt_aux (t, abs (s,ty,m)) = mk_term_of_upt_aux(Abs(s,ty,t),m);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   174
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   175
  fun mk_term_of_upterm (t, ut) = mk_term_of_upt_aux (MinTermS.encode t, ut);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   176
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   177
  (* a function version of the above, given an upterm it returns:
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   178
     a function on that given a term puts it in the context of the upterm,
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   179
     and a list of bound variables that are in the upterm, added as 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   180
     we go up - so in reverse order from that typiclaly used by top-down
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   181
     parsing of terms. *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   182
  fun mk_termf_of_upt_aux (f, bs, root) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   183
      (bs, fn t => MinTermS.decode (f t))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   184
    | mk_termf_of_upt_aux (f, bs, appl (l,m)) =  
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   185
      mk_termf_of_upt_aux (fn t => l $ (f t), bs, m)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   186
    | mk_termf_of_upt_aux (f, bs, appr (r,m)) =  
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   187
      mk_termf_of_upt_aux (fn t => (f t) $ r, bs, m)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   188
    | mk_termf_of_upt_aux (f, bs, abs (s,ty,m)) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   189
      mk_termf_of_upt_aux (fn t => Abs(s,ty,(f t)), (s,ty)::bs, m);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   190
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   191
  fun mk_termf_of_upterm ut = mk_termf_of_upt_aux (MinTermS.encode, [], ut);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   192
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   193
 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   194
  fun upterm_of (focus(t, u)) = u;
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
  (* put a new upterm at the start of our current term *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   197
  fun add_upterm u2 (focus(t, u)) = focus(t, UpTermLib.apply u2 u);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   198
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   199
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   200
  (* As above but work on the pure originial upterm type *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   201
  fun pure_mk_termf_of_upterm ut = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   202
      mk_termf_of_upt_aux 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   203
        (encode, [], (map_to_upterm_parts (encode, I) ut));
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
  fun fcterm_of_term t = focus(encode t, root);
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
  fun term_of_fcterm (focus (t, m)) = mk_term_of_upt_aux (t, m);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   208
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   209
(*  fun term_of_fcterm (focus (t, root)) = mk_term_of_upt_aux (t, root)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   210
    | term_of_fcterm _ = raise focus_term_exp "term_of_fcterm: something bad has happened here."; *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   211
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   212
  fun focus_of_fcterm (focus(t, _)) = MinTermS.decode t;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   213
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   214
  (* replace the focus term with a new term... *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   215
  fun set_focus_of_fcterm (focus(_, m)) nt = focus(MinTermS.encode nt, m);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   216
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   217
  fun focus_left (focus(a $ b, m)) = focus(a, appr(b, m))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   218
    | focus_left (focus(Abs(s,ty,t), m)) = focus_left (focus(t, abs(s,ty,m)))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   219
    | focus_left (focus(l, m)) = raise out_of_term_exception "focus_left";
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   220
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15632
diff changeset
   221
  fun focus_strict_left (focus(a $ b, m)) = focus(a, appr(b, m))
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15632
diff changeset
   222
    | focus_strict_left (focus(l, m)) = raise out_of_term_exception "focus_left";
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15632
diff changeset
   223
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   224
(* Note: ":" is standard for faked bound vars... see: EQStepTacTermTools *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   225
  fun focus_fake_abs (focus(Abs(s,ty,t), m)) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   226
      let val t' = MinTermS.encode (fakebounds (s,ty) (MinTermS.decode t))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   227
      in focus(t', abs(s,ty,m)) end
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   228
    | focus_fake_abs (focus(l, m)) = raise out_of_term_exception "focus_fake_abs";
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   229
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   230
  fun focus_abs (focus(Abs(s,ty,t), m)) = focus(t, abs(s,ty,m))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   231
    | focus_abs (focus(l, m)) = raise out_of_term_exception "focus_abs";
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   232
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   233
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   234
  fun focus_right (focus(a $ b, m)) = focus(b, appl(a, m))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   235
    | focus_right (focus(Abs(s,ty,t), m)) = focus_right (focus(t, abs(s,ty,m)))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   236
    | focus_right (focus(l, m)) = raise out_of_term_exception "focus_right";
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   237
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   238
  fun focus_up (focus(t, appl(l,m))) = focus(l$t, m)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   239
    | focus_up (focus(t, appr(r,m))) = focus(t$r, m)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   240
    | focus_up (focus(t, abs(s,ty,m))) = focus_up (focus(Abs(s,ty,t), m))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   241
    | focus_up (focus(t, root)) = raise out_of_term_exception "focus_up";
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
  fun focus_to_top t = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   244
    focus_to_top (focus_up t) handle out_of_term_exception _ => t;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   245
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   246
  (* focus up until you can move right, and then do so *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   247
  fun focus_up_right (focus(t, appl(l,m))) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   248
      focus_up_right (focus(l$t, m))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   249
    | focus_up_right (focus(t, appr(r,m))) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   250
      focus(r, appl(t,m))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   251
    | focus_up_right (focus(t, abs(s,ty,m))) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   252
      focus_up_right (focus(Abs(s,ty,t), m))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   253
    | focus_up_right (focus(t, root)) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   254
      raise out_of_term_exception "focus_up_right";
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   255
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   256
  (* as above, but do not move up over a left application *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   257
  fun focus_up_right1 (focus(t, appl(l,m))) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   258
      raise out_of_term_exception "focus_up_right1"
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   259
    | focus_up_right1 (focus(t, appr(r,m))) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   260
      focus(r, appl(t,m))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   261
    | focus_up_right1 (focus(t, abs(s,ty,m))) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   262
      focus_up_right (focus(Abs(s,ty,t), m))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   263
    | focus_up_right1 (focus(t, root)) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   264
      raise out_of_term_exception "focus_up_right1";
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   265
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15632
diff changeset
   266
  (* move focus to the bottom left through abstractions *)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   267
  fun focus_bot_left_leaf (ft as focus(t, _)) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   268
        focus_bot_left_leaf (focus_left ft) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   269
        handle out_of_term_exception  _=> ft;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   270
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15632
diff changeset
   271
  (* move focus to the bottom left, but not into abstractions *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15632
diff changeset
   272
  fun focus_bot_left_nonabs_leaf (ft as focus(t, _)) = 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15632
diff changeset
   273
        focus_bot_left_nonabs_leaf (focus_strict_left ft) 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15632
diff changeset
   274
        handle out_of_term_exception  _=> ft;
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15632
diff changeset
   275
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   276
  (* focus tools for working directly with the focus representation *)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   277
  fun focus_up_appr (focus(t, appl(l,m))) = NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   278
    | focus_up_appr (focus(t, appr(r,m))) = SOME (focus(t$r, m))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   279
    | focus_up_appr (focus(t, abs(s,ty,m))) = NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   280
    | focus_up_appr (focus(t, root)) = NONE;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   281
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   282
  fun focus_up_appl (focus(t, appl(l,m))) = SOME (focus(l$t, m))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   283
    | focus_up_appl (focus(t, appr(r,m))) = NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   284
    | focus_up_appl (focus(t, abs(s,ty,m))) = NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   285
    | focus_up_appl (focus(t, root)) = NONE;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   286
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   287
  fun focus_up_abs (focus(t, appl(l,m))) = NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   288
    | focus_up_abs (focus(t, appr(r,m))) = NONE
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   289
    | focus_up_abs (focus(t, abs(s,ty,m))) = 
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   290
      SOME (focus_up (focus(Abs(s,ty,t), m)))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   291
    | focus_up_abs (focus(t, root)) = NONE;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   292
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   293
  fun focus_up_abs_or_appr (focus(t, appl(l,m))) = NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   294
    | focus_up_abs_or_appr (focus(t, appr(r,m))) = SOME (focus(t$r, m))
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   295
    | focus_up_abs_or_appr (focus(t, abs(s,ty,m))) = 
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   296
      SOME (focus_up (focus(Abs(s,ty,t), m)))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   297
    | focus_up_abs_or_appr (focus(t, root)) = NONE;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   298
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   299
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   300
  fun tyenv_of_focus (focus(t, u)) = tyenv_of_upterm u;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   301
  fun tyenv_of_focus' (focus(t, u)) = tyenv_of_upterm' u;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   302
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   303
  (* return the Term.type of the focus term, computes bound vars type,
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   304
     does a quick check only. *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   305
(*  fun fastype_of_focus (focus(t, m)) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   306
       let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   307
         fun boundtypes_of_upterm (abs(s,ty,m)) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   308
         ty::(boundtypes_of_upterm m)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   309
           | boundtypes_of_upterm (appl(t,m)) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   310
         boundtypes_of_upterm m
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   311
           | boundtypes_of_upterm (appr(t,m)) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   312
         boundtypes_of_upterm m
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   313
           | boundtypes_of_upterm (root) = []
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   314
       in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   315
         fastype_of1 ((boundtypes_of_upterm m), t)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   316
       end; *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   317
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   318
  (* gets the next left hand side term, or throws an exception *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   319
  fun next_leaf_fcterm ft = focus_bot_left_leaf (focus_up_right ft);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   320
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   321
  fun next_leaf_of_fcterm_seq_aux t () = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   322
    let val nextt = next_leaf_fcterm t
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   323
    in 
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   324
        SOME (nextt, Seq.make (next_leaf_of_fcterm_seq_aux nextt))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   325
    end handle out_of_term_exception _ => NONE;
15481
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
  (* sequence of next upterms from start upterm... 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   328
     ie a sequence of all leafs after this one*)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   329
  fun next_leaf_of_fcterm_seq in_t = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   330
    Seq.make (next_leaf_of_fcterm_seq_aux in_t);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   331
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   332
  (* returns a sequence of all leaf up terms from term, ie DFS on 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   333
   leafs of term, ie uses time O(n), n = sizeof term. *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   334
  fun leaf_seq_of_fcterm in_t = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   335
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   336
      val botleft = (focus_bot_left_leaf in_t)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   337
    in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   338
      Seq.cons (botleft, (Seq.make (next_leaf_of_fcterm_seq_aux botleft)) )
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   339
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   340
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   341
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   342
  fun upsize_of_fcterm (focus(t, ut)) = upsize_of_upterm ut;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   343
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   344
  fun pretty_fcterm ft = Pretty.str "no yet implemented";
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   345
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   346
end; (* local *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   347
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   348
end; (* functor *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   349
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
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
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   355
(* focus term encoding sturcture for isabelle terms *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   356
structure EncodeIsaFTerm : F_ENCODE_TERM_SIG = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   357
struct
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   358
  infix 9 $;
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
  type term = Term.term;
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
  type TypeT = Term.typ;
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
  datatype LeafT = lConst of string * Term.typ
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   365
     | lVar of ((string * int) * Term.typ)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   366
     | lFree of (string * Term.typ)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   367
     | lBound of int;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   368
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   369
  datatype TermT = op $ of TermT * TermT
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   370
    | Abs of string * TypeT * TermT
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   371
    | lf of LeafT;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   372
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   373
  fun fakebounds (s, ty) t = subst_bound (Free(":" ^ s, ty), t);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   374
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   375
  fun encode (Term.$(a,b)) = (encode a) $ (encode b)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   376
    | encode (Term.Abs(s,ty,t)) = Abs(s,ty,(encode t))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   377
    | encode (Term.Const(s,ty)) = lf(lConst(s,ty))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   378
    | encode (Term.Var((s,i),ty)) = lf(lVar((s,i),ty))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   379
    | encode (Term.Free(s,ty)) = lf(lFree(s,ty))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   380
    | encode (Term.Bound i) = lf(lBound i);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   381
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   382
  fun decode (a $ b) = Term.$(decode a, decode b)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   383
    | decode (Abs(s,ty,t)) = (Term.Abs(s,ty,decode t))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   384
    | decode (lf(lConst(s,ty))) = Term.Const(s,ty)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   385
    | decode (lf(lVar((s,i),ty))) = Term.Var((s,i),ty)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   386
    | decode (lf(lFree(s,ty))) = (Term.Free(s,ty))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   387
    | decode (lf(lBound i)) = (Term.Bound i);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   388
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   389
end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   390