src/Provers/eqsubst.ML
author dixon
Tue, 26 Apr 2005 20:38:39 +0200
changeset 15855 55e443aa711d
parent 15814 d65f461c8672
child 15915 b0e8b37642a4
permissions -rw-r--r--
lucas - updated to reflect isand.ML update
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     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
     2
(*  Title:      Provers/eqsubst.ML
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     3
    Author:     Lucas Dixon, University of Edinburgh
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     4
                lucas.dixon@ed.ac.uk
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
     5
    Modified:   18 Feb 2005 - Lucas - 
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     6
    Created:    29 Jan 2005
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
    A Tactic to perform a substiution using an equation.
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
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
    16
(* Logic specific data stub *)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    17
signature EQRULE_DATA =
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    18
sig
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
    19
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    20
  (* to make a meta equality theorem in the current logic *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    21
  val prep_meta_eq : thm -> thm list
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
    22
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    23
end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    24
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
    25
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    26
(* the signature of an instance of the SQSUBST tactic *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    27
signature EQSUBST_TAC = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    28
sig
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
    29
15550
806214035275 lucas - added more comments and an extra type to clarify the code.
dixon
parents: 15538
diff changeset
    30
  type match = 
806214035275 lucas - added more comments and an extra type to clarify the code.
dixon
parents: 15538
diff changeset
    31
       ((Term.indexname * Term.typ) list (* type instantiations *)
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
    32
        * (Term.indexname * Term.term) list) (* term instantiations *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
    33
       * (string * Term.typ) list (* fake named type abs env *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
    34
       * (string * Term.typ) list (* type abs env *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
    35
       * Term.term (* outer term *)
15550
806214035275 lucas - added more comments and an extra type to clarify the code.
dixon
parents: 15538
diff changeset
    36
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
    37
  val prep_subst_in_asm :
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
    38
      (Sign.sg (* sign for matching *)
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
    39
       -> int (* maxidx *)
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
    40
       -> 'a (* input object kind *)
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
    41
       -> BasicIsaFTerm.FcTerm (* focusterm to search under *)
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
    42
       -> 'b) (* result type *)
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
    43
      -> int (* subgoal to subst in *)
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
    44
      -> Thm.thm (* target theorem with subgoals *)
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
    45
      -> int (* premise to subst in *)
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
    46
      -> (Thm.cterm list (* certified free var placeholders for vars *) 
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
    47
          * int (* premice no. to subst *)
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
    48
          * int (* number of assumptions of premice *)
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
    49
          * Thm.thm) (* premice as a new theorem for forward reasoning *)
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
    50
         * ('a -> 'b) (* matchf *)
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
    51
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
    52
  val prep_subst_in_asms :
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
    53
      (Sign.sg -> int -> 'a -> BasicIsaFTerm.FcTerm -> 'b) 
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
    54
      -> int (* subgoal to subst in *)
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
    55
      -> Thm.thm (* target theorem with subgoals *)
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
    56
      -> ((Thm.cterm list (* certified free var placeholders for vars *) 
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
    57
          * int (* premice no. to subst *)
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
    58
          * int (* number of assumptions of premice *)
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
    59
          * Thm.thm) (* premice as a new theorem for forward reasoning *)
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
    60
         * ('a -> 'b)) (* matchf *)
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
    61
                       Seq.seq
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
    62
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
    63
  val apply_subst_in_asm :
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
    64
      int (* subgoal *)
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
    65
      -> Thm.thm (* overall theorem *)
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
    66
      -> (Thm.cterm list (* certified free var placeholders for vars *) 
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
    67
          * int (* assump no being subst *)
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
    68
          * int (* num of premises of asm *) 
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
    69
          * Thm.thm) (* premthm *)
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
    70
      -> Thm.thm (* rule *)
15550
806214035275 lucas - added more comments and an extra type to clarify the code.
dixon
parents: 15538
diff changeset
    71
      -> match
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
    72
      -> Thm.thm Seq.seq
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
    73
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
    74
  val prep_concl_subst :
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
    75
      (Sign.sg -> int -> 'a -> BasicIsaFTerm.FcTerm -> 'b) (* searchf *) 
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
    76
      -> int (* subgoal *)
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
    77
      -> Thm.thm (* overall goal theorem *)
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
    78
      -> (Thm.cterm list * Thm.thm) * ('a -> 'b) (* (cvfs, conclthm), matchf *)
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
    79
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
    80
  val apply_subst_in_concl :
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
    81
        int (* subgoal *)
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
    82
        -> Thm.thm (* thm with all goals *)
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
    83
        -> Thm.cterm list (* certified free var placeholders for vars *)
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
    84
           * Thm.thm  (* trivial thm of goal concl *)
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
    85
            (* possible matches/unifiers *)
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
    86
        -> Thm.thm (* rule *)
15550
806214035275 lucas - added more comments and an extra type to clarify the code.
dixon
parents: 15538
diff changeset
    87
        -> match
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
    88
        -> Thm.thm Seq.seq (* substituted goal *)
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
    89
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
    90
  val searchf_tlr_unify_all : 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
    91
      (Sign.sg -> int ->
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
    92
       Term.term ->
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
    93
       BasicIsaFTerm.FcTerm ->
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
    94
       match Seq.seq)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
    95
  val searchf_tlr_unify_valid : 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
    96
      (Sign.sg -> int ->
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
    97
       Term.term ->
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
    98
       BasicIsaFTerm.FcTerm ->
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
    99
       match Seq.seq)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   100
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   101
  val eqsubst_asm_meth : Thm.thm list -> Proof.method
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   102
  val eqsubst_asm_tac : Thm.thm list -> int -> Thm.thm -> Thm.thm Seq.seq
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   103
  val eqsubst_asm_tac' : 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   104
      (Sign.sg -> int ->
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   105
       Term.term ->
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   106
       BasicIsaFTerm.FcTerm ->
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   107
       match Seq.seq) -> Thm.thm -> int -> Thm.thm -> Thm.thm Seq.seq
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
   108
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   109
  val eqsubst_meth : Thm.thm list -> Proof.method
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   110
  val eqsubst_tac : Thm.thm list -> int -> Thm.thm -> Thm.thm Seq.seq
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   111
  val eqsubst_tac' : 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   112
      (Sign.sg -> int ->
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   113
       Term.term ->
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   114
       BasicIsaFTerm.FcTerm ->
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   115
       match Seq.seq) -> Thm.thm -> int -> Thm.thm -> Thm.thm Seq.seq
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
   116
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   117
  val meth : bool * Thm.thm list -> Proof.context -> Proof.method
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   118
  val setup : (Theory.theory -> Theory.theory) list
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   119
end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   120
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   121
functor EQSubstTacFUN (structure EqRuleData : EQRULE_DATA) 
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
   122
  : EQSUBST_TAC
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   123
= struct
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   124
15550
806214035275 lucas - added more comments and an extra type to clarify the code.
dixon
parents: 15538
diff changeset
   125
  (* a type abriviation for match infomration *)
806214035275 lucas - added more comments and an extra type to clarify the code.
dixon
parents: 15538
diff changeset
   126
  type match = 
806214035275 lucas - added more comments and an extra type to clarify the code.
dixon
parents: 15538
diff changeset
   127
       ((Term.indexname * Term.typ) list (* type instantiations *)
806214035275 lucas - added more comments and an extra type to clarify the code.
dixon
parents: 15538
diff changeset
   128
         * (Term.indexname * Term.term) list) (* term instantiations *)
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   129
        * (string * Term.typ) list (* fake named type abs env *)
15550
806214035275 lucas - added more comments and an extra type to clarify the code.
dixon
parents: 15538
diff changeset
   130
        * (string * Term.typ) list (* type abs env *)
806214035275 lucas - added more comments and an extra type to clarify the code.
dixon
parents: 15538
diff changeset
   131
        * Term.term (* outer term *)
806214035275 lucas - added more comments and an extra type to clarify the code.
dixon
parents: 15538
diff changeset
   132
806214035275 lucas - added more comments and an extra type to clarify the code.
dixon
parents: 15538
diff changeset
   133
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
   134
(* 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
   135
type trace_subst_errT = int (* subgoal *)
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
   136
        * Thm.thm (* thm with all goals *)
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
   137
        * (Thm.cterm list (* certified free var placeholders for vars *)
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
   138
           * Thm.thm)  (* trivial thm of goal concl *)
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
   139
            (* possible matches/unifiers *)
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
   140
        * Thm.thm (* rule *)
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
   141
        * (((Term.indexname * Term.typ) list (* type instantiations *)
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
   142
              * (Term.indexname * Term.term) list ) (* term instantiations *)
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
   143
             * (string * Term.typ) list (* Type abs env *)
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
   144
             * Term.term) (* outer term *);
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
   145
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
   146
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
   147
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
   148
exception trace_subst_exp of trace_subst_errT;
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
   149
 *)
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
   150
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
   151
(* also defined in /HOL/Tools/inductive_codegen.ML, 
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
   152
   maybe move this to seq.ML ? *)
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
   153
infix 5 :->;
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
   154
fun s :-> f = Seq.flat (Seq.map f s);
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
   155
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   156
(* search from top, left to right, then down *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   157
fun search_tlr_all_f f ft = 
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   158
    let
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   159
      fun maux ft = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   160
          let val t' = (IsaFTerm.focus_of_fcterm ft) 
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
   161
            (* val _ = 
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
   162
                if !trace_subst_search then 
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
   163
                  (writeln ("Examining: " ^ (TermLib.string_of_term t'));
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
   164
                   TermLib.writeterm t'; ())
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
   165
                else (); *)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   166
          in 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   167
          (case t' of 
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
   168
            (_ $ _) => Seq.append(maux (IsaFTerm.focus_left ft), 
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
   169
                       Seq.append(f ft, 
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   170
                                  maux (IsaFTerm.focus_right ft)))
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   171
          | (Abs _) => Seq.append(f ft, maux (IsaFTerm.focus_abs ft))
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   172
          | leaf => f ft) end
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   173
    in maux ft end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   174
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   175
(* search from top, left to right, then down *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   176
fun search_tlr_valid_f f ft = 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   177
    let
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   178
      fun maux ft = 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   179
          let 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   180
            val hereseq = if IsaFTerm.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: 15550
diff changeset
   181
          in 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   182
          (case (IsaFTerm.focus_of_fcterm ft) of 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   183
            (_ $ _) => Seq.append(maux (IsaFTerm.focus_left ft), 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   184
                       Seq.append(hereseq, 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   185
                                  maux (IsaFTerm.focus_right ft)))
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   186
          | (Abs _) => Seq.append(hereseq, maux (IsaFTerm.focus_abs ft))
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   187
          | leaf => hereseq)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   188
          end
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   189
    in maux ft end;
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   190
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   191
(* search all unifications *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   192
fun searchf_tlr_unify_all sgn maxidx lhs  = 
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   193
    IsaFTerm.find_fcterm_matches 
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   194
      search_tlr_all_f 
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   195
      (IsaFTerm.clean_unify_ft sgn maxidx lhs);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   196
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   197
(* search only for 'valid' unifiers (non abs subterms and non vars) *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   198
fun searchf_tlr_unify_valid sgn maxidx lhs  = 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   199
    IsaFTerm.find_fcterm_matches 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   200
      search_tlr_valid_f 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   201
      (IsaFTerm.clean_unify_ft sgn maxidx lhs);
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   202
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   203
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
   204
(* 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
   205
(* 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
   206
(* 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
   207
(* 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
   208
(* rule is the equation for substitution *)
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
   209
fun apply_subst_in_concl i th (cfvs, conclthm) rule m = 
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
   210
    (RWInst.rw m rule conclthm)
15855
55e443aa711d lucas - updated to reflect isand.ML update
dixon
parents: 15814
diff changeset
   211
      |> IsaND.unfix_frees cfvs
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
   212
      |> RWInst.beta_eta_contract_tac
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
   213
      |> (fn r => Tactic.rtac r i th);
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   214
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
   215
(*
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   216
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
   217
 |> (fn r => Thm.bicompose false (false, r, Thm.nprems_of r) i 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
   218
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
   219
*)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   220
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   221
(* 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
   222
equation rule. Note that we assume rule has var indicies zero'd *)
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
   223
fun prep_concl_subst searchf i gth = 
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   224
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   225
      val th = Thm.incr_indexes 1 gth;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   226
      val tgt_term = Thm.prop_of th;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   227
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   228
      val sgn = Thm.sign_of_thm th;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   229
      val ctermify = Thm.cterm_of sgn;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   230
      val trivify = Thm.trivial o ctermify;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   231
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   232
      val (fixedbody, fvs) = IsaND.fix_alls_term i tgt_term;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   233
      val cfvs = rev (map ctermify fvs);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   234
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
   235
      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
   236
      val conclthm = trivify conclterm;
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
   237
      val maxidx = Term.maxidx_of_term conclterm;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   238
    in
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
   239
      ((cfvs, conclthm), 
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
   240
       (fn lhs => searchf sgn maxidx lhs 
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
   241
                          ((IsaFTerm.focus_right  
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
   242
                            o IsaFTerm.focus_left
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
   243
                            o IsaFTerm.fcterm_of_term 
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
   244
                            o Thm.prop_of) conclthm)))
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   245
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   246
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
   247
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   248
(* substitute using an object or meta level equality *)
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   249
fun eqsubst_tac' searchf instepthm i 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
   250
    let 
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
   251
      val (cvfsconclthm, findmatchf) = 
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   252
          prep_concl_subst searchf i 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
   253
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
   254
      val stepthms = 
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
   255
          Seq.map Drule.zero_var_indexes 
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
   256
                  (Seq.of_list (EqRuleData.prep_meta_eq instepthm));
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
   257
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
   258
      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
   259
          let val (lhs,_) = Logic.dest_equals (Thm.concl_of 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
   260
          in (findmatchf lhs)
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
   261
             :-> (apply_subst_in_concl i th cvfsconclthm r) end;
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
   262
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
   263
    in (stepthms :-> rewrite_with_thm) end;
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
   264
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
   265
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   266
(* substitute using one of the given theorems *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   267
fun eqsubst_tac instepthms i 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
   268
    if Thm.nprems_of th < i then Seq.empty else
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   269
    (Seq.of_list instepthms) 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   270
    :-> (fn r => eqsubst_tac' searchf_tlr_unify_valid r i th);
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   271
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   272
(* inthms are the given arguments in Isar, and treated as eqstep with
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   273
   the first one, then the second etc *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   274
fun eqsubst_meth inthms =
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   275
    Method.METHOD 
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
      (fn facts =>
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
   277
          HEADGOAL ( Method.insert_tac facts THEN' eqsubst_tac inthms ));
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   278
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   279
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
   280
fun apply_subst_in_asm i th (cfvs, j, nprems, pth) rule m = 
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
   281
    (RWInst.rw m rule pth)
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
      |> Thm.permute_prems 0 ~1
15855
55e443aa711d lucas - updated to reflect isand.ML update
dixon
parents: 15814
diff changeset
   283
      |> IsaND.unfix_frees cfvs
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
   284
      |> RWInst.beta_eta_contract_tac
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
      |> (fn r => Tactic.dtac r i 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
   286
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
   287
(*
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
   288
? should I be using bicompose what if we match more than one
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
   289
assumption, even after instantiation ? (back will work, but it would
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
   290
be nice to avoid the redudent search)
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
   291
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
something like... 
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
   293
 |> Thm.lift_rule (th, i)
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
   294
 |> (fn r => Thm.bicompose false (false, r, Thm.nprems_of r - nprems) i 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
   295
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
   296
*)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   297
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   298
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
   299
(* 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
   300
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
   301
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
   302
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
   303
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
   304
multiple assumptions.  *)
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
   305
fun prep_subst_in_asm searchf i gth j = 
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   306
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   307
      val th = Thm.incr_indexes 1 gth;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   308
      val tgt_term = Thm.prop_of th;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   309
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   310
      val sgn = Thm.sign_of_thm th;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   311
      val ctermify = Thm.cterm_of sgn;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   312
      val trivify = Thm.trivial o ctermify;
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
      val (fixedbody, fvs) = IsaND.fix_alls_term i tgt_term;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   315
      val cfvs = rev (map ctermify fvs);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   316
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
   317
      val asmt = Library.nth_elem(j - 1,(Logic.strip_imp_prems 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
   318
      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
   319
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
   320
      val pth = trivify 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
   321
      val maxidx = Term.maxidx_of_term 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
   322
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   323
    in
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
   324
      ((cfvs, j, asm_nprems, pth), 
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
   325
       (fn lhs => (searchf sgn maxidx lhs
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
   326
                           ((IsaFTerm.focus_right 
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
   327
                             o IsaFTerm.fcterm_of_term 
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
   328
                             o Thm.prop_of) pth))))
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   329
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   330
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
   331
(* prepare subst in every possible assumption *)
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
   332
fun prep_subst_in_asms searchf i 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
   333
    Seq.map 
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
   334
      (prep_subst_in_asm searchf i 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
   335
      (Seq.of_list (IsaPLib.mk_num_list
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
   336
                      (length (Logic.prems_of_goal (Thm.prop_of gth) i))));
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
   337
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
   338
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
   339
(* substitute in an assumption using an object or meta level equality *)
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   340
fun eqsubst_asm_tac' searchf instepthm i 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
   341
    let 
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   342
      val asmpreps = prep_subst_in_asms searchf i 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
   343
      val stepthms = 
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
   344
          Seq.map Drule.zero_var_indexes 
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
   345
                  (Seq.of_list (EqRuleData.prep_meta_eq instepthm))
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
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
      fun rewrite_with_thm (asminfo, findmatchf) 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
   348
          let val (lhs,_) = Logic.dest_equals (Thm.concl_of 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
   349
          in (findmatchf lhs)
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
             :-> (apply_subst_in_asm i th asminfo r) end;
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
   351
    in
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
      (asmpreps :-> (fn a => stepthms :-> rewrite_with_thm a))
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   353
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   354
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   355
(* substitute using one of the given theorems *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   356
fun eqsubst_asm_tac instepthms i 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
   357
    if Thm.nprems_of th < i then Seq.empty else
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   358
    (Seq.of_list instepthms) 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15550
diff changeset
   359
    :-> (fn r => eqsubst_asm_tac' searchf_tlr_unify_valid r i th);
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   360
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   361
(* inthms are the given arguments in Isar, and treated as eqstep with
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   362
   the first one, then the second etc *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   363
fun eqsubst_asm_meth inthms =
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   364
    Method.METHOD 
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
   365
      (fn facts =>
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
   366
          HEADGOAL (Method.insert_tac facts THEN' eqsubst_asm_tac inthms ));
15481
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
(* combination method that takes a flag (true indicates that subst
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   369
should be done to an assumption, false = apply to the conclusion of
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   370
the goal) as well as the theorems to use *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   371
fun meth (asmflag, inthms) ctxt = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   372
    if asmflag then eqsubst_asm_meth inthms else eqsubst_meth inthms;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   373
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   374
(* syntax for options, given "(asm)" will give back true, without
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   375
   gives back false *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   376
val options_syntax =
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   377
    (Args.parens (Args.$$$ "asm") >> (K true)) ||
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   378
     (Scan.succeed false);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   379
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   380
(* method syntax, first take options, then theorems *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   381
fun meth_syntax meth src ctxt =
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   382
    meth (snd (Method.syntax ((Scan.lift options_syntax) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   383
                                -- Attrib.local_thms) src ctxt)) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   384
         ctxt;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   385
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   386
(* setup function for adding method to theory. *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   387
val setup = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   388
    [Method.add_method ("subst", meth_syntax meth, "Substiution with an equation. Use \"(asm)\" option to substitute in an assumption.")];
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   389
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   390
end;