src/Pure/fact_index.ML
author wenzelm
Thu, 15 Sep 2005 17:16:56 +0200
changeset 17412 e26cb20ef0cc
parent 17221 6cd180204582
child 18026 ccf2972f6f89
permissions -rw-r--r--
TableFun/Symtab: curried lookup and update;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13270
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/fact_index.ML
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
     3
    Author:     Markus Wenzel, TU Muenchen
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
     4
16020
ace2c610b5be major tuning;
wenzelm
parents: 15570
diff changeset
     5
Facts indexed by consts or frees.
13270
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
     6
*)
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
     7
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
     8
signature FACT_INDEX =
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
     9
sig
16020
ace2c610b5be major tuning;
wenzelm
parents: 15570
diff changeset
    10
  type spec
ace2c610b5be major tuning;
wenzelm
parents: 15570
diff changeset
    11
  val index_term: (string -> bool) -> term -> spec -> spec
ace2c610b5be major tuning;
wenzelm
parents: 15570
diff changeset
    12
  val index_thm: (string -> bool) -> thm -> spec -> spec
13270
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    13
  type T
16020
ace2c610b5be major tuning;
wenzelm
parents: 15570
diff changeset
    14
  val facts: T -> (string * thm list) list
13270
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    15
  val empty: T
16020
ace2c610b5be major tuning;
wenzelm
parents: 15570
diff changeset
    16
  val add: (string -> bool) -> (string * thm list) -> T -> T
ace2c610b5be major tuning;
wenzelm
parents: 15570
diff changeset
    17
  val find: T -> spec -> (string * thm list) list
13270
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    18
end;
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    19
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    20
structure FactIndex: FACT_INDEX =
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    21
struct
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    22
16020
ace2c610b5be major tuning;
wenzelm
parents: 15570
diff changeset
    23
type spec = string list * string list;
ace2c610b5be major tuning;
wenzelm
parents: 15570
diff changeset
    24
ace2c610b5be major tuning;
wenzelm
parents: 15570
diff changeset
    25
13270
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    26
(* indexing items *)
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    27
16874
wenzelm
parents: 16491
diff changeset
    28
val add_consts = fold_aterms
wenzelm
parents: 16491
diff changeset
    29
  (fn Const (c, _) => if c = Term.dummy_patternN then I else insert (op =) c | _ => I);
13270
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    30
16874
wenzelm
parents: 16491
diff changeset
    31
fun add_frees pred = fold_aterms
wenzelm
parents: 16491
diff changeset
    32
  (fn Free (x, _) => if pred x then insert (op =) x else I | _ => I);
wenzelm
parents: 16491
diff changeset
    33
wenzelm
parents: 16491
diff changeset
    34
fun index_term pred t (cs, xs) = (add_consts t cs, add_frees pred t xs);
13270
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    35
16020
ace2c610b5be major tuning;
wenzelm
parents: 15570
diff changeset
    36
fun index_thm pred th idx =
ace2c610b5be major tuning;
wenzelm
parents: 15570
diff changeset
    37
  let val {tpairs, hyps, prop, ...} = Thm.rep_thm th in
ace2c610b5be major tuning;
wenzelm
parents: 15570
diff changeset
    38
    idx
ace2c610b5be major tuning;
wenzelm
parents: 15570
diff changeset
    39
    |> index_term pred prop
ace2c610b5be major tuning;
wenzelm
parents: 15570
diff changeset
    40
    |> fold (index_term pred) hyps
16874
wenzelm
parents: 16491
diff changeset
    41
    |> fold (fn (t, u) => index_term pred t #> index_term pred u) tpairs
16020
ace2c610b5be major tuning;
wenzelm
parents: 15570
diff changeset
    42
  end;
13270
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    43
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    44
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    45
(* build index *)
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    46
13283
1051aa66cbf3 proper treatment of border cases;
wenzelm
parents: 13270
diff changeset
    47
datatype T = Index of {next: int, facts: (string * thm list) list,
13270
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    48
  consts: (int * (string * thm list)) list Symtab.table,
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    49
  frees: (int * (string * thm list)) list Symtab.table};
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    50
16020
ace2c610b5be major tuning;
wenzelm
parents: 15570
diff changeset
    51
fun facts (Index {facts, ...}) = facts;
ace2c610b5be major tuning;
wenzelm
parents: 15570
diff changeset
    52
13270
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    53
val empty =
13283
1051aa66cbf3 proper treatment of border cases;
wenzelm
parents: 13270
diff changeset
    54
  Index {next = 0, facts = [], consts = Symtab.empty, frees = Symtab.empty};
13270
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    55
16020
ace2c610b5be major tuning;
wenzelm
parents: 15570
diff changeset
    56
fun add pred (name, ths) (Index {next, facts, consts, frees}) =
13270
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    57
  let
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17221
diff changeset
    58
    fun upd a = Symtab.update_multi (a, (next, (name, ths)));
16020
ace2c610b5be major tuning;
wenzelm
parents: 15570
diff changeset
    59
    val (cs, xs) = fold (index_thm pred) ths ([], []);
13283
1051aa66cbf3 proper treatment of border cases;
wenzelm
parents: 13270
diff changeset
    60
  in
1051aa66cbf3 proper treatment of border cases;
wenzelm
parents: 13270
diff changeset
    61
    Index {next = next + 1, facts = (name, ths) :: facts,
16020
ace2c610b5be major tuning;
wenzelm
parents: 15570
diff changeset
    62
      consts = fold upd cs consts, frees = fold upd xs frees}
13283
1051aa66cbf3 proper treatment of border cases;
wenzelm
parents: 13270
diff changeset
    63
  end;
13270
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    64
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    65
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    66
(* find facts *)
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    67
16491
7310d0a36599 OrdList.inter;
wenzelm
parents: 16020
diff changeset
    68
fun fact_ord ((i, _), (j, _)) = int_ord (j, i);
13270
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    69
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    70
fun intersects [xs] = xs
16491
7310d0a36599 OrdList.inter;
wenzelm
parents: 16020
diff changeset
    71
  | intersects xss =
7310d0a36599 OrdList.inter;
wenzelm
parents: 16020
diff changeset
    72
      if exists null xss then []
7310d0a36599 OrdList.inter;
wenzelm
parents: 16020
diff changeset
    73
      else fold (OrdList.inter fact_ord) (tl xss) (hd xss);
13270
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    74
16020
ace2c610b5be major tuning;
wenzelm
parents: 15570
diff changeset
    75
fun find idx ([], []) = facts idx
ace2c610b5be major tuning;
wenzelm
parents: 15570
diff changeset
    76
  | find (Index {consts, frees, ...}) (cs, xs) =
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17221
diff changeset
    77
      (map (Symtab.lookup_multi consts) cs @
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17221
diff changeset
    78
        map (Symtab.lookup_multi frees) xs) |> intersects |> map #2;
13270
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    79
d7f35250cbad Facts indexed by consts or (some) frees.
wenzelm
parents:
diff changeset
    80
end;