1 (* Title: nominal_inductive.ML |
1 (* Title: nominal_inductive.ML |
2 Author: Christian Urban |
2 Author: Christian Urban |
|
3 Author: Tjark Weber |
3 |
4 |
4 Infrastructure for proving strong induction theorems |
5 Infrastructure for proving strong induction theorems |
5 for inductive predicates involving nominal datatypes. |
6 for inductive predicates involving nominal datatypes. |
6 |
7 |
7 Code based on an earlier version by Stefan Berghofer. |
8 Code based on an earlier version by Stefan Berghofer. |
8 *) |
9 *) |
9 |
10 |
10 |
|
11 signature NOMINAL_INDUCTIVE = |
11 signature NOMINAL_INDUCTIVE = |
12 sig |
12 sig |
13 val prove_strong_inductive: string list -> string list -> term list list -> thm -> thm list -> |
13 val prove_strong_inductive: string list -> string list -> term list list -> thm -> thm list -> |
14 Proof.context -> Proof.state |
14 Proof.context -> Proof.state |
15 |
|
16 val prove_strong_inductive_cmd: xstring * (string * string list) list -> Proof.context -> Proof.state |
15 val prove_strong_inductive_cmd: xstring * (string * string list) list -> Proof.context -> Proof.state |
17 end |
16 end |
18 |
17 |
19 structure Nominal_Inductive : NOMINAL_INDUCTIVE = |
18 structure Nominal_Inductive : NOMINAL_INDUCTIVE = |
20 struct |
19 struct |
21 |
20 |
22 |
21 fun mk_cplus p q = |
23 fun mk_cplus p q = Thm.apply (Thm.apply @{cterm "plus :: perm => perm => perm"} p) q |
22 Thm.apply (Thm.apply @{cterm "plus :: perm => perm => perm"} p) q |
24 |
23 |
25 fun mk_cminus p = Thm.apply @{cterm "uminus :: perm => perm"} p |
24 fun mk_cminus p = |
26 |
25 Thm.apply @{cterm "uminus :: perm => perm"} p |
27 fun minus_permute_intro_tac p = |
26 |
|
27 fun minus_permute_intro_tac p = |
28 rtac (Drule.instantiate' [] [SOME (mk_cminus p)] @{thm permute_boolE}) |
28 rtac (Drule.instantiate' [] [SOME (mk_cminus p)] @{thm permute_boolE}) |
29 |
29 |
30 fun minus_permute_elim p thm = |
30 fun minus_permute_elim p thm = |
31 thm RS (Drule.instantiate' [] [NONE, SOME (mk_cminus p)] @{thm permute_boolI}) |
31 thm RS (Drule.instantiate' [] [NONE, SOME (mk_cminus p)] @{thm permute_boolI}) |
32 |
32 |
33 (* fixme: move to nominal_library *) |
33 (* fixme: move to nominal_library *) |
34 fun real_head_of (@{term Trueprop} $ t) = real_head_of t |
34 fun real_head_of (@{term Trueprop} $ t) = real_head_of t |
35 | real_head_of (Const ("==>", _) $ _ $ t) = real_head_of t |
35 | real_head_of (Const ("==>", _) $ _ $ t) = real_head_of t |
36 | real_head_of (Const (@{const_name all}, _) $ Abs (_, _, t)) = real_head_of t |
36 | real_head_of (Const (@{const_name all}, _) $ Abs (_, _, t)) = real_head_of t |
37 | real_head_of (Const (@{const_name All}, _) $ Abs (_, _, t)) = real_head_of t |
37 | real_head_of (Const (@{const_name All}, _) $ Abs (_, _, t)) = real_head_of t |
38 | real_head_of (Const ("HOL.induct_forall", _) $ Abs (_, _, t)) = real_head_of t |
38 | real_head_of (Const ("HOL.induct_forall", _) $ Abs (_, _, t)) = real_head_of t |
39 | real_head_of t = head_of t |
39 | real_head_of t = head_of t |
40 |
40 |
41 |
41 |
42 fun mk_vc_compat (avoid, avoid_trm) prems concl_args params = |
42 fun mk_vc_compat (avoid, avoid_trm) prems concl_args params = |
43 let |
43 if null avoid then |
44 val vc_goal = concl_args |
44 [] |
45 |> HOLogic.mk_tuple |
45 else |
46 |> mk_fresh_star avoid_trm |
46 let |
47 |> HOLogic.mk_Trueprop |
47 val vc_goal = concl_args |
48 |> (curry Logic.list_implies) prems |
48 |> HOLogic.mk_tuple |
49 |> fold_rev (Logic.all o Free) params |
49 |> mk_fresh_star avoid_trm |
50 val finite_goal = avoid_trm |
50 |> HOLogic.mk_Trueprop |
51 |> mk_finite |
51 |> (curry Logic.list_implies) prems |
52 |> HOLogic.mk_Trueprop |
52 |> fold_rev (Logic.all o Free) params |
53 |> (curry Logic.list_implies) prems |
53 val finite_goal = avoid_trm |
54 |> fold_rev (Logic.all o Free) params |
54 |> mk_finite |
55 in |
55 |> HOLogic.mk_Trueprop |
56 if null avoid then [] else [vc_goal, finite_goal] |
56 |> (curry Logic.list_implies) prems |
57 end |
57 |> fold_rev (Logic.all o Free) params |
|
58 in |
|
59 [vc_goal, finite_goal] |
|
60 end |
58 |
61 |
59 (* fixme: move to nominal_library *) |
62 (* fixme: move to nominal_library *) |
60 fun map_term prop f trm = |
63 fun map_term prop f trm = |
61 if prop trm |
64 if prop trm |
62 then f trm |
65 then f trm |
75 list_comb (Free (P_name, (c_ty :: ty_args) ---> bool), c :: args') |
78 list_comb (Free (P_name, (c_ty :: ty_args) ---> bool), c :: args') |
76 |> (fn t => HOLogic.all_const c_ty $ lambda c t ) |
79 |> (fn t => HOLogic.all_const c_ty $ lambda c t ) |
77 |> (fn t => HOLogic.all_const @{typ perm} $ lambda p t) |
80 |> (fn t => HOLogic.all_const @{typ perm} $ lambda p t) |
78 end |
81 end |
79 |
82 |
80 fun induct_forall_const T = Const ("HOL.induct_forall", (T --> @{typ bool}) --> @{typ bool}) |
83 fun induct_forall_const T = |
81 fun mk_induct_forall (a, T) t = induct_forall_const T $ Abs (a, T, t) |
84 Const ("HOL.induct_forall", (T --> @{typ bool}) --> @{typ bool}) |
|
85 |
|
86 fun mk_induct_forall (a, T) t = |
|
87 induct_forall_const T $ Abs (a, T, t) |
82 |
88 |
83 fun add_c_prop qnt Ps (c, c_name, c_ty) trm = |
89 fun add_c_prop qnt Ps (c, c_name, c_ty) trm = |
84 let |
90 let |
85 fun add t = |
91 fun add t = |
86 let |
92 let |
87 val (P, args) = strip_comb t |
93 val (P, args) = strip_comb t |
88 val (P_name, P_ty) = dest_Free P |
94 val (P_name, P_ty) = dest_Free P |
89 val (ty_args, bool) = strip_type P_ty |
95 val (ty_args, bool) = strip_type P_ty |
90 val args' = args |
96 val args' = args |
91 |> qnt ? map (incr_boundvars 1) |
97 |> qnt ? map (incr_boundvars 1) |
92 in |
98 in |
93 list_comb (Free (P_name, (c_ty :: ty_args) ---> bool), c :: args') |
99 list_comb (Free (P_name, (c_ty :: ty_args) ---> bool), c :: args') |
94 |> qnt ? mk_induct_forall (c_name, c_ty) |
100 |> qnt ? mk_induct_forall (c_name, c_ty) |
95 end |
101 end |
96 in |
102 in |
97 map_term (member (op =) Ps o head_of) add trm |
103 map_term (member (op =) Ps o head_of) add trm |
98 end |
104 end |
99 |
105 |
100 fun prep_prem Ps c_name c_ty (avoid, avoid_trm) (params, prems, concl) = |
106 fun prep_prem Ps c_name c_ty (avoid, avoid_trm) (params, prems, concl) = |
101 let |
107 let |
102 val prems' = prems |
108 val prems' = prems |
103 |> map (incr_boundvars 1) |
109 |> map (incr_boundvars 1) |
104 |> map (add_c_prop true Ps (Bound 0, c_name, c_ty)) |
110 |> map (add_c_prop true Ps (Bound 0, c_name, c_ty)) |
105 |
111 |
106 val avoid_trm' = avoid_trm |
112 val avoid_trm' = avoid_trm |
107 |> fold_rev absfree (params @ [(c_name, c_ty)]) |
113 |> fold_rev absfree (params @ [(c_name, c_ty)]) |
108 |> strip_abs_body |
114 |> strip_abs_body |
109 |> (fn t => mk_fresh_star_ty c_ty t (Bound 0)) |
115 |> (fn t => mk_fresh_star_ty c_ty t (Bound 0)) |
110 |> HOLogic.mk_Trueprop |
116 |> HOLogic.mk_Trueprop |
111 |
117 |
112 val prems'' = |
118 val prems'' = |
113 if null avoid |
119 if null avoid |
114 then prems' |
120 then prems' |
115 else avoid_trm' :: prems' |
121 else avoid_trm' :: prems' |
116 |
122 |
117 val concl' = concl |
123 val concl' = concl |
118 |> incr_boundvars 1 |
124 |> incr_boundvars 1 |
119 |> add_c_prop false Ps (Bound 0, c_name, c_ty) |
125 |> add_c_prop false Ps (Bound 0, c_name, c_ty) |
120 in |
126 in |
121 mk_full_horn (params @ [(c_name, c_ty)]) prems'' concl' |
127 mk_full_horn (params @ [(c_name, c_ty)]) prems'' concl' |
122 end |
128 end |
123 |
129 |
124 (* fixme: move to nominal_library *) |
130 (* fixme: move to nominal_library *) |
127 | same_name (Const (a1, _), Const (a2, _)) = (a1 = a2) |
133 | same_name (Const (a1, _), Const (a2, _)) = (a1 = a2) |
128 | same_name _ = false |
134 | same_name _ = false |
129 |
135 |
130 (* fixme: move to nominal_library *) |
136 (* fixme: move to nominal_library *) |
131 fun map7 _ [] [] [] [] [] [] [] = [] |
137 fun map7 _ [] [] [] [] [] [] [] = [] |
132 | map7 f (x :: xs) (y :: ys) (z :: zs) (u :: us) (v :: vs) (r :: rs) (s :: ss) = |
138 | map7 f (x :: xs) (y :: ys) (z :: zs) (u :: us) (v :: vs) (r :: rs) (s :: ss) = |
133 f x y z u v r s :: map7 f xs ys zs us vs rs ss |
139 f x y z u v r s :: map7 f xs ys zs us vs rs ss |
134 |
140 |
135 (* local abbreviations *) |
141 (* local abbreviations *) |
136 |
142 |
137 local |
143 local |
138 open Nominal_Permeq |
144 open Nominal_Permeq |
139 in |
145 in |
140 (* by default eqvt_strict_config contains unwanted @{thm permute_pure} *) |
146 |
141 |
147 (* by default eqvt_strict_config contains unwanted @{thm permute_pure} *) |
142 val eqvt_sconfig = eqvt_strict_config addpres @{thms permute_minus_cancel} |
148 |
143 |
149 val eqvt_sconfig = eqvt_strict_config addpres @{thms permute_minus_cancel} |
144 fun eqvt_stac ctxt = eqvt_tac ctxt eqvt_sconfig |
150 |
145 fun eqvt_srule ctxt = eqvt_rule ctxt eqvt_sconfig |
151 fun eqvt_stac ctxt = eqvt_tac ctxt eqvt_sconfig |
|
152 fun eqvt_srule ctxt = eqvt_rule ctxt eqvt_sconfig |
146 |
153 |
147 end |
154 end |
148 |
155 |
149 val all_elims = |
156 val all_elims = |
150 let |
157 let |
151 fun spec' ct = Drule.instantiate' [SOME (ctyp_of_term ct)] [NONE, SOME ct] @{thm spec} |
158 fun spec' ct = |
|
159 Drule.instantiate' [SOME (ctyp_of_term ct)] [NONE, SOME ct] @{thm spec} |
152 in |
160 in |
153 fold (fn ct => fn th => th RS spec' ct) |
161 fold (fn ct => fn th => th RS spec' ct) |
154 end |
162 end |
155 |
163 |
156 fun helper_tac flag prm p ctxt = |
164 fun helper_tac flag prm p ctxt = |
381 Proof.theorem NONE (after_qed ctxt) ((map o map) (rpair []) vc_compat_goals) ctxt'' |
391 Proof.theorem NONE (after_qed ctxt) ((map o map) (rpair []) vc_compat_goals) ctxt'' |
382 end |
392 end |
383 |
393 |
384 fun prove_strong_inductive_cmd (pred_name, avoids) ctxt = |
394 fun prove_strong_inductive_cmd (pred_name, avoids) ctxt = |
385 let |
395 let |
386 val thy = Proof_Context.theory_of ctxt; |
|
387 val ({names, ...}, {raw_induct, intrs, ...}) = |
396 val ({names, ...}, {raw_induct, intrs, ...}) = |
388 Inductive.the_inductive ctxt (Sign.intern_const thy pred_name); |
397 Inductive.the_inductive ctxt (long_name ctxt pred_name) |
389 |
398 |
390 val rule_names = |
399 val rule_names = hd names |
391 hd names |
|
392 |> the o Induct.lookup_inductP ctxt |
400 |> the o Induct.lookup_inductP ctxt |
393 |> fst o Rule_Cases.get |
401 |> fst o Rule_Cases.get |
394 |> map (fst o fst) |
402 |> map (fst o fst) |
395 |
403 |
396 val _ = (case duplicates (op = o pairself fst) avoids of |
404 val case_names = map fst avoids |
|
405 val _ = case duplicates (op =) case_names of |
397 [] => () |
406 [] => () |
398 | xs => error ("Duplicate case names: " ^ commas_quote (map fst xs))) |
407 | xs => error ("Duplicate case names: " ^ commas_quote xs) |
399 |
408 val _ = case subtract (op =) rule_names case_names of |
400 val _ = (case subtract (op =) rule_names (map fst avoids) of |
|
401 [] => () |
409 [] => () |
402 | xs => error ("No such case(s) in inductive definition: " ^ commas_quote xs)) |
410 | xs => error ("No such case(s) in inductive definition: " ^ commas_quote xs) |
403 |
411 |
404 val avoids_ordered = order_default (op =) [] rule_names avoids |
412 val avoids_ordered = order_default (op =) [] rule_names avoids |
405 |
413 |
406 fun read_avoids avoid_trms intr = |
414 fun read_avoids avoid_trms intr = |
407 let |
415 let |
408 (* fixme hack *) |
416 (* fixme hack *) |
409 val (((_, ctrms), _), ctxt') = Variable.import true [intr] ctxt |
417 val (((_, ctrms), _), ctxt') = Variable.import true [intr] ctxt |
410 val trms = map (term_of o snd) ctrms |
418 val trms = map (term_of o snd) ctrms |
411 val ctxt'' = fold Variable.declare_term trms ctxt' |
419 val ctxt'' = fold Variable.declare_term trms ctxt' |
412 in |
420 in |
413 map (Syntax.read_term ctxt'') avoid_trms |
421 map (Syntax.read_term ctxt'') avoid_trms |
414 end |
422 end |
415 |
423 |
416 val avoid_trms = map2 read_avoids avoids_ordered intrs |
424 val avoid_trms = map2 read_avoids avoids_ordered intrs |
417 in |
425 in |
418 prove_strong_inductive names rule_names avoid_trms raw_induct intrs ctxt |
426 prove_strong_inductive names rule_names avoid_trms raw_induct intrs ctxt |
419 end |
427 end |
420 |
428 |
421 (* outer syntax *) |
429 (* outer syntax *) |
422 local |
430 local |
423 |
431 val single_avoid_parser = |
424 val single_avoid_parser = |
|
425 Parse.name -- (@{keyword ":"} |-- Parse.and_list1 Parse.term) |
432 Parse.name -- (@{keyword ":"} |-- Parse.and_list1 Parse.term) |
426 |
433 |
427 val avoids_parser = |
434 val avoids_parser = |
428 Scan.optional (@{keyword "avoids"} |-- Parse.enum1 "|" single_avoid_parser) [] |
435 Scan.optional (@{keyword "avoids"} |-- Parse.enum1 "|" single_avoid_parser) [] |
429 |
436 |
430 val main_parser = Parse.xname -- avoids_parser |
437 val main_parser = Parse.xname -- avoids_parser |
431 in |
438 in |
432 val _ = |
439 val _ = |