src/Pure/Isar/instance.ML
author wenzelm
Mon, 25 Feb 2008 16:31:20 +0100
changeset 26132 c927c3ed82c9
parent 25864 11f531354852
child 26269 5bb50f58a113
permissions -rw-r--r--
implicit use of LocalTheory.group etc.;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
24592
dfea1edbf711 added rudimentary instantiation stub
haftmann
parents:
diff changeset
     1
(*  Title:      Pure/Isar/instance.ML
dfea1edbf711 added rudimentary instantiation stub
haftmann
parents:
diff changeset
     2
    ID:         $Id$
dfea1edbf711 added rudimentary instantiation stub
haftmann
parents:
diff changeset
     3
    Author:     Florian Haftmann, TU Muenchen
dfea1edbf711 added rudimentary instantiation stub
haftmann
parents:
diff changeset
     4
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25381
diff changeset
     5
A primitive instance command, based on instantiation target.
24592
dfea1edbf711 added rudimentary instantiation stub
haftmann
parents:
diff changeset
     6
*)
dfea1edbf711 added rudimentary instantiation stub
haftmann
parents:
diff changeset
     7
dfea1edbf711 added rudimentary instantiation stub
haftmann
parents:
diff changeset
     8
signature INSTANCE =
dfea1edbf711 added rudimentary instantiation stub
haftmann
parents:
diff changeset
     9
sig
25536
01753a944433 improved
haftmann
parents: 25514
diff changeset
    10
  val instantiation_cmd: xstring list * sort * xstring -> theory -> local_theory
01753a944433 improved
haftmann
parents: 25514
diff changeset
    11
  val instance_cmd: xstring * sort * xstring -> ((bstring * Attrib.src list) * xstring) list
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25381
diff changeset
    12
    -> theory -> Proof.state
24592
dfea1edbf711 added rudimentary instantiation stub
haftmann
parents:
diff changeset
    13
end;
dfea1edbf711 added rudimentary instantiation stub
haftmann
parents:
diff changeset
    14
dfea1edbf711 added rudimentary instantiation stub
haftmann
parents:
diff changeset
    15
structure Instance : INSTANCE =
dfea1edbf711 added rudimentary instantiation stub
haftmann
parents:
diff changeset
    16
struct
dfea1edbf711 added rudimentary instantiation stub
haftmann
parents:
diff changeset
    17
25864
11f531354852 explicit type variables for instantiation
haftmann
parents: 25829
diff changeset
    18
fun read_single_arity thy (raw_tyco, raw_sorts, raw_sort) =
11f531354852 explicit type variables for instantiation
haftmann
parents: 25829
diff changeset
    19
  let
11f531354852 explicit type variables for instantiation
haftmann
parents: 25829
diff changeset
    20
    val (tyco, sorts, sort) = Sign.read_arity thy (raw_tyco, raw_sorts, raw_sort);
11f531354852 explicit type variables for instantiation
haftmann
parents: 25829
diff changeset
    21
    val vs = Name.names Name.context Name.aT sorts;
11f531354852 explicit type variables for instantiation
haftmann
parents: 25829
diff changeset
    22
  in (tyco, vs, sort) end;
11f531354852 explicit type variables for instantiation
haftmann
parents: 25829
diff changeset
    23
25536
01753a944433 improved
haftmann
parents: 25514
diff changeset
    24
fun read_multi_arity thy (raw_tycos, raw_sorts, raw_sort) =
01753a944433 improved
haftmann
parents: 25514
diff changeset
    25
  let
01753a944433 improved
haftmann
parents: 25514
diff changeset
    26
    val all_arities = map (fn raw_tyco => Sign.read_arity thy
01753a944433 improved
haftmann
parents: 25514
diff changeset
    27
      (raw_tyco, raw_sorts, raw_sort)) raw_tycos;
01753a944433 improved
haftmann
parents: 25514
diff changeset
    28
    val tycos = map #1 all_arities;
01753a944433 improved
haftmann
parents: 25514
diff changeset
    29
    val (_, sorts, sort) = hd all_arities;
25864
11f531354852 explicit type variables for instantiation
haftmann
parents: 25829
diff changeset
    30
    val vs = Name.names Name.context Name.aT sorts;
11f531354852 explicit type variables for instantiation
haftmann
parents: 25829
diff changeset
    31
  in (tycos, vs, sort) end;
25536
01753a944433 improved
haftmann
parents: 25514
diff changeset
    32
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25381
diff changeset
    33
fun instantiation_cmd raw_arities thy =
25536
01753a944433 improved
haftmann
parents: 25514
diff changeset
    34
  TheoryTarget.instantiation (read_multi_arity thy raw_arities) thy;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25381
diff changeset
    35
25864
11f531354852 explicit type variables for instantiation
haftmann
parents: 25829
diff changeset
    36
fun instance_cmd raw_arities defs thy =
24592
dfea1edbf711 added rudimentary instantiation stub
haftmann
parents:
diff changeset
    37
  let
25864
11f531354852 explicit type variables for instantiation
haftmann
parents: 25829
diff changeset
    38
    val (tyco, vs, sort) = read_single_arity thy raw_arities;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25381
diff changeset
    39
    fun mk_def ctxt ((name, raw_attr), raw_t) =
24592
dfea1edbf711 added rudimentary instantiation stub
haftmann
parents:
diff changeset
    40
      let
25864
11f531354852 explicit type variables for instantiation
haftmann
parents: 25829
diff changeset
    41
        val attr = map (Attrib.intern_src thy) raw_attr;
11f531354852 explicit type variables for instantiation
haftmann
parents: 25829
diff changeset
    42
        val t = Syntax.parse_prop ctxt raw_t;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25381
diff changeset
    43
      in (NONE, ((name, attr), t)) end;
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    44
    fun define def ctxt =
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    45
      let
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    46
        val def' = (apsnd o apsnd) (Syntax.check_prop ctxt) def;
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    47
      in Specification.definition def' ctxt end;
25569
c597835d5de4 dropped Instance.instantiate
haftmann
parents: 25536
diff changeset
    48
    val _ = if not (null defs) then Output.legacy_feature
25829
4b44d945702f improved warning
haftmann
parents: 25569
diff changeset
    49
      "Instance command with attached definitions; use instantiation instead." else ();
25536
01753a944433 improved
haftmann
parents: 25514
diff changeset
    50
  in if not (null defs) orelse forall (Class.is_class thy) sort
25502
9200b36280c0 instance command as rudimentary class target
haftmann
parents: 25485
diff changeset
    51
  then
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25381
diff changeset
    52
    thy
25864
11f531354852 explicit type variables for instantiation
haftmann
parents: 25829
diff changeset
    53
    |> TheoryTarget.instantiation ([tyco], vs, sort)
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25381
diff changeset
    54
    |> `(fn ctxt => map (mk_def ctxt) defs)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25381
diff changeset
    55
    |-> (fn defs => fold_map Specification.definition defs)
25864
11f531354852 explicit type variables for instantiation
haftmann
parents: 25829
diff changeset
    56
    |> snd
11f531354852 explicit type variables for instantiation
haftmann
parents: 25829
diff changeset
    57
    |> LocalTheory.reinit
11f531354852 explicit type variables for instantiation
haftmann
parents: 25829
diff changeset
    58
    |> Class.instantiation_instance Class.conclude_instantiation
25502
9200b36280c0 instance command as rudimentary class target
haftmann
parents: 25485
diff changeset
    59
  else
9200b36280c0 instance command as rudimentary class target
haftmann
parents: 25485
diff changeset
    60
    thy
25864
11f531354852 explicit type variables for instantiation
haftmann
parents: 25829
diff changeset
    61
    |> Class.instance_arity I (tyco, map snd vs, sort)
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25381
diff changeset
    62
  end;
24592
dfea1edbf711 added rudimentary instantiation stub
haftmann
parents:
diff changeset
    63
dfea1edbf711 added rudimentary instantiation stub
haftmann
parents:
diff changeset
    64
end;