src/HOL/Tools/Quickcheck/abstract_generators.ML
author wenzelm
Wed, 06 Dec 2017 20:43:09 +0100
changeset 67149 e61557884799
parent 62979 1e527c40ae40
permissions -rw-r--r--
prefer control symbol antiquotations;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
45925
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
     1
(*  Title:      HOL/Tools/Quickcheck/abstract_generators.ML
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
     2
    Author:     Lukas Bulwahn, TU Muenchen
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
     3
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
     4
Quickcheck generators for abstract types.
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
     5
*)
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
     6
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
     7
signature ABSTRACT_GENERATORS =
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
     8
sig
46564
daa915508f63 adding parsing of an optional predicate with quickcheck_generator command
bulwahn
parents: 45940
diff changeset
     9
  val quickcheck_generator : string -> term option -> term list -> theory -> theory
62979
1e527c40ae40 misc tuning and standardization;
wenzelm
parents: 62391
diff changeset
    10
end
45925
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
    11
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
    12
structure Abstract_Generators : ABSTRACT_GENERATORS =
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
    13
struct
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
    14
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
    15
fun check_constrs ctxt tyco constrs =
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
    16
  let
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
    17
    fun check_type c =
62979
1e527c40ae40 misc tuning and standardization;
wenzelm
parents: 62391
diff changeset
    18
      (case try (dest_Type o body_type o fastype_of) c of
45925
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
    19
        NONE => error (Syntax.string_of_term ctxt c ^ " has mismatching result type.")
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
    20
      | SOME (tyco', vs) => if not (tyco' = tyco) then
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
    21
            error (Syntax.string_of_term ctxt c ^ " has mismatching result type. " ^
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
    22
              "Expected type constructor " ^ quote tyco ^ ", but found " ^ quote tyco' ^ ".")
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
    23
          else
62979
1e527c40ae40 misc tuning and standardization;
wenzelm
parents: 62391
diff changeset
    24
            (case try (map dest_TFree) vs of
45925
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
    25
              NONE => error (Syntax.string_of_term ctxt c ^ " has mismatching result type.")
62979
1e527c40ae40 misc tuning and standardization;
wenzelm
parents: 62391
diff changeset
    26
            | SOME _ => ()))
1e527c40ae40 misc tuning and standardization;
wenzelm
parents: 62391
diff changeset
    27
  in map check_type constrs end
1e527c40ae40 misc tuning and standardization;
wenzelm
parents: 62391
diff changeset
    28
46564
daa915508f63 adding parsing of an optional predicate with quickcheck_generator command
bulwahn
parents: 45940
diff changeset
    29
fun gen_quickcheck_generator prep_tyco prep_term tyco_raw opt_pred_raw constrs_raw thy =
45925
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
    30
  let
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
    31
    val ctxt = Proof_Context.init_global thy
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
    32
    val tyco = prep_tyco ctxt tyco_raw
46565
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46564
diff changeset
    33
    val opt_pred = Option.map (prep_term ctxt) opt_pred_raw
45925
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
    34
    val constrs = map (prep_term ctxt) constrs_raw
46565
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46564
diff changeset
    35
    val _ = check_constrs ctxt tyco constrs
45925
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
    36
    val vs = map dest_TFree (snd (dest_Type (body_type (fastype_of (hd constrs)))))
62979
1e527c40ae40 misc tuning and standardization;
wenzelm
parents: 62391
diff changeset
    37
    val name = Long_Name.base_name tyco
45925
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
    38
    fun mk_dconstrs (Const (s, T)) =
62979
1e527c40ae40 misc tuning and standardization;
wenzelm
parents: 62391
diff changeset
    39
          (s, map (Old_Datatype_Aux.dtyp_of_typ [(tyco, vs)]) (binder_types T))
1e527c40ae40 misc tuning and standardization;
wenzelm
parents: 62391
diff changeset
    40
      | mk_dconstrs t =
1e527c40ae40 misc tuning and standardization;
wenzelm
parents: 62391
diff changeset
    41
          error (Syntax.string_of_term ctxt t ^
1e527c40ae40 misc tuning and standardization;
wenzelm
parents: 62391
diff changeset
    42
            " is not a valid constructor for quickcheck_generator, which expects a constant.")
46564
daa915508f63 adding parsing of an optional predicate with quickcheck_generator command
bulwahn
parents: 45940
diff changeset
    43
    fun the_descr _ _ =
62979
1e527c40ae40 misc tuning and standardization;
wenzelm
parents: 62391
diff changeset
    44
      let val descr = [(0, (tyco, map Old_Datatype_Aux.DtTFree vs, map mk_dconstrs constrs))]
1e527c40ae40 misc tuning and standardization;
wenzelm
parents: 62391
diff changeset
    45
      in (descr, vs, [tyco], name, ([name], []), ([Type (tyco, map TFree vs)], [])) end
45940
71970a26a269 quickcheck_generator command also creates random generators
bulwahn
parents: 45925
diff changeset
    46
    fun ensure_sort (sort, instantiate) =
62979
1e527c40ae40 misc tuning and standardization;
wenzelm
parents: 62391
diff changeset
    47
      Quickcheck_Common.ensure_sort
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 62979
diff changeset
    48
        (((\<^sort>\<open>typerep\<close>, \<^sort>\<open>term_of\<close>), sort), (the_descr, instantiate))
62979
1e527c40ae40 misc tuning and standardization;
wenzelm
parents: 62391
diff changeset
    49
        Old_Datatype_Aux.default_config [tyco]
45925
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
    50
  in
45940
71970a26a269 quickcheck_generator command also creates random generators
bulwahn
parents: 45925
diff changeset
    51
    thy
62979
1e527c40ae40 misc tuning and standardization;
wenzelm
parents: 62391
diff changeset
    52
    |> ensure_sort
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 62979
diff changeset
    53
        (\<^sort>\<open>full_exhaustive\<close>, Exhaustive_Generators.instantiate_full_exhaustive_datatype)
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 62979
diff changeset
    54
    |> ensure_sort (\<^sort>\<open>exhaustive\<close>, Exhaustive_Generators.instantiate_exhaustive_datatype)
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 62979
diff changeset
    55
    |> ensure_sort (\<^sort>\<open>random\<close>, Random_Generators.instantiate_random_datatype)
62979
1e527c40ae40 misc tuning and standardization;
wenzelm
parents: 62391
diff changeset
    56
    |> (case opt_pred of NONE => I
1e527c40ae40 misc tuning and standardization;
wenzelm
parents: 62391
diff changeset
    57
       | SOME t => Context.theory_map (Quickcheck_Common.register_predicate (t, tyco)))
45925
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
    58
  end
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
    59
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
    60
val quickcheck_generator = gen_quickcheck_generator (K I) (K I)
62979
1e527c40ae40 misc tuning and standardization;
wenzelm
parents: 62391
diff changeset
    61
55951
c07d184aebe9 tuned signature -- more uniform check_type_name/read_type_name;
wenzelm
parents: 46961
diff changeset
    62
val quickcheck_generator_cmd =
c07d184aebe9 tuned signature -- more uniform check_type_name/read_type_name;
wenzelm
parents: 46961
diff changeset
    63
  gen_quickcheck_generator
56002
2028467b4df4 tuned signature;
wenzelm
parents: 55952
diff changeset
    64
    ((#1 o dest_Type) oo Proof_Context.read_type_name {proper = true, strict = true})
55951
c07d184aebe9 tuned signature -- more uniform check_type_name/read_type_name;
wenzelm
parents: 46961
diff changeset
    65
    Syntax.read_term
62979
1e527c40ae40 misc tuning and standardization;
wenzelm
parents: 62391
diff changeset
    66
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46949
diff changeset
    67
val _ =
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 62979
diff changeset
    68
  Outer_Syntax.command \<^command_keyword>\<open>quickcheck_generator\<close>
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46949
diff changeset
    69
    "define quickcheck generators for abstract types"
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46949
diff changeset
    70
    ((Parse.type_const --
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 62979
diff changeset
    71
      Scan.option (Args.$$$ "predicate" |-- \<^keyword>\<open>:\<close> |-- Parse.term)) --
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 62979
diff changeset
    72
      (Args.$$$ "constructors" |-- \<^keyword>\<open>:\<close> |-- Parse.list1 Parse.term)
46564
daa915508f63 adding parsing of an optional predicate with quickcheck_generator command
bulwahn
parents: 45940
diff changeset
    73
      >> (fn ((tyco, opt_pred), constrs) =>
daa915508f63 adding parsing of an optional predicate with quickcheck_generator command
bulwahn
parents: 45940
diff changeset
    74
        Toplevel.theory (quickcheck_generator_cmd tyco opt_pred constrs)))
45925
cd4243c025bb quickcheck generators for abstract types; tuned
bulwahn
parents:
diff changeset
    75
62979
1e527c40ae40 misc tuning and standardization;
wenzelm
parents: 62391
diff changeset
    76
end