src/HOL/Tools/quickcheck_generators.ML
author wenzelm
Fri, 12 Jun 2009 20:20:41 +0200
changeset 31581 907616b9536c
parent 31485 259a3c90016e
child 31595 bd2f7211a420
permissions -rw-r--r--
more isatests; macbroy23: nice at-sml-dev-e;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31260
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
     1
(* Author: Florian Haftmann, TU Muenchen
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
     2
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
     3
Quickcheck generators for various types.
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
     4
*)
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
     5
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
     6
signature QUICKCHECK_GENERATORS =
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
     7
sig
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
     8
  val compile_generator_expr: theory -> term -> int -> term list option
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
     9
  type seed = Random_Engine.seed
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    10
  val random_fun: typ -> typ -> ('a -> 'a -> bool) -> ('a -> term)
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    11
    -> (seed -> ('b * (unit -> term)) * seed) -> (seed -> seed * seed)
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    12
    -> seed -> (('a -> 'b) * (unit -> Term.term)) * seed
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    13
  val ensure_random_typecopy: string -> theory -> theory
31485
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
    14
  val random_aux_specification: string -> term list -> local_theory -> local_theory
31260
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    15
  val eval_ref: (unit -> int -> int * int -> term list option * (int * int)) option ref
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    16
  val setup: theory -> theory
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    17
end;
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    18
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    19
structure Quickcheck_Generators : QUICKCHECK_GENERATORS =
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    20
struct
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    21
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    22
(** building and compiling generator expressions **)
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    23
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    24
val eval_ref : (unit -> int -> int * int -> term list option * (int * int)) option ref = ref NONE;
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    25
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    26
val target = "Quickcheck";
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    27
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    28
fun mk_generator_expr thy prop tys =
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    29
  let
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    30
    val bound_max = length tys - 1;
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    31
    val bounds = map_index (fn (i, ty) =>
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    32
      (2 * (bound_max - i) + 1, 2 * (bound_max - i), 2 * i, ty)) tys;
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    33
    val result = list_comb (prop, map (fn (i, _, _, _) => Bound i) bounds);
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    34
    val terms = HOLogic.mk_list @{typ term} (map (fn (_, i, _, _) => Bound i $ @{term "()"}) bounds);
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    35
    val check = @{term "If :: bool => term list option => term list option => term list option"}
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    36
      $ result $ @{term "None :: term list option"} $ (@{term "Some :: term list => term list option "} $ terms);
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    37
    val return = @{term "Pair :: term list option => Random.seed => term list option * Random.seed"};
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    38
    fun liftT T sT = sT --> HOLogic.mk_prodT (T, sT);
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    39
    fun mk_termtyp ty = HOLogic.mk_prodT (ty, @{typ "unit => term"});
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    40
    fun mk_scomp T1 T2 sT f g = Const (@{const_name scomp},
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    41
      liftT T1 sT --> (T1 --> liftT T2 sT) --> liftT T2 sT) $ f $ g;
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    42
    fun mk_split ty = Sign.mk_const thy
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    43
      (@{const_name split}, [ty, @{typ "unit => term"}, liftT @{typ "term list option"} @{typ Random.seed}]);
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    44
    fun mk_scomp_split ty t t' =
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    45
      mk_scomp (mk_termtyp ty) @{typ "term list option"} @{typ Random.seed} t
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    46
        (mk_split ty $ Abs ("", ty, Abs ("", @{typ "unit => term"}, t')));
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    47
    fun mk_bindclause (_, _, i, ty) = mk_scomp_split ty
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    48
      (Sign.mk_const thy (@{const_name random}, [ty]) $ Bound i);
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    49
  in Abs ("n", @{typ code_numeral}, fold_rev mk_bindclause bounds (return $ check)) end;
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    50
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    51
fun compile_generator_expr thy t =
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    52
  let
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    53
    val tys = (map snd o fst o strip_abs) t;
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    54
    val t' = mk_generator_expr thy t tys;
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    55
    val f = Code_ML.eval (SOME target) ("Quickcheck_Generators.eval_ref", eval_ref)
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    56
      (fn proc => fn g => fn s => g s #>> (Option.map o map) proc) thy t' [];
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    57
  in f #> Random_Engine.run end;
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    58
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    59
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    60
(** typ "'a => 'b" **)
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    61
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    62
type seed = Random_Engine.seed;
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    63
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    64
fun random_fun (T1 : typ) (T2 : typ) (eq : 'a -> 'a -> bool) (term_of : 'a -> term)
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    65
    (random : seed -> ('b * (unit -> term)) * seed)
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    66
    (random_split : seed -> seed * seed)
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    67
    (seed : seed) =
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    68
  let
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    69
    val (seed', seed'') = random_split seed;
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    70
    val state = ref (seed', [], Const (@{const_name undefined}, T1 --> T2));
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    71
    val fun_upd = Const (@{const_name fun_upd},
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    72
      (T1 --> T2) --> T1 --> T2 --> T1 --> T2);
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    73
    fun random_fun' x =
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    74
      let
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    75
        val (seed, fun_map, f_t) = ! state;
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    76
      in case AList.lookup (uncurry eq) fun_map x
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    77
       of SOME y => y
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    78
        | NONE => let
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    79
              val t1 = term_of x;
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    80
              val ((y, t2), seed') = random seed;
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    81
              val fun_map' = (x, y) :: fun_map;
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    82
              val f_t' = fun_upd $ f_t $ t1 $ t2 ();
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    83
              val _ = state := (seed', fun_map', f_t');
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    84
            in y end
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    85
      end;
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    86
    fun term_fun' () = #3 (! state);
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    87
  in ((random_fun', term_fun'), seed'') end;
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    88
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    89
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    90
(** type copies **)
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    91
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    92
fun mk_random_typecopy tyco vs constr typ thy =
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    93
  let
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    94
    val Ts = map TFree vs;  
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    95
    val T = Type (tyco, Ts);
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    96
    fun mk_termifyT T = HOLogic.mk_prodT (T, @{typ "unit => term"})
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    97
    val Ttm = mk_termifyT T;
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    98
    val typtm = mk_termifyT typ;
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
    99
    fun mk_const c Ts = Const (c, Sign.const_instance thy (c, Ts));
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   100
    fun mk_random T = mk_const @{const_name random} [T];
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   101
    val size = @{term "j::code_numeral"};
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   102
    val v = "x";
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   103
    val t_v = Free (v, typtm);
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   104
    val t_constr = mk_const constr Ts;
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   105
    val lhs = mk_random T $ size;
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   106
    val rhs = HOLogic.mk_ST [(((mk_random typ) $ size, @{typ Random.seed}), SOME (v, typtm))]
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   107
      (HOLogic.mk_return Ttm @{typ Random.seed}
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   108
      (mk_const "Code_Eval.valapp" [typ, T]
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   109
        $ HOLogic.mk_prod (t_constr, Abs ("u", @{typ unit}, HOLogic.reflect_term t_constr)) $ t_v))
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   110
      @{typ Random.seed} (SOME Ttm, @{typ Random.seed});
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   111
    val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   112
  in   
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   113
    thy
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   114
    |> TheoryTarget.instantiation ([tyco], vs, @{sort random})
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   115
    |> `(fn lthy => Syntax.check_term lthy eq)
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   116
    |-> (fn eq => Specification.definition (NONE, (Attrib.empty_binding, eq)))
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   117
    |> snd
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   118
    |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   119
  end;
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   120
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   121
fun ensure_random_typecopy tyco thy =
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   122
  let
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   123
    val SOME { vs = raw_vs, constr, typ = raw_typ, ... } =
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   124
      TypecopyPackage.get_info thy tyco;
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   125
    val constrain = curry (Sorts.inter_sort (Sign.classes_of thy));
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   126
    val typ = map_atyps (fn TFree (v, sort) =>
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   127
      TFree (v, constrain sort @{sort random})) raw_typ;
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   128
    val vs' = Term.add_tfreesT typ [];
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   129
    val vs = map (fn (v, sort) =>
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   130
      (v, the_default (constrain sort @{sort typerep}) (AList.lookup (op =) vs' v))) raw_vs;
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   131
    val do_inst = Sign.of_sort thy (typ, @{sort random});
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   132
  in if do_inst then mk_random_typecopy tyco vs constr typ thy else thy end;
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   133
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   134
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   135
(** datatypes **)
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   136
31485
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   137
(* definitional scheme for random instances on datatypes *)
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   138
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   139
(*FIXME avoid this low-level proving*)
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   140
val rct = Thm.cprop_of @{thm random_aux_rec} |> Thm.dest_arg |> Thm.dest_arg
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   141
  |> Thm.dest_arg |> Thm.dest_arg1 |> Thm.dest_fun;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   142
fun dest_ctyp_nth k cT = nth (Thm.dest_ctyp cT) k;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   143
val aT = rct |> Thm.ctyp_of_term |> dest_ctyp_nth 1;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   144
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   145
fun random_aux_primrec eq lthy =
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   146
  let
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   147
    val thy = ProofContext.theory_of lthy;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   148
    val rews = map mk_meta_eq [@{thm code_numeral_zero_minus_one},
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   149
      @{thm Suc_code_numeral_minus_one}, @{thm select_weight_cons_zero}, @{thm beyond_zero}];
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   150
    val (rt as Free (random_aux, T)) $ (vt as Free (v, _)) =
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   151
      (fst o HOLogic.dest_eq o HOLogic.dest_Trueprop) eq;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   152
    val Type (_, [_, iT]) = T;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   153
    val icT = Thm.ctyp_of thy iT;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   154
    fun subst_v t' = map_aterms (fn t as Free (w, _) => if v = w then t' else t | t => t);
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   155
    val eqs0 = [subst_v @{term "0::code_numeral"} eq, subst_v (@{term "Suc_code_numeral"} $ vt) eq];
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   156
    val eqs1 = map (Pattern.rewrite_term thy (map (Logic.dest_equals o Thm.prop_of) rews) []) eqs0
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   157
    val ((_, eqs2), lthy') = PrimrecPackage.add_primrec_simple
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   158
      [((Binding.name random_aux, T), NoSyn)] eqs1 lthy;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   159
    val eq_tac = ALLGOALS Goal.conjunction_tac THEN ALLGOALS (simp_tac (HOL_ss addsimps rews))
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   160
      THEN (ALLGOALS (ProofContext.fact_tac (flat eqs2)));
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   161
    val eqs3 = Goal.prove_multi lthy' [v] [] eqs0 (K eq_tac);
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   162
    val rct' = Thm.instantiate_cterm ([(aT, icT)], []) rct
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   163
    val rule = @{thm random_aux_rec}
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   164
      |> Drule.instantiate ([(aT, icT)], [(rct', Thm.cterm_of thy rt)])
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   165
      |> (fn thm => thm OF eqs3)
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   166
    val tac = ALLGOALS (rtac rule);
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   167
    val simp = Goal.prove lthy' [v] [] eq (K tac);
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   168
  in (simp, lthy') end;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   169
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   170
fun random_aux_primrec_multi prefix [eq] lthy =
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   171
      lthy
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   172
      |> random_aux_primrec eq
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   173
      |>> (fn simp => [simp])
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   174
  | random_aux_primrec_multi prefix (eqs as _ :: _ :: _) lthy =
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   175
      let
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   176
        val thy = ProofContext.theory_of lthy;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   177
        val (lhss, rhss) = map_split (HOLogic.dest_eq o HOLogic.dest_Trueprop) eqs;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   178
        val (vs, (arg as Free (v, _)) :: _) = map_split (fn (t1 $ t2) => (t1, t2)) lhss;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   179
        val Ts = map fastype_of lhss;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   180
        val tupleT = foldr1 HOLogic.mk_prodT Ts;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   181
        val aux_lhs = Free ("mutual_" ^ prefix, fastype_of arg --> tupleT) $ arg;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   182
        val aux_eq = (HOLogic.mk_Trueprop o HOLogic.mk_eq)
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   183
          (aux_lhs, foldr1 HOLogic.mk_prod rhss);
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   184
        fun mk_proj t [T] = [t]
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   185
          | mk_proj t (Ts as T :: (Ts' as _ :: _)) =
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   186
              Const (@{const_name fst}, foldr1 HOLogic.mk_prodT Ts --> T) $ t
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   187
                :: mk_proj (Const (@{const_name snd},
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   188
                  foldr1 HOLogic.mk_prodT Ts --> foldr1 HOLogic.mk_prodT Ts') $ t) Ts';
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   189
        val projs = mk_proj (aux_lhs) Ts;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   190
        val proj_eqs = map2 (fn v => fn proj => (v, lambda arg proj)) vs projs;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   191
        val proj_defs = map2 (fn Free (name, _) => fn (_, rhs) =>
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   192
          ((Binding.name name, NoSyn), (Attrib.empty_binding, rhs))) vs proj_eqs;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   193
        val aux_eq' = Pattern.rewrite_term thy proj_eqs [] aux_eq;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   194
        fun prove_eqs aux_simp proj_defs lthy = 
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   195
          let
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   196
            val proj_simps = map (snd o snd) proj_defs;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   197
            fun tac { context = ctxt, ... } = ALLGOALS Goal.conjunction_tac
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   198
              THEN ALLGOALS (simp_tac (HOL_ss addsimps proj_simps))
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   199
              THEN ALLGOALS (EqSubst.eqsubst_tac ctxt [0] [aux_simp])
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   200
              THEN ALLGOALS (simp_tac (HOL_ss addsimps [fst_conv, snd_conv]));
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   201
          in (Goal.prove_multi lthy [v] [] eqs tac, lthy) end;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   202
      in
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   203
        lthy
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   204
        |> random_aux_primrec aux_eq'
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   205
        ||>> fold_map (LocalTheory.define Thm.definitionK) proj_defs
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   206
        |-> (fn (aux_simp, proj_defs) => prove_eqs aux_simp proj_defs)
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   207
      end;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   208
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   209
fun random_aux_specification prefix eqs lthy =
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   210
  let
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   211
    val _ $ Free (v, _) $ Free (w, _) =
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   212
      (fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o hd) eqs;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   213
    fun mk_proto_eq eq =
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   214
      let
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   215
        val (head $ arg, rhs) = (HOLogic.dest_eq o HOLogic.dest_Trueprop) eq;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   216
      in ((HOLogic.mk_Trueprop o HOLogic.mk_eq) (head, lambda arg rhs)) end;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   217
    val proto_eqs = map mk_proto_eq eqs;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   218
    fun prove_simps proto_simps lthy =
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   219
      let
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   220
        val ext_simps = map (fn thm => fun_cong OF [thm]) proto_simps;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   221
        val tac = ALLGOALS Goal.conjunction_tac
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   222
          THEN ALLGOALS (ProofContext.fact_tac ext_simps);
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   223
      in (Goal.prove_multi lthy [v, w] [] eqs (K tac), lthy) end;
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   224
    val b = Binding.qualify true prefix (Binding.name "simps");
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   225
  in
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   226
    lthy
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   227
    |> random_aux_primrec_multi prefix proto_eqs
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   228
    |-> (fn proto_simps => prove_simps proto_simps)
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   229
    |-> (fn simps => LocalTheory.note Thm.generatedK ((b,
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   230
           Code.add_default_eqn_attrib :: map (Attrib.internal o K)
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   231
          [Simplifier.simp_add, Nitpick_Const_Simp_Thms.add, Quickcheck_RecFun_Simp_Thms.add]),
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   232
            simps))
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   233
    |> snd
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   234
  end
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   235
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   236
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   237
(* constructing random instances on datatypes *)
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   238
259a3c90016e added infrastructure for definitorial construction of generators for datatypes
haftmann
parents: 31260
diff changeset
   239
(*still under construction*)
31260
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   240
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   241
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   242
(** setup **)
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   243
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   244
val setup = Code_Target.extend_target (target, (Code_ML.target_Eval, K I))
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   245
  #> Quickcheck.add_generator ("code", compile_generator_expr o ProofContext.theory_of)
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   246
  #> TypecopyPackage.interpretation ensure_random_typecopy;
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   247
4d273d043d59 separate module for quickcheck generators
haftmann
parents:
diff changeset
   248
end;