author | wenzelm |
Tue, 22 Nov 2005 19:34:46 +0100 | |
changeset 18226 | 8fde30f5cca6 |
parent 18210 | ad4b8567f6eb |
child 18637 | 33a6f6caa617 |
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 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
5 |
Declaration of rules for cases and induction. |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
6 |
*) |
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 |
signature INDUCT_ATTRIB = |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
9 |
sig |
11730 | 10 |
val vars_of: term -> term list |
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
11 |
val dest_global_rules: theory -> |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
12 |
{type_cases: (string * thm) list, set_cases: (string * thm) list, |
18210 | 13 |
type_induct: (string * thm) list, set_induct: (string * thm) list, |
14 |
type_coinduct: (string * thm) list, set_coinduct: (string * thm) list} |
|
11656
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, |
18210 | 18 |
type_induct: (string * thm) list, set_induct: (string * thm) list, |
19 |
type_coinduct: (string * thm) list, set_coinduct: (string * thm) list} |
|
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
20 |
val print_local_rules: Proof.context -> unit |
11730 | 21 |
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
|
22 |
val lookup_casesS : Proof.context -> string -> thm option |
11730 | 23 |
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
|
24 |
val lookup_inductS : Proof.context -> string -> thm option |
18210 | 25 |
val lookup_coinductT : Proof.context -> string -> thm option |
26 |
val lookup_coinductS : Proof.context -> string -> thm option |
|
11730 | 27 |
val find_casesT: Proof.context -> typ -> thm list |
18226 | 28 |
val find_casesS: Proof.context -> term -> thm list |
11730 | 29 |
val find_inductT: Proof.context -> typ -> thm list |
18226 | 30 |
val find_inductS: Proof.context -> term -> thm list |
18210 | 31 |
val find_coinductT: Proof.context -> typ -> thm list |
18226 | 32 |
val find_coinductS: Proof.context -> term -> thm list |
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
33 |
val cases_type_global: string -> theory attribute |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
34 |
val cases_set_global: string -> theory attribute |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
35 |
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
|
36 |
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
|
37 |
val induct_type_global: string -> theory attribute |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
38 |
val induct_set_global: string -> theory attribute |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
39 |
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
|
40 |
val induct_set_local: string -> Proof.context attribute |
18210 | 41 |
val coinduct_type_global: string -> theory attribute |
42 |
val coinduct_set_global: string -> theory attribute |
|
43 |
val coinduct_type_local: string -> Proof.context attribute |
|
44 |
val coinduct_set_local: string -> Proof.context attribute |
|
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
45 |
val casesN: string |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
46 |
val inductN: string |
18210 | 47 |
val coinductN: string |
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
48 |
val typeN: string |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
49 |
val setN: string |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
50 |
end; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
51 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
52 |
structure InductAttrib: INDUCT_ATTRIB = |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
53 |
struct |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
54 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
55 |
|
11730 | 56 |
(** misc utils **) |
57 |
||
58 |
(* encode_type -- for indexing purposes *) |
|
59 |
||
60 |
fun encode_type (Type (c, Ts)) = Term.list_comb (Const (c, dummyT), map encode_type Ts) |
|
61 |
| encode_type (TFree (a, _)) = Free (a, dummyT) |
|
62 |
| encode_type (TVar (a, _)) = Var (a, dummyT); |
|
63 |
||
64 |
||
65 |
(* variables -- ordered left-to-right, preferring right *) |
|
66 |
||
18210 | 67 |
fun vars_of tm = |
68 |
Term.fold_aterms (fn (t as Var _) => cons t | _ => I) tm [] |
|
69 |
|> Library.distinct |
|
70 |
|> rev; |
|
11730 | 71 |
|
18210 | 72 |
local |
11730 | 73 |
|
74 |
val mk_var = encode_type o #2 o Term.dest_Var; |
|
75 |
||
18210 | 76 |
fun concl_var which thm = mk_var (which (vars_of (Thm.concl_of thm))) handle Empty => |
77 |
raise THM ("No variables in conclusion of rule", 0, [thm]); |
|
78 |
||
11730 | 79 |
in |
80 |
||
18210 | 81 |
fun left_var_prem thm = mk_var (hd (vars_of (hd (Thm.prems_of thm)))) handle Empty => |
82 |
raise THM ("No variables in major premise of rule", 0, [thm]); |
|
11730 | 83 |
|
18210 | 84 |
val left_var_concl = concl_var hd; |
85 |
val right_var_concl = concl_var List.last; |
|
11730 | 86 |
|
87 |
end; |
|
88 |
||
89 |
||
90 |
||
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
91 |
(** global and local induct data **) |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
92 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
93 |
(* rules *) |
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 |
type rules = (string * thm) NetRules.T; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
96 |
|
12381 | 97 |
val init_rules = |
13105
3d1e7a199bdc
use eq_thm_prop instead of slightly inadequate eq_thm;
wenzelm
parents:
12381
diff
changeset
|
98 |
NetRules.init (fn ((s1: string, th1), (s2, th2)) => s1 = s2 andalso |
3d1e7a199bdc
use eq_thm_prop instead of slightly inadequate eq_thm;
wenzelm
parents:
12381
diff
changeset
|
99 |
Drule.eq_thm_prop (th1, th2)); |
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
100 |
|
17184 | 101 |
fun lookup_rule (rs: rules) = AList.lookup (op =) (NetRules.rules rs); |
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
102 |
|
18210 | 103 |
fun pretty_rules prt kind x rs = |
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
104 |
let val thms = map snd (NetRules.rules rs) |
18210 | 105 |
in Pretty.big_list kind (map (prt x) thms) end; |
12055 | 106 |
|
18210 | 107 |
fun print_all_rules prt x ((casesT, casesS), (inductT, inductS), (coinductT, coinductS)) = |
108 |
[pretty_rules prt "coinduct type:" x coinductT, |
|
109 |
pretty_rules prt "coinduct set:" x coinductS, |
|
110 |
pretty_rules prt "induct type:" x inductT, |
|
111 |
pretty_rules prt "induct set:" x inductS, |
|
112 |
pretty_rules prt "cases type:" x casesT, |
|
113 |
pretty_rules prt "cases set:" x casesS] |
|
114 |
|> Pretty.chunks |> Pretty.writeln; |
|
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
115 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
116 |
|
18210 | 117 |
(* theory data *) |
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 |
structure GlobalInductArgs = |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
120 |
struct |
18210 | 121 |
val name = "Isar/induct"; |
122 |
type T = (rules * rules) * (rules * rules) * (rules * rules); |
|
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
123 |
|
11730 | 124 |
val empty = |
18210 | 125 |
((init_rules (left_var_prem o #2), init_rules (Thm.major_prem_of o #2)), |
126 |
(init_rules (right_var_concl o #2), init_rules (Thm.major_prem_of o #2)), |
|
127 |
(init_rules (left_var_concl o #2), init_rules (Thm.concl_of o #2))); |
|
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
128 |
val copy = I; |
16424 | 129 |
val extend = I; |
18210 | 130 |
fun merge _ (((casesT1, casesS1), (inductT1, inductS1), (coinductT1, coinductS1)), |
131 |
((casesT2, casesS2), (inductT2, inductS2), (coinductT2, coinductS2))) = |
|
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
132 |
((NetRules.merge (casesT1, casesT2), NetRules.merge (casesS1, casesS2)), |
18210 | 133 |
(NetRules.merge (inductT1, inductT2), NetRules.merge (inductS1, inductS2)), |
134 |
(NetRules.merge (coinductT1, coinductT2), NetRules.merge (coinductS1, coinductS2))); |
|
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
135 |
|
12102 | 136 |
fun print x = print_all_rules Display.pretty_thm_sg x; |
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
137 |
|
18210 | 138 |
fun dest ((casesT, casesS), (inductT, inductS), (coinductT, coinductS)) = |
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
139 |
{type_cases = NetRules.rules casesT, |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
140 |
set_cases = NetRules.rules casesS, |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
141 |
type_induct = NetRules.rules inductT, |
18210 | 142 |
set_induct = NetRules.rules inductS, |
143 |
type_coinduct = NetRules.rules coinductT, |
|
144 |
set_coinduct = NetRules.rules coinductS}; |
|
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
145 |
end; |
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 |
structure GlobalInduct = TheoryDataFun(GlobalInductArgs); |
15801 | 148 |
val _ = Context.add_setup [GlobalInduct.init]; |
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
149 |
val print_global_rules = GlobalInduct.print; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
150 |
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
|
151 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
152 |
|
18210 | 153 |
(* proof data *) |
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
154 |
|
16424 | 155 |
structure LocalInduct = ProofDataFun |
156 |
(struct |
|
18210 | 157 |
val name = "Isar/induct"; |
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
158 |
type T = GlobalInductArgs.T; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
159 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
160 |
fun init thy = GlobalInduct.get thy; |
12102 | 161 |
fun print x = print_all_rules ProofContext.pretty_thm x; |
16424 | 162 |
end); |
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
163 |
|
15801 | 164 |
val _ = Context.add_setup [LocalInduct.init]; |
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
165 |
val print_local_rules = LocalInduct.print; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
166 |
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
|
167 |
|
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 |
(* access rules *) |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
170 |
|
11730 | 171 |
val lookup_casesT = lookup_rule o #1 o #1 o LocalInduct.get; |
172 |
val lookup_casesS = lookup_rule o #2 o #1 o LocalInduct.get; |
|
173 |
val lookup_inductT = lookup_rule o #1 o #2 o LocalInduct.get; |
|
174 |
val lookup_inductS = lookup_rule o #2 o #2 o LocalInduct.get; |
|
18210 | 175 |
val lookup_coinductT = lookup_rule o #1 o #3 o LocalInduct.get; |
176 |
val lookup_coinductS = lookup_rule o #2 o #3 o LocalInduct.get; |
|
11730 | 177 |
|
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
178 |
|
11730 | 179 |
fun find_rules which how ctxt x = |
11784 | 180 |
map snd (NetRules.retrieve (which (LocalInduct.get ctxt)) (how x)); |
11730 | 181 |
|
182 |
val find_casesT = find_rules (#1 o #1) encode_type; |
|
18226 | 183 |
val find_casesS = find_rules (#2 o #1) I; |
11730 | 184 |
val find_inductT = find_rules (#1 o #2) encode_type; |
18226 | 185 |
val find_inductS = find_rules (#2 o #2) I; |
18210 | 186 |
val find_coinductT = find_rules (#1 o #3) encode_type; |
18226 | 187 |
val find_coinductS = find_rules (#2 o #3) I; |
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
188 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
189 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
190 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
191 |
(** attributes **) |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
192 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
193 |
local |
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 |
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
|
196 |
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
|
197 |
|
18210 | 198 |
fun map1 f (x, y, z) = (f x, y, z); |
199 |
fun map2 f (x, y, z) = (x, f y, z); |
|
200 |
fun map3 f (x, y, z) = (x, y, f z); |
|
201 |
||
202 |
fun add_casesT rule x = map1 (apfst (NetRules.insert rule)) x; |
|
203 |
fun add_casesS rule x = map1 (apsnd (NetRules.insert rule)) x; |
|
204 |
fun add_inductT rule x = map2 (apfst (NetRules.insert rule)) x; |
|
205 |
fun add_inductS rule x = map2 (apsnd (NetRules.insert rule)) x; |
|
206 |
fun add_coinductT rule x = map3 (apfst (NetRules.insert rule)) x; |
|
207 |
fun add_coinductS rule x = map3 (apsnd (NetRules.insert rule)) x; |
|
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
208 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
209 |
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
|
210 |
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
|
211 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
212 |
in |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
213 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
214 |
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
|
215 |
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
|
216 |
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
|
217 |
val induct_set_global = mk_att GlobalInduct.map add_inductS consumes1; |
18210 | 218 |
val coinduct_type_global = mk_att GlobalInduct.map add_coinductT consumes0; |
219 |
val coinduct_set_global = mk_att GlobalInduct.map add_coinductS consumes1; |
|
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
220 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
221 |
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
|
222 |
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
|
223 |
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
|
224 |
val induct_set_local = mk_att LocalInduct.map add_inductS consumes1; |
18210 | 225 |
val coinduct_type_local = mk_att LocalInduct.map add_coinductT consumes0; |
226 |
val coinduct_set_local = mk_att LocalInduct.map add_coinductS consumes1; |
|
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
227 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
228 |
end; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
229 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
230 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
231 |
(** concrete syntax **) |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
232 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
233 |
val casesN = "cases"; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
234 |
val inductN = "induct"; |
18210 | 235 |
val coinductN = "coinduct"; |
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
236 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
237 |
val typeN = "type"; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
238 |
val setN = "set"; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
239 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
240 |
local |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
241 |
|
15703 | 242 |
fun spec k arg = |
243 |
Scan.lift (Args.$$$ k -- Args.colon) |-- arg || |
|
244 |
Scan.lift (Args.$$$ k) >> K ""; |
|
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
245 |
|
15703 | 246 |
fun attrib tyname const add_type add_set = |
247 |
Attrib.syntax (spec typeN tyname >> add_type || spec setN const >> add_set); |
|
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
248 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
249 |
in |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
250 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
251 |
val cases_attr = |
15703 | 252 |
(attrib Args.global_tyname Args.global_const cases_type_global cases_set_global, |
253 |
attrib Args.local_tyname Args.local_const cases_type_local cases_set_local); |
|
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
254 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
255 |
val induct_attr = |
15703 | 256 |
(attrib Args.global_tyname Args.global_const induct_type_global induct_set_global, |
257 |
attrib Args.local_tyname Args.local_const induct_type_local induct_set_local); |
|
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
258 |
|
18210 | 259 |
val coinduct_attr = |
260 |
(attrib Args.global_tyname Args.global_const coinduct_type_global coinduct_set_global, |
|
261 |
attrib Args.local_tyname Args.local_const coinduct_type_local coinduct_set_local); |
|
262 |
||
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
263 |
end; |
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
264 |
|
15801 | 265 |
val _ = Context.add_setup |
266 |
[Attrib.add_attributes |
|
267 |
[(casesN, cases_attr, "declaration of cases rule for type or set"), |
|
18210 | 268 |
(inductN, induct_attr, "declaration of induction rule for type or set"), |
269 |
(coinductN, coinduct_attr, "declaration of coinduction rule for type or set")]]; |
|
11656
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
270 |
|
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff
changeset
|
271 |
end; |