author | Christian Urban <christian dot urban at kcl dot ac dot uk> |
Sat, 19 Mar 2016 21:06:48 +0000 | |
branch | Nominal2-Isabelle2016 |
changeset 3243 | c4f31f1564b7 |
parent 3239 | 67370521c09c |
child 3245 | 017e33849f4d |
permissions | -rw-r--r-- |
1079 | 1 |
(* Title: nominal_atoms/ML |
2 |
Authors: Brian Huffman, Christian Urban |
|
3 |
||
4 |
Command for defining concrete atom types. |
|
5 |
||
6 |
At the moment, only single-sorted atom types |
|
7 |
are supported. |
|
8 |
*) |
|
9 |
||
10 |
signature ATOM_DECL = |
|
11 |
sig |
|
12 |
val add_atom_decl: (binding * (binding option)) -> theory -> theory |
|
13 |
end; |
|
14 |
||
3243
c4f31f1564b7
updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3239
diff
changeset
|
15 |
structure Atom_Decl : ATOM_DECL = |
1079 | 16 |
struct |
17 |
||
2891
304dfe6cc83a
the simplifier can simplify "sort (atom a)" if a is a concrete atom type declared with atom_decl
Christian Urban <urbanc@in.tum.de>
parents:
2568
diff
changeset
|
18 |
val simp_attr = Attrib.internal (K Simplifier.simp_add) |
304dfe6cc83a
the simplifier can simplify "sort (atom a)" if a is a concrete atom type declared with atom_decl
Christian Urban <urbanc@in.tum.de>
parents:
2568
diff
changeset
|
19 |
|
1079 | 20 |
fun atom_decl_set (str : string) : term = |
21 |
let |
|
1962
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
Christian Urban <urbanc@in.tum.de>
parents:
1774
diff
changeset
|
22 |
val a = Free ("a", @{typ atom}); |
1079 | 23 |
val s = Const (@{const_name "Sort"}, @{typ "string => atom_sort list => atom_sort"}) |
24 |
$ HOLogic.mk_string str $ HOLogic.nil_const @{typ "atom_sort"}; |
|
25 |
in |
|
1962
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
Christian Urban <urbanc@in.tum.de>
parents:
1774
diff
changeset
|
26 |
HOLogic.mk_Collect ("a", @{typ atom}, HOLogic.mk_eq (mk_sort_of a, s)) |
1079 | 27 |
end |
28 |
||
29 |
fun add_atom_decl (name : binding, arg : binding option) (thy : theory) = |
|
30 |
let |
|
31 |
val str = Sign.full_name thy name; |
|
32 |
||
33 |
(* typedef *) |
|
34 |
val set = atom_decl_set str; |
|
3243
c4f31f1564b7
updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3239
diff
changeset
|
35 |
fun tac ctxt = resolve_tac ctxt @{thms exists_eq_simple_sort} 1; |
3239
67370521c09c
updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3233
diff
changeset
|
36 |
val ((full_tname, info as ({Rep_name, Abs_name, ...}, {type_definition, ...})), thy) = |
3243
c4f31f1564b7
updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3239
diff
changeset
|
37 |
Typedef.add_typedef_global {overloaded = false} (name, [], NoSyn) set NONE tac thy; |
1079 | 38 |
|
39 |
(* definition of atom and permute *) |
|
1689
8c0eef2b84e7
fixed a problem due to a change in type-def (needs new Isabelle)
Christian Urban <urbanc@in.tum.de>
parents:
1460
diff
changeset
|
40 |
val newT = #abs_type (fst info); |
1962
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
Christian Urban <urbanc@in.tum.de>
parents:
1774
diff
changeset
|
41 |
val RepC = Const (Rep_name, newT --> @{typ atom}); |
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
Christian Urban <urbanc@in.tum.de>
parents:
1774
diff
changeset
|
42 |
val AbsC = Const (Abs_name, @{typ atom} --> newT); |
1079 | 43 |
val a = Free ("a", newT); |
1962
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
Christian Urban <urbanc@in.tum.de>
parents:
1774
diff
changeset
|
44 |
val p = Free ("p", @{typ perm}); |
1079 | 45 |
val atom_eqn = |
46 |
HOLogic.mk_Trueprop (HOLogic.mk_eq (mk_atom a, RepC $ a)); |
|
47 |
val permute_eqn = |
|
48 |
HOLogic.mk_Trueprop (HOLogic.mk_eq |
|
1962
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
Christian Urban <urbanc@in.tum.de>
parents:
1774
diff
changeset
|
49 |
(mk_perm p a, AbsC $ (mk_perm p (RepC $ a)))); |
1079 | 50 |
val atom_def_name = |
51 |
Binding.prefix_name "atom_" (Binding.suffix_name "_def" name); |
|
2891
304dfe6cc83a
the simplifier can simplify "sort (atom a)" if a is a concrete atom type declared with atom_decl
Christian Urban <urbanc@in.tum.de>
parents:
2568
diff
changeset
|
52 |
val sort_thm_name = |
304dfe6cc83a
the simplifier can simplify "sort (atom a)" if a is a concrete atom type declared with atom_decl
Christian Urban <urbanc@in.tum.de>
parents:
2568
diff
changeset
|
53 |
Binding.prefix_name "atom_" (Binding.suffix_name "_sort" name); |
1079 | 54 |
val permute_def_name = |
55 |
Binding.prefix_name "permute_" (Binding.suffix_name "_def" name); |
|
56 |
||
57 |
(* at class instance *) |
|
58 |
val lthy = |
|
2396
f2f611daf480
updated to Isabelle 12th Aug
Christian Urban <urbanc@in.tum.de>
parents:
2168
diff
changeset
|
59 |
Class.instantiation ([full_tname], [], @{sort at}) thy; |
1079 | 60 |
val ((_, (_, permute_ldef)), lthy) = |
61 |
Specification.definition (NONE, ((permute_def_name, []), permute_eqn)) lthy; |
|
62 |
val ((_, (_, atom_ldef)), lthy) = |
|
63 |
Specification.definition (NONE, ((atom_def_name, []), atom_eqn)) lthy; |
|
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:
2891
diff
changeset
|
64 |
val ctxt_thy = Proof_Context.init_global (Proof_Context.theory_of lthy); |
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:
2891
diff
changeset
|
65 |
val permute_def = singleton (Proof_Context.export lthy ctxt_thy) permute_ldef; |
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:
2891
diff
changeset
|
66 |
val atom_def = singleton (Proof_Context.export lthy ctxt_thy) atom_ldef; |
1079 | 67 |
val class_thm = @{thm at_class} OF [type_definition, atom_def, permute_def]; |
2891
304dfe6cc83a
the simplifier can simplify "sort (atom a)" if a is a concrete atom type declared with atom_decl
Christian Urban <urbanc@in.tum.de>
parents:
2568
diff
changeset
|
68 |
val sort_thm = @{thm at_class_sort} OF [type_definition, atom_def] |
1079 | 69 |
val thy = lthy |
2891
304dfe6cc83a
the simplifier can simplify "sort (atom a)" if a is a concrete atom type declared with atom_decl
Christian Urban <urbanc@in.tum.de>
parents:
2568
diff
changeset
|
70 |
|> snd o (Local_Theory.note ((sort_thm_name, [simp_attr]), [sort_thm])) |
3243
c4f31f1564b7
updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3239
diff
changeset
|
71 |
|> Class.prove_instantiation_instance (fn ctxt => resolve_tac ctxt [class_thm] 1) |
1079 | 72 |
|> Local_Theory.exit_global; |
73 |
in |
|
74 |
thy |
|
75 |
end; |
|
76 |
||
77 |
(** outer syntax **) |
|
78 |
val _ = |
|
3239
67370521c09c
updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3233
diff
changeset
|
79 |
Outer_Syntax.command @{command_keyword atom_decl} |
3135
92b9b8d2888d
updated to new Isabelle (20 March)
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
80 |
"declaration of a concrete atom type" |
92b9b8d2888d
updated to new Isabelle (20 March)
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
81 |
((Parse.binding -- Scan.option (Args.parens (Parse.binding))) >> |
3233
9ae285ce814e
updated changes from upstream (AFP)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3202
diff
changeset
|
82 |
(Toplevel.theory o add_atom_decl)) |
1079 | 83 |
|
84 |
end; |