src/Tools/eqsubst.ML
author wenzelm
Sat, 28 Feb 2009 14:02:12 +0100
changeset 30160 5f7b17941730
parent 29269 src/Provers/eqsubst.ML@5c25a2012975
child 30318 3d03190d2864
permissions -rw-r--r--
moved some generic tools to src/Tools/ -- src/Provers is essentially obsolete;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
30160
5f7b17941730 moved some generic tools to src/Tools/ -- src/Provers is essentially obsolete;
wenzelm
parents: 29269
diff changeset
     1
(*  Title:      Tools/eqsubst.ML
29269
5c25a2012975 moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
wenzelm
parents: 27809
diff changeset
     2
    Author:     Lucas Dixon, University of Edinburgh
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     3
18598
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
     4
A proof method to perform a substiution using an equation.
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
     5
*)
15538
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
     6
18591
04b9f2bf5a48 tuned EqSubst setup;
wenzelm
parents: 18011
diff changeset
     7
signature EQSUBST =
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     8
sig
29269
5c25a2012975 moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
wenzelm
parents: 27809
diff changeset
     9
  (* a type abbreviation for match information *)
19871
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    10
  type match =
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    11
       ((indexname * (sort * typ)) list (* type instantiations *)
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    12
        * (indexname * (typ * term)) list) (* term instantiations *)
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    13
       * (string * typ) list (* fake named type abs env *)
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    14
       * (string * typ) list (* type abs env *)
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    15
       * term (* outer term *)
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    16
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    17
  type searchinfo =
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    18
       theory
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    19
       * int (* maxidx *)
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    20
       * Zipper.T (* focusterm to search under *)
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    21
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    22
    exception eqsubst_occL_exp of
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    23
       string * int list * Thm.thm list * int * Thm.thm
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    24
    
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    25
    (* low level substitution functions *)
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    26
    val apply_subst_in_asm :
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    27
       int ->
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    28
       Thm.thm ->
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    29
       Thm.thm ->
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    30
       (Thm.cterm list * int * 'a * Thm.thm) * match -> Thm.thm Seq.seq
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    31
    val apply_subst_in_concl :
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    32
       int ->
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    33
       Thm.thm ->
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    34
       Thm.cterm list * Thm.thm ->
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    35
       Thm.thm -> match -> Thm.thm Seq.seq
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    36
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    37
    (* matching/unification within zippers *)
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    38
    val clean_match_z :
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    39
       Context.theory -> Term.term -> Zipper.T -> match option
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    40
    val clean_unify_z :
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    41
       Context.theory -> int -> Term.term -> Zipper.T -> match Seq.seq
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    42
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    43
    (* skipping things in seq seq's *)
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    44
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    45
   (* skipping non-empty sub-sequences but when we reach the end
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    46
      of the seq, remembering how much we have left to skip. *)
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    47
    datatype 'a skipseq = SkipMore of int
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    48
      | SkipSeq of 'a Seq.seq Seq.seq;
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    49
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    50
    val skip_first_asm_occs_search :
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    51
       ('a -> 'b -> 'c Seq.seq Seq.seq) ->
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    52
       'a -> int -> 'b -> 'c skipseq
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    53
    val skip_first_occs_search :
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    54
       int -> ('a -> 'b -> 'c Seq.seq Seq.seq) -> 'a -> 'b -> 'c Seq.seq
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    55
    val skipto_skipseq : int -> 'a Seq.seq Seq.seq -> 'a skipseq
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    56
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    57
    (* tactics *)
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    58
    val eqsubst_asm_tac :
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    59
       Proof.context ->
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    60
       int list -> Thm.thm list -> int -> Thm.thm -> Thm.thm Seq.seq
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    61
    val eqsubst_asm_tac' :
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    62
       Proof.context ->
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    63
       (searchinfo -> int -> Term.term -> match skipseq) ->
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    64
       int -> Thm.thm -> int -> Thm.thm -> Thm.thm Seq.seq
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    65
    val eqsubst_tac :
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    66
       Proof.context ->
22727
473c7f67c64f Improved comments.
dixon
parents: 22578
diff changeset
    67
       int list -> (* list of occurences to rewrite, use [0] for any *)
473c7f67c64f Improved comments.
dixon
parents: 22578
diff changeset
    68
       Thm.thm list -> int -> Thm.thm -> Thm.thm Seq.seq
19871
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    69
    val eqsubst_tac' :
22727
473c7f67c64f Improved comments.
dixon
parents: 22578
diff changeset
    70
       Proof.context -> (* proof context *)
473c7f67c64f Improved comments.
dixon
parents: 22578
diff changeset
    71
       (searchinfo -> Term.term -> match Seq.seq) (* search function *)
473c7f67c64f Improved comments.
dixon
parents: 22578
diff changeset
    72
       -> Thm.thm (* equation theorem to rewrite with *)
473c7f67c64f Improved comments.
dixon
parents: 22578
diff changeset
    73
       -> int (* subgoal number in goal theorem *)
473c7f67c64f Improved comments.
dixon
parents: 22578
diff changeset
    74
       -> Thm.thm (* goal theorem *)
473c7f67c64f Improved comments.
dixon
parents: 22578
diff changeset
    75
       -> Thm.thm Seq.seq (* rewritten goal theorem *)
19871
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    76
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    77
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    78
    val fakefree_badbounds :
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    79
       (string * Term.typ) list ->
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    80
       Term.term ->
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    81
       (string * Term.typ) list * (string * Term.typ) list * Term.term
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    82
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    83
    val mk_foo_match :
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    84
       (Term.term -> Term.term) ->
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    85
       ('a * Term.typ) list -> Term.term -> Term.term
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    86
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    87
    (* preparing substitution *)
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    88
    val prep_meta_eq : Proof.context -> Thm.thm -> Thm.thm list
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    89
    val prep_concl_subst :
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    90
       int -> Thm.thm -> (Thm.cterm list * Thm.thm) * searchinfo
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    91
    val prep_subst_in_asm :
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    92
       int -> Thm.thm -> int ->
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    93
       (Thm.cterm list * int * int * Thm.thm) * searchinfo
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    94
    val prep_subst_in_asms :
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    95
       int -> Thm.thm ->
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    96
       ((Thm.cterm list * int * int * Thm.thm) * searchinfo) list
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    97
    val prep_zipper_match :
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    98
       Zipper.T -> Term.term * ((string * Term.typ) list * (string * Term.typ) list * Term.term)
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
    99
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
   100
    (* search for substitutions *)
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
   101
    val valid_match_start : Zipper.T -> bool
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
   102
    val search_lr_all : Zipper.T -> Zipper.T Seq.seq
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
   103
    val search_lr_valid : (Zipper.T -> bool) -> Zipper.T -> Zipper.T Seq.seq
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
   104
    val searchf_lr_unify_all :
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
   105
       searchinfo -> Term.term -> match Seq.seq Seq.seq
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
   106
    val searchf_lr_unify_valid :
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
   107
       searchinfo -> Term.term -> match Seq.seq Seq.seq
23064
6ee131d1a618 add a bottom up search function
narboux
parents: 22727
diff changeset
   108
    val searchf_bt_unify_valid :
6ee131d1a618 add a bottom up search function
narboux
parents: 22727
diff changeset
   109
       searchinfo -> Term.term -> match Seq.seq Seq.seq
19871
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
   110
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
   111
    (* syntax tools *)
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
   112
    val ith_syntax : Args.T list -> int list * Args.T list
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
   113
    val options_syntax : Args.T list -> bool * Args.T list
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
   114
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
   115
    (* Isar level hooks *)
20289
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 20071
diff changeset
   116
    val eqsubst_asm_meth : Proof.context -> int list -> Thm.thm list -> Proof.method
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 20071
diff changeset
   117
    val eqsubst_meth : Proof.context -> int list -> Thm.thm list -> Proof.method
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 20071
diff changeset
   118
    val subst_meth : Method.src -> Proof.context -> Proof.method
19871
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
   119
    val setup : theory -> theory
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
   120
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   121
end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   122
19835
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   123
structure EqSubst
19871
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
   124
: EQSUBST
19835
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   125
= struct
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   126
19835
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   127
structure Z = Zipper;
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   128
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   129
(* changes object "=" to meta "==" which prepares a given rewrite rule *)
18598
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
   130
fun prep_meta_eq ctxt =
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
   131
  let val (_, {mk_rews = {mk, ...}, ...}) = Simplifier.rep_ss (Simplifier.local_simpset_of ctxt)
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
   132
  in mk #> map Drule.zero_var_indexes end;
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
   133
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   134
15915
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15855
diff changeset
   135
  (* a type abriviation for match information *)
16978
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   136
  type match =
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   137
       ((indexname * (sort * typ)) list (* type instantiations *)
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   138
        * (indexname * (typ * term)) list) (* term instantiations *)
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   139
       * (string * typ) list (* fake named type abs env *)
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   140
       * (string * typ) list (* type abs env *)
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   141
       * term (* outer term *)
15550
806214035275 lucas - added more comments and an extra type to clarify the code.
dixon
parents: 15538
diff changeset
   142
16978
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   143
  type searchinfo =
18598
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
   144
       theory
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   145
       * int (* maxidx *)
19835
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   146
       * Zipper.T (* focusterm to search under *)
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   147
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   148
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   149
(* skipping non-empty sub-sequences but when we reach the end
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   150
   of the seq, remembering how much we have left to skip. *)
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   151
datatype 'a skipseq = SkipMore of int
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   152
  | SkipSeq of 'a Seq.seq Seq.seq;
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   153
(* given a seqseq, skip the first m non-empty seq's, note deficit *)
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   154
fun skipto_skipseq m s = 
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   155
    let 
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   156
      fun skip_occs n sq = 
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   157
          case Seq.pull sq of 
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   158
            NONE => SkipMore n
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   159
          | SOME (h,t) => 
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   160
            (case Seq.pull h of NONE => skip_occs n t
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   161
             | SOME _ => if n <= 1 then SkipSeq (Seq.cons h t)
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   162
                         else skip_occs (n - 1) t)
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   163
    in (skip_occs m s) end;
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   164
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   165
(* note: outerterm is the taget with the match replaced by a bound 
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   166
         variable : ie: "P lhs" beocmes "%x. P x" 
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   167
         insts is the types of instantiations of vars in lhs
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   168
         and typinsts is the type instantiations of types in the lhs
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   169
         Note: Final rule is the rule lifted into the ontext of the 
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   170
         taget thm. *)
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   171
fun mk_foo_match mkuptermfunc Ts t = 
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   172
    let 
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   173
      val ty = Term.type_of t
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   174
      val bigtype = (rev (map snd Ts)) ---> ty
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   175
      fun mk_foo 0 t = t
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   176
        | mk_foo i t = mk_foo (i - 1) (t $ (Bound (i - 1)))
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   177
      val num_of_bnds = (length Ts)
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   178
      (* foo_term = "fooabs y0 ... yn" where y's are local bounds *)
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   179
      val foo_term = mk_foo num_of_bnds (Bound num_of_bnds)
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   180
    in Abs("fooabs", bigtype, mkuptermfunc foo_term) end;
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   181
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   182
(* T is outer bound vars, n is number of locally bound vars *)
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   183
(* THINK: is order of Ts correct...? or reversed? *)
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   184
fun fakefree_badbounds Ts t = 
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   185
    let val (FakeTs,Ts,newnames) = 
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   186
            List.foldr (fn ((n,ty),(FakeTs,Ts,usednames)) => 
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19975
diff changeset
   187
                           let val newname = Name.variant usednames n
19835
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   188
                           in ((RWTools.mk_fake_bound_name newname,ty)::FakeTs,
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   189
                               (newname,ty)::Ts, 
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   190
                               newname::usednames) end)
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   191
                       ([],[],[])
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   192
                       Ts
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   193
    in (FakeTs, Ts, Term.subst_bounds (map Free FakeTs, t)) end;
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   194
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   195
(* before matching we need to fake the bound vars that are missing an
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   196
abstraction. In this function we additionally construct the
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   197
abstraction environment, and an outer context term (with the focus
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   198
abstracted out) for use in rewriting with RWInst.rw *)
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   199
fun prep_zipper_match z = 
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   200
    let 
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   201
      val t = Z.trm z  
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   202
      val c = Z.ctxt z
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   203
      val Ts = Z.C.nty_ctxt c
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   204
      val (FakeTs', Ts', t') = fakefree_badbounds Ts t
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   205
      val absterm = mk_foo_match (Z.C.apply c) Ts' t'
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   206
    in
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   207
      (t', (FakeTs', Ts', absterm))
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   208
    end;
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   209
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   210
(* Matching and Unification with exception handled *)
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   211
fun clean_match thy (a as (pat, t)) =
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   212
  let val (tyenv, tenv) = Pattern.match thy a (Vartab.empty, Vartab.empty)
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   213
  in SOME (Vartab.dest tyenv, Vartab.dest tenv)
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   214
  end handle Pattern.MATCH => NONE;
27033
6ef5134fc631 fixed bug: maxidx was wrongly calculuated from term, now calculated
dixon
parents: 23064
diff changeset
   215
19835
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   216
(* given theory, max var index, pat, tgt; returns Seq of instantiations *)
27033
6ef5134fc631 fixed bug: maxidx was wrongly calculuated from term, now calculated
dixon
parents: 23064
diff changeset
   217
fun clean_unify thry ix (a as (pat, tgt)) =
19835
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   218
    let
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   219
      (* type info will be re-derived, maybe this can be cached
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   220
         for efficiency? *)
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   221
      val pat_ty = Term.type_of pat;
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   222
      val tgt_ty = Term.type_of tgt;
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   223
      (* is it OK to ignore the type instantiation info?
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   224
         or should I be using it? *)
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   225
      val typs_unify =
29269
5c25a2012975 moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
wenzelm
parents: 27809
diff changeset
   226
          SOME (Sign.typ_unify thry (pat_ty, tgt_ty) (Vartab.empty, ix))
19835
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   227
            handle Type.TUNIFY => NONE;
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   228
    in
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   229
      case typs_unify of
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   230
        SOME (typinsttab, ix2) =>
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   231
        let
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   232
      (* is it right to throw away the flexes?
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   233
         or should I be using them somehow? *)
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   234
          fun mk_insts env =
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   235
            (Vartab.dest (Envir.type_env env),
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   236
             Envir.alist_of env);
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   237
          val initenv = Envir.Envir {asol = Vartab.empty,
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   238
                                     iTs = typinsttab, maxidx = ix2};
27033
6ef5134fc631 fixed bug: maxidx was wrongly calculuated from term, now calculated
dixon
parents: 23064
diff changeset
   239
          val useq = Unify.smash_unifiers thry [a] initenv
19835
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   240
	            handle UnequalLengths => Seq.empty
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   241
		               | Term.TERM _ => Seq.empty;
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   242
          fun clean_unify' useq () =
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   243
              (case (Seq.pull useq) of
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   244
                 NONE => NONE
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   245
               | SOME (h,t) => SOME (mk_insts h, Seq.make (clean_unify' t)))
27033
6ef5134fc631 fixed bug: maxidx was wrongly calculuated from term, now calculated
dixon
parents: 23064
diff changeset
   246
	            handle UnequalLengths => NONE
6ef5134fc631 fixed bug: maxidx was wrongly calculuated from term, now calculated
dixon
parents: 23064
diff changeset
   247
                   | Term.TERM _ => NONE
19835
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   248
        in
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   249
          (Seq.make (clean_unify' useq))
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   250
        end
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   251
      | NONE => Seq.empty
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   252
    end;
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   253
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   254
(* Matching and Unification for zippers *)
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   255
(* Note: Ts is a modified version of the original names of the outer
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   256
bound variables. New names have been introduced to make sure they are
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   257
unique w.r.t all names in the term and each other. usednames' is
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   258
oldnames + new names. *)
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   259
fun clean_match_z thy pat z = 
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   260
    let val (t, (FakeTs,Ts,absterm)) = prep_zipper_match z in
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   261
      case clean_match thy (pat, t) of 
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   262
        NONE => NONE 
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   263
      | SOME insts => SOME (insts, FakeTs, Ts, absterm) end;
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   264
(* ix = max var index *)
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   265
fun clean_unify_z sgn ix pat z = 
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   266
    let val (t, (FakeTs, Ts,absterm)) = prep_zipper_match z in
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   267
    Seq.map (fn insts => (insts, FakeTs, Ts, absterm)) 
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   268
            (clean_unify sgn ix (t, pat)) end;
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   269
15550
806214035275 lucas - added more comments and an extra type to clarify the code.
dixon
parents: 15538
diff changeset
   270
15538
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   271
(* FOR DEBUGGING...
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   272
type trace_subst_errT = int (* subgoal *)
16978
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   273
        * thm (* thm with all goals *)
15538
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   274
        * (Thm.cterm list (* certified free var placeholders for vars *)
16978
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   275
           * thm)  (* trivial thm of goal concl *)
15538
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   276
            (* possible matches/unifiers *)
16978
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   277
        * thm (* rule *)
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   278
        * (((indexname * typ) list (* type instantiations *)
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   279
              * (indexname * term) list ) (* term instantiations *)
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   280
             * (string * typ) list (* Type abs env *)
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   281
             * term) (* outer term *);
15538
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   282
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   283
val trace_subst_err = (ref NONE : trace_subst_errT option ref);
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   284
val trace_subst_search = ref false;
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   285
exception trace_subst_exp of trace_subst_errT;
19835
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   286
*)
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   287
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   288
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   289
fun bot_left_leaf_of (l $ r) = bot_left_leaf_of l
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   290
  | bot_left_leaf_of (Abs(s,ty,t)) = bot_left_leaf_of t
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   291
  | bot_left_leaf_of x = x;
15538
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   292
19975
ecd684d62808 fix to subst in order to allow subst when head of a term is a bound variable.
dixon
parents: 19871
diff changeset
   293
(* Avoid considering replacing terms which have a var at the head as
ecd684d62808 fix to subst in order to allow subst when head of a term is a bound variable.
dixon
parents: 19871
diff changeset
   294
   they always succeed trivially, and uninterestingly. *)
19835
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   295
fun valid_match_start z =
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   296
    (case bot_left_leaf_of (Z.trm z) of 
19975
ecd684d62808 fix to subst in order to allow subst when head of a term is a bound variable.
dixon
parents: 19871
diff changeset
   297
      Var _ => false 
ecd684d62808 fix to subst in order to allow subst when head of a term is a bound variable.
dixon
parents: 19871
diff changeset
   298
      | _ => true);
ecd684d62808 fix to subst in order to allow subst when head of a term is a bound variable.
dixon
parents: 19871
diff changeset
   299
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   300
(* search from top, left to right, then down *)
19871
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
   301
val search_lr_all = ZipperSearch.all_bl_ur;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   302
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   303
(* search from top, left to right, then down *)
19871
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
   304
fun search_lr_valid validf =
19835
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   305
    let 
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   306
      fun sf_valid_td_lr z = 
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   307
          let val here = if validf z then [Z.Here z] else [] in
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   308
            case Z.trm z 
19871
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
   309
             of _ $ _ => [Z.LookIn (Z.move_down_left z)] 
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
   310
                         @ here 
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
   311
                         @ [Z.LookIn (Z.move_down_right z)]
19835
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   312
              | Abs _ => here @ [Z.LookIn (Z.move_down_abs z)]
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   313
              | _ => here
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   314
          end;
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   315
    in Z.lzy_search sf_valid_td_lr end;
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   316
23064
6ee131d1a618 add a bottom up search function
narboux
parents: 22727
diff changeset
   317
(* search from bottom to top, left to right *)
6ee131d1a618 add a bottom up search function
narboux
parents: 22727
diff changeset
   318
6ee131d1a618 add a bottom up search function
narboux
parents: 22727
diff changeset
   319
fun search_bt_valid validf =
6ee131d1a618 add a bottom up search function
narboux
parents: 22727
diff changeset
   320
    let 
6ee131d1a618 add a bottom up search function
narboux
parents: 22727
diff changeset
   321
      fun sf_valid_td_lr z = 
6ee131d1a618 add a bottom up search function
narboux
parents: 22727
diff changeset
   322
          let val here = if validf z then [Z.Here z] else [] in
6ee131d1a618 add a bottom up search function
narboux
parents: 22727
diff changeset
   323
            case Z.trm z 
6ee131d1a618 add a bottom up search function
narboux
parents: 22727
diff changeset
   324
             of _ $ _ => [Z.LookIn (Z.move_down_left z), 
6ee131d1a618 add a bottom up search function
narboux
parents: 22727
diff changeset
   325
                          Z.LookIn (Z.move_down_right z)] @ here
6ee131d1a618 add a bottom up search function
narboux
parents: 22727
diff changeset
   326
              | Abs _ => [Z.LookIn (Z.move_down_abs z)] @ here
6ee131d1a618 add a bottom up search function
narboux
parents: 22727
diff changeset
   327
              | _ => here
6ee131d1a618 add a bottom up search function
narboux
parents: 22727
diff changeset
   328
          end;
6ee131d1a618 add a bottom up search function
narboux
parents: 22727
diff changeset
   329
    in Z.lzy_search sf_valid_td_lr end;
6ee131d1a618 add a bottom up search function
narboux
parents: 22727
diff changeset
   330
6ee131d1a618 add a bottom up search function
narboux
parents: 22727
diff changeset
   331
fun searchf_unify_gen f (sgn, maxidx, z) lhs =
6ee131d1a618 add a bottom up search function
narboux
parents: 22727
diff changeset
   332
    Seq.map (clean_unify_z sgn maxidx lhs) 
6ee131d1a618 add a bottom up search function
narboux
parents: 22727
diff changeset
   333
            (Z.limit_apply f z);
6ee131d1a618 add a bottom up search function
narboux
parents: 22727
diff changeset
   334
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   335
(* search all unifications *)
23064
6ee131d1a618 add a bottom up search function
narboux
parents: 22727
diff changeset
   336
val searchf_lr_unify_all =
6ee131d1a618 add a bottom up search function
narboux
parents: 22727
diff changeset
   337
    searchf_unify_gen search_lr_all;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   338
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   339
(* search only for 'valid' unifiers (non abs subterms and non vars) *)
23064
6ee131d1a618 add a bottom up search function
narboux
parents: 22727
diff changeset
   340
val searchf_lr_unify_valid = 
6ee131d1a618 add a bottom up search function
narboux
parents: 22727
diff changeset
   341
    searchf_unify_gen (search_lr_valid valid_match_start);
15929
68bd1e16552a lucas - added option to select occurance to rewrite e.g. (occ 4)
dixon
parents: 15915
diff changeset
   342
23064
6ee131d1a618 add a bottom up search function
narboux
parents: 22727
diff changeset
   343
val searchf_bt_unify_valid =
6ee131d1a618 add a bottom up search function
narboux
parents: 22727
diff changeset
   344
    searchf_unify_gen (search_bt_valid valid_match_start);
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   345
15538
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   346
(* apply a substitution in the conclusion of the theorem th *)
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   347
(* cfvs are certified free var placeholders for goal params *)
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   348
(* conclthm is a theorem of for just the conclusion *)
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   349
(* m is instantiation/match information *)
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   350
(* rule is the equation for substitution *)
16978
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   351
fun apply_subst_in_concl i th (cfvs, conclthm) rule m =
15538
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   352
    (RWInst.rw m rule conclthm)
15855
55e443aa711d lucas - updated to reflect isand.ML update
dixon
parents: 15814
diff changeset
   353
      |> IsaND.unfix_frees cfvs
15915
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15855
diff changeset
   354
      |> RWInst.beta_eta_contract
15538
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   355
      |> (fn r => Tactic.rtac r i th);
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   356
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   357
(* substitute within the conclusion of goal i of gth, using a meta
15538
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   358
equation rule. Note that we assume rule has var indicies zero'd *)
16978
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   359
fun prep_concl_subst i gth =
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   360
    let
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   361
      val th = Thm.incr_indexes 1 gth;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   362
      val tgt_term = Thm.prop_of th;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   363
22578
b0eb5652f210 removed obsolete sign_of/sign_of_thm;
wenzelm
parents: 21879
diff changeset
   364
      val sgn = Thm.theory_of_thm th;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   365
      val ctermify = Thm.cterm_of sgn;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   366
      val trivify = Thm.trivial o ctermify;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   367
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   368
      val (fixedbody, fvs) = IsaND.fix_alls_term i tgt_term;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   369
      val cfvs = rev (map ctermify fvs);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   370
15538
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   371
      val conclterm = Logic.strip_imp_concl fixedbody;
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   372
      val conclthm = trivify conclterm;
27033
6ef5134fc631 fixed bug: maxidx was wrongly calculuated from term, now calculated
dixon
parents: 23064
diff changeset
   373
      val maxidx = Thm.maxidx_of th;
19835
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   374
      val ft = ((Z.move_down_right (* ==> *)
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   375
                 o Z.move_down_left (* Trueprop *)
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   376
                 o Z.mktop
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   377
                 o Thm.prop_of) conclthm)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   378
    in
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   379
      ((cfvs, conclthm), (sgn, maxidx, ft))
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   380
    end;
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
(* substitute using an object or meta level equality *)
18598
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
   383
fun eqsubst_tac' ctxt searchf instepthm i th =
16978
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   384
    let
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   385
      val (cvfsconclthm, searchinfo) = prep_concl_subst i th;
18598
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
   386
      val stepthms = Seq.of_list (prep_meta_eq ctxt instepthm);
15538
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   387
      fun rewrite_with_thm r =
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   388
          let val (lhs,_) = Logic.dest_equals (Thm.concl_of r);
18598
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
   389
          in searchf searchinfo lhs
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
   390
             |> Seq.maps (apply_subst_in_concl i th cvfsconclthm r) end;
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
   391
    in stepthms |> Seq.maps rewrite_with_thm end;
15538
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   392
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   393
19047
670ce193b618 used Tactic.distinct_subgoals_tac;
wenzelm
parents: 18988
diff changeset
   394
(* distinct subgoals *)
15959
366d39e95d3c lucas - fixed bug with uninstantiated type contexts in eqsubst and added the automatic removal of duplicate subgoals (when there are no flex-flex constraints)
dixon
parents: 15936
diff changeset
   395
fun distinct_subgoals th =
19047
670ce193b618 used Tactic.distinct_subgoals_tac;
wenzelm
parents: 18988
diff changeset
   396
  the_default th (SINGLE distinct_subgoals_tac th);
15538
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   397
19047
670ce193b618 used Tactic.distinct_subgoals_tac;
wenzelm
parents: 18988
diff changeset
   398
(* General substitution of multiple occurances using one of
15936
817ac93ee786 lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents: 15929
diff changeset
   399
   the given theorems*)
19835
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   400
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   401
16978
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   402
exception eqsubst_occL_exp of
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   403
          string * (int list) * (thm list) * int * thm;
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   404
fun skip_first_occs_search occ srchf sinfo lhs =
19835
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   405
    case (skipto_skipseq occ (srchf sinfo lhs)) of
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   406
      SkipMore _ => Seq.empty
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   407
    | SkipSeq ss => Seq.flat ss;
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   408
22727
473c7f67c64f Improved comments.
dixon
parents: 22578
diff changeset
   409
(* The occL is a list of integers indicating which occurence
473c7f67c64f Improved comments.
dixon
parents: 22578
diff changeset
   410
w.r.t. the search order, to rewrite. Backtracking will also find later
473c7f67c64f Improved comments.
dixon
parents: 22578
diff changeset
   411
occurences, but all earlier ones are skipped. Thus you can use [0] to
473c7f67c64f Improved comments.
dixon
parents: 22578
diff changeset
   412
just find all rewrites. *)
473c7f67c64f Improved comments.
dixon
parents: 22578
diff changeset
   413
18598
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
   414
fun eqsubst_tac ctxt occL thms i th =
15936
817ac93ee786 lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents: 15929
diff changeset
   415
    let val nprems = Thm.nprems_of th in
817ac93ee786 lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents: 15929
diff changeset
   416
      if nprems < i then Seq.empty else
16978
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   417
      let val thmseq = (Seq.of_list thms)
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   418
        fun apply_occ occ th =
18598
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
   419
            thmseq |> Seq.maps
19835
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   420
                    (fn r => eqsubst_tac' 
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   421
                               ctxt 
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   422
                               (skip_first_occs_search
19871
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
   423
                                  occ searchf_lr_unify_valid) r
15936
817ac93ee786 lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents: 15929
diff changeset
   424
                                 (i + ((Thm.nprems_of th) - nprems))
817ac93ee786 lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents: 15929
diff changeset
   425
                                 th);
16978
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   426
        val sortedoccL =
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   427
            Library.sort (Library.rev_order o Library.int_ord) occL;
15936
817ac93ee786 lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents: 15929
diff changeset
   428
      in
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   429
        Seq.map distinct_subgoals (Seq.EVERY (map apply_occ sortedoccL) th)
15936
817ac93ee786 lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents: 15929
diff changeset
   430
      end
15959
366d39e95d3c lucas - fixed bug with uninstantiated type contexts in eqsubst and added the automatic removal of duplicate subgoals (when there are no flex-flex constraints)
dixon
parents: 15936
diff changeset
   431
    end
366d39e95d3c lucas - fixed bug with uninstantiated type contexts in eqsubst and added the automatic removal of duplicate subgoals (when there are no flex-flex constraints)
dixon
parents: 15936
diff changeset
   432
    handle THM _ => raise eqsubst_occL_exp ("THM",occL,thms,i,th);
366d39e95d3c lucas - fixed bug with uninstantiated type contexts in eqsubst and added the automatic removal of duplicate subgoals (when there are no flex-flex constraints)
dixon
parents: 15936
diff changeset
   433
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   434
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   435
(* inthms are the given arguments in Isar, and treated as eqstep with
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   436
   the first one, then the second etc *)
18598
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
   437
fun eqsubst_meth ctxt occL inthms =
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 20289
diff changeset
   438
    Method.SIMPLE_METHOD' (eqsubst_tac ctxt occL inthms);
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   439
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   440
(* apply a substitution inside assumption j, keeps asm in the same place *)
16978
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   441
fun apply_subst_in_asm i th rule ((cfvs, j, ngoalprems, pth),m) =
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   442
    let
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   443
      val th2 = Thm.rotate_rule (j - 1) i th; (* put premice first *)
16978
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   444
      val preelimrule =
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   445
          (RWInst.rw m rule pth)
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21588
diff changeset
   446
            |> (Seq.hd o prune_params_tac)
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   447
            |> Thm.permute_prems 0 ~1 (* put old asm first *)
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   448
            |> IsaND.unfix_frees cfvs (* unfix any global params *)
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   449
            |> RWInst.beta_eta_contract; (* normal form *)
16978
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   450
  (*    val elimrule =
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   451
          preelimrule
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   452
            |> Tactic.make_elim (* make into elim rule *)
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   453
            |> Thm.lift_rule (th2, i); (* lift into context *)
16007
4dcccaa11a13 lucas - bugfix to subst in assumptions: fixed index error for conditional rules.
dixon
parents: 16004
diff changeset
   454
   *)
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   455
    in
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   456
      (* ~j because new asm starts at back, thus we subtract 1 *)
16007
4dcccaa11a13 lucas - bugfix to subst in assumptions: fixed index error for conditional rules.
dixon
parents: 16004
diff changeset
   457
      Seq.map (Thm.rotate_rule (~j) ((Thm.nprems_of rule) + i))
4dcccaa11a13 lucas - bugfix to subst in assumptions: fixed index error for conditional rules.
dixon
parents: 16004
diff changeset
   458
      (Tactic.dtac preelimrule i th2)
4dcccaa11a13 lucas - bugfix to subst in assumptions: fixed index error for conditional rules.
dixon
parents: 16004
diff changeset
   459
16978
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   460
      (* (Thm.bicompose
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   461
                 false (* use unification *)
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   462
                 (true, (* elim resolution *)
16007
4dcccaa11a13 lucas - bugfix to subst in assumptions: fixed index error for conditional rules.
dixon
parents: 16004
diff changeset
   463
                  elimrule, (2 + (Thm.nprems_of rule)) - ngoalprems)
4dcccaa11a13 lucas - bugfix to subst in assumptions: fixed index error for conditional rules.
dixon
parents: 16004
diff changeset
   464
                 i th2) *)
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   465
    end;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   466
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   467
15538
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   468
(* prepare to substitute within the j'th premise of subgoal i of gth,
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   469
using a meta-level equation. Note that we assume rule has var indicies
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   470
zero'd. Note that we also assume that premt is the j'th premice of
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   471
subgoal i of gth. Note the repetition of work done for each
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   472
assumption, i.e. this can be made more efficient for search over
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   473
multiple assumptions.  *)
16978
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   474
fun prep_subst_in_asm i gth j =
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   475
    let
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   476
      val th = Thm.incr_indexes 1 gth;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   477
      val tgt_term = Thm.prop_of th;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   478
22578
b0eb5652f210 removed obsolete sign_of/sign_of_thm;
wenzelm
parents: 21879
diff changeset
   479
      val sgn = Thm.theory_of_thm th;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   480
      val ctermify = Thm.cterm_of sgn;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   481
      val trivify = Thm.trivial o ctermify;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   482
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   483
      val (fixedbody, fvs) = IsaND.fix_alls_term i tgt_term;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   484
      val cfvs = rev (map ctermify fvs);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   485
18011
685d95c793ff cleaned up nth, nth_update, nth_map and nth_string functions
haftmann
parents: 17795
diff changeset
   486
      val asmt = nth (Logic.strip_imp_prems fixedbody) (j - 1);
15538
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   487
      val asm_nprems = length (Logic.strip_imp_prems asmt);
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   488
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   489
      val pth = trivify asmt;
27033
6ef5134fc631 fixed bug: maxidx was wrongly calculuated from term, now calculated
dixon
parents: 23064
diff changeset
   490
      val maxidx = Thm.maxidx_of th;
15538
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   491
19835
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   492
      val ft = ((Z.move_down_right (* trueprop *)
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   493
                 o Z.mktop
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   494
                 o Thm.prop_of) pth)
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   495
    in ((cfvs, j, asm_nprems, pth), (sgn, maxidx, ft)) end;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   496
15538
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   497
(* prepare subst in every possible assumption *)
16978
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   498
fun prep_subst_in_asms i gth =
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   499
    map (prep_subst_in_asm i gth)
19835
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   500
        ((fn l => Library.upto (1, length l))
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   501
           (Logic.prems_of_goal (Thm.prop_of gth) i));
15538
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   502
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   503
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   504
(* substitute in an assumption using an object or meta level equality *)
18598
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
   505
fun eqsubst_asm_tac' ctxt searchf skipocc instepthm i th =
16978
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   506
    let
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   507
      val asmpreps = prep_subst_in_asms i th;
18598
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
   508
      val stepthms = Seq.of_list (prep_meta_eq ctxt instepthm);
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   509
      fun rewrite_with_thm r =
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   510
          let val (lhs,_) = Logic.dest_equals (Thm.concl_of r)
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   511
            fun occ_search occ [] = Seq.empty
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   512
              | occ_search occ ((asminfo, searchinfo)::moreasms) =
16978
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   513
                (case searchf searchinfo occ lhs of
19835
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   514
                   SkipMore i => occ_search i moreasms
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   515
                 | SkipSeq ss =>
19861
620d90091788 tuned Seq/Envir/Unify interfaces;
wenzelm
parents: 19835
diff changeset
   516
                   Seq.append (Seq.map (Library.pair asminfo) (Seq.flat ss))
620d90091788 tuned Seq/Envir/Unify interfaces;
wenzelm
parents: 19835
diff changeset
   517
                               (occ_search 1 moreasms))
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   518
                              (* find later substs also *)
16978
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   519
          in
18598
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
   520
            occ_search skipocc asmpreps |> Seq.maps (apply_subst_in_asm i th r)
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   521
          end;
18598
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
   522
    in stepthms |> Seq.maps rewrite_with_thm end;
15538
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   523
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   524
16978
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   525
fun skip_first_asm_occs_search searchf sinfo occ lhs =
19835
81d6dc597559 added updated version of IsaPlanner and substitution.
dixon
parents: 19473
diff changeset
   526
    skipto_skipseq occ (searchf sinfo lhs);
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   527
18598
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
   528
fun eqsubst_asm_tac ctxt occL thms i th =
16978
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   529
    let val nprems = Thm.nprems_of th
15538
d8edf54cc28c lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents: 15486
diff changeset
   530
    in
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   531
      if nprems < i then Seq.empty else
16978
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   532
      let val thmseq = (Seq.of_list thms)
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   533
        fun apply_occ occK th =
18598
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
   534
            thmseq |> Seq.maps
16978
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   535
                    (fn r =>
18598
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
   536
                        eqsubst_asm_tac' ctxt (skip_first_asm_occs_search
19871
88e8f6173bab Corrected search order for zippers.
dixon
parents: 19861
diff changeset
   537
                                            searchf_lr_unify_valid) occK r
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   538
                                         (i + ((Thm.nprems_of th) - nprems))
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   539
                                         th);
16978
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   540
        val sortedoccs =
16004
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   541
            Library.sort (Library.rev_order o Library.int_ord) occL
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   542
      in
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   543
        Seq.map distinct_subgoals
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   544
                (Seq.EVERY (map apply_occ sortedoccs) th)
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   545
      end
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   546
    end
031f56012483 lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents: 15959
diff changeset
   547
    handle THM _ => raise eqsubst_occL_exp ("THM",occL,thms,i,th);
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   548
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   549
(* inthms are the given arguments in Isar, and treated as eqstep with
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   550
   the first one, then the second etc *)
18598
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
   551
fun eqsubst_asm_meth ctxt occL inthms =
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 20289
diff changeset
   552
    Method.SIMPLE_METHOD' (eqsubst_asm_tac ctxt occL inthms);
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   553
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   554
(* syntax for options, given "(asm)" will give back true, without
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   555
   gives back false *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   556
val options_syntax =
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   557
    (Args.parens (Args.$$$ "asm") >> (K true)) ||
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   558
     (Scan.succeed false);
15936
817ac93ee786 lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents: 15929
diff changeset
   559
15929
68bd1e16552a lucas - added option to select occurance to rewrite e.g. (occ 4)
dixon
parents: 15915
diff changeset
   560
val ith_syntax =
27809
a1e409db516b unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27033
diff changeset
   561
    Scan.optional (Args.parens (Scan.repeat OuterParse.nat)) [0];
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   562
18598
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
   563
(* combination method that takes a flag (true indicates that subst
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
   564
should be done to an assumption, false = apply to the conclusion of
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
   565
the goal) as well as the theorems to use *)
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
   566
fun subst_meth src =
18988
d6e5fa2ba8b8 Args/Attrib syntax: Context.generic;
wenzelm
parents: 18833
diff changeset
   567
  Method.syntax ((Scan.lift options_syntax) -- (Scan.lift ith_syntax) -- Attrib.thms) src
21879
a3efbae45735 switched argument order in *.syntax lifters
haftmann
parents: 21708
diff changeset
   568
  #> (fn (((asmflag, occL), inthms), ctxt) =>
18598
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
   569
    (if asmflag then eqsubst_asm_meth else eqsubst_meth) ctxt occL inthms);
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   570
18598
94d658871c98 prep_meta_eq: reuse mk_rews of local simpset;
wenzelm
parents: 18591
diff changeset
   571
16978
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   572
val setup =
18833
bead1a4e966b tuned comment;
wenzelm
parents: 18708
diff changeset
   573
  Method.add_method ("subst", subst_meth, "single-step substitution");
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   574
16978
e35b518bffc9 tuned signature;
wenzelm
parents: 16434
diff changeset
   575
end;