-
Notifications
You must be signed in to change notification settings - Fork 0
/
TSubst.thy
71 lines (59 loc) · 2.16 KB
/
TSubst.thy
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
(*
* Copyright 2020, Data61, CSIRO (ABN 41 687 119 230)
*
* SPDX-License-Identifier: BSD-2-Clause
*)
(*
* Term subst method: substitution with a given term as the equation.
* The equation will be added as a subgoal.
* See example at the end of this file.
*)
theory TSubst
imports
Main
begin
method_setup tsubst = \<open>
Scan.lift (Args.mode "asm" --
Scan.optional (Args.parens (Scan.repeat Parse.nat)) [0] --
Parse.term)
>> (fn ((asm,occs),t) => (fn ctxt =>
Method.SIMPLE_METHOD (Subgoal.FOCUS_PARAMS (fn focus => (fn thm =>
let
(* This code used to use Thm.certify_inst in 2014, which was removed.
The following is just a best guess for what it did. *)
fun certify_inst ctxt (typ_insts, term_insts) =
(typ_insts
|> map (fn (tvar, inst) =>
(Thm.ctyp_of ctxt (TVar tvar),
Thm.ctyp_of ctxt inst)),
term_insts
|> map (fn (var, inst) =>
(Thm.cterm_of ctxt (Var var),
Thm.cterm_of ctxt inst)))
val ctxt' = #context focus
val ((_, schematic_terms), ctxt2) =
Variable.import_inst true [(#concl focus) |> Thm.term_of] ctxt'
|>> certify_inst ctxt'
val ctxt3 = fold (fn (t,t') => Variable.bind_term (Thm.term_of t |> Term.dest_Var |> fst, (t' |> Thm.term_of))) schematic_terms ctxt2
val athm = Syntax.read_term ctxt3 t
|> Object_Logic.ensure_propT ctxt'
|> Thm.cterm_of ctxt'
|> Thm.trivial
val thm' = Thm.instantiate ([], map (apfst (Thm.term_of #> dest_Var)) schematic_terms) thm
in
(if asm then EqSubst.eqsubst_asm_tac else EqSubst.eqsubst_tac)
ctxt3 occs [athm] 1 thm'
|> Seq.map (singleton (Variable.export ctxt3 ctxt'))
end)) ctxt 1)))
\<close> "subst, with term instead of theorem as equation"
schematic_goal
assumes a: "\<And>x y. P x \<Longrightarrow> P y"
fixes x :: 'b
shows "\<And>x ::'a :: type. ?Q x \<Longrightarrow> P x \<and> ?Q x"
apply (tsubst (asm) "?Q x = (P x \<and> P x)")
apply (rule refl)
apply (tsubst "P x = P y",simp add:a)+
apply (tsubst (2) "P y = P x", simp add:a)
apply (clarsimp simp: a)
done
end