| author | blanchet | 
| Thu, 18 Sep 2014 16:47:40 +0200 | |
| changeset 58371 | 7f30ec82fe40 | 
| parent 58189 | 9d714be4f028 | 
| child 58659 | 6c9821c32dd5 | 
| permissions | -rw-r--r-- | 
| 56942 | 1  | 
(* Title: HOL/Library/bnf_axiomatization.ML  | 
2  | 
Author: Dmitriy Traytel, TU Muenchen  | 
|
3  | 
Copyright 2013  | 
|
4  | 
||
5  | 
Axiomatic declaration of bounded natural functors.  | 
|
6  | 
*)  | 
|
7  | 
||
8  | 
signature BNF_AXIOMATIZATION =  | 
|
9  | 
sig  | 
|
| 
58189
 
9d714be4f028
added 'plugins' option to control which hooks are enabled
 
blanchet 
parents: 
58188 
diff
changeset
 | 
10  | 
val bnf_axiomatization: (string -> bool) -> (binding option * (typ * sort)) list -> binding ->  | 
| 
 
9d714be4f028
added 'plugins' option to control which hooks are enabled
 
blanchet 
parents: 
58188 
diff
changeset
 | 
11  | 
mixfix -> binding -> binding -> typ list -> local_theory -> BNF_Def.bnf * local_theory  | 
| 56942 | 12  | 
end  | 
13  | 
||
14  | 
structure BNF_Axiomatization : BNF_AXIOMATIZATION =  | 
|
15  | 
struct  | 
|
16  | 
||
17  | 
open BNF_Util  | 
|
18  | 
open BNF_Def  | 
|
19  | 
||
| 
58189
 
9d714be4f028
added 'plugins' option to control which hooks are enabled
 
blanchet 
parents: 
58188 
diff
changeset
 | 
20  | 
fun prepare_decl prepare_constraint prepare_typ plugins raw_vars b mx user_mapb user_relb user_witTs  | 
| 
 
9d714be4f028
added 'plugins' option to control which hooks are enabled
 
blanchet 
parents: 
58188 
diff
changeset
 | 
21  | 
lthy =  | 
| 56942 | 22  | 
let  | 
23  | 
fun prepare_type_arg (set_opt, (ty, c)) =  | 
|
24  | 
let val s = fst (dest_TFree (prepare_typ lthy ty)) in  | 
|
25  | 
(set_opt, (s, prepare_constraint lthy c))  | 
|
26  | 
end;  | 
|
27  | 
val ((user_setbs, vars), raw_vars') =  | 
|
28  | 
map prepare_type_arg raw_vars  | 
|
29  | 
|> `split_list  | 
|
30  | 
|>> apfst (map_filter I);  | 
|
31  | 
val deads = map_filter (fn (NONE, x) => SOME x | _ => NONE) raw_vars';  | 
|
32  | 
||
33  | 
fun mk_b name user_b =  | 
|
34  | 
(if Binding.is_empty user_b then Binding.prefix_name (name ^ "_") b else user_b)  | 
|
35  | 
|> Binding.qualify false (Binding.name_of b);  | 
|
36  | 
val (Tname, lthy) = Typedecl.basic_typedecl (b, length vars, mx) lthy;  | 
|
37  | 
val (bd_type_Tname, lthy) =  | 
|
38  | 
Typedecl.basic_typedecl (mk_b "bd_type" Binding.empty, length deads, NoSyn) lthy;  | 
|
39  | 
val T = Type (Tname, map TFree vars);  | 
|
40  | 
val bd_type_T = Type (bd_type_Tname, map TFree deads);  | 
|
41  | 
val lives = map TFree (filter_out (member (op =) deads) vars);  | 
|
42  | 
val live = length lives;  | 
|
43  | 
val _ = "Trying to declare a BNF with no live variables" |> null lives ? error;  | 
|
44  | 
val (lives', _) = BNF_Util.mk_TFrees (length lives)  | 
|
45  | 
(fold Variable.declare_typ (map TFree vars) lthy);  | 
|
46  | 
val T' = Term.typ_subst_atomic (lives ~~ lives') T;  | 
|
47  | 
val mapT = map2 (curry op -->) lives lives' ---> T --> T';  | 
|
48  | 
val setTs = map (fn U => T --> HOLogic.mk_setT U) lives;  | 
|
49  | 
val bdT = BNF_Util.mk_relT (bd_type_T, bd_type_T);  | 
|
| 
58189
 
9d714be4f028
added 'plugins' option to control which hooks are enabled
 
blanchet 
parents: 
58188 
diff
changeset
 | 
50  | 
val mapb = mk_b mapN user_mapb;  | 
| 56942 | 51  | 
val bdb = mk_b "bd" Binding.empty;  | 
| 
58189
 
9d714be4f028
added 'plugins' option to control which hooks are enabled
 
blanchet 
parents: 
58188 
diff
changeset
 | 
52  | 
val setbs = map2 (fn b => fn i => mk_b (mk_setN i) b) user_setbs  | 
| 56942 | 53  | 
(if live = 1 then [0] else 1 upto live);  | 
54  | 
||
55  | 
val witTs = map (prepare_typ lthy) user_witTs;  | 
|
56  | 
val nwits = length witTs;  | 
|
| 
58189
 
9d714be4f028
added 'plugins' option to control which hooks are enabled
 
blanchet 
parents: 
58188 
diff
changeset
 | 
57  | 
val witbs = map (fn i => mk_b (mk_witN i) Binding.empty)  | 
| 56942 | 58  | 
(if nwits = 1 then [0] else 1 upto nwits);  | 
59  | 
||
60  | 
val lthy = Local_Theory.background_theory  | 
|
61  | 
(Sign.add_consts ((mapb, mapT, NoSyn) :: (bdb, bdT, NoSyn) ::  | 
|
62  | 
map2 (fn b => fn T => (b, T, NoSyn)) setbs setTs @  | 
|
63  | 
map2 (fn b => fn T => (b, T, NoSyn)) witbs witTs))  | 
|
64  | 
lthy;  | 
|
65  | 
val Fmap = Const (Local_Theory.full_name lthy mapb, mapT);  | 
|
66  | 
val Fsets = map2 (fn setb => fn setT =>  | 
|
67  | 
Const (Local_Theory.full_name lthy setb, setT)) setbs setTs;  | 
|
68  | 
val Fbd = Const (Local_Theory.full_name lthy bdb, bdT);  | 
|
69  | 
val Fwits = map2 (fn witb => fn witT =>  | 
|
70  | 
Const (Local_Theory.full_name lthy witb, witT)) witbs witTs;  | 
|
71  | 
val (key, goals, (triv_tac_opt, wit_goalss), after_qed, lthy, _) =  | 
|
72  | 
prepare_def Do_Inline (user_policy Note_Some) false I (K I) (K I) (SOME (map TFree deads))  | 
|
73  | 
user_mapb user_relb user_setbs ((((((Binding.empty, T), Fmap), Fsets), Fbd), Fwits), NONE)  | 
|
74  | 
lthy;  | 
|
75  | 
||
76  | 
fun mk_wits_tac ctxt set_maps = TRYALL Goal.conjunction_tac THEN the triv_tac_opt ctxt set_maps;  | 
|
77  | 
val wit_goals = map Logic.mk_conjunction_balanced wit_goalss;  | 
|
78  | 
val all_goalss = map single goals @ (if nwits > 0 then wit_goalss else []);  | 
|
79  | 
||
80  | 
val (((_, [raw_thms])), (lthy_old, lthy)) = Local_Theory.background_theory_result  | 
|
81  | 
(Specification.axiomatization [] [((mk_b "axioms" Binding.empty, []), flat all_goalss)]) lthy  | 
|
82  | 
||> `Local_Theory.restore;  | 
|
83  | 
||
84  | 
fun mk_wit_thms set_maps =  | 
|
85  | 
Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced wit_goals)  | 
|
86  | 
        (fn {context = ctxt, prems = _} => mk_wits_tac ctxt set_maps)
 | 
|
87  | 
|> Conjunction.elim_balanced (length wit_goals)  | 
|
88  | 
|> map2 (Conjunction.elim_balanced o length) wit_goalss  | 
|
89  | 
|> map (map (Thm.close_derivation o Thm.forall_elim_vars 0));  | 
|
90  | 
val phi = Proof_Context.export_morphism lthy_old lthy;  | 
|
91  | 
val thms = unflat all_goalss (Morphism.fact phi raw_thms);  | 
|
92  | 
||
93  | 
val (bnf, lthy') = after_qed mk_wit_thms thms lthy  | 
|
94  | 
in  | 
|
| 
58189
 
9d714be4f028
added 'plugins' option to control which hooks are enabled
 
blanchet 
parents: 
58188 
diff
changeset
 | 
95  | 
(bnf, register_bnf plugins key bnf lthy')  | 
| 56942 | 96  | 
end;  | 
97  | 
||
98  | 
val bnf_axiomatization = prepare_decl (K I) (K I);  | 
|
99  | 
||
100  | 
fun read_constraint _ NONE = @{sort type}
 | 
|
101  | 
| read_constraint ctxt (SOME s) = Syntax.read_sort ctxt s;  | 
|
102  | 
||
103  | 
val bnf_axiomatization_cmd = prepare_decl read_constraint Syntax.read_typ;  | 
|
104  | 
||
105  | 
val parse_witTs =  | 
|
106  | 
  @{keyword "["} |-- (Parse.short_ident --| @{keyword ":"} -- Scan.repeat Parse.typ
 | 
|
107  | 
    >> (fn ("wits", Ts) => Ts
 | 
|
108  | 
         | (s, _) => error ("Unknown label " ^ quote s ^ " (expected \"wits\")"))) --|
 | 
|
109  | 
  @{keyword "]"} || Scan.succeed [];
 | 
|
110  | 
||
| 
58189
 
9d714be4f028
added 'plugins' option to control which hooks are enabled
 
blanchet 
parents: 
58188 
diff
changeset
 | 
111  | 
val parse_bnf_axiomatization_options =  | 
| 
 
9d714be4f028
added 'plugins' option to control which hooks are enabled
 
blanchet 
parents: 
58188 
diff
changeset
 | 
112  | 
  Scan.optional (@{keyword "("} |-- parse_plugins --| @{keyword ")"}) (K true);
 | 
| 
 
9d714be4f028
added 'plugins' option to control which hooks are enabled
 
blanchet 
parents: 
58188 
diff
changeset
 | 
113  | 
|
| 56942 | 114  | 
val parse_bnf_axiomatization =  | 
| 
58189
 
9d714be4f028
added 'plugins' option to control which hooks are enabled
 
blanchet 
parents: 
58188 
diff
changeset
 | 
115  | 
parse_bnf_axiomatization_options -- parse_type_args_named_constrained -- Parse.binding --  | 
| 
 
9d714be4f028
added 'plugins' option to control which hooks are enabled
 
blanchet 
parents: 
58188 
diff
changeset
 | 
116  | 
parse_witTs -- Parse.opt_mixfix -- parse_map_rel_bindings;  | 
| 56942 | 117  | 
|
118  | 
val _ =  | 
|
119  | 
  Outer_Syntax.local_theory @{command_spec "bnf_axiomatization"} "bnf declaration"
 | 
|
120  | 
(parse_bnf_axiomatization >>  | 
|
| 
58189
 
9d714be4f028
added 'plugins' option to control which hooks are enabled
 
blanchet 
parents: 
58188 
diff
changeset
 | 
121  | 
(fn (((((plugins, bsTs), b), witTs), mx), (mapb, relb)) =>  | 
| 
 
9d714be4f028
added 'plugins' option to control which hooks are enabled
 
blanchet 
parents: 
58188 
diff
changeset
 | 
122  | 
bnf_axiomatization_cmd plugins bsTs b mx mapb relb witTs #> snd));  | 
| 56942 | 123  | 
|
124  | 
end;  |