author | Christian Urban <christian dot urban at kcl dot ac dot uk> |
Sun, 15 Dec 2013 15:14:40 +1100 | |
changeset 3226 | 780b7a2c50b6 |
parent 3222 | 8c53bcd5c0ae |
child 3229 | b52e8651591f |
permissions | -rw-r--r-- |
2631
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1 |
(* Author: Christian Urban and Makarius |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
2 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
3 |
The nominal induct proof method. |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
4 |
*) |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
5 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
6 |
structure NominalInduct: |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
7 |
sig |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
8 |
val nominal_induct_tac: Proof.context -> bool -> (binding option * (term * bool)) option list list -> |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
9 |
(string * typ) list -> (string * typ) list list -> thm list -> thm list -> int -> Rule_Cases.cases_tactic |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
10 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
11 |
val nominal_induct_method: (Proof.context -> Proof.method) context_parser |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
12 |
end = |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
13 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
14 |
struct |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
15 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
16 |
(* proper tuples -- nested left *) |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
17 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
18 |
fun tupleT Ts = HOLogic.unitT |> fold (fn T => fn U => HOLogic.mk_prodT (U, T)) Ts; |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
19 |
fun tuple ts = HOLogic.unit |> fold (fn t => fn u => HOLogic.mk_prod (u, t)) ts; |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
20 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
21 |
fun tuple_fun Ts (xi, T) = |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
22 |
Library.funpow (length Ts) HOLogic.mk_split |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
23 |
(Var (xi, (HOLogic.unitT :: Ts) ---> Term.range_type T)); |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
24 |
|
3218
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3046
diff
changeset
|
25 |
fun split_all_tuples ctxt = |
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3046
diff
changeset
|
26 |
Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimps |
2632
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
2631
diff
changeset
|
27 |
@{thms split_conv split_paired_all unit_all_eq1} @ |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
2631
diff
changeset
|
28 |
@{thms fresh_Unit_elim fresh_Pair_elim} @ |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
2631
diff
changeset
|
29 |
@{thms fresh_star_Unit_elim fresh_star_Pair_elim fresh_star_Un_elim} @ |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
2631
diff
changeset
|
30 |
@{thms fresh_star_insert_elim fresh_star_empty_elim}) |
2631
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
31 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
32 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
33 |
(* prepare rule *) |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
34 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
35 |
fun inst_mutual_rule ctxt insts avoiding rules = |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
36 |
let |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
37 |
val (nconcls, joined_rule) = Rule_Cases.strict_mutual_rule ctxt rules; |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
38 |
val concls = Logic.dest_conjunctions (Thm.concl_of joined_rule); |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
39 |
val (cases, consumes) = Rule_Cases.get joined_rule; |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
40 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
41 |
val l = length rules; |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
42 |
val _ = |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
43 |
if length insts = l then () |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
44 |
else error ("Bad number of instantiations for " ^ string_of_int l ^ " rules"); |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
45 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
46 |
fun subst inst concl = |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
47 |
let |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
48 |
val vars = Induct.vars_of concl; |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
49 |
val m = length vars and n = length inst; |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
50 |
val _ = if m >= n + 2 then () else error "Too few variables in conclusion of rule"; |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
51 |
val P :: x :: ys = vars; |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
52 |
val zs = drop (m - n - 2) ys; |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
53 |
in |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
54 |
(P, tuple_fun (map #2 avoiding) (Term.dest_Var P)) :: |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
55 |
(x, tuple (map Free avoiding)) :: |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
56 |
map_filter (fn (z, SOME t) => SOME (z, t) | _ => NONE) (zs ~~ inst) |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
57 |
end; |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
58 |
val substs = |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
59 |
map2 subst insts concls |> flat |> distinct (op =) |
3045
d0ad264f8c4f
updated to Isabelle 3 Nov; it includes a hack to work around a bug in the localised version of the quotient package
Christian Urban <urbanc@in.tum.de>
parents:
2781
diff
changeset
|
60 |
|> map (pairself (Thm.cterm_of (Proof_Context.theory_of ctxt))); |
2631
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
61 |
in |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
62 |
(((cases, nconcls), consumes), Drule.cterm_instantiate substs joined_rule) |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
63 |
end; |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
64 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
65 |
fun rename_params_rule internal xs rule = |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
66 |
let |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
67 |
val tune = |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
68 |
if internal then Name.internal |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
69 |
else fn x => the_default x (try Name.dest_internal x); |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
70 |
val n = length xs; |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
71 |
fun rename prem = |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
72 |
let |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
73 |
val ps = Logic.strip_params prem; |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
74 |
val p = length ps; |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
75 |
val ys = |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
76 |
if p < n then [] |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
77 |
else map (tune o #1) (take (p - n) ps) @ xs; |
3046
9b0324e1293b
all examples work again after quotient package has been "de-localised"
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
78 |
in Logic.list_rename_params ys prem end; |
2631
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
79 |
fun rename_prems prop = |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
80 |
let val (As, C) = Logic.strip_horn prop |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
81 |
in Logic.list_implies (map rename As, C) end; |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
82 |
in Thm.equal_elim (Thm.reflexive (Drule.cterm_fun rename_prems (Thm.cprop_of rule))) rule end; |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
83 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
84 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
85 |
(* nominal_induct_tac *) |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
86 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
87 |
fun nominal_induct_tac ctxt simp def_insts avoiding fixings rules facts = |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
88 |
let |
3045
d0ad264f8c4f
updated to Isabelle 3 Nov; it includes a hack to work around a bug in the localised version of the quotient package
Christian Urban <urbanc@in.tum.de>
parents:
2781
diff
changeset
|
89 |
val thy = Proof_Context.theory_of ctxt; |
2631
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
90 |
val cert = Thm.cterm_of thy; |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
91 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
92 |
val ((insts, defs), defs_ctxt) = fold_map Induct.add_defs def_insts ctxt |>> split_list; |
3226
780b7a2c50b6
updated to changes in Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3222
diff
changeset
|
93 |
val atomized_defs = map (map (Conv.fconv_rule (Induct.atomize_cterm defs_ctxt))) defs; |
2631
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
94 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
95 |
val finish_rule = |
3218
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3046
diff
changeset
|
96 |
split_all_tuples defs_ctxt |
2631
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
97 |
#> rename_params_rule true |
2781
542ff50555f5
updated to new Isabelle (> 9 May)
Christian Urban <urbanc@in.tum.de>
parents:
2632
diff
changeset
|
98 |
(map (Name.clean o Variable.revert_fixed defs_ctxt o fst) avoiding); |
2631
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
99 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
100 |
fun rule_cases ctxt r = |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
101 |
let val r' = if simp then Induct.simplified_rule ctxt r else r |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
102 |
in Rule_Cases.make_nested (Thm.prop_of r') (Induct.rulified_term r') end; |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
103 |
in |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
104 |
(fn i => fn st => |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
105 |
rules |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
106 |
|> inst_mutual_rule ctxt insts avoiding |
3226
780b7a2c50b6
updated to changes in Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3222
diff
changeset
|
107 |
|> Rule_Cases.consume ctxt (flat defs) facts |
2631
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
108 |
|> Seq.maps (fn (((cases, concls), (more_consumes, more_facts)), rule) => |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
109 |
(PRECISE_CONJUNCTS (length concls) (ALLGOALS (fn j => |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
110 |
(CONJUNCTS (ALLGOALS |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
111 |
let |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
112 |
val adefs = nth_list atomized_defs (j - 1); |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
113 |
val frees = fold (Term.add_frees o prop_of) adefs []; |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
114 |
val xs = nth_list fixings (j - 1); |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
115 |
val k = nth concls (j - 1) + more_consumes |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
116 |
in |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
117 |
Method.insert_tac (more_facts @ adefs) THEN' |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
118 |
(if simp then |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
119 |
Induct.rotate_tac k (length adefs) THEN' |
3045
d0ad264f8c4f
updated to Isabelle 3 Nov; it includes a hack to work around a bug in the localised version of the quotient package
Christian Urban <urbanc@in.tum.de>
parents:
2781
diff
changeset
|
120 |
Induct.arbitrary_tac defs_ctxt k |
2631
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
121 |
(List.partition (member op = frees) xs |> op @) |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
122 |
else |
3045
d0ad264f8c4f
updated to Isabelle 3 Nov; it includes a hack to work around a bug in the localised version of the quotient package
Christian Urban <urbanc@in.tum.de>
parents:
2781
diff
changeset
|
123 |
Induct.arbitrary_tac defs_ctxt k xs) |
2631
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
124 |
end) |
3226
780b7a2c50b6
updated to changes in Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3222
diff
changeset
|
125 |
THEN' Induct.inner_atomize_tac ctxt) j)) |
780b7a2c50b6
updated to changes in Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3222
diff
changeset
|
126 |
THEN' Induct.atomize_tac ctxt) i st |> Seq.maps (fn st' => |
2631
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
127 |
Induct.guess_instance ctxt |
3226
780b7a2c50b6
updated to changes in Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3222
diff
changeset
|
128 |
(finish_rule (Induct.internalize ctxt more_consumes rule)) i st' |
2631
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
129 |
|> Seq.maps (fn rule' => |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
130 |
CASES (rule_cases ctxt rule' cases) |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
131 |
(Tactic.rtac (rename_params_rule false [] rule') i THEN |
3045
d0ad264f8c4f
updated to Isabelle 3 Nov; it includes a hack to work around a bug in the localised version of the quotient package
Christian Urban <urbanc@in.tum.de>
parents:
2781
diff
changeset
|
132 |
PRIMITIVE (singleton (Proof_Context.export defs_ctxt ctxt))) st')))) |
2631
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
133 |
THEN_ALL_NEW_CASES |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
134 |
((if simp then Induct.simplify_tac ctxt THEN' (TRY o Induct.trivial_tac) |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
135 |
else K all_tac) |
3226
780b7a2c50b6
updated to changes in Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3222
diff
changeset
|
136 |
THEN_ALL_NEW Induct.rulify_tac ctxt) |
2631
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
137 |
end; |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
138 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
139 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
140 |
(* concrete syntax *) |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
141 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
142 |
local |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
143 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
144 |
val avoidingN = "avoiding"; |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
145 |
val fixingN = "arbitrary"; (* to be consistent with induct; hopefully this changes again *) |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
146 |
val ruleN = "rule"; |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
147 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
148 |
val inst = Scan.lift (Args.$$$ "_") >> K NONE || |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
149 |
Args.term >> (SOME o rpair false) || |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
150 |
Scan.lift (Args.$$$ "(") |-- (Args.term >> (SOME o rpair true)) --| |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
151 |
Scan.lift (Args.$$$ ")"); |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
152 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
153 |
val def_inst = |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
154 |
((Scan.lift (Args.binding --| (Args.$$$ "\<equiv>" || Args.$$$ "==")) >> SOME) |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
155 |
-- (Args.term >> rpair false)) >> SOME || |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
156 |
inst >> Option.map (pair NONE); |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
157 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
158 |
val free = Args.context -- Args.term >> (fn (_, Free v) => v | (ctxt, t) => |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
159 |
error ("Bad free variable: " ^ Syntax.string_of_term ctxt t)); |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
160 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
161 |
fun unless_more_args scan = Scan.unless (Scan.lift |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
162 |
((Args.$$$ avoidingN || Args.$$$ fixingN || Args.$$$ ruleN) -- Args.colon)) scan; |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
163 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
164 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
165 |
val avoiding = Scan.optional (Scan.lift (Args.$$$ avoidingN -- Args.colon) |-- |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
166 |
Scan.repeat (unless_more_args free)) []; |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
167 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
168 |
val fixing = Scan.optional (Scan.lift (Args.$$$ fixingN -- Args.colon) |-- |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
169 |
Parse.and_list' (Scan.repeat (unless_more_args free))) []; |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
170 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
171 |
val rule_spec = Scan.lift (Args.$$$ "rule" -- Args.colon) |-- Attrib.thms; |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
172 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
173 |
in |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
174 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
175 |
val nominal_induct_method = |
3222
8c53bcd5c0ae
updated to lates Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3218
diff
changeset
|
176 |
Scan.lift (Args.mode Induct.no_simpN) -- |
8c53bcd5c0ae
updated to lates Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3218
diff
changeset
|
177 |
(Parse.and_list' (Scan.repeat (unless_more_args def_inst)) -- |
8c53bcd5c0ae
updated to lates Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3218
diff
changeset
|
178 |
avoiding -- fixing -- rule_spec) >> |
2631
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
179 |
(fn (no_simp, (((x, y), z), w)) => fn ctxt => |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
180 |
RAW_METHOD_CASES (fn facts => |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
181 |
HEADGOAL (nominal_induct_tac ctxt (not no_simp) x y z w facts))); |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
182 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
183 |
end |
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
184 |
|
e73bd379e839
removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
185 |
end; |