src/HOL/Tools/Quickcheck/quickcheck_common.ML
author bulwahn
Wed, 30 Mar 2011 09:44:16 +0200
changeset 42159 234ec7011e5d
parent 42110 17e0cd9bc259
child 42195 1e7b62c93f5d
permissions -rw-r--r--
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41927
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
     1
(*  Title:      HOL/Tools/Quickcheck/quickcheck_common.ML
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
     2
    Author:     Florian Haftmann, Lukas Bulwahn, TU Muenchen
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
     3
41938
645cca858c69 tuned headers;
wenzelm
parents: 41935
diff changeset
     4
Common functions for quickcheck's generators.
41927
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
     5
*)
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
     6
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
     7
signature QUICKCHECK_COMMON =
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
     8
sig
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
     9
  val perhaps_constrain: theory -> (typ * sort) list -> (string * sort) list
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    10
    -> (string * sort -> string * sort) option
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    11
  val ensure_sort_datatype:
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    12
    sort * (Datatype.config -> Datatype.descr -> (string * sort) list -> string list -> string ->
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    13
      string list * string list -> typ list * typ list -> theory -> theory)
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    14
    -> Datatype.config -> string list -> theory -> theory
42159
234ec7011e5d generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents: 42110
diff changeset
    15
  val gen_mk_parametric_generator_expr :
234ec7011e5d generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents: 42110
diff changeset
    16
   (((Proof.context -> term * term list -> term) * term) * typ) -> Proof.context -> (term * term list) list -> term
41935
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    17
  val post_process_term : term -> term
41927
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    18
end;
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    19
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    20
structure Quickcheck_Common : QUICKCHECK_COMMON =
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    21
struct
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    22
41935
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    23
(** ensuring sort constraints **)
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    24
41927
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    25
fun perhaps_constrain thy insts raw_vs =
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    26
  let
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    27
    fun meet (T, sort) = Sorts.meet_sort (Sign.classes_of thy) 
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    28
      (Logic.varifyT_global T, sort);
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    29
    val vtab = Vartab.empty
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    30
      |> fold (fn (v, sort) => Vartab.update ((v, 0), sort)) raw_vs
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    31
      |> fold meet insts;
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    32
  in SOME (fn (v, _) => (v, (the o Vartab.lookup vtab) (v, 0)))
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    33
  end handle Sorts.CLASS_ERROR _ => NONE;
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    34
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    35
fun ensure_sort_datatype (sort, instantiate_datatype) config raw_tycos thy =
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    36
  let
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    37
    val algebra = Sign.classes_of thy;
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    38
    val (descr, raw_vs, tycos, prfx, (names, auxnames), raw_TUs) =
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    39
      Datatype.the_descr thy raw_tycos;
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    40
    val typerep_vs = (map o apsnd)
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    41
      (curry (Sorts.inter_sort algebra) @{sort typerep}) raw_vs;
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    42
    val sort_insts = (map (rpair sort) o flat o maps snd o maps snd)
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    43
      (Datatype_Aux.interpret_construction descr typerep_vs
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    44
        { atyp = single, dtyp = (K o K o K) [] });
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    45
    val term_of_insts = (map (rpair @{sort term_of}) o flat o maps snd o maps snd)
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    46
      (Datatype_Aux.interpret_construction descr typerep_vs
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    47
        { atyp = K [], dtyp = K o K });
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    48
    val has_inst = exists (fn tyco =>
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    49
      can (Sorts.mg_domain algebra tyco) sort) tycos;
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    50
  in if has_inst then thy
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    51
    else case perhaps_constrain thy (sort_insts @ term_of_insts) typerep_vs
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    52
     of SOME constrain => instantiate_datatype config descr
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    53
          (map constrain typerep_vs) tycos prfx (names, auxnames)
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    54
            ((pairself o map o map_atyps) (fn TFree v => TFree (constrain v)) raw_TUs) thy
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    55
      | NONE => thy
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    56
  end;
41935
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    57
  
42159
234ec7011e5d generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents: 42110
diff changeset
    58
(** generic parametric compilation **)
234ec7011e5d generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents: 42110
diff changeset
    59
234ec7011e5d generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents: 42110
diff changeset
    60
fun gen_mk_parametric_generator_expr ((mk_generator_expr, out_of_bounds), T) ctxt ts =
234ec7011e5d generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents: 42110
diff changeset
    61
  let
234ec7011e5d generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents: 42110
diff changeset
    62
    val if_t = Const (@{const_name "If"}, @{typ bool} --> T --> T --> T)
234ec7011e5d generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents: 42110
diff changeset
    63
    fun mk_if (index, (t, eval_terms)) else_t =
234ec7011e5d generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents: 42110
diff changeset
    64
      if_t $ (HOLogic.eq_const @{typ code_numeral} $ Bound 0 $ HOLogic.mk_number @{typ code_numeral} index) $
234ec7011e5d generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents: 42110
diff changeset
    65
        (mk_generator_expr ctxt (t, eval_terms)) $ else_t
234ec7011e5d generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents: 42110
diff changeset
    66
  in
234ec7011e5d generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents: 42110
diff changeset
    67
    absdummy (@{typ "code_numeral"}, fold_rev mk_if (1 upto (length ts) ~~ ts) out_of_bounds)
234ec7011e5d generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents: 42110
diff changeset
    68
  end
234ec7011e5d generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents: 42110
diff changeset
    69
41935
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    70
(** post-processing of function terms **)
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    71
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    72
fun dest_fun_upd (Const (@{const_name fun_upd}, _) $ t0 $ t1 $ t2) = (t0, (t1, t2))
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    73
  | dest_fun_upd t = raise TERM ("dest_fun_upd", [t])
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    74
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    75
fun mk_fun_upd T1 T2 (t1, t2) t = 
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    76
  Const (@{const_name fun_upd}, (T1 --> T2) --> T1 --> T2 --> T1 --> T2) $ t $ t1 $ t2
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    77
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    78
fun dest_fun_upds t =
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    79
  case try dest_fun_upd t of
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    80
    NONE =>
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    81
      (case t of
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    82
        Abs (_, _, _) => ([], t) 
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    83
      | _ => raise TERM ("dest_fun_upds", [t]))
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    84
  | SOME (t0, (t1, t2)) => apfst (cons (t1, t2)) (dest_fun_upds t0)
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    85
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    86
fun make_fun_upds T1 T2 (tps, t) = fold_rev (mk_fun_upd T1 T2) tps t
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    87
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    88
fun make_set T1 [] = Const (@{const_abbrev Set.empty}, T1 --> @{typ bool})
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    89
  | make_set T1 ((_, @{const False}) :: tps) = make_set T1 tps
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    90
  | make_set T1 ((t1, @{const True}) :: tps) =
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    91
    Const (@{const_name insert}, T1 --> (T1 --> @{typ bool}) --> T1 --> @{typ bool})
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    92
      $ t1 $ (make_set T1 tps)
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    93
  | make_set T1 ((_, t) :: tps) = raise TERM ("make_set", [t])
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    94
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    95
fun make_coset T [] = Const (@{const_abbrev UNIV}, T --> @{typ bool})
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    96
  | make_coset T tps = 
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    97
    let
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    98
      val U = T --> @{typ bool}
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    99
      fun invert @{const False} = @{const True}
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   100
        | invert @{const True} = @{const False}
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   101
    in
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   102
      Const (@{const_name "Groups.minus_class.minus"}, U --> U --> U)
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   103
        $ Const (@{const_abbrev UNIV}, U) $ make_set T (map (apsnd invert) tps)
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   104
    end
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   105
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   106
fun make_map T1 T2 [] = Const (@{const_abbrev Map.empty}, T1 --> T2)
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   107
  | make_map T1 T2 ((_, Const (@{const_name None}, _)) :: tps) = make_map T1 T2 tps
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   108
  | make_map T1 T2 ((t1, t2) :: tps) = mk_fun_upd T1 T2 (t1, t2) (make_map T1 T2 tps)
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   109
  
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   110
fun post_process_term t =
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   111
  let
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   112
    fun map_Abs f t =
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   113
      case t of Abs (x, T, t') => Abs (x, T, f t') | _ => raise TERM ("map_Abs", [t]) 
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   114
    fun process_args t = case strip_comb t of
42110
17e0cd9bc259 fixed postprocessing for term presentation in quickcheck; tuned spacing
bulwahn
parents: 41938
diff changeset
   115
      (c as Const (_, _), ts) => list_comb (c, map post_process_term ts)
41935
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   116
  in
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   117
    case fastype_of t of
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   118
      Type (@{type_name fun}, [T1, T2]) =>
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   119
        (case try dest_fun_upds t of
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   120
          SOME (tps, t) =>
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   121
            (map (pairself post_process_term) tps, map_Abs post_process_term t)
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   122
            |> (case T2 of
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   123
              @{typ bool} => 
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   124
                (case t of
42110
17e0cd9bc259 fixed postprocessing for term presentation in quickcheck; tuned spacing
bulwahn
parents: 41938
diff changeset
   125
                   Abs(_, _, @{const False}) => fst #> rev #> make_set T1
17e0cd9bc259 fixed postprocessing for term presentation in quickcheck; tuned spacing
bulwahn
parents: 41938
diff changeset
   126
                 | Abs(_, _, @{const True}) => fst #> rev #> make_coset T1
41935
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   127
                 | Abs(_, _, Const (@{const_name undefined}, _)) => fst #> rev #> make_set T1
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   128
                 | _ => raise TERM ("post_process_term", [t]))
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   129
            | Type (@{type_name option}, _) =>
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   130
                (case t of
42110
17e0cd9bc259 fixed postprocessing for term presentation in quickcheck; tuned spacing
bulwahn
parents: 41938
diff changeset
   131
                  Abs(_, _, Const (@{const_name None}, _)) => fst #> make_map T1 T2
41935
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   132
                | Abs(_, _, Const (@{const_name undefined}, _)) => fst #> make_map T1 T2
42110
17e0cd9bc259 fixed postprocessing for term presentation in quickcheck; tuned spacing
bulwahn
parents: 41938
diff changeset
   133
                | _ => make_fun_upds T1 T2)
41935
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   134
            | _ => make_fun_upds T1 T2)
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   135
        | NONE => process_args t)
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   136
    | _ => process_args t
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   137
  end
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   138
  
41927
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
   139
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
   140
end;