Nominal/nominal_atoms.ML
author Christian Urban <urbanc@in.tum.de>
Tue, 07 Aug 2012 18:54:52 +0100
changeset 3197 25d11b449e92
parent 3193 87d1e815aa59
child 3202 3611bc56c177
permissions -rw-r--r--
definition of an auxiliary graph in nominal-primrec definitions
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
(*  Title:      nominal_atoms/ML
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
    Authors:    Brian Huffman, Christian Urban
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
    Command for defining concrete atom types. 
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
    
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
    At the moment, only single-sorted atom types
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
    are supported. 
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
*)
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
signature ATOM_DECL =
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
sig
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
  val add_atom_decl: (binding * (binding option)) -> theory -> theory
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
end;
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
structure Atom_Decl :> ATOM_DECL =
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
struct
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    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
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
fun atom_decl_set (str : string) : term =
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    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
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
    val s = Const (@{const_name "Sort"}, @{typ "string => atom_sort list => atom_sort"})
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
              $ HOLogic.mk_string str $ HOLogic.nil_const @{typ "atom_sort"};
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    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
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
  end
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
fun add_atom_decl (name : binding, arg : binding option) (thy : theory) =
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
  let
2467
67b3933c3190 got rid of Nominal_Atoms (folded into Nominal2_Base)
Christian Urban <urbanc@in.tum.de>
parents: 2396
diff changeset
    31
    val _ = Theory.requires thy "Nominal2_Base" "nominal logic";
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
    val str = Sign.full_name thy name;
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
    (* typedef *)
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
    val set = atom_decl_set str;
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
    val tac = rtac @{thm exists_eq_simple_sort} 1;
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
    37
    val ((full_tname, info  as ({Rep_name, Abs_name, ...}, {type_definition, ...})), thy) =
1460
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
    38
      Typedef.add_typedef_global false NONE (name, [], NoSyn) set NONE tac thy;
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
    (* 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
    41
    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
    42
    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
    43
    val AbsC = Const (Abs_name, @{typ atom} --> newT);
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
    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
    45
    val p = Free ("p", @{typ perm});
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
    val atom_eqn =
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
      HOLogic.mk_Trueprop (HOLogic.mk_eq (mk_atom a, RepC $ a));
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
    val permute_eqn =
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
      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
    50
        (mk_perm p a, AbsC $ (mk_perm p (RepC $ a))));
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
    val atom_def_name =
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
      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
    53
    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
    54
      Binding.prefix_name "atom_" (Binding.suffix_name "_sort" name);
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
    val permute_def_name =
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
      Binding.prefix_name "permute_" (Binding.suffix_name "_def" name);
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
    (* at class instance *)
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
    val lthy =
2396
f2f611daf480 updated to Isabelle 12th Aug
Christian Urban <urbanc@in.tum.de>
parents: 2168
diff changeset
    60
      Class.instantiation ([full_tname], [], @{sort at}) thy;
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
    val ((_, (_, permute_ldef)), lthy) =
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
      Specification.definition (NONE, ((permute_def_name, []), permute_eqn)) lthy;
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
    val ((_, (_, atom_ldef)), lthy) =
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
      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
    65
    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
    66
    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
    67
    val atom_def = singleton (Proof_Context.export lthy ctxt_thy) atom_ldef;
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
    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
    69
    val sort_thm = @{thm at_class_sort} OF [type_definition, atom_def]
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
    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
    71
      |> snd o (Local_Theory.note ((sort_thm_name, [simp_attr]), [sort_thm]))
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
    72
      |> Class.prove_instantiation_instance (K (rtac class_thm 1))
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
      |> Local_Theory.exit_global;
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
  in
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
    thy
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
  end;
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
(** outer syntax **)
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
val _ =
3193
87d1e815aa59 command_spec antiquotation.
webertj
parents: 3135
diff changeset
    80
  Outer_Syntax.command @{command_spec "atom_decl"}
3135
92b9b8d2888d updated to new Isabelle (20 March)
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
    81
    "declaration of a concrete atom type" 
92b9b8d2888d updated to new Isabelle (20 March)
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
    82
      ((Parse.binding -- Scan.option (Args.parens (Parse.binding))) >>
92b9b8d2888d updated to new Isabelle (20 March)
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
    83
        (Toplevel.print oo (Toplevel.theory o add_atom_decl)))
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
end;