src/HOL/Tools/Predicate_Compile/pred_compile_set.ML
author haftmann
Fri, 25 Sep 2009 09:50:31 +0200
changeset 32705 04ce6bb14d85
parent 32668 b2de45007537
child 33112 6672184a736b
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
     1
(* Author: Lukas Bulwahn, TU Muenchen
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
     2
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
     3
Preprocessing sets to predicates
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
     4
*)
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
     5
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
     6
signature PRED_COMPILE_SET =
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
     7
sig
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
     8
(*
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
     9
  val preprocess_intro : thm -> theory -> thm * theory
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    10
  val preprocess_clause : term -> theory -> term * theory
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    11
*)
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    12
  val unfold_set_notation : thm -> thm;
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    13
end;
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    14
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    15
structure Pred_Compile_Set : PRED_COMPILE_SET =
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    16
struct
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    17
(*FIXME: unfolding Ball in pretty adhoc here *)
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    18
val unfold_set_lemmas = [@{thm Collect_def}, @{thm mem_def}, @{thm Ball_def}]
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    19
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    20
val unfold_set_notation = Simplifier.rewrite_rule unfold_set_lemmas
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    21
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    22
(*
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    23
fun find_set_theorems ctxt cname =
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    24
  let
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    25
    val _ = cname 
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    26
*)
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    27
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    28
(*
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    29
fun preprocess_term t ctxt =
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    30
  case t of
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    31
    Const ("op :", _) $ x $ A => 
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    32
      case strip_comb A of
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    33
        (Const (cname, T), params) =>
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    34
          let 
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    35
            val _ = find_set_theorems
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    36
          in
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    37
            (t, ctxt)
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    38
          end
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    39
      | _ => (t, ctxt)  
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    40
  | _ => (t, ctxt)
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    41
*) 
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    42
(*
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    43
fun preprocess_intro th thy =
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    44
  let
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    45
    val cnames = find_heads_of_prems
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    46
    val set_cname = filter (has_set_definition
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    47
    val _ = define_preds
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    48
    val _ = prep_pred_def
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    49
  in
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    50
*)
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    51
end;