author | wenzelm |
Mon, 15 Oct 2001 20:36:48 +0200 | |
changeset 11784 | b66b198ee29a |
parent 11730 | 418533653668 |
child 11791 | 2c201f3b018e |
permissions | -rw-r--r-- |
11658 | 1 |
(* Title: Pure/Isar/induct_attrib.ML |
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
2 |
ID: $Id$ |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
3 |
Author: Markus Wenzel, TU Muenchen |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
4 |
License: GPL (GNU GENERAL PUBLIC LICENSE) |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
5 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
6 |
Declaration of rules for cases and induction. |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
7 |
*) |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
8 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
9 |
signature INDUCT_ATTRIB = |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
10 |
sig |
11730 | 11 |
val vars_of: term -> term list |
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
12 |
val dest_global_rules: theory -> |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
13 |
{type_cases: (string * thm) list, set_cases: (string * thm) list, |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
14 |
type_induct: (string * thm) list, set_induct: (string * thm) list} |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
15 |
val print_global_rules: theory -> unit |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
16 |
val dest_local_rules: Proof.context -> |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
17 |
{type_cases: (string * thm) list, set_cases: (string * thm) list, |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
18 |
type_induct: (string * thm) list, set_induct: (string * thm) list} |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
19 |
val print_local_rules: Proof.context -> unit |
11730 | 20 |
val lookup_casesT : Proof.context -> string -> thm option |
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
21 |
val lookup_casesS : Proof.context -> string -> thm option |
11730 | 22 |
val lookup_inductT : Proof.context -> string -> thm option |
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
23 |
val lookup_inductS : Proof.context -> string -> thm option |
11730 | 24 |
val find_casesT: Proof.context -> typ -> thm list |
25 |
val find_casesS: Proof.context -> thm -> thm list |
|
26 |
val find_inductT: Proof.context -> typ -> thm list |
|
27 |
val find_inductS: Proof.context -> thm -> thm list |
|
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
28 |
val cases_type_global: string -> theory attribute |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
29 |
val cases_set_global: string -> theory attribute |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
30 |
val cases_type_local: string -> Proof.context attribute |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
31 |
val cases_set_local: string -> Proof.context attribute |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
32 |
val induct_type_global: string -> theory attribute |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
33 |
val induct_set_global: string -> theory attribute |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
34 |
val induct_type_local: string -> Proof.context attribute |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
35 |
val induct_set_local: string -> Proof.context attribute |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
36 |
val casesN: string |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
37 |
val inductN: string |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
38 |
val typeN: string |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
39 |
val setN: string |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
40 |
val setup: (theory -> theory) list |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
41 |
end; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
42 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
43 |
structure InductAttrib: INDUCT_ATTRIB = |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
44 |
struct |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
45 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
46 |
|
11730 | 47 |
(** misc utils **) |
48 |
||
49 |
(* encode_type -- for indexing purposes *) |
|
50 |
||
51 |
fun encode_type (Type (c, Ts)) = Term.list_comb (Const (c, dummyT), map encode_type Ts) |
|
52 |
| encode_type (TFree (a, _)) = Free (a, dummyT) |
|
53 |
| encode_type (TVar (a, _)) = Var (a, dummyT); |
|
54 |
||
55 |
||
56 |
(* variables -- ordered left-to-right, preferring right *) |
|
57 |
||
58 |
local |
|
59 |
||
60 |
fun rev_vars_of tm = |
|
61 |
Term.foldl_aterms (fn (ts, t as Var _) => t :: ts | (ts, _) => ts) ([], tm) |
|
62 |
|> Library.distinct; |
|
63 |
||
64 |
val mk_var = encode_type o #2 o Term.dest_Var; |
|
65 |
||
66 |
in |
|
67 |
||
68 |
val vars_of = rev o rev_vars_of; |
|
69 |
||
70 |
fun first_var thm = mk_var (hd (vars_of (hd (Thm.prems_of thm)))) handle LIST _ => |
|
71 |
raise THM ("No variables in first premise of rule", 0, [thm]); |
|
72 |
||
73 |
fun last_var thm = mk_var (hd (rev_vars_of (Thm.concl_of thm))) handle LIST _ => |
|
74 |
raise THM ("No variables in conclusion of rule", 0, [thm]); |
|
75 |
||
76 |
end; |
|
77 |
||
78 |
||
79 |
||
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
80 |
(** global and local induct data **) |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
81 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
82 |
(* rules *) |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
83 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
84 |
type rules = (string * thm) NetRules.T; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
85 |
|
11730 | 86 |
val init_rules = NetRules.init (fn ((s1: string, th1), (s2, th2)) => s1 = s2 |
11784 | 87 |
andalso Thm.eq_thm (th1, th2)) (K 0); |
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
88 |
|
11730 | 89 |
fun lookup_rule (rs: rules) name = Library.assoc (NetRules.rules rs, name); |
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
90 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
91 |
fun print_rules kind sg rs = |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
92 |
let val thms = map snd (NetRules.rules rs) |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
93 |
in Pretty.writeln (Pretty.big_list kind (map (Display.pretty_thm_sg sg) thms)) end; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
94 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
95 |
|
11658 | 96 |
(* theory data kind 'Isar/induction' *) |
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
97 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
98 |
structure GlobalInductArgs = |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
99 |
struct |
11658 | 100 |
val name = "Isar/induction"; |
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
101 |
type T = (rules * rules) * (rules * rules); |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
102 |
|
11730 | 103 |
val empty = |
104 |
((init_rules (first_var o #2), init_rules (Thm.major_prem_of o #2)), |
|
105 |
(init_rules (last_var o #2), init_rules (Thm.major_prem_of o #2))); |
|
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
106 |
val copy = I; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
107 |
val prep_ext = I; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
108 |
fun merge (((casesT1, casesS1), (inductT1, inductS1)), |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
109 |
((casesT2, casesS2), (inductT2, inductS2))) = |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
110 |
((NetRules.merge (casesT1, casesT2), NetRules.merge (casesS1, casesS2)), |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
111 |
(NetRules.merge (inductT1, inductT2), NetRules.merge (inductS1, inductS2))); |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
112 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
113 |
fun print sg ((casesT, casesS), (inductT, inductS)) = |
11665 | 114 |
(print_rules "induct type:" sg inductT; |
115 |
print_rules "induct set:" sg inductS; |
|
116 |
print_rules "cases type:" sg casesT; |
|
117 |
print_rules "cases set:" sg casesS); |
|
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
118 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
119 |
fun dest ((casesT, casesS), (inductT, inductS)) = |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
120 |
{type_cases = NetRules.rules casesT, |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
121 |
set_cases = NetRules.rules casesS, |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
122 |
type_induct = NetRules.rules inductT, |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
123 |
set_induct = NetRules.rules inductS}; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
124 |
end; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
125 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
126 |
structure GlobalInduct = TheoryDataFun(GlobalInductArgs); |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
127 |
val print_global_rules = GlobalInduct.print; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
128 |
val dest_global_rules = GlobalInductArgs.dest o GlobalInduct.get; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
129 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
130 |
|
11658 | 131 |
(* proof data kind 'Isar/induction' *) |
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
132 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
133 |
structure LocalInductArgs = |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
134 |
struct |
11658 | 135 |
val name = "Isar/induction"; |
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
136 |
type T = GlobalInductArgs.T; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
137 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
138 |
fun init thy = GlobalInduct.get thy; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
139 |
fun print x = GlobalInductArgs.print (ProofContext.sign_of x); |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
140 |
end; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
141 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
142 |
structure LocalInduct = ProofDataFun(LocalInductArgs); |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
143 |
val print_local_rules = LocalInduct.print; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
144 |
val dest_local_rules = GlobalInductArgs.dest o LocalInduct.get; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
145 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
146 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
147 |
(* access rules *) |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
148 |
|
11730 | 149 |
val lookup_casesT = lookup_rule o #1 o #1 o LocalInduct.get; |
150 |
val lookup_casesS = lookup_rule o #2 o #1 o LocalInduct.get; |
|
151 |
val lookup_inductT = lookup_rule o #1 o #2 o LocalInduct.get; |
|
152 |
val lookup_inductS = lookup_rule o #2 o #2 o LocalInduct.get; |
|
153 |
||
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
154 |
|
11730 | 155 |
fun find_rules which how ctxt x = |
11784 | 156 |
map snd (NetRules.retrieve (which (LocalInduct.get ctxt)) (how x)); |
11730 | 157 |
|
158 |
val find_casesT = find_rules (#1 o #1) encode_type; |
|
159 |
val find_casesS = find_rules (#2 o #1) Thm.concl_of; |
|
160 |
val find_inductT = find_rules (#1 o #2) encode_type; |
|
161 |
val find_inductS = find_rules (#2 o #2) Thm.concl_of; |
|
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
162 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
163 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
164 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
165 |
(** attributes **) |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
166 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
167 |
local |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
168 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
169 |
fun mk_att f g h name arg = |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
170 |
let val (x, thm) = h arg in (f (g (name, thm)) x, thm) end; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
171 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
172 |
fun add_casesT rule x = apfst (apfst (NetRules.insert rule)) x; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
173 |
fun add_casesS rule x = apfst (apsnd (NetRules.insert rule)) x; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
174 |
fun add_inductT rule x = apsnd (apfst (NetRules.insert rule)) x; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
175 |
fun add_inductS rule x = apsnd (apsnd (NetRules.insert rule)) x; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
176 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
177 |
fun consumes0 x = RuleCases.consumes_default 0 x; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
178 |
fun consumes1 x = RuleCases.consumes_default 1 x; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
179 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
180 |
in |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
181 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
182 |
val cases_type_global = mk_att GlobalInduct.map add_casesT consumes0; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
183 |
val cases_set_global = mk_att GlobalInduct.map add_casesS consumes1; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
184 |
val induct_type_global = mk_att GlobalInduct.map add_inductT consumes0; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
185 |
val induct_set_global = mk_att GlobalInduct.map add_inductS consumes1; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
186 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
187 |
val cases_type_local = mk_att LocalInduct.map add_casesT consumes0; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
188 |
val cases_set_local = mk_att LocalInduct.map add_casesS consumes1; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
189 |
val induct_type_local = mk_att LocalInduct.map add_inductT consumes0; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
190 |
val induct_set_local = mk_att LocalInduct.map add_inductS consumes1; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
191 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
192 |
end; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
193 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
194 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
195 |
(** concrete syntax **) |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
196 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
197 |
val casesN = "cases"; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
198 |
val inductN = "induct"; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
199 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
200 |
val typeN = "type"; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
201 |
val setN = "set"; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
202 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
203 |
local |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
204 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
205 |
fun spec k = (Args.$$$ k -- Args.colon) |-- Args.!!! Args.name; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
206 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
207 |
fun attrib sign_of add_type add_set = Scan.depend (fn x => |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
208 |
let val sg = sign_of x in |
11730 | 209 |
spec typeN >> (add_type o Sign.certify_tyname sg o Sign.intern_tycon sg) || |
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
210 |
spec setN >> (add_set o Sign.certify_const sg o Sign.intern_const sg) |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
211 |
end >> pair x); |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
212 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
213 |
in |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
214 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
215 |
val cases_attr = |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
216 |
(Attrib.syntax (attrib Theory.sign_of cases_type_global cases_set_global), |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
217 |
Attrib.syntax (attrib ProofContext.sign_of cases_type_local cases_set_local)); |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
218 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
219 |
val induct_attr = |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
220 |
(Attrib.syntax (attrib Theory.sign_of induct_type_global induct_set_global), |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
221 |
Attrib.syntax (attrib ProofContext.sign_of induct_type_local induct_set_local)); |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
222 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
223 |
end; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
224 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
225 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
226 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
227 |
(** theory setup **) |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
228 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
229 |
val setup = |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
230 |
[GlobalInduct.init, LocalInduct.init, |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
231 |
Attrib.add_attributes |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
232 |
[(casesN, cases_attr, "declaration of cases rule for type or set"), |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
233 |
(inductN, induct_attr, "declaration of induction rule for type or set")]]; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
234 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
235 |
end; |