src/Pure/IsaPlanner/isa_fterm.ML
author wenzelm
Sat, 11 Mar 2006 21:23:10 +0100
changeset 19250 932a50e2332f
parent 17795 5b18c3343028
child 19475 8aa2b380614a
permissions -rw-r--r--
got rid of type Sign.sg;
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: 15915
diff changeset
     2
(*  Title:      Pure/IsaPlanner/isa_fterm.ML
fa7e70be26b0 header;
wenzelm
parents: 15915
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) instantiation for Isabelle terms. 
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
17795
5b18c3343028 minor tweaks for Poplog/PML;
wenzelm
parents: 17203
diff changeset
    16
signature ISA_ENCODE_TERM =  (* cf. F_ENCODE_TERM_SIG *)
5b18c3343028 minor tweaks for Poplog/PML;
wenzelm
parents: 17203
diff changeset
    17
sig
5b18c3343028 minor tweaks for Poplog/PML;
wenzelm
parents: 17203
diff changeset
    18
  type term
5b18c3343028 minor tweaks for Poplog/PML;
wenzelm
parents: 17203
diff changeset
    19
  type typ
5b18c3343028 minor tweaks for Poplog/PML;
wenzelm
parents: 17203
diff changeset
    20
  type LeafT
5b18c3343028 minor tweaks for Poplog/PML;
wenzelm
parents: 17203
diff changeset
    21
  datatype TermT = $ of TermT * TermT
5b18c3343028 minor tweaks for Poplog/PML;
wenzelm
parents: 17203
diff changeset
    22
    | Abs of string * typ * TermT
5b18c3343028 minor tweaks for Poplog/PML;
wenzelm
parents: 17203
diff changeset
    23
    | lf of LeafT
5b18c3343028 minor tweaks for Poplog/PML;
wenzelm
parents: 17203
diff changeset
    24
  val fakebounds : string * typ -> term -> term
5b18c3343028 minor tweaks for Poplog/PML;
wenzelm
parents: 17203
diff changeset
    25
  val encode : term -> TermT
5b18c3343028 minor tweaks for Poplog/PML;
wenzelm
parents: 17203
diff changeset
    26
  val decode : TermT -> term
5b18c3343028 minor tweaks for Poplog/PML;
wenzelm
parents: 17203
diff changeset
    27
end;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    28
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    29
(* signature BASIC_ISA_FTERM = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    30
FOCUS_TERM_SIG where type Term = ISA_ENCODE_TERM.Term *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    31
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    32
signature ISA_FTERM =
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    33
  sig
17795
5b18c3343028 minor tweaks for Poplog/PML;
wenzelm
parents: 17203
diff changeset
    34
    type Term (*= EncodeIsaFTerm.TermT*)
5b18c3343028 minor tweaks for Poplog/PML;
wenzelm
parents: 17203
diff changeset
    35
    type Type (*= Term.typ*)
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    36
    type UpTerm
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    37
    datatype FcTerm = focus of Term * UpTerm
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    38
    structure MinTermS : F_ENCODE_TERM_SIG
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    39
    val add_upterm :
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    40
       UpTerm -> FcTerm -> FcTerm
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    41
    val clean_match_ft :
17203
29b2563f5c11 refer to theory instead of low-level tsig;
wenzelm
parents: 17045
diff changeset
    42
       theory ->
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    43
       Term.term ->
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    44
       FcTerm ->
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    45
       (
15915
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
    46
       ((Term.indexname * (Term.sort * Term.typ)) list 
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
    47
        * (Term.indexname * (Term.typ * Term.term)) list)
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
    48
       * (string * Term.typ) list 
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
    49
       * (string * Term.typ) list 
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
    50
       * Term.term)
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    51
       option
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    52
    val clean_unify_ft :
19250
932a50e2332f got rid of type Sign.sg;
wenzelm
parents: 17795
diff changeset
    53
       theory ->
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    54
       int ->
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    55
       Term.term ->
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    56
       FcTerm ->
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    57
       (
15915
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
    58
       ((Term.indexname * (Term.sort * Term.typ)) list 
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
    59
        * (Term.indexname * (Term.typ * Term.term)) list)
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    60
       * (string * Term.typ) list * (string * Term.typ) list * Term.term)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    61
       Seq.seq
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    62
    val enc_appl :
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    63
       EncodeIsaFTerm.term * UpTerm -> UpTerm
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    64
    val enc_appr :
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    65
       EncodeIsaFTerm.term * UpTerm -> UpTerm
17045
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
    66
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    67
    val fcterm_of_term : EncodeIsaFTerm.term -> FcTerm
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    68
    val find_fcterm_matches :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    69
       ((FcTerm -> 'a) -> FcTerm -> 'b) ->
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    70
       (FcTerm -> 'a) -> FcTerm -> 'b
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    71
    val find_sg_concl_matches :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    72
       ((FcTerm -> 'a) -> FcTerm -> 'b) ->
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    73
       (FcTerm -> 'a) -> int -> EncodeIsaFTerm.term -> 'b
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    74
    val find_sg_concl_thm_matches :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    75
       ((FcTerm -> 'a) -> FcTerm -> 'b) ->
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    76
       (FcTerm -> 'a) -> int -> Thm.thm -> 'b
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    77
    val find_sg_matches :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    78
       ((FcTerm -> 'a) -> FcTerm -> 'b) ->
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    79
       (FcTerm -> 'a) -> int -> EncodeIsaFTerm.term -> 'b
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    80
    val find_sg_thm_matches :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    81
       ((FcTerm -> 'a) -> FcTerm -> 'b) ->
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    82
       (FcTerm -> 'a) -> int -> Thm.thm -> 'b
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    83
    val focus_abs : FcTerm -> FcTerm
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    84
    val focus_bot_left_leaf : FcTerm -> FcTerm
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    85
    val focus_bot_left_nonabs_leaf :
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    86
       FcTerm -> FcTerm
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    87
    val focus_fake_abs : FcTerm -> FcTerm
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    88
    val focus_left : FcTerm -> FcTerm
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    89
    val focus_of_fcterm : FcTerm -> EncodeIsaFTerm.term
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    90
    val focus_right : FcTerm -> FcTerm
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    91
    val focus_strict_left : FcTerm -> FcTerm
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    92
    exception focus_term_exp of string
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    93
    val focus_to_concl : FcTerm -> FcTerm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    94
    val focus_to_concl_of_term : EncodeIsaFTerm.term -> FcTerm
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    95
    val focus_to_dest_impl :
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
    96
       FcTerm -> FcTerm * FcTerm
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    97
    val focus_to_subgoal :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    98
       int -> FcTerm -> FcTerm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    99
    val focus_to_subgoal_of_term :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   100
       int -> EncodeIsaFTerm.term -> FcTerm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   101
    val focus_to_term_goal_prem :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   102
       int * int -> EncodeIsaFTerm.term -> FcTerm
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   103
    val focus_to_top : FcTerm -> FcTerm
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   104
    val focus_up : FcTerm -> FcTerm
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   105
    val focus_up_abs : FcTerm -> FcTerm option
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   106
    val focus_up_abs_or_appr :
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   107
       FcTerm -> FcTerm option
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   108
    val focus_up_appl : FcTerm -> FcTerm option
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   109
    val focus_up_appr : FcTerm -> FcTerm option
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   110
    val focus_up_right : FcTerm -> FcTerm
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   111
    val focus_up_right1 : FcTerm -> FcTerm
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   112
    val focuseq_to_subgoals :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   113
       FcTerm -> FcTerm Seq.seq
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   114
    exception isa_focus_term_exp of string
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   115
    val leaf_seq_of_fcterm :
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   116
       FcTerm -> FcTerm Seq.seq
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   117
    val mk_term_of_upterm :
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   118
       EncodeIsaFTerm.term * UpTerm -> EncodeIsaFTerm.term
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   119
    val mk_termf_of_upterm :
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   120
       UpTerm ->
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   121
       (string * Type) list *
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   122
       (EncodeIsaFTerm.term -> EncodeIsaFTerm.term)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   123
    val next_leaf_fcterm : FcTerm -> FcTerm
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   124
    val next_leaf_of_fcterm_seq :
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   125
       FcTerm -> FcTerm Seq.seq
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   126
    exception out_of_term_exception of string
17045
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   127
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   128
    val pure_mk_termf_of_upterm :
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   129
       (EncodeIsaFTerm.term, Type) UpTermLib.T ->
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   130
       (string * Type) list *
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   131
       (EncodeIsaFTerm.term -> EncodeIsaFTerm.term)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   132
    val search_all_bl_ru_f :
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   133
       (FcTerm -> 'a Seq.seq) ->
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   134
       FcTerm -> 'a Seq.seq
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   135
    val search_all_bl_ur_f :
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   136
       (FcTerm -> 'a Seq.seq) ->
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   137
       FcTerm -> 'a Seq.seq
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   138
    val search_tlr_all_f :
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   139
       (FcTerm -> 'a Seq.seq) ->
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   140
       FcTerm -> 'a Seq.seq
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   141
    val search_tlr_valid_f :
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   142
       (FcTerm -> 'a Seq.seq) ->
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   143
       FcTerm -> 'a Seq.seq
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   144
    val search_valid_bl_ru_f :
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   145
       (FcTerm -> 'a Seq.seq) ->
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   146
       FcTerm -> 'a Seq.seq
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   147
    val search_valid_bl_ur_f :
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   148
       (FcTerm -> 'a Seq.seq) ->
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   149
       FcTerm -> 'a Seq.seq
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   150
    val set_focus_of_fcterm :
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   151
       FcTerm -> EncodeIsaFTerm.term -> FcTerm
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   152
    val term_of_fcterm : FcTerm -> EncodeIsaFTerm.term
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   153
    val tyenv_of_focus :
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   154
       FcTerm -> (string * Type) list
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   155
    val tyenv_of_focus' : FcTerm -> Type list
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   156
    val upsize_of_fcterm : FcTerm -> int
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   157
    val upterm_of : FcTerm -> UpTerm
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   158
    val valid_match_start : FcTerm -> bool
17045
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   159
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   160
     (* pre-matching/pre-unification *)
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   161
     val fakefree_badbounds : 
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   162
       (string * Term.typ) list -> Term.term ->
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   163
       (string * Term.typ) list * (string * Term.typ) list * Term.term
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   164
     val mk_foo_match :
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   165
       (Term.term -> Term.term) ->
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   166
       ('a * Term.typ) list -> Term.term -> Term.term
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   167
     val prepmatch :
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   168
       FcTerm ->
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   169
       Term.term *
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   170
       ((string * Term.typ) list * (string * Term.typ) list * Term.term)
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   171
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   172
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   173
    val pretty : Theory.theory -> FcTerm -> Pretty.T
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   174
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   175
  end;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   176
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   177
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   178
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   179
structure BasicIsaFTerm = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   180
FocusTermFUN( structure EncodeTerm = EncodeIsaFTerm);
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
(* HOL Dependent *)
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   183
structure IsaFTerm : ISA_FTERM=
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   184
struct 
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
(* include BasicIsaFTerm *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   187
(* open FocusTermFUN( structure EncodeTerm = EncodeIsaFTerm ); *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   188
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   189
open BasicIsaFTerm;
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
15661
9ef583b08647 reverted renaming of Some/None in comments and strings;
wenzelm
parents: 15531
diff changeset
   192
(* Some general search within a focus term... *)
15481
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
(* 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
   195
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
   196
important if we switch to a unification net! in particular to avoid
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   197
vars. *)
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
fun valid_match_start ft =
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   200
    (case TermLib.bot_left_leaf_of (focus_of_fcterm ft) of 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   201
       Const _ => true
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   202
     | Free _ => true
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   203
     | Abs _ => true (* allowed to look inside abs... search decides if we actually consider the abstraction itself *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   204
     | _ => false); (* avoid vars - always suceeds uninterestingly. *)
15481
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
(* 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
   207
   then moving right to the next leaf and up again etc *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   208
(* FIXME: make properly lazy! *)
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   209
fun search_valid_bl_ur_f f ft = 
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   210
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   211
      val fts = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   212
          Seq.filter valid_match_start
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   213
                     (leaf_seq_of_fcterm ft)
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
      fun mk_match_up_seq ft =
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   216
          (* avoid abstractions? - possibly infinite unifiers? *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   217
          let val hereseq = case (focus_of_fcterm ft) of 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   218
                              Abs _ => Seq.empty
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   219
                            | _ => f ft
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   220
          in
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   221
            case focus_up_abs_or_appr ft of 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   222
              SOME ft' => Seq.append(hereseq, mk_match_up_seq ft')
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   223
            | NONE => hereseq
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   224
          end
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   225
    in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   226
      Seq.flat (Seq.map mk_match_up_seq fts)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   227
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   228
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   229
fun search_all_bl_ur_f f ft = 
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   230
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   231
      val fts = (leaf_seq_of_fcterm ft)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   232
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   233
      fun mk_match_up_seq ft =            
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   234
          case focus_up_abs_or_appr ft of 
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   235
            SOME ft' => Seq.append(f ft, mk_match_up_seq ft')
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   236
          | NONE => f ft
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   237
    in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   238
      Seq.flat (Seq.map mk_match_up_seq fts)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   239
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   240
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
(* FIXME: make properly lazy! *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   243
(* FIXME: make faking of bound vars local - for speeeeed *)
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   244
fun search_valid_bl_ru_f f ft = 
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   245
    let
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   246
      fun mauxtop ft = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   247
          if (valid_match_start ft)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   248
            then maux ft else Seq.empty
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   249
      and maux ft = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   250
          let val t' = (focus_of_fcterm ft) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   251
          (* val _ = writeln ("Examining: " ^ (TermLib.string_of_term t')) *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   252
          in 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   253
          (case t' of 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   254
            (_ $ _) => Seq.append (maux (focus_left ft), 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   255
                         Seq.append(mauxtop (focus_right ft), 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   256
                                    f ft))
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   257
          | (Abs _) => maux (focus_abs ft)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   258
          | leaf => f ft) end
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   259
    in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   260
      mauxtop ft
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   261
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   262
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   263
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   264
fun search_all_bl_ru_f f ft = 
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   265
    let
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   266
      fun maux ft = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   267
          let val t' = (focus_of_fcterm ft) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   268
          (*   val _ = writeln ("Examining: " ^ (TermLib.string_of_term t')) *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   269
          in 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   270
          (case t' of 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   271
            (_ $ _) => Seq.append (maux (focus_left ft), 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   272
                         Seq.append(maux (focus_right ft), 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   273
                                    f ft))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   274
          | (Abs _) => Seq.append (maux (focus_abs ft), 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   275
                                   f ft)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   276
          | leaf => f ft) end
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   277
    in maux ft end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   278
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   279
(* search from top, left to right, then down *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   280
fun search_tlr_all_f f ft = 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   281
    let
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   282
      fun maux ft = 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   283
          let val t' = (focus_of_fcterm ft) 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   284
            (* val _ = 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   285
                if !trace_subst_search then 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   286
                  (writeln ("Examining: " ^ (TermLib.string_of_term t'));
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   287
                   TermLib.writeterm t'; ())
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   288
                else (); *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   289
          in 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   290
          (case t' of 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   291
            (_ $ _) => Seq.append(maux (focus_left ft), 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   292
                       Seq.append(f ft, 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   293
                                  maux (focus_right ft)))
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   294
          | (Abs _) => Seq.append(f ft, maux (focus_abs ft))
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   295
          | leaf => f ft) end
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   296
    in maux ft end;
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   297
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   298
fun search_tlr_valid_f f ft = 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   299
    let
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   300
      fun maux ft = 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   301
          let 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   302
            val hereseq = if valid_match_start ft then f ft else Seq.empty
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   303
          in 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   304
          (case (focus_of_fcterm ft) of 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   305
            (_ $ _) => Seq.append(maux (focus_left ft), 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   306
                       Seq.append(hereseq, 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   307
                                  maux (focus_right ft)))
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   308
          | (Abs _) => Seq.append(hereseq, maux (focus_abs ft))
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   309
          | leaf => hereseq)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   310
          end
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   311
    in maux ft end;
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   312
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   313
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   314
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   315
exception isa_focus_term_exp of string;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   316
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
fun focus_to_dest_impl ft = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   319
    let val (lhs, rhs) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   320
            Logic.dest_implies (focus_of_fcterm ft) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   321
    in (focus_left ft, focus_right ft) end
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   322
    handle Term.TERM _ => 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   323
           raise isa_focus_term_exp 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   324
                   "focus_to_dest_impl: applied to non implication";
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   325
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
(* move focus to the conlusion *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   328
fun focus_to_concl ft = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   329
    let val (lhs, rhs) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   330
            Logic.dest_implies (focus_of_fcterm ft) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   331
    in focus_to_concl (focus_right ft) end
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   332
    handle Term.TERM _ =>  ft;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   333
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   334
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
   335
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
(* give back sequence of focuses at different subgoals *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   339
(* FIXME: make truly lazy *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   340
fun focuseq_to_subgoals ft = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   341
    if (Logic.is_implies (focus_of_fcterm ft)) then 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   342
      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
   343
    else
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   344
      Seq.empty;
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
(* move focus to a specific subgoal, 0 is first  *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   347
fun focus_to_subgoal j ft = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   348
    let fun focus_to_subgoal' (ft, 0) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   349
            let val (lhs, rhs) = Logic.dest_implies (focus_of_fcterm ft) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   350
            in ft |> focus_left |> focus_right end
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   351
          | focus_to_subgoal' (ft, i) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   352
            let val (lhs, rhs) = Logic.dest_implies (focus_of_fcterm ft) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   353
            in focus_to_subgoal' (focus_right ft, i - 1) end
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   354
    in focus_to_subgoal' (ft, j - 1) end
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   355
    handle Term.TERM _ => 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   356
           raise isa_focus_term_exp 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   357
                   ("focus_to_subgoal: No such subgoal: " ^ 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   358
                    (string_of_int j));
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
fun focus_to_subgoal_of_term i t = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   361
    focus_to_subgoal i (fcterm_of_term t)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   362
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   363
(* move focus to a specific premise *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   364
(* fun focus_past_params i ft = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   365
    (focus_to_subgoal (focus_right ft, i))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   366
    handle isa_focus_term_exp _ => 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   367
           raise isa_focus_term_exp 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   368
             ("focus_to_prmise: No such premise: " ^ (string_of_int i)); *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   369
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   370
fun focus_to_term_goal_prem (premid,gaolid) t = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   371
    focus_to_subgoal premid (focus_to_subgoal_of_term gaolid t);
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
17045
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   374
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   375
(* FIXME: make a sturcture for holding free variable names and making
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   376
them distinct --- ? maybe part of the new prooftree datatype ?
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   377
Alos: use this to fix below... *)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   378
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   379
(* T is outer bound vars, n is number of locally bound vars *)
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   380
(* THINK: is order of Ts correct...? or reversed? *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   381
fun fakefree_badbounds Ts t = 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   382
    let val (FakeTs,Ts,newnames) = 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   383
            List.foldr (fn ((n,ty),(FakeTs,Ts,usednames)) => 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   384
                           let val newname = Term.variant usednames n
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   385
                           in ((RWTools.mk_fake_bound_name newname,ty)::FakeTs,
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   386
                               (newname,ty)::Ts, 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   387
                               newname::usednames) end)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   388
                       ([],[],[])
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   389
                       Ts
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   390
    in (FakeTs, Ts, Term.subst_bounds (map Free FakeTs, t)) end;
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   391
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   392
(* fun fakefree_badbounds usednames T n (a $ b) =  *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   393
(*     let val (usednames', T', a') = fakefree_badbounds usednames T n a *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   394
(*       val (usednames'', T'', b') = fakefree_badbounds usednames' T n b in  *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   395
(*       (usednames'', T'', a' $ b') end *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   396
(*   | fakefree_badbounds usednames T n (Abs(s,ty,t)) =  *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   397
(*     let val (usednames', T', t') = fakefree_badbounds usednames T (n + 1) t in *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   398
(*     (usednames', T', Abs(s,ty, t')) *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   399
(*   | fakefree_badbounds usednames T n (b as Bound i) =  *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   400
(*     let fun mkfake_bound j [] =  *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   401
(*             raise ERROR_MESSAGE "fakefree_badbounds: bound is outside of the known types!" *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   402
(*           | mkfake_bound 0 ((s,ty)::Ts) =  *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   403
(*             let val newname = Term.variant s usednames in *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   404
(*             usednames, (newname,ty) :: T,  Free (newname,ty) *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   405
(*           | mkfake_bound j (d::Ts) = mkfake_bound (j - 1) Ts *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   406
(*     in if n <= i then mkfake_bound (i - n) T else b end *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   407
(*   | fakefree_badbounds usednames T t = t; *)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   408
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   409
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   410
(* note: outerterm is the taget with the match replaced by a bound 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   411
         variable : ie: "P lhs" beocmes "%x. P x" 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   412
         insts is the types of instantiations of vars in lhs
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   413
         and typinsts is the type instantiations of types in the lhs
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   414
         Note: Final rule is the rule lifted into the ontext of the 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   415
         taget thm. *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   416
fun mk_foo_match mkuptermfunc Ts t = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   417
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   418
      val ty = Term.type_of t
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   419
      val bigtype = (rev (map snd Ts)) ---> ty
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   420
      fun mk_foo 0 t = t
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   421
        | mk_foo i t = mk_foo (i - 1) (t $ (Bound (i - 1)))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   422
      val num_of_bnds = (length Ts)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   423
      (* foo_term = "fooabs y0 ... yn" where y's are local bounds *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   424
      val foo_term = mk_foo num_of_bnds (Bound num_of_bnds)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   425
    in Abs("fooabs", bigtype, mkuptermfunc foo_term) end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   426
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   427
(* before matching we need to fake the bound vars that are missing an
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   428
abstraction. In this function we additionally construct the
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   429
abstraction environment, and an outer context term (with the focus
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   430
abstracted out) for use in rewriting with RWInst.rw *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   431
fun prepmatch ft = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   432
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   433
      val t = focus_of_fcterm ft  
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   434
      val Ts = tyenv_of_focus ft
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   435
      val (FakeTs', Ts', t') = fakefree_badbounds Ts t
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   436
      fun mktermf t = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   437
          term_of_fcterm (set_focus_of_fcterm ft t)
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   438
      val absterm = mk_foo_match mktermf Ts' t'
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   439
    in
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   440
      (t', (FakeTs', Ts', absterm))
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   441
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   442
17045
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   443
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   444
fun pretty thy ft = 
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   445
    let val (t', (_,_,absterm)) = prepmatch ft in 
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   446
      Pretty.chunks
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   447
        [ Pretty.block [Pretty.str "(Abs:",
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   448
                        Pretty.quote (Display.pretty_cterm (Thm.cterm_of thy absterm)),
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   449
                        Pretty.str ","],
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   450
          Pretty.block [Pretty.str " Foc:",
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   451
                        Pretty.quote (Display.pretty_cterm (Thm.cterm_of thy t')),
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   452
                        Pretty.str ")" ] ]
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   453
    end;
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   454
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   455
(* 
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   456
    Pretty.str "no yet implemented";
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   457
      Display.pretty_cterm (Thm.cterm_of t)
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   458
      Term.Free ("FOCUS", Term.type_of t)
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   459
*)
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   460
e108cd5b6986 lucas - added pretty printing function and cleaned up signature a little.
dixon
parents: 16179
diff changeset
   461
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   462
(* matching and unification for a focus term's focus *)
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   463
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   464
(* Note: Ts is a modified version of the original names of the outer
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   465
bound variables. New names have been introduced to make sure they are
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   466
unique w.r.t all names in the term and each other. usednames' is
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   467
oldnames + new names. *)
17203
29b2563f5c11 refer to theory instead of low-level tsig;
wenzelm
parents: 17045
diff changeset
   468
fun clean_match_ft thy pat ft = 
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   469
    let val (t, (FakeTs,Ts,absterm)) = prepmatch ft in
17203
29b2563f5c11 refer to theory instead of low-level tsig;
wenzelm
parents: 17045
diff changeset
   470
      case TermLib.clean_match thy (pat, t) of 
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   471
        NONE => NONE 
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   472
      | SOME insts => SOME (insts, FakeTs, Ts, absterm) end;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   473
(* ix = max var index *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   474
fun clean_unify_ft sgn ix pat ft = 
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   475
    let val (t, (FakeTs, Ts,absterm)) = prepmatch ft in
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   476
    Seq.map (fn insts => (insts, FakeTs, Ts, absterm)) 
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   477
            (TermLib.clean_unify sgn ix (t, pat)) end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   478
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   479
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   480
(* THINK: ? do we not need to incremement bound indices? *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   481
(* THINK: it should be the search which localisaes the search to the
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   482
current focus, not this hack in the matcher... ? *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   483
(* find matches below this particular focus term *) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   484
(* The search function is to find a match within a term... *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   485
(* 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
   486
searchf and the results are flattened to form a lazy list. *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   487
fun find_fcterm_matches searchf matcher ft = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   488
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   489
      val ftupterm = upterm_of ft
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   490
      val focusft = focus_of_fcterm ft
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   491
      val add_uptermf = add_upterm ftupterm 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   492
    in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   493
    searchf
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   494
      (fn ft' => matcher (add_uptermf ft'))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   495
      (fcterm_of_term focusft)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   496
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   497
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   498
(* FIXME: move argument orders for efficiency... 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   499
i.e. wenzel style val foofunc = x o y;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   500
*)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   501
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   502
(* find the matches inside subgoal i of th *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   503
fun find_sg_matches searchf matcher i t = 
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   504
    let 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15661
diff changeset
   505
      val subgoal_fcterm = focus_to_subgoal_of_term i t
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   506
    in find_fcterm_matches searchf matcher subgoal_fcterm end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   507
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   508
(* find the matches inside subgoal i of th *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   509
fun find_sg_thm_matches searchf matcher i th = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   510
    find_sg_matches searchf matcher i (Thm.prop_of th);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   511
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   512
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   513
(* find the matches inside subgoal i's conclusion of th *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   514
fun find_sg_concl_matches searchf matcher i t = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   515
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   516
      val subgoal_fcterm = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   517
          focus_to_concl (focus_to_subgoal_of_term i t)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   518
    in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   519
      find_fcterm_matches searchf matcher subgoal_fcterm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   520
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   521
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   522
(* find the matches inside subgoal i's conclusion of th *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   523
fun find_sg_concl_thm_matches searchf matcher i th = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   524
    find_sg_concl_matches searchf matcher i (Thm.prop_of th);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   525
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   526
end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   527
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   528
(* 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   529
test... 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   530
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   531
f_encode_isatermS.encode (read "P a");
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   532
isafocustermS.fcterm_of_term (read "f a");
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   533
isafocustermS.term_of_fcterm it;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   534
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   535
Goal "P b ==> P (suc b)";
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   536
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   537
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
   538
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   539
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
   540
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   541
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
   542
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   543
*)