author | nipkow |
Sun, 23 Sep 2018 15:42:19 +0200 | |
changeset 69038 | 2ce9bc515a64 |
parent 62178 | c3c98ed94b0f |
permissions | -rw-r--r-- |
30092
9c3b1c136d1f
Added lemmas for normalizing freshness results involving fresh_star.
berghofe
parents:
29581
diff
changeset
|
1 |
(* Author: Christian Urban and Makarius |
18283
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
2 |
|
18288
feb79a6b274b
proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents:
18283
diff
changeset
|
3 |
The nominal induct proof method. |
18283
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
4 |
*) |
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
5 |
|
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
6 |
structure NominalInduct: |
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
7 |
sig |
61841
4d3527b94f2a
more general types Proof.method / context_tactic;
wenzelm
parents:
61834
diff
changeset
|
8 |
val nominal_induct_tac: bool -> (binding option * (term * bool)) option list list -> |
18583
96e1ef2f806f
proper handling of simultaneous goals and mutual rules;
wenzelm
parents:
18311
diff
changeset
|
9 |
(string * typ) list -> (string * typ) list list -> thm list -> |
61841
4d3527b94f2a
more general types Proof.method / context_tactic;
wenzelm
parents:
61834
diff
changeset
|
10 |
thm list -> int -> context_tactic |
30549 | 11 |
val nominal_induct_method: (Proof.context -> Proof.method) context_parser |
18283
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
12 |
end = |
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
13 |
struct |
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
14 |
|
18288
feb79a6b274b
proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents:
18283
diff
changeset
|
15 |
(* proper tuples -- nested left *) |
18283
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
16 |
|
18288
feb79a6b274b
proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents:
18283
diff
changeset
|
17 |
fun tupleT Ts = HOLogic.unitT |> fold (fn T => fn U => HOLogic.mk_prodT (U, T)) Ts; |
feb79a6b274b
proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents:
18283
diff
changeset
|
18 |
fun tuple ts = HOLogic.unit |> fold (fn t => fn u => HOLogic.mk_prod (u, t)) ts; |
feb79a6b274b
proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents:
18283
diff
changeset
|
19 |
|
feb79a6b274b
proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents:
18283
diff
changeset
|
20 |
fun tuple_fun Ts (xi, T) = |
61424
c3658c18b7bc
prod_case as canonical name for product type eliminator
haftmann
parents:
60787
diff
changeset
|
21 |
Library.funpow (length Ts) HOLogic.mk_case_prod |
18288
feb79a6b274b
proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents:
18283
diff
changeset
|
22 |
(Var (xi, (HOLogic.unitT :: Ts) ---> Term.range_type T)); |
18283
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
23 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
45328
diff
changeset
|
24 |
fun split_all_tuples ctxt = |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
45328
diff
changeset
|
25 |
Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimps |
37137 | 26 |
[@{thm split_conv}, @{thm split_paired_all}, @{thm unit_all_eq1}, @{thm fresh_unit_elim}, @{thm fresh_prod_elim}] @ |
30092
9c3b1c136d1f
Added lemmas for normalizing freshness results involving fresh_star.
berghofe
parents:
29581
diff
changeset
|
27 |
@{thms fresh_star_unit_elim} @ @{thms fresh_star_prod_elim}); |
18283
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
28 |
|
18288
feb79a6b274b
proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents:
18283
diff
changeset
|
29 |
|
18297
116fe71fad51
fresh: frees instead of terms, rename corresponding params in rule;
wenzelm
parents:
18288
diff
changeset
|
30 |
(* prepare rule *) |
18288
feb79a6b274b
proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents:
18283
diff
changeset
|
31 |
|
19903 | 32 |
fun inst_mutual_rule ctxt insts avoiding rules = |
18283
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
33 |
let |
33368 | 34 |
val (nconcls, joined_rule) = Rule_Cases.strict_mutual_rule ctxt rules; |
22072 | 35 |
val concls = Logic.dest_conjunctions (Thm.concl_of joined_rule); |
33368 | 36 |
val (cases, consumes) = Rule_Cases.get joined_rule; |
18583
96e1ef2f806f
proper handling of simultaneous goals and mutual rules;
wenzelm
parents:
18311
diff
changeset
|
37 |
|
96e1ef2f806f
proper handling of simultaneous goals and mutual rules;
wenzelm
parents:
18311
diff
changeset
|
38 |
val l = length rules; |
96e1ef2f806f
proper handling of simultaneous goals and mutual rules;
wenzelm
parents:
18311
diff
changeset
|
39 |
val _ = |
96e1ef2f806f
proper handling of simultaneous goals and mutual rules;
wenzelm
parents:
18311
diff
changeset
|
40 |
if length insts = l then () |
96e1ef2f806f
proper handling of simultaneous goals and mutual rules;
wenzelm
parents:
18311
diff
changeset
|
41 |
else error ("Bad number of instantiations for " ^ string_of_int l ^ " rules"); |
18288
feb79a6b274b
proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents:
18283
diff
changeset
|
42 |
|
22072 | 43 |
fun subst inst concl = |
18583
96e1ef2f806f
proper handling of simultaneous goals and mutual rules;
wenzelm
parents:
18311
diff
changeset
|
44 |
let |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
23591
diff
changeset
|
45 |
val vars = Induct.vars_of concl; |
18583
96e1ef2f806f
proper handling of simultaneous goals and mutual rules;
wenzelm
parents:
18311
diff
changeset
|
46 |
val m = length vars and n = length inst; |
96e1ef2f806f
proper handling of simultaneous goals and mutual rules;
wenzelm
parents:
18311
diff
changeset
|
47 |
val _ = if m >= n + 2 then () else error "Too few variables in conclusion of rule"; |
96e1ef2f806f
proper handling of simultaneous goals and mutual rules;
wenzelm
parents:
18311
diff
changeset
|
48 |
val P :: x :: ys = vars; |
33957 | 49 |
val zs = drop (m - n - 2) ys; |
18583
96e1ef2f806f
proper handling of simultaneous goals and mutual rules;
wenzelm
parents:
18311
diff
changeset
|
50 |
in |
96e1ef2f806f
proper handling of simultaneous goals and mutual rules;
wenzelm
parents:
18311
diff
changeset
|
51 |
(P, tuple_fun (map #2 avoiding) (Term.dest_Var P)) :: |
96e1ef2f806f
proper handling of simultaneous goals and mutual rules;
wenzelm
parents:
18311
diff
changeset
|
52 |
(x, tuple (map Free avoiding)) :: |
32952 | 53 |
map_filter (fn (z, SOME t) => SOME (z, t) | _ => NONE) (zs ~~ inst) |
18583
96e1ef2f806f
proper handling of simultaneous goals and mutual rules;
wenzelm
parents:
18311
diff
changeset
|
54 |
end; |
96e1ef2f806f
proper handling of simultaneous goals and mutual rules;
wenzelm
parents:
18311
diff
changeset
|
55 |
val substs = |
32952 | 56 |
map2 subst insts concls |> flat |> distinct (op =) |
60787 | 57 |
|> map (fn (t, u) => (#1 (dest_Var t), Thm.cterm_of ctxt u)); |
22072 | 58 |
in |
60787 | 59 |
(((cases, nconcls), consumes), infer_instantiate ctxt substs joined_rule) |
22072 | 60 |
end; |
18283
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
61 |
|
18299
af72dfc4b9f9
added rename_params_rule: recover orginal fresh names in subgoals/cases;
wenzelm
parents:
18297
diff
changeset
|
62 |
fun rename_params_rule internal xs rule = |
18297
116fe71fad51
fresh: frees instead of terms, rename corresponding params in rule;
wenzelm
parents:
18288
diff
changeset
|
63 |
let |
18299
af72dfc4b9f9
added rename_params_rule: recover orginal fresh names in subgoals/cases;
wenzelm
parents:
18297
diff
changeset
|
64 |
val tune = |
20072 | 65 |
if internal then Name.internal |
55947 | 66 |
else perhaps (try Name.dest_internal); |
18299
af72dfc4b9f9
added rename_params_rule: recover orginal fresh names in subgoals/cases;
wenzelm
parents:
18297
diff
changeset
|
67 |
val n = length xs; |
af72dfc4b9f9
added rename_params_rule: recover orginal fresh names in subgoals/cases;
wenzelm
parents:
18297
diff
changeset
|
68 |
fun rename prem = |
af72dfc4b9f9
added rename_params_rule: recover orginal fresh names in subgoals/cases;
wenzelm
parents:
18297
diff
changeset
|
69 |
let |
af72dfc4b9f9
added rename_params_rule: recover orginal fresh names in subgoals/cases;
wenzelm
parents:
18297
diff
changeset
|
70 |
val ps = Logic.strip_params prem; |
af72dfc4b9f9
added rename_params_rule: recover orginal fresh names in subgoals/cases;
wenzelm
parents:
18297
diff
changeset
|
71 |
val p = length ps; |
af72dfc4b9f9
added rename_params_rule: recover orginal fresh names in subgoals/cases;
wenzelm
parents:
18297
diff
changeset
|
72 |
val ys = |
af72dfc4b9f9
added rename_params_rule: recover orginal fresh names in subgoals/cases;
wenzelm
parents:
18297
diff
changeset
|
73 |
if p < n then [] |
33957 | 74 |
else map (tune o #1) (take (p - n) ps) @ xs; |
45328 | 75 |
in Logic.list_rename_params ys prem end; |
18299
af72dfc4b9f9
added rename_params_rule: recover orginal fresh names in subgoals/cases;
wenzelm
parents:
18297
diff
changeset
|
76 |
fun rename_prems prop = |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
77 |
let val (As, C) = Logic.strip_horn prop |
18299
af72dfc4b9f9
added rename_params_rule: recover orginal fresh names in subgoals/cases;
wenzelm
parents:
18297
diff
changeset
|
78 |
in Logic.list_implies (map rename As, C) end; |
60313 | 79 |
in Thm.renamed_prop (rename_prems (Thm.prop_of rule)) rule end; |
18297
116fe71fad51
fresh: frees instead of terms, rename corresponding params in rule;
wenzelm
parents:
18288
diff
changeset
|
80 |
|
18283
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
81 |
|
18288
feb79a6b274b
proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents:
18283
diff
changeset
|
82 |
(* nominal_induct_tac *) |
18283
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
83 |
|
61841
4d3527b94f2a
more general types Proof.method / context_tactic;
wenzelm
parents:
61834
diff
changeset
|
84 |
fun nominal_induct_tac simp def_insts avoiding fixings rules facts i (ctxt, st) = |
18283
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
85 |
let |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
23591
diff
changeset
|
86 |
val ((insts, defs), defs_ctxt) = fold_map Induct.add_defs def_insts ctxt |>> split_list; |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
87 |
val atomized_defs = map (map (Conv.fconv_rule (Induct.atomize_cterm ctxt))) defs; |
18283
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
88 |
|
19115 | 89 |
val finish_rule = |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
45328
diff
changeset
|
90 |
split_all_tuples defs_ctxt |
26712
e2dcda7b0401
adapted to ProofContext.revert_skolem: extra Name.clean required;
wenzelm
parents:
25985
diff
changeset
|
91 |
#> rename_params_rule true |
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42361
diff
changeset
|
92 |
(map (Name.clean o Variable.revert_fixed defs_ctxt o fst) avoiding); |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
93 |
|
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
94 |
fun rule_cases ctxt r = |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
95 |
let val r' = if simp then Induct.simplified_rule ctxt r else r |
59970 | 96 |
in Rule_Cases.make_nested ctxt (Thm.prop_of r') (Induct.rulified_term ctxt r') end; |
61841
4d3527b94f2a
more general types Proof.method / context_tactic;
wenzelm
parents:
61834
diff
changeset
|
97 |
|
4d3527b94f2a
more general types Proof.method / context_tactic;
wenzelm
parents:
61834
diff
changeset
|
98 |
fun context_tac _ _ = |
18583
96e1ef2f806f
proper handling of simultaneous goals and mutual rules;
wenzelm
parents:
18311
diff
changeset
|
99 |
rules |
19903 | 100 |
|> inst_mutual_rule ctxt insts avoiding |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
101 |
|> Rule_Cases.consume ctxt (flat defs) facts |
18583
96e1ef2f806f
proper handling of simultaneous goals and mutual rules;
wenzelm
parents:
18311
diff
changeset
|
102 |
|> Seq.maps (fn (((cases, concls), (more_consumes, more_facts)), rule) => |
96e1ef2f806f
proper handling of simultaneous goals and mutual rules;
wenzelm
parents:
18311
diff
changeset
|
103 |
(PRECISE_CONJUNCTS (length concls) (ALLGOALS (fn j => |
96e1ef2f806f
proper handling of simultaneous goals and mutual rules;
wenzelm
parents:
18311
diff
changeset
|
104 |
(CONJUNCTS (ALLGOALS |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
105 |
let |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
106 |
val adefs = nth_list atomized_defs (j - 1); |
59582 | 107 |
val frees = fold (Term.add_frees o Thm.prop_of) adefs []; |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
108 |
val xs = nth_list fixings (j - 1); |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
109 |
val k = nth concls (j - 1) + more_consumes |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
110 |
in |
61841
4d3527b94f2a
more general types Proof.method / context_tactic;
wenzelm
parents:
61834
diff
changeset
|
111 |
Method.insert_tac ctxt (more_facts @ adefs) THEN' |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
112 |
(if simp then |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
113 |
Induct.rotate_tac k (length adefs) THEN' |
45132 | 114 |
Induct.arbitrary_tac defs_ctxt k (List.partition (member op = frees) xs |> op @) |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
115 |
else |
45132 | 116 |
Induct.arbitrary_tac defs_ctxt k xs) |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
117 |
end) |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
118 |
THEN' Induct.inner_atomize_tac defs_ctxt) j)) |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
119 |
THEN' Induct.atomize_tac ctxt) i st |> Seq.maps (fn st' => |
26940 | 120 |
Induct.guess_instance ctxt |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
121 |
(finish_rule (Induct.internalize ctxt more_consumes rule)) i st' |
18583
96e1ef2f806f
proper handling of simultaneous goals and mutual rules;
wenzelm
parents:
18311
diff
changeset
|
122 |
|> Seq.maps (fn rule' => |
61841
4d3527b94f2a
more general types Proof.method / context_tactic;
wenzelm
parents:
61834
diff
changeset
|
123 |
CONTEXT_CASES (rule_cases ctxt rule' cases) |
60754 | 124 |
(resolve_tac ctxt [rename_params_rule false [] rule'] i THEN |
61841
4d3527b94f2a
more general types Proof.method / context_tactic;
wenzelm
parents:
61834
diff
changeset
|
125 |
PRIMITIVE (singleton (Proof_Context.export defs_ctxt ctxt))) (ctxt, st')))); |
4d3527b94f2a
more general types Proof.method / context_tactic;
wenzelm
parents:
61834
diff
changeset
|
126 |
in |
4d3527b94f2a
more general types Proof.method / context_tactic;
wenzelm
parents:
61834
diff
changeset
|
127 |
(context_tac CONTEXT_THEN_ALL_NEW |
58957 | 128 |
((if simp then Induct.simplify_tac ctxt THEN' (TRY o Induct.trivial_tac ctxt) |
61841
4d3527b94f2a
more general types Proof.method / context_tactic;
wenzelm
parents:
61834
diff
changeset
|
129 |
else K all_tac) THEN_ALL_NEW Induct.rulify_tac ctxt)) i (ctxt, st) |
18283
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
130 |
end; |
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
131 |
|
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
132 |
|
18288
feb79a6b274b
proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents:
18283
diff
changeset
|
133 |
(* concrete syntax *) |
17870 | 134 |
|
135 |
local |
|
136 |
||
18583
96e1ef2f806f
proper handling of simultaneous goals and mutual rules;
wenzelm
parents:
18311
diff
changeset
|
137 |
val avoidingN = "avoiding"; |
20998
714a08286899
To be consistent with "induct", I renamed "fixing" to "arbitrary".
urbanc
parents:
20288
diff
changeset
|
138 |
val fixingN = "arbitrary"; (* to be consistent with induct; hopefully this changes again *) |
18283
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
139 |
val ruleN = "rule"; |
17870 | 140 |
|
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
141 |
val inst = Scan.lift (Args.$$$ "_") >> K NONE || |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
142 |
Args.term >> (SOME o rpair false) || |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
143 |
Scan.lift (Args.$$$ "(") |-- (Args.term >> (SOME o rpair true)) --| |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
144 |
Scan.lift (Args.$$$ ")"); |
17870 | 145 |
|
18283
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
146 |
val def_inst = |
28083
103d9282a946
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27809
diff
changeset
|
147 |
((Scan.lift (Args.binding --| (Args.$$$ "\<equiv>" || Args.$$$ "==")) >> SOME) |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
148 |
-- (Args.term >> rpair false)) >> SOME || |
18283
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
149 |
inst >> Option.map (pair NONE); |
18099 | 150 |
|
27370 | 151 |
val free = Args.context -- Args.term >> (fn (_, Free v) => v | (ctxt, t) => |
152 |
error ("Bad free variable: " ^ Syntax.string_of_term ctxt t)); |
|
18283
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
153 |
|
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
154 |
fun unless_more_args scan = Scan.unless (Scan.lift |
18583
96e1ef2f806f
proper handling of simultaneous goals and mutual rules;
wenzelm
parents:
18311
diff
changeset
|
155 |
((Args.$$$ avoidingN || Args.$$$ fixingN || Args.$$$ ruleN) -- Args.colon)) scan; |
18283
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
156 |
|
17870 | 157 |
|
18583
96e1ef2f806f
proper handling of simultaneous goals and mutual rules;
wenzelm
parents:
18311
diff
changeset
|
158 |
val avoiding = Scan.optional (Scan.lift (Args.$$$ avoidingN -- Args.colon) |-- |
18297
116fe71fad51
fresh: frees instead of terms, rename corresponding params in rule;
wenzelm
parents:
18288
diff
changeset
|
159 |
Scan.repeat (unless_more_args free)) []; |
17870 | 160 |
|
18283
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
161 |
val fixing = Scan.optional (Scan.lift (Args.$$$ fixingN -- Args.colon) |-- |
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
34907
diff
changeset
|
162 |
Parse.and_list' (Scan.repeat (unless_more_args free))) []; |
17870 | 163 |
|
19036 | 164 |
val rule_spec = Scan.lift (Args.$$$ "rule" -- Args.colon) |-- Attrib.thms; |
17870 | 165 |
|
166 |
in |
|
167 |
||
62178 | 168 |
val nominal_induct_method : (Proof.context -> Proof.method) context_parser = |
53168 | 169 |
Scan.lift (Args.mode Induct.no_simpN) -- |
170 |
(Parse.and_list' (Scan.repeat (unless_more_args def_inst)) -- |
|
171 |
avoiding -- fixing -- rule_spec) >> |
|
61841
4d3527b94f2a
more general types Proof.method / context_tactic;
wenzelm
parents:
61834
diff
changeset
|
172 |
(fn (no_simp, (((x, y), z), w)) => fn _ => fn facts => |
4d3527b94f2a
more general types Proof.method / context_tactic;
wenzelm
parents:
61834
diff
changeset
|
173 |
(nominal_induct_tac (not no_simp) x y z w facts 1)); |
17870 | 174 |
|
175 |
end; |
|
18283
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
176 |
|
f8a49f09a202
reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents:
18265
diff
changeset
|
177 |
end; |