src/Pure/Isar/spec_rules.ML
author wenzelm
Wed, 02 Dec 2009 12:04:07 +0100
changeset 33930 6a973bd43949
parent 33702 9e6b6594da6b
child 34948 2d5f2a9f7601
permissions -rw-r--r--
slightly less ambitious settings, to avoid potential out-of-memory problem;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33374
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Isar/spec_rules.ML
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
     3
33454
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
     4
Rules that characterize specifications, with rough classification.
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
     5
NB: In the face of arbitrary morphisms, the original shape of
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
     6
specifications may get lost.
33374
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
     7
*)
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
     8
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
     9
signature SPEC_RULES =
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    10
sig
33454
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    11
  datatype rough_classification = Unknown | Equational | Inductive | Co_Inductive
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    12
  type spec = rough_classification * (term list * thm list)
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    13
  val get: Proof.context -> spec list
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    14
  val get_global: theory -> spec list
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    15
  val add: rough_classification -> term list * thm list -> local_theory -> local_theory
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    16
  val add_global: rough_classification -> term list * thm list -> theory -> theory
33374
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    17
end;
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    18
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    19
structure Spec_Rules: SPEC_RULES =
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    20
struct
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    21
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    22
(* maintain rules *)
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    23
33454
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    24
datatype rough_classification = Unknown | Equational | Inductive | Co_Inductive;
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    25
type spec = rough_classification * (term list * thm list);
33374
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    26
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 33457
diff changeset
    27
structure Rules = Generic_Data
33374
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    28
(
33454
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    29
  type T = spec Item_Net.T;
33374
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    30
  val empty : T =
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    31
    Item_Net.init
33454
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    32
      (fn ((class1, (ts1, ths1)), (class2, (ts2, ths2))) =>
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    33
        class1 = class2 andalso
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    34
        eq_list (op aconv) (ts1, ts2) andalso
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    35
        eq_list Thm.eq_thm_prop (ths1, ths2))
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    36
      (#1 o #2);
33374
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    37
  val extend = I;
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 33457
diff changeset
    38
  fun merge data = Item_Net.merge data;
33374
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    39
);
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    40
33454
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    41
val get = Item_Net.content o Rules.get o Context.Proof;
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    42
val get_global = Item_Net.content o Rules.get o Context.Theory;
33374
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    43
33702
9e6b6594da6b use simultaneous Morphics.fact;
wenzelm
parents: 33671
diff changeset
    44
fun add class (ts, ths) lthy =
9e6b6594da6b use simultaneous Morphics.fact;
wenzelm
parents: 33671
diff changeset
    45
  let
9e6b6594da6b use simultaneous Morphics.fact;
wenzelm
parents: 33671
diff changeset
    46
    val cts = map (Thm.cterm_of (ProofContext.theory_of lthy)) ts;
9e6b6594da6b use simultaneous Morphics.fact;
wenzelm
parents: 33671
diff changeset
    47
  in
9e6b6594da6b use simultaneous Morphics.fact;
wenzelm
parents: 33671
diff changeset
    48
    lthy |> Local_Theory.declaration true
9e6b6594da6b use simultaneous Morphics.fact;
wenzelm
parents: 33671
diff changeset
    49
      (fn phi =>
9e6b6594da6b use simultaneous Morphics.fact;
wenzelm
parents: 33671
diff changeset
    50
        let
9e6b6594da6b use simultaneous Morphics.fact;
wenzelm
parents: 33671
diff changeset
    51
          val (ts', ths') =
9e6b6594da6b use simultaneous Morphics.fact;
wenzelm
parents: 33671
diff changeset
    52
            Morphism.fact phi (map Drule.mk_term cts @ ths)
9e6b6594da6b use simultaneous Morphics.fact;
wenzelm
parents: 33671
diff changeset
    53
            |> chop (length cts)
9e6b6594da6b use simultaneous Morphics.fact;
wenzelm
parents: 33671
diff changeset
    54
            |>> map (Thm.term_of o Drule.dest_term);
9e6b6594da6b use simultaneous Morphics.fact;
wenzelm
parents: 33671
diff changeset
    55
        in Rules.map (Item_Net.update (class, (ts', ths'))) end)
9e6b6594da6b use simultaneous Morphics.fact;
wenzelm
parents: 33671
diff changeset
    56
  end;
33374
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    57
33454
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    58
fun add_global class spec =
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    59
  Context.theory_map (Rules.map (Item_Net.update (class, spec)));
33374
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    60
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    61
end;