src/Pure/IsaPlanner/isand.ML
author huffman
Mon, 07 Mar 2005 23:54:01 +0100
changeset 15587 f363e6e080e7
parent 15574 b1d1b5bfc464
child 15854 1ae0a47dcccd
permissions -rw-r--r--
added subsections and text for document generation
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     1
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     2
(*  Title:      sys/isand.ML
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     3
    Author:     Lucas Dixon, University of Edinburgh
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     4
                lucas.dixon@ed.ac.uk
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     5
    Date:       6 Aug 2004
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     6
*)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     7
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     8
(*  DESCRIPTION:
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     9
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    10
    Natural Deduction tools
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    11
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    12
    For working with Isbaelle theorem in a natural detuction style,
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    13
    ie, not having to deal with meta level quantified varaibles,
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    14
    instead, we work with newly introduced frees, and hide the
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    15
    "all"'s, exporting results from theorems proved with the frees, to
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    16
    solve the all cases of the previous goal. 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    17
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    18
    Note: A nice idea: allow esxporting to solve any subgoal, thus
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    19
    allowing the interleaving of proof, or provide a structure for the
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    20
    ordering of proof, thus allowing proof attempts in parrelle, but
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    21
    recording the order to apply things in.
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    22
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    23
    debugging tools:
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    24
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    25
    fun asm_mk t = (assume (cterm_of (Theory.sign_of (the_context())) t)); 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    26
    fun asm_read s =  
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    27
      (assume (read_cterm (Theory.sign_of (Context.the_context())) (s,propT))); 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    28
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    29
    THINK: are we really ok with our varify name w.r.t the prop - do
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    30
    we alos need to avoid names in the hidden hyps?
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    31
*)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    32
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    33
signature ISA_ND =
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    34
sig
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    35
  datatype export = export of
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    36
           {gth: Thm.thm, (* initial goal theorem *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    37
            sgid: int, (* subgoal id which has been fixed etc *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    38
            fixes: Thm.cterm list, (* frees *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    39
            assumes: Thm.cterm list} (* assumptions *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    40
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    41
  val fixes_of_exp : export -> Thm.cterm list
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    42
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    43
  val export_back : export -> Thm.thm -> Thm.thm Seq.seq
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    44
  val export_solution : export -> Thm.thm -> Thm.thm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    45
  val export_solutions : export list * Thm.thm -> Thm.thm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    46
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    47
  val allify_conditions :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    48
      (Term.term -> Thm.cterm) ->
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    49
      (string * Term.typ) list -> Thm.thm -> Thm.thm * Thm.cterm list
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    50
  val allify_conditions' :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    51
      (string * Term.typ) list -> Thm.thm -> Thm.thm * Thm.cterm list
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    52
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    53
  val assume_prems :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    54
      int -> Thm.thm -> Thm.thm list * Thm.thm * Thm.cterm list
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    55
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    56
  val fix_alls_term : int -> Term.term -> Term.term * Term.term list
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    57
  val fix_alls_cterm : int -> Thm.thm -> Thm.cterm * Thm.cterm list
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    58
  val fix_alls' : int -> Thm.thm -> Thm.thm * Thm.cterm list
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    59
  val fix_alls : int -> Thm.thm -> Thm.thm * export
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    60
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    61
  val fix_vars_to_frees : Thm.thm -> (Thm.cterm * Thm.cterm) list * Thm.thm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    62
  val schemify_frees_to_vars : Thm.cterm list -> Thm.thm -> Thm.thm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    63
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    64
  val fixed_subgoal_thms : Thm.thm -> Thm.thm list * (Thm.thm list -> Thm.thm)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    65
  val fixes_and_assumes : int -> Thm.thm -> Thm.thm list * Thm.thm * export
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    66
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    67
  val hide_other_goals : Thm.thm -> Thm.thm * Thm.cterm list
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    68
  val hide_prems : Thm.thm -> Thm.thm * Thm.cterm list
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    69
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    70
  (* abstracts cterms (vars) to locally meta-all bounds *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    71
  val prepare_goal_export : string list * Thm.cterm list -> Thm.thm 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    72
                            -> int * Thm.thm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    73
  val solve_with : Thm.thm -> Thm.thm -> Thm.thm Seq.seq
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    74
  val subgoal_thms : Thm.thm -> Thm.thm list * (Thm.thm list -> Thm.thm)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    75
end
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    76
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    77
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    78
structure IsaND : ISA_ND =
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    79
struct
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    80
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    81
(* Solve *some* subgoal of "th" directly by "sol" *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    82
(* Note: this is probably what Markus ment to do upon export of a
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    83
"show" but maybe he used RS/rtac instead, which would wrongly lead to
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    84
failing if there are premices to the shown goal. *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    85
fun solve_with sol th = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    86
    let fun solvei 0 = Seq.empty
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    87
          | solvei i = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    88
            Seq.append (bicompose false (false,sol,0) i th, 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    89
                        solvei (i - 1))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    90
    in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    91
      solvei (Thm.nprems_of th)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    92
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    93
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    94
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    95
(* Given ctertmify function, (string,type) pairs capturing the free
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    96
vars that need to be allified in the assumption, and a theorem with
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    97
assumptions possibly containing the free vars, then we give back the
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    98
assumptions allified as hidden hyps. *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    99
(* 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   100
Given: vs 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   101
th: A vs ==> B vs 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   102
Results in: "B vs" [!!x. A x]
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   103
*)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   104
fun allify_conditions ctermify Ts th = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   105
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   106
      val premts = Thm.prems_of th;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   107
    
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   108
      fun allify_prem_var (vt as (n,ty),t)  = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   109
          (Term.all ty) $ (Abs(n,ty,Term.abstract_over (Free vt, t)))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   110
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   111
      fun allify_prem Ts p = foldr allify_prem_var p Ts
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   112
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   113
      val cTs = map (ctermify o Free) Ts
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   114
      val cterm_asms = map (ctermify o allify_prem Ts) premts
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   115
      val allifyied_asm_thms = map (Drule.forall_elim_list cTs o Thm.assume) cterm_asms
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   116
    in 
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   117
      (Library.foldl (fn (x,y) => y COMP x) (th, allifyied_asm_thms), cterm_asms)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   118
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   119
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   120
fun allify_conditions' Ts th = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   121
    allify_conditions (Thm.cterm_of (Thm.sign_of_thm th)) Ts th;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   122
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   123
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   124
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   125
(* change schematic vars to fresh free vars *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   126
fun fix_vars_to_frees th = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   127
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   128
      val ctermify = Thm.cterm_of (Thm.sign_of_thm th)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   129
      val prop = Thm.prop_of th
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   130
      val vars = map Term.dest_Var (Term.term_vars prop)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   131
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   132
      val names = Term.add_term_names (prop, [])
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   133
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   134
      val (insts,names2) = 
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   135
          Library.foldl (fn ((insts,names),v as ((n,i),ty)) => 
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   136
                    let val n2 = Term.variant names n in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   137
                      ((ctermify (Var v), ctermify (Free(n2,ty))) :: insts, 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   138
                       n2 :: names)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   139
                    end)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   140
                (([],names), vars)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   141
    in (insts, Thm.instantiate ([], insts) th) end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   142
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   143
(* make free vars into schematic vars with index zero *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   144
 fun schemify_frees_to_vars frees = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   145
     apply (map (K (Drule.forall_elim_var 0)) frees) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   146
     o Drule.forall_intr_list frees;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   147
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   148
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   149
(* datatype to capture an exported result, ie a fix or assume. *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   150
datatype export = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   151
         export of {fixes : Thm.cterm list, (* fixed vars *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   152
                    assumes : Thm.cterm list, (* hidden hyps/assumed prems *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   153
                    sgid : int,
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   154
                    gth :  Thm.thm}; (* subgoal/goalthm *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   155
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   156
fun fixes_of_exp (export rep) = #fixes rep;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   157
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   158
(* export the result of the new goal thm, ie if we reduced teh
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   159
subgoal, then we get a new reduced subtgoal with the old
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   160
all-quantified variables *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   161
local 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   162
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   163
(* allify puts in a meta level univ quantifier for a free variavble *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   164
fun allify_term (v, t) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   165
    let val vt = #t (Thm.rep_cterm v)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   166
      val (n,ty) = Term.dest_Free vt
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   167
    in (Term.all ty) $ (Abs(n,ty,Term.abstract_over (vt, t))) end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   168
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   169
fun allify_for_sg_term ctermify vs t =
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   170
    let val t_alls = foldr allify_term t vs;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   171
        val ct_alls = ctermify t_alls; 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   172
    in 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   173
      (ct_alls, Drule.forall_elim_list vs (Thm.assume ct_alls))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   174
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   175
(* lookup type of a free var name from a list *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   176
fun lookupfree vs vn  = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   177
    case Library.find_first (fn (n,ty) => n = vn) vs of 
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   178
      NONE => raise ERROR_MESSAGE ("prepare_goal_export:lookupfree: " 
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   179
                    ^ vn ^ " does not occur in the term")
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   180
    | SOME x => x;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   181
in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   182
fun export_back (export {fixes = vs, assumes = hprems, 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   183
                         sgid = i, gth = gth}) newth = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   184
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   185
      val sgn = Thm.sign_of_thm newth;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   186
      val ctermify = Thm.cterm_of sgn;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   187
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   188
      val sgs = prems_of newth;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   189
      val (sgallcts, sgthms) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   190
          Library.split_list (map (allify_for_sg_term ctermify vs) sgs);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   191
      val minimal_newth = 
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   192
          (Library.foldl (fn ( newth', sgthm) => 
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   193
                          Drule.compose_single (sgthm, 1, newth'))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   194
                      (newth, sgthms));
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   195
      val allified_newth = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   196
          minimal_newth 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   197
            |> Drule.implies_intr_list hprems
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   198
            |> Drule.forall_intr_list vs 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   199
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   200
      val newth' = Drule.implies_intr_list sgallcts allified_newth
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   201
    in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   202
      bicompose false (false, newth', (length sgallcts)) i gth
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   203
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   204
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   205
(* 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   206
Given "vs" : names of free variables to abstract over,
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   207
Given cterms : premices to abstract over (P1... Pn) in terms of vs,
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   208
Given a thm of the form: 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   209
P1 vs; ...; Pn vs ==> Goal(C vs)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   210
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   211
Gives back: 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   212
(n, length of given cterms which have been allified
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   213
 [| !! vs. P1 vs; !! vs. Pn vs |] ==> !! C vs) the allified thm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   214
*)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   215
(* note: C may contain further premices etc 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   216
Note that cterms is the assumed facts, ie prems of "P1" that are
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   217
reintroduced in allified form.
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   218
*)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   219
fun prepare_goal_export (vs, cterms) th = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   220
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   221
      val sgn = Thm.sign_of_thm th;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   222
      val ctermify = Thm.cterm_of sgn;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   223
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   224
      val allfrees = map Term.dest_Free (Term.term_frees (Thm.prop_of th))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   225
      val cfrees = map (ctermify o Free o lookupfree allfrees) vs
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   226
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   227
      val sgs = prems_of th;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   228
      val (sgallcts, sgthms) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   229
          Library.split_list (map (allify_for_sg_term ctermify cfrees) sgs);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   230
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   231
      val minimal_th = 
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   232
          (Library.foldl (fn ( th', sgthm) => 
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   233
                          Drule.compose_single (sgthm, 1, th'))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   234
                      (th, sgthms)) RS Drule.rev_triv_goal;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   235
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   236
      val allified_th = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   237
          minimal_th 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   238
            |> Drule.implies_intr_list cterms
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   239
            |> Drule.forall_intr_list cfrees 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   240
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   241
      val th' = Drule.implies_intr_list sgallcts allified_th
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   242
    in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   243
      ((length sgallcts), th')
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   244
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   245
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   246
end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   247
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   248
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   249
(* exporting function that takes a solution to the fixed/assumed goal,
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   250
and uses this to solve the subgoal in the main theorem *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   251
fun export_solution (export {fixes = cfvs, assumes = hcprems,
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   252
                             sgid = i, gth = gth}) solth = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   253
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   254
      val solth' = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   255
          solth |> Drule.implies_intr_list hcprems
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   256
                |> Drule.forall_intr_list cfvs
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   257
    in Drule.compose_single (solth', i, gth) end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   258
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   259
fun export_solutions (xs,th) = foldr (uncurry export_solution) th xs;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   260
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   261
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   262
(* fix parameters of a subgoal "i", as free variables, and create an
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   263
exporting function that will use the result of this proved goal to
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   264
show the goal in the original theorem. 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   265
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   266
Note, an advantage of this over Isar is that it supports instantiation
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   267
of unkowns in the earlier theorem, ie we can do instantiation of meta
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   268
vars! *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   269
(* loosely corresponds to:
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   270
Given "[| SG0; ... !! x. As ==> SGi x; ... SGm |] ==> G" : thm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   271
Result: 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   272
  ("(As ==> SGi x') ==> (As ==> SGi x')" : thm, 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   273
   expf : 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   274
     ("As ==> SGi x'" : thm) -> 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   275
     ("[| SG0; ... SGi-1; SGi+1; ... SGm |] ==> G") : thm)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   276
*)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   277
fun fix_alls_term i t = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   278
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   279
      val names = Term.add_term_names (t,[]);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   280
      val gt = Logic.get_goal t i;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   281
      val body = Term.strip_all_body gt;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   282
      val alls = rev (Term.strip_all_vars gt);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   283
      val fvs = map Free 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   284
                    ((Term.variantlist (map fst alls, names)) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   285
                       ~~ (map snd alls));
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   286
    in ((subst_bounds (fvs,body)), fvs) end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   287
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   288
fun fix_alls_cterm i th = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   289
    let
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   290
      val ctermify = Thm.cterm_of (Thm.sign_of_thm th);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   291
      val (fixedbody, fvs) = fix_alls_term i (Thm.prop_of th);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   292
      val cfvs = rev (map ctermify fvs);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   293
      val ct_body = ctermify fixedbody
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   294
    in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   295
      (ct_body, cfvs)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   296
    end;
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
fun fix_alls' i = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   299
     (apfst Thm.trivial) o (fix_alls_cterm i);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   300
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   301
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   302
(* hide other goals *) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   303
(* note the export goal is rotated by (i - 1) and will have to be
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   304
unrotated to get backto the originial position(s) *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   305
fun hide_other_goals th = 
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
      (* tl beacuse fst sg is the goal we are interested in *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   308
      val cprems = tl (Drule.cprems_of th)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   309
      val aprems = map Thm.assume cprems
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   310
    in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   311
      (Drule.implies_elim_list (Drule.rotate_prems 1 th) aprems, 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   312
       cprems)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   313
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   314
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   315
(* a nicer version of the above that leaves only a single subgoal (the
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   316
other subgoals are hidden hyps, that the exporter suffles about)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   317
namely the subgoal that we were trying to solve. *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   318
(* loosely corresponds to:
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   319
Given "[| SG0; ... !! x. As ==> SGi x; ... SGm |] ==> G" : thm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   320
Result: 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   321
  ("(As ==> SGi x') ==> SGi x'" : thm, 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   322
   expf : 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   323
     ("SGi x'" : thm) -> 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   324
     ("[| SG0; ... SGi-1; SGi+1; ... SGm |] ==> G") : thm)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   325
*)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   326
fun fix_alls i th = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   327
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   328
      val (fixed_gth, fixedvars) = fix_alls' i th
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   329
      val (sml_gth, othergoals) = hide_other_goals fixed_gth
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   330
    in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   331
      (sml_gth, export {fixes = fixedvars, 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   332
                        assumes = othergoals, 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   333
                        sgid = i, gth = th})
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   334
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   335
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   336
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   337
(* assume the premises of subgoal "i", this gives back a list of
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   338
assumed theorems that are the premices of subgoal i, it also gives
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   339
back a new goal thm and an exporter, the new goalthm is as the old
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   340
one, but without the premices, and the exporter will use a proof of
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   341
the new goalthm, possibly using the assumed premices, to shoe the
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   342
orginial goal. *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   343
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   344
(* Note: Dealing with meta vars, need to meta-level-all them in the
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   345
shyps, which we can later instantiate with a specific value.... ? 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   346
think about this... maybe need to introduce some new fixed vars and
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   347
then remove them again at the end... like I do with rw_inst. *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   348
(* loosely corresponds to:
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   349
Given "[| SG0; ... [| A0; ... An |] ==> SGi; ... SGm |] ==> G" : thm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   350
Result: 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   351
(["A0" [A0], ... ,"An" [An]] : thm list, -- assumptions
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   352
 "SGi ==> SGi" : thm, -- new goal 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   353
 "SGi" ["A0" ... "An"] : thm ->   -- export function
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   354
    ("[| SG0 ... SGi-1, SGi+1, SGm |] ==> G" : thm) list)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   355
*)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   356
fun assume_prems i th =
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   357
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   358
      val t = (prop_of th); 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   359
      val gt = Logic.get_goal t i;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   360
      val _ = case Term.strip_all_vars gt of [] => () 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   361
              | _ => raise ERROR_MESSAGE "assume_prems: goal has params"
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   362
      val body = gt;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   363
      val prems = Logic.strip_imp_prems body;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   364
      val concl = Logic.strip_imp_concl body;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   365
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   366
      val sgn = Thm.sign_of_thm th;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   367
      val ctermify = Thm.cterm_of sgn;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   368
      val cprems = map ctermify prems;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   369
      val aprems = map Thm.assume cprems;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   370
      val gthi = Thm.trivial (ctermify concl);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   371
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   372
      (* fun explortf thi = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   373
          Drule.compose (Drule.implies_intr_list cprems thi, 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   374
                         i, th) *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   375
    in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   376
      (aprems, gthi, cprems)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   377
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   378
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   379
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   380
(* first fix the variables, then assume the assumptions *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   381
(* loosely corresponds to:
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   382
Given 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   383
  "[| SG0; ... 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   384
      !! xs. [| A0 xs; ... An xs |] ==> SGi xs; 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   385
      ... SGm |] ==> G" : thm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   386
Result: 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   387
(["A0 xs'" [A0 xs'], ... ,"An xs'" [An xs']] : thm list, -- assumptions
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   388
 "SGi xs' ==> SGi xs'" : thm,  -- new goal 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   389
 "SGi xs'" ["A0 xs'" ... "An xs'"] : thm ->   -- export function
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   390
    ("[| SG0 ... SGi-1, SGi+1, SGm |] ==> G" : thm) list)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   391
*)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   392
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   393
(* Note: the fix_alls actually pulls through all the assumptions which
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   394
means that the second export is not needed. *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   395
fun fixes_and_assumes i th = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   396
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   397
      val (fixgth, exp1) = fix_alls i th
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   398
      val (assumps, goalth, _) = assume_prems 1 fixgth
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   399
    in 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   400
      (assumps, goalth, exp1)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   401
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   402
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   403
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   404
(* Fixme: allow different order of subgoals given to expf *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   405
(* make each subgoal into a separate thm that needs to be proved *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   406
(* loosely corresponds to:
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   407
Given 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   408
  "[| SG0; ... SGm |] ==> G" : thm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   409
Result: 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   410
(["SG0 ==> SG0", ... ,"SGm ==> SGm"] : thm list, -- goals
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   411
 ["SG0", ..., "SGm"] : thm list ->   -- export function
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   412
   "G" : thm)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   413
*)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   414
fun subgoal_thms th = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   415
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   416
      val t = (prop_of th); 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   417
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   418
      val prems = Logic.strip_imp_prems t;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   419
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   420
      val sgn = Thm.sign_of_thm th;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   421
      val ctermify = Thm.cterm_of sgn;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   422
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   423
      val aprems = map (Thm.trivial o ctermify) prems;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   424
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   425
      fun explortf premths = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   426
          Drule.implies_elim_list th premths
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   427
    in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   428
      (aprems, explortf)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   429
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   430
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   431
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   432
(* make all the premices of a theorem hidden, and provide an unhide
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   433
function, that will bring them back out at a later point. This is
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   434
useful if you want to get back these premices, after having used the
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   435
theorem with the premices hidden *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   436
(* loosely corresponds to:
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   437
Given "As ==> G" : thm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   438
Result: ("G [As]" : thm, 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   439
         "G [As]" : thm -> "As ==> G" : thm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   440
*)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   441
fun hide_prems th = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   442
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   443
      val sgn = Thm.sign_of_thm th;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   444
      val ctermify = Thm.cterm_of sgn;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   445
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   446
      val t = (prop_of th);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   447
      val prems = Logic.strip_imp_prems t;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   448
      val cprems = map ctermify prems;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   449
      val aprems = map Thm.trivial cprems;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   450
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   451
    (*   val unhidef = Drule.implies_intr_list cprems; *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   452
    in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   453
      (Drule.implies_elim_list th aprems, cprems)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   454
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   455
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   456
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   457
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   458
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   459
(* Fixme: allow different order of subgoals in exportf *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   460
(* as above, but also fix all parameters in all subgoals, and uses
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   461
fix_alls, not fix_alls', ie doesn't leave extra asumptions as apparent
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   462
subgoals. *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   463
(* loosely corresponds to:
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   464
Given 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   465
  "[| !! x0s. A0s x0s ==> SG0 x0s; 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   466
      ...; !! xms. Ams xms ==> SGm xms|] ==> G" : thm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   467
Result: 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   468
(["(A0s x0s' ==> SG0 x0s') ==> SG0 x0s'", 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   469
  ... ,"(Ams xms' ==> SGm xms') ==> SGm xms'"] : thm list, -- goals
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   470
 ["SG0 x0s'", ..., "SGm xms'"] : thm list ->   -- export function
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   471
   "G" : thm)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   472
*)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   473
(* requires being given solutions! *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   474
fun fixed_subgoal_thms th = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   475
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   476
      val (subgoals, expf) = subgoal_thms th;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   477
(*       fun export_sg (th, exp) = exp th; *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   478
      fun export_sgs expfs solthms = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   479
          expf (map2 (op |>) (solthms, expfs));
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   480
(*           expf (map export_sg (ths ~~ expfs)); *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   481
    in 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   482
      apsnd export_sgs (Library.split_list (map (apsnd export_solution o 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   483
                                                 fix_alls 1) subgoals))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   484
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   485
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   486
end;