src/HOL/Tools/Predicate_Compile/predicate_compile.ML
author bulwahn
Sat, 24 Oct 2009 16:55:42 +0200
changeset 33123 3c7c4372f9ad
parent 33122 7d01480cc8e3
child 33124 5378e61add1a
permissions -rw-r--r--
cleaned up debugging messages; added options to code_pred command
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
*)
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
     4
signature PREDICATE_COMPILE =
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
     5
sig
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
     6
  val setup : theory -> theory
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
     7
  val preprocess : Predicate_Compile_Aux.options -> string -> theory -> theory
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
     8
end;
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
     9
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    10
structure Predicate_Compile : PREDICATE_COMPILE =
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    11
struct
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    12
33108
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33107
diff changeset
    13
(* options *)
33122
7d01480cc8e3 added first support for higher-order function translation
bulwahn
parents: 33121
diff changeset
    14
val fail_safe_mode = false
33108
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33107
diff changeset
    15
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    16
open Predicate_Compile_Aux;
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    17
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    18
val priority = tracing;
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    19
33113
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    20
(* tuple processing *)
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    21
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    22
fun expand_tuples thy intro =
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    23
  let
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    24
    fun rewrite_args [] (pats, intro_t, ctxt) = (pats, intro_t, ctxt)
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    25
      | rewrite_args (arg::args) (pats, intro_t, ctxt) = 
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    26
      (case HOLogic.strip_tupleT (fastype_of arg) of
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    27
        (Ts as _ :: _ :: _) =>
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    28
        let
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    29
          fun rewrite_arg' (Const ("Pair", _) $ _ $ t2, Type ("*", [_, T2]))
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    30
            (args, (pats, intro_t, ctxt)) = rewrite_arg' (t2, T2) (args, (pats, intro_t, ctxt))
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    31
            | rewrite_arg' (t, Type ("*", [T1, T2])) (args, (pats, intro_t, ctxt)) =
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    32
              let
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    33
                val ([x, y], ctxt') = Variable.variant_fixes ["x", "y"] ctxt
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    34
                val pat = (t, HOLogic.mk_prod (Free (x, T1), Free (y, T2)))
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    35
                  (*val _ = tracing ("Rewriting term " ^
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    36
                    (Syntax.string_of_term_global thy (fst pat)) ^ " to " ^
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    37
                    (Syntax.string_of_term_global thy (snd pat)) ^ " in " ^
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    38
                  (Syntax.string_of_term_global thy intro_t))*)
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    39
                val intro_t' = Pattern.rewrite_term thy [pat] [] intro_t
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    40
                val args' = map (Pattern.rewrite_term thy [pat] []) args
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    41
              in
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    42
                rewrite_arg' (Free (y, T2), T2) (args', (pat::pats, intro_t', ctxt'))
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    43
              end
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    44
            | rewrite_arg' _ (args, (pats, intro_t, ctxt)) = (args, (pats, intro_t, ctxt))
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    45
          val (args', (pats, intro_t', ctxt')) = rewrite_arg' (arg, fastype_of arg)
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    46
            (args, (pats, intro_t, ctxt))
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    47
        in
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    48
          rewrite_args args' (pats, intro_t', ctxt')
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    49
        end
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    50
      | _ => rewrite_args args (pats, intro_t, ctxt))
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    51
    fun rewrite_prem atom =
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    52
      let
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    53
        val (_, args) = strip_comb atom
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    54
      in rewrite_args args end
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    55
    val ctxt = ProofContext.init thy
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    56
    val (((T_insts, t_insts), [intro']), ctxt1) = Variable.import false [intro] ctxt
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    57
    val intro_t = prop_of intro'
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    58
    val concl = Logic.strip_imp_concl intro_t
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    59
    val (p, args) = strip_comb (HOLogic.dest_Trueprop concl)
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    60
    val (pats', intro_t', ctxt2) = rewrite_args args ([], intro_t, ctxt1)
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    61
    val (pats', intro_t', ctxt3) = 
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    62
      fold_atoms rewrite_prem intro_t' (pats', intro_t', ctxt2)
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    63
    (*val _ = Output.tracing ("intro_t': " ^ (Syntax.string_of_term_global thy intro_t'))
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    64
    val _ = Output.tracing ("pats : " ^ (commas (map
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    65
      (fn (t1, t2) => (Syntax.string_of_term_global thy t1) ^ " -> " ^
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    66
      Syntax.string_of_term_global thy t2) pats'))*)
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    67
    fun rewrite_pat (ct1, ct2) =
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    68
      (ct1, cterm_of thy (Pattern.rewrite_term thy pats' [] (term_of ct2)))
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    69
    val t_insts' = map rewrite_pat t_insts
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    70
      (*val _ = Output.tracing ("t_insts':" ^ (commas (map
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    71
      (fn (ct1, ct2) => (Syntax.string_of_term_global thy (term_of ct1) ^ " -> " ^
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    72
    Syntax.string_of_term_global thy (term_of ct2))) t_insts')))*)
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    73
    val intro'' = Thm.instantiate (T_insts, t_insts') intro
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    74
      (*val _ = Output.tracing ("intro'':" ^ (Display.string_of_thm_global thy intro''))*)
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    75
    val [intro'''] = Variable.export ctxt3 ctxt [intro'']
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    76
    (*val _ = Output.tracing ("intro''':" ^ (Display.string_of_thm_global thy intro'''))*)
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    77
  in
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    78
    intro'''
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    79
  end 
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    80
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    81
  (* eliminating fst/snd functions *)
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    82
val simplify_fst_snd = Simplifier.full_simplify
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    83
  (HOL_basic_ss addsimps [@{thm fst_conv}, @{thm snd_conv}, @{thm Pair_eq}])
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    84
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    85
(* Some last processing *)
33113
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33111
diff changeset
    86
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    87
fun remove_pointless_clauses intro =
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    88
  if Logic.strip_imp_prems (prop_of intro) = [@{prop "False"}] then
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    89
    []
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    90
  else [intro]
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
    91
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
    92
fun tracing s = ()
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
    93
33120
ca77d8c34ce2 cleaned up
bulwahn
parents: 33114
diff changeset
    94
fun print_intross thy msg intross =
ca77d8c34ce2 cleaned up
bulwahn
parents: 33114
diff changeset
    95
  tracing (msg ^ 
ca77d8c34ce2 cleaned up
bulwahn
parents: 33114
diff changeset
    96
    (space_implode "; " (map 
ca77d8c34ce2 cleaned up
bulwahn
parents: 33114
diff changeset
    97
      (fn intros => commas (map (Display.string_of_thm_global thy) intros)) intross)))
ca77d8c34ce2 cleaned up
bulwahn
parents: 33114
diff changeset
    98
33122
7d01480cc8e3 added first support for higher-order function translation
bulwahn
parents: 33121
diff changeset
    99
fun print_specs thy specs =
7d01480cc8e3 added first support for higher-order function translation
bulwahn
parents: 33121
diff changeset
   100
  map (fn (c, thms) => "Constant " ^ c ^ " has specification:\n"
7d01480cc8e3 added first support for higher-order function translation
bulwahn
parents: 33121
diff changeset
   101
    ^ (space_implode "\n" (map (Display.string_of_thm_global thy) thms)) ^ "\n") specs
7d01480cc8e3 added first support for higher-order function translation
bulwahn
parents: 33121
diff changeset
   102
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   103
fun process_specification options specs thy' =
33121
9b10dc5da0e0 added to process higher-order arguments by adding new constants
bulwahn
parents: 33120
diff changeset
   104
  let
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   105
    val _ = print_step options "Compiling predicates to flat introrules..."
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   106
    val specs = map (apsnd (map
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   107
      (fn th => if is_equationlike th then Pred_Compile_Data.normalize_equation thy' th else th))) specs
33121
9b10dc5da0e0 added to process higher-order arguments by adding new constants
bulwahn
parents: 33120
diff changeset
   108
    val (intross1, thy'') = apfst flat (fold_map Predicate_Compile_Pred.preprocess specs thy')
9b10dc5da0e0 added to process higher-order arguments by adding new constants
bulwahn
parents: 33120
diff changeset
   109
    val _ = print_intross thy'' "Flattened introduction rules: " intross1
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   110
    val _ =  "Replacing functions in introrules..."
33121
9b10dc5da0e0 added to process higher-order arguments by adding new constants
bulwahn
parents: 33120
diff changeset
   111
    val intross2 =
9b10dc5da0e0 added to process higher-order arguments by adding new constants
bulwahn
parents: 33120
diff changeset
   112
      if fail_safe_mode then
9b10dc5da0e0 added to process higher-order arguments by adding new constants
bulwahn
parents: 33120
diff changeset
   113
        case try (burrow (maps (Predicate_Compile_Fun.rewrite_intro thy''))) intross1 of
9b10dc5da0e0 added to process higher-order arguments by adding new constants
bulwahn
parents: 33120
diff changeset
   114
          SOME intross => intross
9b10dc5da0e0 added to process higher-order arguments by adding new constants
bulwahn
parents: 33120
diff changeset
   115
        | NONE => let val _ = warning "Function replacement failed!" in intross1 end
9b10dc5da0e0 added to process higher-order arguments by adding new constants
bulwahn
parents: 33120
diff changeset
   116
      else burrow (maps (Predicate_Compile_Fun.rewrite_intro thy'')) intross1
9b10dc5da0e0 added to process higher-order arguments by adding new constants
bulwahn
parents: 33120
diff changeset
   117
    val _ = print_intross thy'' "Introduction rules with replaced functions: " intross2
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   118
    val _ = print_step options "Introducing new constants for abstractions at higher-order argument positions..."
33121
9b10dc5da0e0 added to process higher-order arguments by adding new constants
bulwahn
parents: 33120
diff changeset
   119
    val (intross3, (new_defs, thy''')) = Predicate_Compile_Pred.flat_higher_order_arguments (intross2, thy'')
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   120
    val (new_intross, thy'''')  =
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   121
      if not (null new_defs) then
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   122
      let
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   123
        val _ = print_step options "Recursively obtaining introduction rules for new definitions..."
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   124
      in process_specification options new_defs thy''' end
33121
9b10dc5da0e0 added to process higher-order arguments by adding new constants
bulwahn
parents: 33120
diff changeset
   125
    else ([], thy''')
9b10dc5da0e0 added to process higher-order arguments by adding new constants
bulwahn
parents: 33120
diff changeset
   126
  in
9b10dc5da0e0 added to process higher-order arguments by adding new constants
bulwahn
parents: 33120
diff changeset
   127
    (intross3 @ new_intross, thy'''')
9b10dc5da0e0 added to process higher-order arguments by adding new constants
bulwahn
parents: 33120
diff changeset
   128
  end
9b10dc5da0e0 added to process higher-order arguments by adding new constants
bulwahn
parents: 33120
diff changeset
   129
9b10dc5da0e0 added to process higher-order arguments by adding new constants
bulwahn
parents: 33120
diff changeset
   130
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   131
fun preprocess_strong_conn_constnames options gr constnames thy =
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
   132
  let
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
   133
    val get_specs = map (fn k => (k, Graph.get_node gr k))
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   134
    val _ = print_step options ("Preprocessing scc of " ^ commas constnames)
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
   135
    val (prednames, funnames) = List.partition (is_pred thy) constnames
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
   136
    (* untangle recursion by defining predicates for all functions *)
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   137
    val _ = print_step options
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   138
      ("Compiling functions (" ^ commas funnames ^ ") to predicates...")
33122
7d01480cc8e3 added first support for higher-order function translation
bulwahn
parents: 33121
diff changeset
   139
    val (fun_pred_specs, thy') =
7d01480cc8e3 added first support for higher-order function translation
bulwahn
parents: 33121
diff changeset
   140
      if not (null funnames) then Predicate_Compile_Fun.define_predicates
7d01480cc8e3 added first support for higher-order function translation
bulwahn
parents: 33121
diff changeset
   141
      (get_specs funnames) thy else ([], thy)
7d01480cc8e3 added first support for higher-order function translation
bulwahn
parents: 33121
diff changeset
   142
    val _ = print_specs thy' fun_pred_specs
7d01480cc8e3 added first support for higher-order function translation
bulwahn
parents: 33121
diff changeset
   143
    val specs = (get_specs prednames) @ fun_pred_specs
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   144
    val (intross3, thy''') = process_specification options specs thy'
33121
9b10dc5da0e0 added to process higher-order arguments by adding new constants
bulwahn
parents: 33120
diff changeset
   145
    val _ = print_intross thy''' "Introduction rules with new constants: " intross3
9b10dc5da0e0 added to process higher-order arguments by adding new constants
bulwahn
parents: 33120
diff changeset
   146
    val intross4 = map (maps remove_pointless_clauses) intross3
9b10dc5da0e0 added to process higher-order arguments by adding new constants
bulwahn
parents: 33120
diff changeset
   147
    val _ = print_intross thy''' "After removing pointless clauses: " intross4
9b10dc5da0e0 added to process higher-order arguments by adding new constants
bulwahn
parents: 33120
diff changeset
   148
    val intross5 = burrow (map (AxClass.overload thy''')) intross4
9b10dc5da0e0 added to process higher-order arguments by adding new constants
bulwahn
parents: 33120
diff changeset
   149
    val intross6 = burrow (map (simplify_fst_snd o expand_tuples thy''')) intross5
9b10dc5da0e0 added to process higher-order arguments by adding new constants
bulwahn
parents: 33120
diff changeset
   150
    val _ = print_intross thy''' "introduction rules before registering: " intross6
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   151
    val _ = print_step options "Registering introduction rules..."
33121
9b10dc5da0e0 added to process higher-order arguments by adding new constants
bulwahn
parents: 33120
diff changeset
   152
    val thy'''' = fold Predicate_Compile_Core.register_intros intross6 thy'''
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
   153
  in
33121
9b10dc5da0e0 added to process higher-order arguments by adding new constants
bulwahn
parents: 33120
diff changeset
   154
    thy''''
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
   155
  end;
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
   156
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   157
fun preprocess options const thy =
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32669
diff changeset
   158
  let
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   159
    val _ = print_step options "Fetching definitions from theory..."
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32669
diff changeset
   160
    val table = Pred_Compile_Data.make_const_spec_table thy
33107
6aa76ce59ef2 added filtering of case constants in the definition retrieval of the predicate compiler
bulwahn
parents: 32672
diff changeset
   161
    val gr = Pred_Compile_Data.obtain_specification_graph thy table const
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32669
diff changeset
   162
    val gr = Graph.subgraph (member (op =) (Graph.all_succs gr [const])) gr
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   163
  in fold_rev (preprocess_strong_conn_constnames options gr)
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32669
diff changeset
   164
    (Graph.strong_conn gr) thy
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32669
diff changeset
   165
  end
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32669
diff changeset
   166
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   167
fun extract_options ((modes, raw_options), raw_const) =
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   168
  let
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   169
    fun chk s = member (op =) raw_options s
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   170
  in
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   171
    Options {
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   172
      show_steps = chk "show_steps",
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   173
      show_mode_inference = chk "show_mode_inference",
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   174
      inductify = chk "inductify",
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   175
      rpred = chk "rpred"
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   176
    }
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   177
  end
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   178
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   179
fun code_pred_cmd ((modes, raw_options), raw_const) lthy =
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   180
  let
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   181
    val options = extract_options ((modes, raw_options), raw_const)
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   182
  in  
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   183
    if (is_inductify options) then
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   184
      let
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   185
        val thy = ProofContext.theory_of lthy
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   186
        val const = Code.read_const thy raw_const
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   187
        val lthy' = LocalTheory.theory (preprocess options const) lthy
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   188
          |> LocalTheory.checkpoint
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   189
          
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   190
        val const = case Predicate_Compile_Fun.pred_of_function (ProofContext.theory_of lthy') const of
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   191
            SOME c => c
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   192
          | NONE => const
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   193
        val _ = print_step options "Starting Predicate Compile Core..."
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   194
      in
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   195
        Predicate_Compile_Core.code_pred options modes (is_rpred options) const lthy'
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   196
      end
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   197
    else
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   198
      Predicate_Compile_Core.code_pred_cmd options modes (is_rpred options) raw_const lthy
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   199
  end
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
   200
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
   201
val setup = Predicate_Compile_Fun.setup_oracle #> Predicate_Compile_Core.setup
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
   202
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   203
val bool_options = ["show_steps", "show_mode_inference", "inductify", "rpred"]
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   204
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   205
val _ = List.app OuterKeyword.keyword ("mode" :: bool_options)
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
   206
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
   207
local structure P = OuterParse
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
   208
in
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
   209
33114
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   210
val opt_modes =
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   211
  Scan.optional (P.$$$ "(" |-- P.$$$ "mode" |-- P.$$$ ":" |--
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   212
   P.enum1 "," (P.$$$ "[" |-- P.enum "," P.nat --| P.$$$ "]")
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   213
  --| P.$$$ ")" >> SOME) NONE
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   214
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   215
val scan_params =
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   216
  let
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   217
    val scan_bool_param = foldl1 (op ||) (map P.$$$ bool_options)
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   218
  in
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   219
    Scan.optional (P.$$$ "[" |-- P.enum1 "," scan_bool_param --| P.$$$ "]") []
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   220
  end
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   221
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
   222
val _ = OuterSyntax.local_theory_to_proof "code_pred"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
   223
  "prove equations for predicate specified by intro/elim rules"
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   224
  OuterKeyword.thy_goal (opt_modes -- scan_params -- P.term_group >>
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33122
diff changeset
   225
    code_pred_cmd)
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
   226
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
   227
end
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
   228
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents:
diff changeset
   229
end