src/Pure/Isar/spec_rules.ML
author wenzelm
Tue, 25 Jun 2013 11:41:16 +0200
changeset 52443 725916b7dee5
parent 45291 57cd50f98fdc
child 59573 d09cc83cdce9
permissions -rw-r--r--
more formal isabelle_admin_build; tuned;
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
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33702
diff changeset
    15
  val retrieve: Proof.context -> term -> spec list
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33702
diff changeset
    16
  val retrieve_global: theory -> term -> spec list
33454
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    17
  val add: rough_classification -> term list * thm list -> local_theory -> local_theory
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    18
  val add_global: rough_classification -> term list * thm list -> theory -> theory
33374
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    19
end;
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    20
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    21
structure Spec_Rules: SPEC_RULES =
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    22
struct
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    23
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    24
(* maintain rules *)
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    25
33454
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    26
datatype rough_classification = Unknown | Equational | Inductive | Co_Inductive;
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    27
type spec = rough_classification * (term list * thm list);
33374
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    28
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 33457
diff changeset
    29
structure Rules = Generic_Data
33374
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    30
(
33454
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    31
  type T = spec Item_Net.T;
33374
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    32
  val empty : T =
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    33
    Item_Net.init
33454
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    34
      (fn ((class1, (ts1, ths1)), (class2, (ts2, ths2))) =>
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    35
        class1 = class2 andalso
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    36
        eq_list (op aconv) (ts1, ts2) andalso
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    37
        eq_list Thm.eq_thm_prop (ths1, ths2))
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    38
      (#1 o #2);
33374
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    39
  val extend = I;
41472
f6ab14e61604 misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents: 34948
diff changeset
    40
  val merge = Item_Net.merge;
33374
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    41
);
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    42
33454
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    43
val get = Item_Net.content o Rules.get o Context.Proof;
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    44
val get_global = Item_Net.content o Rules.get o Context.Theory;
33374
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    45
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33702
diff changeset
    46
val retrieve = Item_Net.retrieve o Rules.get o Context.Proof;
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33702
diff changeset
    47
val retrieve_global = Item_Net.retrieve o Rules.get o Context.Theory;
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33702
diff changeset
    48
33702
9e6b6594da6b use simultaneous Morphics.fact;
wenzelm
parents: 33671
diff changeset
    49
fun add class (ts, ths) lthy =
9e6b6594da6b use simultaneous Morphics.fact;
wenzelm
parents: 33671
diff changeset
    50
  let
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 41472
diff changeset
    51
    val cts = map (Thm.cterm_of (Proof_Context.theory_of lthy)) ts;
33702
9e6b6594da6b use simultaneous Morphics.fact;
wenzelm
parents: 33671
diff changeset
    52
  in
45291
57cd50f98fdc uniform Local_Theory.declaration with explicit params;
wenzelm
parents: 42360
diff changeset
    53
    lthy |> Local_Theory.declaration {syntax = false, pervasive = true}
33702
9e6b6594da6b use simultaneous Morphics.fact;
wenzelm
parents: 33671
diff changeset
    54
      (fn phi =>
9e6b6594da6b use simultaneous Morphics.fact;
wenzelm
parents: 33671
diff changeset
    55
        let
9e6b6594da6b use simultaneous Morphics.fact;
wenzelm
parents: 33671
diff changeset
    56
          val (ts', ths') =
9e6b6594da6b use simultaneous Morphics.fact;
wenzelm
parents: 33671
diff changeset
    57
            Morphism.fact phi (map Drule.mk_term cts @ ths)
9e6b6594da6b use simultaneous Morphics.fact;
wenzelm
parents: 33671
diff changeset
    58
            |> chop (length cts)
9e6b6594da6b use simultaneous Morphics.fact;
wenzelm
parents: 33671
diff changeset
    59
            |>> map (Thm.term_of o Drule.dest_term);
9e6b6594da6b use simultaneous Morphics.fact;
wenzelm
parents: 33671
diff changeset
    60
        in Rules.map (Item_Net.update (class, (ts', ths'))) end)
9e6b6594da6b use simultaneous Morphics.fact;
wenzelm
parents: 33671
diff changeset
    61
  end;
33374
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    62
33454
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    63
fun add_global class spec =
485fd398dd33 misc tuning and clarification;
wenzelm
parents: 33374
diff changeset
    64
  Context.theory_map (Rules.map (Item_Net.update (class, spec)));
33374
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    65
8099185908a4 Rules that characterize functional/relational specifications.
wenzelm
parents:
diff changeset
    66
end;