src/Tools/misc_legacy.ML
author wenzelm
Wed, 12 Jun 2024 16:58:55 +0200
changeset 80358 45b434464cd8
parent 74525 c960bfcb91db
child 82697 cc05bc2cfb2f
permissions -rw-r--r--
proper sha1_digest: need to include offset + length;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37781
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
     1
(*  Title:      Tools/misc_legacy.ML
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
     2
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
     3
Misc legacy stuff -- to be phased out eventually.
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
     4
*)
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
     5
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
     6
signature MISC_LEGACY =
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
     7
sig
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
     8
  val add_term_names: term * string list -> string list
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
     9
  val add_term_tvars: term * (indexname * sort) list -> (indexname * sort) list
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    10
  val add_term_tfrees: term * (string * sort) list -> (string * sort) list
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    11
  val typ_tvars: typ -> (indexname * sort) list
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    12
  val term_tfrees: term -> (string * sort) list
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    13
  val term_tvars: term -> (indexname * sort) list
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    14
  val add_term_vars: term * term list -> term list
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    15
  val term_vars: term -> term list
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    16
  val add_term_frees: term * term list -> term list
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    17
  val term_frees: term -> term list
37781
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
    18
  val mk_defpair: term * term -> string * term
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
    19
  val get_def: theory -> xstring -> thm
59165
115965966e15 proper context;
wenzelm
parents: 58950
diff changeset
    20
  val METAHYPS: Proof.context -> (thm list -> tactic) -> int -> tactic
60358
aebfbcab1eb8 clarified context;
wenzelm
parents: 59621
diff changeset
    21
  val freeze_thaw_robust: Proof.context -> thm -> thm * (int -> thm -> thm)
37781
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
    22
end;
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
    23
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
    24
structure Misc_Legacy: MISC_LEGACY =
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
    25
struct
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
    26
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    27
(*iterate a function over all types in a term*)
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    28
fun it_term_types f =
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    29
let fun iter(Const(_,T), a) = f(T,a)
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    30
      | iter(Free(_,T), a) = f(T,a)
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    31
      | iter(Var(_,T), a) = f(T,a)
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    32
      | iter(Abs(_,T,t), a) = iter(t,f(T,a))
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    33
      | iter(f$u, a) = iter(f, iter(u, a))
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    34
      | iter(Bound _, a) = a
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    35
in iter end
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    36
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    37
(*Accumulates the names in the term, suppressing duplicates.
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    38
  Includes Frees and Consts.  For choosing unambiguous bound var names.*)
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    39
fun add_term_names (Const(a,_), bs) = insert (op =) (Long_Name.base_name a) bs
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    40
  | add_term_names (Free(a,_), bs) = insert (op =) a bs
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    41
  | add_term_names (f$u, bs) = add_term_names (f, add_term_names(u, bs))
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    42
  | add_term_names (Abs(_,_,t), bs) = add_term_names(t,bs)
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    43
  | add_term_names (_, bs) = bs;
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    44
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    45
(*Accumulates the TVars in a type, suppressing duplicates.*)
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    46
fun add_typ_tvars(Type(_,Ts),vs) = List.foldr add_typ_tvars vs Ts
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    47
  | add_typ_tvars(TFree(_),vs) = vs
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    48
  | add_typ_tvars(TVar(v),vs) = insert (op =) v vs;
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    49
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    50
(*Accumulates the TFrees in a type, suppressing duplicates.*)
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    51
fun add_typ_tfree_names(Type(_,Ts),fs) = List.foldr add_typ_tfree_names fs Ts
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    52
  | add_typ_tfree_names(TFree(f,_),fs) = insert (op =) f fs
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    53
  | add_typ_tfree_names(TVar(_),fs) = fs;
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    54
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    55
fun add_typ_tfrees(Type(_,Ts),fs) = List.foldr add_typ_tfrees fs Ts
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    56
  | add_typ_tfrees(TFree(f),fs) = insert (op =) f fs
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    57
  | add_typ_tfrees(TVar(_),fs) = fs;
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    58
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    59
(*Accumulates the TVars in a term, suppressing duplicates.*)
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    60
val add_term_tvars = it_term_types add_typ_tvars;
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    61
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    62
(*Accumulates the TFrees in a term, suppressing duplicates.*)
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    63
val add_term_tfrees = it_term_types add_typ_tfrees;
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    64
val add_term_tfree_names = it_term_types add_typ_tfree_names;
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    65
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    66
(*Non-list versions*)
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    67
fun typ_tfrees T = add_typ_tfrees(T,[]);
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    68
fun typ_tvars T = add_typ_tvars(T,[]);
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    69
fun term_tfrees t = add_term_tfrees(t,[]);
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    70
fun term_tvars t = add_term_tvars(t,[]);
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    71
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    72
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    73
(*Accumulates the Vars in the term, suppressing duplicates.*)
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    74
fun add_term_vars (t, vars: term list) = case t of
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    75
    Var   _ => Ord_List.insert Term_Ord.term_ord t vars
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    76
  | Abs (_,_,body) => add_term_vars(body,vars)
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    77
  | f$t =>  add_term_vars (f, add_term_vars(t, vars))
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    78
  | _ => vars;
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    79
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    80
fun term_vars t = add_term_vars(t,[]);
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    81
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    82
(*Accumulates the Frees in the term, suppressing duplicates.*)
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    83
fun add_term_frees (t, frees: term list) = case t of
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    84
    Free   _ => Ord_List.insert Term_Ord.term_ord t frees
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    85
  | Abs (_,_,body) => add_term_frees(body,frees)
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    86
  | f$t =>  add_term_frees (f, add_term_frees(t, frees))
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    87
  | _ => frees;
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    88
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    89
fun term_frees t = add_term_frees(t,[]);
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    90
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    91
37781
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
    92
fun mk_defpair (lhs, rhs) =
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
    93
  (case Term.head_of lhs of
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
    94
    Const (name, _) =>
46909
3c73a121a387 more explicit indication of def names;
wenzelm
parents: 46215
diff changeset
    95
      (Thm.def_name (Long_Name.base_name name), Logic.mk_equals (lhs, rhs))
37781
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
    96
  | _ => raise TERM ("Malformed definition: head of lhs not a constant", [lhs, rhs]));
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
    97
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
    98
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
    99
fun get_def thy = Thm.axiom thy o Name_Space.intern (Theory.axiom_space thy) o Thm.def_name;
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   100
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   101
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   102
(**** METAHYPS -- tactical for using hypotheses as meta-level assumptions
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   103
       METAHYPS (fn prems => tac prems) i
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   104
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   105
converts subgoal i, of the form !!x1...xm. [| A1;...;An] ==> A into a new
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   106
proof state A==>A, supplying A1,...,An as meta-level assumptions (in
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   107
"prems").  The parameters x1,...,xm become free variables.  If the
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   108
resulting proof state is [| B1;...;Bk] ==> C (possibly assuming A1,...,An)
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   109
then it is lifted back into the original context, yielding k subgoals.
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   110
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   111
Replaces unknowns in the context by Frees having the prefix METAHYP_
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   112
New unknowns in [| B1;...;Bk] ==> C are lifted over x1,...,xm.
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   113
DOES NOT HANDLE TYPE UNKNOWNS.
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   114
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   115
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   116
NOTE: This version does not observe the proof context, and thus cannot
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   117
work reliably.  See also Subgoal.SUBPROOF and Subgoal.FOCUS for
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   118
properly localized variants of the same idea.
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   119
****)
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   120
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   121
local
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   122
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   123
(*Strips assumptions in goal yielding  ( [x1,...,xm], [H1,...,Hn], B )
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   124
    H1,...,Hn are the hypotheses;  x1...xm are variants of the parameters.
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   125
  Main difference from strip_assums concerns parameters:
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   126
    it replaces the bound variables by free variables.  *)
74295
9a9326a072bb more antiquotations;
wenzelm
parents: 74282
diff changeset
   127
fun strip_context_aux (params, Hs, \<^Const_>\<open>Pure.imp for H B\<close>) =
37781
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   128
      strip_context_aux (params, H :: Hs, B)
74525
c960bfcb91db discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents: 74408
diff changeset
   129
  | strip_context_aux (params, Hs, \<^Const_>\<open>Pure.all _ for \<open>t as Abs _\<close>\<close>) =
c960bfcb91db discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents: 74408
diff changeset
   130
      let val (v, u) = Term.dest_abs_global t
c960bfcb91db discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents: 74408
diff changeset
   131
      in strip_context_aux (v :: params, Hs, u) end
37781
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   132
  | strip_context_aux (params, Hs, B) = (rev params, rev Hs, B);
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   133
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   134
fun strip_context A = strip_context_aux ([], [], A);
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   135
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   136
(*Left-to-right replacements: ctpairs = [...,(vi,ti),...].
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   137
  Instantiates distinct free variables by terms of same type.*)
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   138
fun free_instantiate ctpairs =
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   139
  forall_elim_list (map snd ctpairs) o forall_intr_list (map fst ctpairs);
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   140
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   141
fun free_of s ((a, i), T) =
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   142
  Free (s ^ (case i of 0 => a | _ => a ^ "_" ^ string_of_int i), T)
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   143
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   144
fun mk_inst v = (Var v, free_of "METAHYP1_" v)
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   145
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   146
fun metahyps_split_prem prem =
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   147
  let (*find all vars in the hyps -- should find tvars also!*)
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   148
      val hyps_vars = fold Term.add_vars (Logic.strip_assums_hyp prem) []
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   149
      val insts = map mk_inst hyps_vars
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   150
      (*replace the hyps_vars by Frees*)
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   151
      val prem' = subst_atomic insts prem
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   152
      val (params,hyps,concl) = strip_context prem'
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   153
  in (insts,params,hyps,concl)  end;
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   154
59165
115965966e15 proper context;
wenzelm
parents: 58950
diff changeset
   155
fun metahyps_aux_tac ctxt tacf (prem,gno) state =
37781
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   156
  let val (insts,params,hyps,concl) = metahyps_split_prem prem
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   157
      val maxidx = Thm.maxidx_of state
60358
aebfbcab1eb8 clarified context;
wenzelm
parents: 59621
diff changeset
   158
      val chyps = map (Thm.cterm_of ctxt) hyps
37781
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   159
      val hypths = map Thm.assume chyps
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   160
      val subprems = map (Thm.forall_elim_vars 0) hypths
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   161
      val fparams = map Free params
60358
aebfbcab1eb8 clarified context;
wenzelm
parents: 59621
diff changeset
   162
      val cparams = map (Thm.cterm_of ctxt) fparams
aebfbcab1eb8 clarified context;
wenzelm
parents: 59621
diff changeset
   163
      fun swap_ctpair (t, u) = apply2 (Thm.cterm_of ctxt) (u, t)
37781
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   164
      (*Subgoal variables: make Free; lift type over params*)
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   165
      fun mk_subgoal_inst concl_vars (v, T) =
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   166
          if member (op =) concl_vars (v, T)
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   167
          then ((v, T), true, free_of "METAHYP2_" (v, T))
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   168
          else ((v, T), false, free_of "METAHYP2_" (v, map #2 params ---> T))
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   169
      (*Instantiate subgoal vars by Free applied to params*)
60642
48dd1cefb4ae simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents: 60358
diff changeset
   170
      fun mk_inst (v, in_concl, u) =
48dd1cefb4ae simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents: 60358
diff changeset
   171
          if in_concl then (v, Thm.cterm_of ctxt u)
48dd1cefb4ae simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents: 60358
diff changeset
   172
          else (v, Thm.cterm_of ctxt (list_comb (u, fparams)))
37781
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   173
      (*Restore Vars with higher type and index*)
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   174
      fun mk_subgoal_swap_ctpair (((a, i), T), in_concl, u as Free (_, U)) =
60358
aebfbcab1eb8 clarified context;
wenzelm
parents: 59621
diff changeset
   175
          if in_concl then apply2 (Thm.cterm_of ctxt) (u, Var ((a, i), T))
aebfbcab1eb8 clarified context;
wenzelm
parents: 59621
diff changeset
   176
          else apply2 (Thm.cterm_of ctxt) (u, Var ((a, i + maxidx), U))
37781
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   177
      (*Embed B in the original context of params and hyps*)
46215
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 45862
diff changeset
   178
      fun embed B = fold_rev Logic.all fparams (Logic.list_implies (hyps, B))
37781
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   179
      (*Strip the context using elimination rules*)
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   180
      fun elim Bhyp = implies_elim_list (forall_elim_list cparams Bhyp) hypths
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   181
      (*A form of lifting that discharges assumptions.*)
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   182
      fun relift st =
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   183
        let val prop = Thm.prop_of st
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   184
            val subgoal_vars = (*Vars introduced in the subgoals*)
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   185
              fold Term.add_vars (Logic.strip_imp_prems prop) []
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   186
            and concl_vars = Term.add_vars (Logic.strip_imp_concl prop) []
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   187
            val subgoal_insts = map (mk_subgoal_inst concl_vars) subgoal_vars
74282
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 74274
diff changeset
   188
            val st' =
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 74274
diff changeset
   189
              Thm.instantiate (TVars.empty, Vars.build (fold (Vars.add o mk_inst) subgoal_insts)) st
60358
aebfbcab1eb8 clarified context;
wenzelm
parents: 59621
diff changeset
   190
            val emBs = map (Thm.cterm_of ctxt o embed) (Thm.prems_of st')
37781
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   191
            val Cth  = implies_elim_list st' (map (elim o Thm.assume) emBs)
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   192
        in  (*restore the unknowns to the hypotheses*)
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   193
            free_instantiate (map swap_ctpair insts @
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   194
                              map mk_subgoal_swap_ctpair subgoal_insts)
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   195
                (*discharge assumptions from state in same order*)
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   196
                (implies_intr_list emBs
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   197
                  (forall_intr_list cparams (implies_intr_list chyps Cth)))
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   198
        end
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   199
      (*function to replace the current subgoal*)
52223
5bb6ae8acb87 tuned signature -- more explicit flags for low-level Thm.bicompose;
wenzelm
parents: 51429
diff changeset
   200
      fun next st =
59165
115965966e15 proper context;
wenzelm
parents: 58950
diff changeset
   201
        Thm.bicompose (SOME ctxt) {flatten = true, match = false, incremented = false}
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59165
diff changeset
   202
          (false, relift st, Thm.nprems_of st) gno state
60358
aebfbcab1eb8 clarified context;
wenzelm
parents: 59621
diff changeset
   203
  in Seq.maps next (tacf subprems (Thm.trivial (Thm.cterm_of ctxt concl))) end;
37781
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   204
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   205
in
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   206
59165
115965966e15 proper context;
wenzelm
parents: 58950
diff changeset
   207
fun METAHYPS ctxt tacf n thm = SUBGOAL (metahyps_aux_tac ctxt tacf) n thm
64548
8b187a7a9776 avoid spurious messages -- potential cause of problems for "meson";
wenzelm
parents: 61914
diff changeset
   208
  handle THM ("assume: variables", _, _) => Seq.empty
37781
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   209
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   210
end;
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   211
47022
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   212
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   213
(* generating identifiers -- often fresh *)
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   214
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   215
local
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   216
(*Maps 0-61 to A-Z, a-z, 0-9; exclude _ or ' to avoid clash with internal/unusual indentifiers*)
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   217
fun gensym_char i =
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   218
  if i<26 then chr (ord "A" + i)
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   219
  else if i<52 then chr (ord "a" + i - 26)
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   220
  else chr (ord "0" + i - 52);
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   221
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   222
val char_vec = Vector.tabulate (62, gensym_char);
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   223
fun newid n = implode (map (fn i => Vector.sub (char_vec, i)) (radixpand (62, n)));
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   224
56147
9589605bcf41 prefer more robust Synchronized.var;
wenzelm
parents: 52223
diff changeset
   225
val gensym_seed = Synchronized.var "gensym_seed" (0: int);
47022
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   226
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   227
in
56147
9589605bcf41 prefer more robust Synchronized.var;
wenzelm
parents: 52223
diff changeset
   228
  fun gensym pre =
9589605bcf41 prefer more robust Synchronized.var;
wenzelm
parents: 52223
diff changeset
   229
    Synchronized.change_result gensym_seed (fn i => (pre ^ newid i, i + 1));
37781
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   230
end;
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
diff changeset
   231
47022
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   232
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   233
(*Convert all Vars in a theorem to Frees.  Also return a function for
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   234
  reversing that operation.  DOES NOT WORK FOR TYPE VARIABLES.*)
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   235
60358
aebfbcab1eb8 clarified context;
wenzelm
parents: 59621
diff changeset
   236
fun freeze_thaw_robust ctxt th =
47022
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   237
 let val fth = Thm.legacy_freezeT th
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   238
 in
74241
eb265f54e3ce more efficient operations: traverse hyps only when required;
wenzelm
parents: 69593
diff changeset
   239
   case Thm.fold_terms {hyps = false} Term.add_vars fth [] of
47576
b32aae03e3d6 dropped dead code;
haftmann
parents: 47022
diff changeset
   240
       [] => (fth, fn _ => fn x => x)   (*No vars: nothing to do!*)
47022
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   241
     | vars =>
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   242
         let fun newName (ix,_) = (ix, gensym (string_of_indexname ix))
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   243
             val alist = map newName vars
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   244
             fun mk_inst (v,T) =
60358
aebfbcab1eb8 clarified context;
wenzelm
parents: 59621
diff changeset
   245
                 apply2 (Thm.cterm_of ctxt)
aebfbcab1eb8 clarified context;
wenzelm
parents: 59621
diff changeset
   246
                  (Var (v, T), Free (the (AList.lookup (op =) alist v), T))
47022
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   247
             val insts = map mk_inst vars
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   248
             fun thaw i th' = (*i is non-negative increment for Var indexes*)
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   249
                 th' |> forall_intr_list (map #2 insts)
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   250
                     |> forall_elim_list (map (Thm.incr_indexes_cterm i o #1) insts)
74282
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 74274
diff changeset
   251
         in
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 74274
diff changeset
   252
           (Thm.instantiate (TVars.empty, Vars.build (fold (Vars.add o apfst (dest_Var o Thm.term_of)) insts)) fth, thaw)
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 74274
diff changeset
   253
         end
47022
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   254
 end;
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   255
8eac39af4ec0 moved some legacy stuff;
wenzelm
parents: 46909
diff changeset
   256
end;