src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML
author wenzelm
Wed, 26 Nov 2014 16:55:43 +0100
changeset 59057 5b649fb2f2e1
parent 57962 0284a7d083be
child 59205 663794ab87e6
permissions -rw-r--r--
added ML antiquotation @{apply n} or @{apply n(k)};
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33265
01c9c6dbd890 proper headers;
wenzelm
parents: 33252
diff changeset
     1
(*  Title:      HOL/Tools/Predicate_Compile/predicate_compile_data.ML
01c9c6dbd890 proper headers;
wenzelm
parents: 33252
diff changeset
     2
    Author:     Lukas Bulwahn, TU Muenchen
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
     3
33265
01c9c6dbd890 proper headers;
wenzelm
parents: 33252
diff changeset
     4
Book-keeping datastructure for the predicate compiler.
01c9c6dbd890 proper headers;
wenzelm
parents: 33252
diff changeset
     5
*)
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
     6
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
     7
signature PREDICATE_COMPILE_DATA =
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
     8
sig
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
     9
  val ignore_consts : string list -> theory -> theory
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    10
  val keep_functions : string list -> theory -> theory
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    11
  val keep_function : theory -> string -> bool
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    12
  val processed_specs : theory -> string -> (string * thm list) list option
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    13
  val store_processed_specs : (string * (string * thm list) list) -> theory -> theory
55437
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
    14
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    15
  val get_specification : Predicate_Compile_Aux.options -> theory -> term -> thm list
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
    16
  val obtain_specification_graph :
35404
de56579ae229 just one copy of structure Term_Graph (in Pure);
wenzelm
parents: 35324
diff changeset
    17
    Predicate_Compile_Aux.options -> theory -> term -> thm list Term_Graph.T
55437
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
    18
35404
de56579ae229 just one copy of structure Term_Graph (in Pure);
wenzelm
parents: 35324
diff changeset
    19
  val present_graph : thm list Term_Graph.T -> unit
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    20
  val normalize_equation : theory -> thm -> thm
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    21
end;
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    22
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    23
structure Predicate_Compile_Data : PREDICATE_COMPILE_DATA =
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    24
struct
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    25
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    26
open Predicate_Compile_Aux;
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    27
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33487
diff changeset
    28
structure Data = Theory_Data
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    29
(
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    30
  type T =
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    31
    {ignore_consts : unit Symtab.table,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    32
     keep_functions : unit Symtab.table,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    33
     processed_specs : ((string * thm list) list) Symtab.table};
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    34
  val empty =
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    35
    {ignore_consts = Symtab.empty,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    36
     keep_functions = Symtab.empty,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    37
     processed_specs =  Symtab.empty};
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    38
  val extend = I;
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33487
diff changeset
    39
  fun merge
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    40
    ({ignore_consts = c1, keep_functions = k1, processed_specs = s1},
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    41
     {ignore_consts = c2, keep_functions = k2, processed_specs = s2}) =
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    42
     {ignore_consts = Symtab.merge (K true) (c1, c2),
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    43
      keep_functions = Symtab.merge (K true) (k1, k2),
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    44
      processed_specs = Symtab.merge (K true) (s1, s2)}
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    45
);
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    46
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    47
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    48
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    49
fun mk_data (c, k, s) = {ignore_consts = c, keep_functions = k, processed_specs = s}
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    50
fun map_data f {ignore_consts = c, keep_functions = k, processed_specs = s} = mk_data (f (c, k, s))
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    51
59057
5b649fb2f2e1 added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents: 57962
diff changeset
    52
fun ignore_consts cs =
5b649fb2f2e1 added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents: 57962
diff changeset
    53
  Data.map (map_data (@{apply 3(1)} (fold (fn c => Symtab.insert (op =) (c, ())) cs)))
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    54
59057
5b649fb2f2e1 added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents: 57962
diff changeset
    55
fun keep_functions cs =
5b649fb2f2e1 added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents: 57962
diff changeset
    56
  Data.map (map_data (@{apply 3(2)} (fold (fn c => Symtab.insert (op =) (c, ())) cs)))
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    57
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    58
fun keep_function thy = Symtab.defined (#keep_functions (Data.get thy))
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    59
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    60
fun processed_specs thy = Symtab.lookup (#processed_specs (Data.get thy))
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    61
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    62
fun store_processed_specs (constname, specs) =
59057
5b649fb2f2e1 added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents: 57962
diff changeset
    63
  Data.map (map_data (@{apply 3(3)} (Symtab.update_new (constname, specs))))
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
    64
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    65
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
    66
fun defining_term_of_introrule_term t =
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    67
  let
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    68
    val _ $ u = Logic.strip_imp_concl t
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
    69
  in fst (strip_comb u) end
55437
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
    70
(*
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    71
  in case pred of
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    72
    Const (c, T) => c
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    73
    | _ => raise TERM ("defining_const_of_introrule_term failed: Not a constant", [t])
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    74
  end
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
    75
*)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
    76
val defining_term_of_introrule = defining_term_of_introrule_term o prop_of
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    77
40142
128f8a1611e6 relaxing the filtering condition for getting specifications from Spec_Rules
bulwahn
parents: 40053
diff changeset
    78
fun defining_const_of_introrule th =
55437
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
    79
  (case defining_term_of_introrule th of
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
    80
    Const (c, _) => c
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
    81
  | _ => raise TERM ("defining_const_of_introrule failed: Not a constant", [prop_of th]))
40142
128f8a1611e6 relaxing the filtering condition for getting specifications from Spec_Rules
bulwahn
parents: 40053
diff changeset
    82
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    83
(*TODO*)
50056
72efd6b4038d dropped dead code
haftmann
parents: 49561
diff changeset
    84
fun is_introlike_term _ = true
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    85
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    86
val is_introlike = is_introlike_term o prop_of
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    87
56245
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 55440
diff changeset
    88
fun check_equation_format_term (t as (Const (@{const_name Pure.eq}, _) $ u $ _)) =
55437
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
    89
      (case strip_comb u of
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
    90
        (Const (_, T), args) =>
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
    91
          if (length (binder_types T) = length args) then
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
    92
            true
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
    93
          else
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
    94
            raise TERM ("check_equation_format_term failed: Number of arguments mismatch", [t])
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
    95
      | _ => raise TERM ("check_equation_format_term failed: Not a constant", [t]))
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    96
  | check_equation_format_term t =
55437
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
    97
      raise TERM ("check_equation_format_term failed: Not an equation", [t])
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    98
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    99
val check_equation_format = check_equation_format_term o prop_of
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   100
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
   101
56245
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 55440
diff changeset
   102
fun defining_term_of_equation_term (Const (@{const_name Pure.eq}, _) $ u $ _) = fst (strip_comb u)
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
   103
  | defining_term_of_equation_term t =
55437
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   104
      raise TERM ("defining_const_of_equation_term failed: Not an equation", [t])
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   105
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
   106
val defining_term_of_equation = defining_term_of_equation_term o prop_of
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
   107
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
   108
fun defining_const_of_equation th =
55437
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   109
  (case defining_term_of_equation th of
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   110
    Const (c, _) => c
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   111
  | _ => raise TERM ("defining_const_of_equation failed: Not a constant", [prop_of th]))
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
   112
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
   113
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
   114
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   115
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   116
(* Normalizing equations *)
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   117
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   118
fun mk_meta_equation th =
55437
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   119
  (case prop_of th of
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   120
    Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ _) =>
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   121
      th RS @{thm eq_reflection}
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   122
  | _ => th)
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   123
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   124
val meta_fun_cong = @{lemma "f == g ==> f x == g x" by simp}
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   125
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   126
fun full_fun_cong_expand th =
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   127
  let
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   128
    val (f, args) = strip_comb (fst (Logic.dest_equals (prop_of th)))
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   129
    val i = length (binder_types (fastype_of f)) - length args
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   130
  in funpow i (fn th => th RS meta_fun_cong) th end;
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   131
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   132
fun declare_names s xs ctxt =
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   133
  let
43329
84472e198515 tuned signature: Name.invent and Name.invent_names;
wenzelm
parents: 43277
diff changeset
   134
    val res = Name.invent_names ctxt s xs
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   135
  in (res, fold Name.declare (map fst res) ctxt) end
55437
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   136
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   137
fun split_all_pairs thy th =
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   138
  let
51552
c713c9505f68 clarified Skip_Proof.cheat_tac: more standard tactic;
wenzelm
parents: 50056
diff changeset
   139
    val ctxt = Proof_Context.init_global thy  (* FIXME proper context!? *)
50056
72efd6b4038d dropped dead code
haftmann
parents: 49561
diff changeset
   140
    val ((_, [th']), _) = Variable.import true [th] ctxt
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   141
    val t = prop_of th'
55437
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   142
    val frees = Term.add_frees t []
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   143
    val freenames = Term.add_free_names t []
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   144
    val nctxt = Name.make_context freenames
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   145
    fun mk_tuple_rewrites (x, T) nctxt =
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   146
      let
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   147
        val Ts = HOLogic.flatten_tupleT T
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   148
        val (xTs, nctxt') = declare_names x Ts nctxt
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   149
        val paths = HOLogic.flat_tupleT_paths T
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   150
      in ((Free (x, T), HOLogic.mk_ptuple paths T (map Free xTs)), nctxt') end
55437
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   151
    val (rewr, _) = fold_map mk_tuple_rewrites frees nctxt
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   152
    val t' = Pattern.rewrite_term thy rewr [] t
51552
c713c9505f68 clarified Skip_Proof.cheat_tac: more standard tactic;
wenzelm
parents: 50056
diff changeset
   153
    val th'' =
c713c9505f68 clarified Skip_Proof.cheat_tac: more standard tactic;
wenzelm
parents: 50056
diff changeset
   154
      Goal.prove ctxt (Term.add_free_names t' []) [] t'
c713c9505f68 clarified Skip_Proof.cheat_tac: more standard tactic;
wenzelm
parents: 50056
diff changeset
   155
        (fn _ => ALLGOALS Skip_Proof.cheat_tac)
35624
c4e29a0bb8c1 modernized structure Local_Defs;
wenzelm
parents: 35405
diff changeset
   156
    val th''' = Local_Defs.unfold ctxt [@{thm split_conv}, @{thm fst_conv}, @{thm snd_conv}] th''
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   157
  in
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   158
    th'''
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   159
  end;
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   160
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   161
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
   162
fun inline_equations thy th =
33404
66746e4b4531 adapted the inlining in the predicate compiler
bulwahn
parents: 33376
diff changeset
   163
  let
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   164
    val ctxt = Proof_Context.init_global thy
57962
0284a7d083be updated to named_theorems;
wenzelm
parents: 56245
diff changeset
   165
    val inline_defs = Named_Theorems.get ctxt @{named_theorems code_pred_inline}
0284a7d083be updated to named_theorems;
wenzelm
parents: 56245
diff changeset
   166
    val th' = Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimps inline_defs) th
55437
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   167
    (*val _ = print_step options
33404
66746e4b4531 adapted the inlining in the predicate compiler
bulwahn
parents: 33376
diff changeset
   168
      ("Inlining " ^ (Syntax.string_of_term_global thy (prop_of th))
66746e4b4531 adapted the inlining in the predicate compiler
bulwahn
parents: 33376
diff changeset
   169
       ^ "with " ^ (commas (map ((Syntax.string_of_term_global thy) o prop_of) inline_defs))
66746e4b4531 adapted the inlining in the predicate compiler
bulwahn
parents: 33376
diff changeset
   170
       ^" to " ^ (Syntax.string_of_term_global thy (prop_of th')))*)
66746e4b4531 adapted the inlining in the predicate compiler
bulwahn
parents: 33376
diff changeset
   171
  in
66746e4b4531 adapted the inlining in the predicate compiler
bulwahn
parents: 33376
diff changeset
   172
    th'
66746e4b4531 adapted the inlining in the predicate compiler
bulwahn
parents: 33376
diff changeset
   173
  end
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   174
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   175
fun normalize_equation thy th =
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   176
  mk_meta_equation th
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   177
  |> full_fun_cong_expand
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   178
  |> split_all_pairs thy
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   179
  |> tap check_equation_format
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   180
  |> inline_equations thy
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   181
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   182
fun normalize_intros thy th =
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   183
  split_all_pairs thy th
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   184
  |> inline_equations thy
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   185
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   186
fun normalize thy th =
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   187
  if is_equationlike th then
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   188
    normalize_equation thy th
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   189
  else
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   190
    normalize_intros thy th
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   191
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   192
fun get_specification options thy t =
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   193
  let
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   194
    (*val (c, T) = dest_Const t
51685
385ef6706252 more standard module name Axclass (according to file name);
wenzelm
parents: 51552
diff changeset
   195
    val t = Const (Axclass.unoverload_const thy (c, T), T)*)
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   196
    val _ = if show_steps options then
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   197
        tracing ("getting specification of " ^ Syntax.string_of_term_global thy t ^
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   198
          " with type " ^ Syntax.string_of_typ_global thy (fastype_of t))
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   199
      else ()
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 40142
diff changeset
   200
    val ctxt = Proof_Context.init_global thy
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
   201
    fun filtering th =
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
   202
      if is_equationlike th andalso
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   203
        defining_const_of_equation (normalize_equation thy th) = fst (dest_Const t) then
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
   204
        SOME (normalize_equation thy th)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
   205
      else
40142
128f8a1611e6 relaxing the filtering condition for getting specifications from Spec_Rules
bulwahn
parents: 40053
diff changeset
   206
        if is_introlike th andalso defining_const_of_introrule th = fst (dest_Const t) then
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
   207
          SOME th
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
   208
        else
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
   209
          NONE
35758
c029f85d3879 adopting predicate compiler to changes in Spec_Rules; removed dependency to Nitpick_Intros
bulwahn
parents: 35624
diff changeset
   210
    fun filter_defs ths = map_filter filtering (map (normalize thy o Thm.transfer thy) ths)
55437
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   211
    val spec =
57962
0284a7d083be updated to named_theorems;
wenzelm
parents: 56245
diff changeset
   212
      (case filter_defs (Named_Theorems.get ctxt @{named_theorems code_pred_def}) of
55437
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   213
        [] =>
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   214
          (case Spec_Rules.retrieve ctxt t of
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   215
            [] => error ("No specification for " ^ Syntax.string_of_term_global thy t)
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   216
          | ((_, (_, ths)) :: _) => filter_defs ths)
57962
0284a7d083be updated to named_theorems;
wenzelm
parents: 56245
diff changeset
   217
      | ths => ths)
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   218
    val _ =
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   219
      if show_intermediate_results options then
43277
1fd31f859fc7 pervasive Output operations;
wenzelm
parents: 42361
diff changeset
   220
        tracing ("Specification for " ^ (Syntax.string_of_term_global thy t) ^ ":\n" ^
40142
128f8a1611e6 relaxing the filtering condition for getting specifications from Spec_Rules
bulwahn
parents: 40053
diff changeset
   221
          commas (map (Display.string_of_thm_global thy) spec))
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   222
      else ()
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   223
  in
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   224
    spec
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   225
  end
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   226
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   227
val logic_operator_names =
56245
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 55440
diff changeset
   228
  [@{const_name Pure.eq},
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 55440
diff changeset
   229
   @{const_name Pure.imp},
38558
32ad17fe2b9c tuned quotes
haftmann
parents: 38549
diff changeset
   230
   @{const_name Trueprop},
32ad17fe2b9c tuned quotes
haftmann
parents: 38549
diff changeset
   231
   @{const_name Not},
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38795
diff changeset
   232
   @{const_name HOL.eq},
38786
e46e7a9cb622 formerly unnamed infix impliciation now named HOL.implies
haftmann
parents: 38731
diff changeset
   233
   @{const_name HOL.implies},
38558
32ad17fe2b9c tuned quotes
haftmann
parents: 38549
diff changeset
   234
   @{const_name All},
55437
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   235
   @{const_name Ex},
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   236
   @{const_name HOL.conj},
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   237
   @{const_name HOL.disj}]
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   238
55437
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   239
fun special_cases (c, _) =
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   240
  member (op =)
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   241
   [@{const_name Product_Type.Unity},
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   242
    @{const_name False},
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   243
    @{const_name Suc}, @{const_name Nat.zero_nat_inst.zero_nat},
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   244
    @{const_name Nat.one_nat_inst.one_nat},
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   245
    @{const_name Orderings.less}, @{const_name Orderings.less_eq},
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   246
    @{const_name Groups.zero},
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   247
    @{const_name Groups.one},  @{const_name Groups.plus},
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   248
    @{const_name Nat.ord_nat_inst.less_eq_nat},
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   249
    @{const_name Nat.ord_nat_inst.less_nat},
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   250
  (* FIXME
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   251
    @{const_name number_nat_inst.number_of_nat},
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   252
  *)
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   253
    @{const_name Num.Bit0},
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   254
    @{const_name Num.Bit1},
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   255
    @{const_name Num.One},
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   256
    @{const_name Int.zero_int_inst.zero_int},
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   257
    @{const_name List.filter},
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   258
    @{const_name HOL.If},
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   259
    @{const_name Groups.minus}] c
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   260
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   261
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
   262
fun obtain_specification_graph options thy t =
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   263
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 40142
diff changeset
   264
    val ctxt = Proof_Context.init_global thy
50056
72efd6b4038d dropped dead code
haftmann
parents: 49561
diff changeset
   265
    fun is_nondefining_const (c, _) = member (op =) logic_operator_names c
72efd6b4038d dropped dead code
haftmann
parents: 49561
diff changeset
   266
    fun has_code_pred_intros (c, _) = can (Core_Data.intros_of ctxt) c
55399
5c8e91f884af ported predicate compiler to 'ctr_sugar'
blanchet
parents: 51717
diff changeset
   267
    fun case_consts (c, _) = is_some (Ctr_Sugar.ctr_sugar_of_case ctxt c)
5c8e91f884af ported predicate compiler to 'ctr_sugar'
blanchet
parents: 51717
diff changeset
   268
    fun is_datatype_constructor (x as (_, T)) =
5c8e91f884af ported predicate compiler to 'ctr_sugar'
blanchet
parents: 51717
diff changeset
   269
      (case body_type T of
5c8e91f884af ported predicate compiler to 'ctr_sugar'
blanchet
parents: 51717
diff changeset
   270
        Type (Tcon, _) => can (Ctr_Sugar.dest_ctr ctxt Tcon) (Const x)
5c8e91f884af ported predicate compiler to 'ctr_sugar'
blanchet
parents: 51717
diff changeset
   271
      | _ => false)
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   272
    fun defiants_of specs =
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
   273
      fold (Term.add_consts o prop_of) specs []
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
   274
      |> filter_out is_datatype_constructor
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
   275
      |> filter_out is_nondefining_const
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   276
      |> filter_out has_code_pred_intros
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
   277
      |> filter_out case_consts
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   278
      |> filter_out special_cases
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   279
      |> filter_out (fn (c, _) => Symtab.defined (#ignore_consts (Data.get thy)) c)
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   280
      |> map (fn (c, _) => (c, Sign.the_const_constraint thy c))
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
   281
      |> map Const
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
   282
      (*
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33756
diff changeset
   283
      |> filter is_defining_constname*)
35405
fc130c5e81ec code simplification by inlining;
wenzelm
parents: 35404
diff changeset
   284
    fun extend t gr =
fc130c5e81ec code simplification by inlining;
wenzelm
parents: 35404
diff changeset
   285
      if can (Term_Graph.get_node gr) t then gr
fc130c5e81ec code simplification by inlining;
wenzelm
parents: 35404
diff changeset
   286
      else
fc130c5e81ec code simplification by inlining;
wenzelm
parents: 35404
diff changeset
   287
        let
fc130c5e81ec code simplification by inlining;
wenzelm
parents: 35404
diff changeset
   288
          val specs = get_specification options thy t
fc130c5e81ec code simplification by inlining;
wenzelm
parents: 35404
diff changeset
   289
          (*val _ = print_specification options thy constname specs*)
fc130c5e81ec code simplification by inlining;
wenzelm
parents: 35404
diff changeset
   290
          val us = defiants_of specs
fc130c5e81ec code simplification by inlining;
wenzelm
parents: 35404
diff changeset
   291
        in
fc130c5e81ec code simplification by inlining;
wenzelm
parents: 35404
diff changeset
   292
          gr
fc130c5e81ec code simplification by inlining;
wenzelm
parents: 35404
diff changeset
   293
          |> Term_Graph.new_node (t, specs)
fc130c5e81ec code simplification by inlining;
wenzelm
parents: 35404
diff changeset
   294
          |> fold extend us
fc130c5e81ec code simplification by inlining;
wenzelm
parents: 35404
diff changeset
   295
          |> fold (fn u => Term_Graph.add_edge (t, u)) us
fc130c5e81ec code simplification by inlining;
wenzelm
parents: 35404
diff changeset
   296
        end
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   297
  in
35405
fc130c5e81ec code simplification by inlining;
wenzelm
parents: 35404
diff changeset
   298
    extend t Term_Graph.empty
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   299
  end;
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   300
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   301
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   302
fun present_graph gr =
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   303
  let
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   304
    fun eq_cname (Const (c1, _), Const (c2, _)) = (c1 = c2)
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   305
    fun string_of_const (Const (c, _)) = c
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   306
      | string_of_const _ = error "string_of_const: unexpected term"
35404
de56579ae229 just one copy of structure Term_Graph (in Pure);
wenzelm
parents: 35324
diff changeset
   307
    val constss = Term_Graph.strong_conn gr;
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   308
    val mapping = Termtab.empty |> fold (fn consts => fold (fn const =>
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   309
      Termtab.update (const, consts)) consts) constss;
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   310
    fun succs consts = consts
44338
700008399ee5 refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
wenzelm
parents: 43329
diff changeset
   311
      |> maps (Term_Graph.immediate_succs gr)
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   312
      |> subtract eq_cname consts
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   313
      |> map (the o Termtab.lookup mapping)
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   314
      |> distinct (eq_list eq_cname);
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   315
    val conn = [] |> fold (fn consts => cons (consts, succs consts)) constss;
55437
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   316
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   317
    fun namify consts = map string_of_const consts
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   318
      |> commas;
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   319
    val prgr = map (fn (consts, constss) =>
55437
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   320
      {name = namify consts, ID = namify consts, dir = "", unfold = true,
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   321
       path = "", parents = map namify constss, content = [] }) conn
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   322
  in Graph_Display.display_graph prgr end
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   323
55437
3fd63b92ea3b tuned whitespace;
wenzelm
parents: 51717
diff changeset
   324
end