src/HOL/Tools/Quickcheck/narrowing_generators.ML
author bulwahn
Tue, 14 Jun 2011 08:30:19 +0200
changeset 43379 8c4b383e5143
parent 43329 84472e198515
child 43585 ea959ab7bbe3
permissions -rw-r--r--
quickcheck_narrowing returns some timing information
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41930
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
     1
(*  Title:      HOL/Tools/Quickcheck/narrowing_generators.ML
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
     2
    Author:     Lukas Bulwahn, TU Muenchen
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
     3
41938
645cca858c69 tuned headers;
wenzelm
parents: 41936
diff changeset
     4
Narrowing-based counterexample generation.
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
     5
*)
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
     6
41930
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
     7
signature NARROWING_GENERATORS =
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
     8
sig
43237
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
     9
  val allow_existentials : bool Config.T
42023
1bd4b6d1c482 adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
bulwahn
parents: 42020
diff changeset
    10
  val finite_functions : bool Config.T
43079
4022892a2f28 improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
bulwahn
parents: 43047
diff changeset
    11
  val overlord : bool Config.T
43237
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
    12
  val test_term: Proof.context -> bool * bool -> term * term list -> Quickcheck.result
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
    13
  datatype counterexample = Universal_Counterexample of (term * counterexample)
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
    14
    | Existential_Counterexample of (term * counterexample) list
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
    15
    | Empty_Assignment
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
    16
  val put_counterexample: (unit -> term list option) -> Proof.context -> Proof.context
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
    17
  val put_existential_counterexample : (unit -> counterexample option) -> Proof.context -> Proof.context
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    18
  val setup: theory -> theory
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    19
end;
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    20
41930
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
    21
structure Narrowing_Generators : NARROWING_GENERATORS =
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    22
struct
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
    23
42023
1bd4b6d1c482 adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
bulwahn
parents: 42020
diff changeset
    24
(* configurations *)
1bd4b6d1c482 adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
bulwahn
parents: 42020
diff changeset
    25
43237
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
    26
val allow_existentials = Attrib.setup_config_bool @{binding quickcheck_allow_existentials} (K true)
42616
92715b528e78 added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents: 42361
diff changeset
    27
val finite_functions = Attrib.setup_config_bool @{binding quickcheck_finite_functions} (K true)
43047
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    28
val overlord = Attrib.setup_config_bool @{binding quickcheck_narrowing_overlord} (K false)
42023
1bd4b6d1c482 adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
bulwahn
parents: 42020
diff changeset
    29
43047
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    30
(* partial_term_of instances *)
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    31
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    32
fun mk_partial_term_of (x, T) =
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    33
  Const (@{const_name Quickcheck_Narrowing.partial_term_of_class.partial_term_of},
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    34
    Term.itselfT T --> @{typ narrowing_term} --> @{typ Code_Evaluation.term})
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    35
      $ Const ("TYPE", Term.itselfT T) $ x
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    36
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    37
(** formal definition **)
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    38
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    39
fun add_partial_term_of tyco raw_vs thy =
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    40
  let
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    41
    val vs = map (fn (v, _) => (v, @{sort typerep})) raw_vs;
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    42
    val ty = Type (tyco, map TFree vs);
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    43
    val lhs = Const (@{const_name partial_term_of},
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    44
        Term.itselfT ty --> @{typ narrowing_term} --> @{typ Code_Evaluation.term})
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    45
      $ Free ("x", Term.itselfT ty) $ Free ("t", @{typ narrowing_term});
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    46
    val rhs = @{term "undefined :: Code_Evaluation.term"};
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    47
    val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    48
    fun triv_name_of t = (fst o dest_Free o fst o strip_comb o fst
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    49
      o HOLogic.dest_eq o HOLogic.dest_Trueprop) t ^ "_triv";
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    50
  in
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    51
    thy
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    52
    |> Class.instantiation ([tyco], vs, @{sort partial_term_of})
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    53
    |> `(fn lthy => Syntax.check_term lthy eq)
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    54
    |-> (fn eq => Specification.definition (NONE, ((Binding.name (triv_name_of eq), []), eq)))
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    55
    |> snd
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    56
    |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    57
  end;
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    58
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    59
fun ensure_partial_term_of (tyco, (raw_vs, _)) thy =
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    60
  let
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    61
    val need_inst = not (can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort partial_term_of})
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    62
      andalso can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort typerep};
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    63
  in if need_inst then add_partial_term_of tyco raw_vs thy else thy end;
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    64
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    65
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    66
(** code equations for datatypes **)
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    67
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    68
fun mk_partial_term_of_eq thy ty (i, (c, (_, tys))) =
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    69
  let
43329
84472e198515 tuned signature: Name.invent and Name.invent_names;
wenzelm
parents: 43317
diff changeset
    70
    val frees = map Free (Name.invent_names Name.context "a" (map (K @{typ narrowing_term}) tys))
43047
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    71
    val narrowing_term = @{term "Quickcheck_Narrowing.Ctr"} $ HOLogic.mk_number @{typ code_int} i
43079
4022892a2f28 improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
bulwahn
parents: 43047
diff changeset
    72
      $ (HOLogic.mk_list @{typ narrowing_term} (rev frees))
43047
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    73
    val rhs = fold (fn u => fn t => @{term "Code_Evaluation.App"} $ t $ u)
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    74
        (map mk_partial_term_of (frees ~~ tys))
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    75
        (@{term "Code_Evaluation.Const"} $ HOLogic.mk_literal c $ HOLogic.mk_typerep (tys ---> ty))
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    76
    val insts =
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    77
      map (SOME o Thm.cterm_of thy o map_types Logic.unvarifyT_global o Logic.varify_global)
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    78
        [Free ("ty", Term.itselfT ty), narrowing_term, rhs]
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    79
    val cty = Thm.ctyp_of thy ty;
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    80
  in
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    81
    @{thm partial_term_of_anything}
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    82
    |> Drule.instantiate' [SOME cty] insts
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    83
    |> Thm.varifyT_global
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    84
  end
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    85
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    86
fun add_partial_term_of_code tyco raw_vs raw_cs thy =
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    87
  let
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    88
    val algebra = Sign.classes_of thy;
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    89
    val vs = map (fn (v, sort) =>
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    90
      (v, curry (Sorts.inter_sort algebra) @{sort typerep} sort)) raw_vs;
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    91
    val ty = Type (tyco, map TFree vs);
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    92
    val cs = (map o apsnd o apsnd o map o map_atyps)
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    93
      (fn TFree (v, _) => TFree (v, (the o AList.lookup (op =) vs) v)) raw_cs;
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
    94
    val const = AxClass.param_of_inst thy (@{const_name partial_term_of}, tyco);
43079
4022892a2f28 improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
bulwahn
parents: 43047
diff changeset
    95
    val var_insts = map (SOME o Thm.cterm_of thy o map_types Logic.unvarifyT_global o Logic.varify_global)
4022892a2f28 improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
bulwahn
parents: 43047
diff changeset
    96
        [Free ("ty", Term.itselfT ty), @{term "Quickcheck_Narrowing.Var p tt"},
4022892a2f28 improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
bulwahn
parents: 43047
diff changeset
    97
          @{term "Code_Evaluation.Free (STR ''_'')"} $ HOLogic.mk_typerep ty]
4022892a2f28 improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
bulwahn
parents: 43047
diff changeset
    98
    val var_eq =
4022892a2f28 improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
bulwahn
parents: 43047
diff changeset
    99
      @{thm partial_term_of_anything}
4022892a2f28 improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
bulwahn
parents: 43047
diff changeset
   100
      |> Drule.instantiate' [SOME (Thm.ctyp_of thy ty)] var_insts
4022892a2f28 improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
bulwahn
parents: 43047
diff changeset
   101
      |> Thm.varifyT_global
4022892a2f28 improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
bulwahn
parents: 43047
diff changeset
   102
    val eqs = var_eq :: map_index (mk_partial_term_of_eq thy ty) cs;
43047
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
   103
 in
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
   104
    thy
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
   105
    |> Code.del_eqns const
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
   106
    |> fold Code.add_eqn eqs
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
   107
  end;
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
   108
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
   109
fun ensure_partial_term_of_code (tyco, (raw_vs, cs)) thy =
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
   110
  let
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
   111
    val has_inst = can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort partial_term_of};
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
   112
  in if has_inst then add_partial_term_of_code tyco raw_vs cs thy else thy end;
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
   113
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
   114
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
   115
(* narrowing generators *)
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
   116
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
   117
(** narrowing specific names and types **)
41963
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   118
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   119
exception FUNCTION_TYPE;
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   120
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   121
val narrowingN = "narrowing";
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   122
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   123
fun narrowingT T =
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   124
  @{typ Quickcheck_Narrowing.code_int} --> Type (@{type_name Quickcheck_Narrowing.cons}, [T])
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   125
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   126
fun mk_empty T = Const (@{const_name Quickcheck_Narrowing.empty}, narrowingT T)
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   127
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   128
fun mk_cons c T = Const (@{const_name Quickcheck_Narrowing.cons}, T --> narrowingT T) $ Const (c, T)
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   129
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   130
fun mk_apply (T, t) (U, u) =
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   131
  let
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   132
    val (_, U') = dest_funT U
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   133
  in
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   134
    (U', Const (@{const_name Quickcheck_Narrowing.apply},
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   135
      narrowingT U --> narrowingT T --> narrowingT U') $ u $ t)
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   136
  end
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   137
  
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   138
fun mk_sum (t, u) =
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   139
  let
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   140
    val T = fastype_of t
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   141
  in
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   142
    Const (@{const_name Quickcheck_Narrowing.sum}, T --> T --> T) $ t $ u
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   143
  end
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   144
43047
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
   145
(** deriving narrowing instances **)
41963
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   146
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   147
fun mk_equations descr vs tycos narrowings (Ts, Us) =
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   148
  let
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   149
    fun mk_call T =
42214
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42184
diff changeset
   150
      (T, Const (@{const_name "Quickcheck_Narrowing.narrowing_class.narrowing"}, narrowingT T))
41963
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   151
    fun mk_aux_call fTs (k, _) (tyco, Ts) =
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   152
      let
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   153
        val T = Type (tyco, Ts)
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   154
        val _ = if not (null fTs) then raise FUNCTION_TYPE else ()
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   155
      in
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   156
        (T, nth narrowings k)
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   157
      end
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   158
    fun mk_consexpr simpleT (c, xs) =
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   159
      let
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   160
        val Ts = map fst xs
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   161
      in snd (fold mk_apply xs (Ts ---> simpleT, mk_cons c (Ts ---> simpleT))) end
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   162
    fun mk_rhs exprs = foldr1 mk_sum exprs
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   163
    val rhss =
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   164
      Datatype_Aux.interpret_construction descr vs
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   165
        { atyp = mk_call, dtyp = mk_aux_call }
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   166
      |> (map o apfst) Type
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   167
      |> map (fn (T, cs) => map (mk_consexpr T) cs)
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   168
      |> map mk_rhs
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   169
    val lhss = narrowings
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   170
    val eqs = map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (lhss ~~ rhss)
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   171
  in
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   172
    eqs
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   173
  end
43237
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   174
    
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   175
fun contains_recursive_type_under_function_types xs =
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   176
  exists (fn (_, (_, _, cs)) => cs |> exists (snd #> exists (fn dT =>
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   177
    (case Datatype_Aux.strip_dtyp dT of (_ :: _, Datatype.DtRec _) => true | _ => false)))) xs
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   178
41963
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   179
fun instantiate_narrowing_datatype config descr vs tycos prfx (names, auxnames) (Ts, Us) thy =
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   180
  let
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   181
    val _ = Datatype_Aux.message config "Creating narrowing generators ...";
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   182
    val narrowingsN = map (prefix (narrowingN ^ "_")) (names @ auxnames);
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   183
  in
43237
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   184
    if not (contains_recursive_type_under_function_types descr) then
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   185
      thy
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   186
      |> Class.instantiation (tycos, vs, @{sort narrowing})
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   187
      |> Quickcheck_Common.define_functions
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   188
        (fn narrowings => mk_equations descr vs tycos narrowings (Ts, Us), NONE)
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   189
        prfx [] narrowingsN (map narrowingT (Ts @ Us))
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   190
      |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   191
    else
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   192
      thy
41963
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   193
  end;
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   194
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   195
(* testing framework *)
d8c3b26b3da4 tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents: 41953
diff changeset
   196
43308
fd6cc1378fec compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
bulwahn
parents: 43240
diff changeset
   197
val target = "Haskell_Quickcheck"
41930
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
   198
43047
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
   199
(** invocation of Haskell interpreter **)
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   200
41930
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
   201
val narrowing_engine = File.read (Path.explode "~~/src/HOL/Tools/Quickcheck/Narrowing_Engine.hs")
43237
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   202
val pnf_narrowing_engine = File.read (Path.explode "~~/src/HOL/Tools/Quickcheck/PNF_Narrowing_Engine.hs")
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   203
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   204
fun exec verbose code =
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   205
  ML_Context.exec (fn () => Secure.use_text ML_Env.local_context (0, "generated code") verbose code)
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   206
43079
4022892a2f28 improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
bulwahn
parents: 43047
diff changeset
   207
fun with_overlord_dir name f =
43047
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
   208
  let
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
   209
    val path = Path.append (Path.explode "~/.isabelle") (Path.basic (name ^ serial_string ()))
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
   210
    val _ = Isabelle_System.mkdirs path;
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
   211
  in Exn.release (Exn.capture f path) end;
43379
8c4b383e5143 quickcheck_narrowing returns some timing information
bulwahn
parents: 43329
diff changeset
   212
8c4b383e5143 quickcheck_narrowing returns some timing information
bulwahn
parents: 43329
diff changeset
   213
fun elapsed_time description e =
8c4b383e5143 quickcheck_narrowing returns some timing information
bulwahn
parents: 43329
diff changeset
   214
  let val ({elapsed, ...}, result) = Timing.timing e ()
8c4b383e5143 quickcheck_narrowing returns some timing information
bulwahn
parents: 43329
diff changeset
   215
  in (result, (description, Time.toMilliseconds elapsed)) end
43047
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
   216
  
43308
fd6cc1378fec compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
bulwahn
parents: 43240
diff changeset
   217
fun value (contains_existentials, (quiet, size)) ctxt (get, put, put_ml) (code, value_name) =
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   218
  let
43308
fd6cc1378fec compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
bulwahn
parents: 43240
diff changeset
   219
    fun message s = if quiet then () else Output.urgent_message s
41930
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
   220
    val tmp_prefix = "Quickcheck_Narrowing"
43079
4022892a2f28 improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
bulwahn
parents: 43047
diff changeset
   221
    val with_tmp_dir =
4022892a2f28 improving overlord option and partial_term_of derivation; changing Narrowing_Engine to print partial terms
bulwahn
parents: 43047
diff changeset
   222
      if Config.get ctxt overlord then with_overlord_dir else Isabelle_System.with_tmp_dir 
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   223
    fun run in_path = 
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   224
      let
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   225
        val code_file = Path.append in_path (Path.basic "Code.hs")
41930
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
   226
        val narrowing_engine_file = Path.append in_path (Path.basic "Narrowing_Engine.hs")
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   227
        val main_file = Path.append in_path (Path.basic "Main.hs")
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   228
        val main = "module Main where {\n\n" ^
43114
b9fca691addd Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents: 43079
diff changeset
   229
          "import System;\n" ^
41933
10f254a4e5b9 adapting Main file generation for Quickcheck_Narrowing
bulwahn
parents: 41932
diff changeset
   230
          "import Narrowing_Engine;\n" ^
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   231
          "import Code;\n\n" ^
43114
b9fca691addd Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents: 43079
diff changeset
   232
          "main = getArgs >>= \\[size] -> Narrowing_Engine.depthCheck (read size) (Code.value ())\n\n" ^
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   233
          "}\n"
41909
383bbdad1650 replacing strings in generated Code resolves the changing names of Typerep in lazysmallcheck prototype
bulwahn
parents: 41908
diff changeset
   234
        val code' = prefix "module Code where {\n\ndata Typerep = Typerep String [Typerep];\n"
383bbdad1650 replacing strings in generated Code resolves the changing names of Typerep in lazysmallcheck prototype
bulwahn
parents: 41908
diff changeset
   235
          (unprefix "module Code where {" code)
383bbdad1650 replacing strings in generated Code resolves the changing names of Typerep in lazysmallcheck prototype
bulwahn
parents: 41908
diff changeset
   236
        val _ = File.write code_file code'
43237
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   237
        val _ = File.write narrowing_engine_file
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   238
          (if contains_existentials then pnf_narrowing_engine else narrowing_engine)
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   239
        val _ = File.write main_file main
43114
b9fca691addd Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents: 43079
diff changeset
   240
        val executable = File.shell_path (Path.append in_path (Path.basic "isabelle_quickcheck_narrowing"))
b9fca691addd Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents: 43079
diff changeset
   241
        val cmd = "exec \"$ISABELLE_GHC\" -fglasgow-exts " ^
41946
380f7f5ff126 proper File.shell_path;
wenzelm
parents: 41940
diff changeset
   242
          (space_implode " " (map File.shell_path [code_file, narrowing_engine_file, main_file])) ^
43114
b9fca691addd Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents: 43079
diff changeset
   243
          " -o " ^ executable ^ ";"
43379
8c4b383e5143 quickcheck_narrowing returns some timing information
bulwahn
parents: 43329
diff changeset
   244
        val (result, compilation_time) = elapsed_time "Haskell compilation" (fn () => bash cmd) 
43237
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   245
        val _ = if bash cmd <> 0 then error "Compilation with GHC failed" else ()
43379
8c4b383e5143 quickcheck_narrowing returns some timing information
bulwahn
parents: 43329
diff changeset
   246
        fun with_size k exec_times =
43308
fd6cc1378fec compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
bulwahn
parents: 43240
diff changeset
   247
          if k > size then
43379
8c4b383e5143 quickcheck_narrowing returns some timing information
bulwahn
parents: 43329
diff changeset
   248
            (NONE, exec_times)
43114
b9fca691addd Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents: 43079
diff changeset
   249
          else
b9fca691addd Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents: 43079
diff changeset
   250
            let
b9fca691addd Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents: 43079
diff changeset
   251
              val _ = message ("Test data size: " ^ string_of_int k)
43379
8c4b383e5143 quickcheck_narrowing returns some timing information
bulwahn
parents: 43329
diff changeset
   252
              val ((response, _), exec_time) = elapsed_time ("execution of size " ^ string_of_int k)
8c4b383e5143 quickcheck_narrowing returns some timing information
bulwahn
parents: 43329
diff changeset
   253
                (fn () => bash_output (executable ^ " " ^ string_of_int k))
43114
b9fca691addd Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents: 43079
diff changeset
   254
            in
43379
8c4b383e5143 quickcheck_narrowing returns some timing information
bulwahn
parents: 43329
diff changeset
   255
              if response = "NONE\n" then with_size (k + 1) (exec_time :: exec_times)
8c4b383e5143 quickcheck_narrowing returns some timing information
bulwahn
parents: 43329
diff changeset
   256
                else (SOME response, exec_time :: exec_times)
43114
b9fca691addd Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents: 43079
diff changeset
   257
            end
43379
8c4b383e5143 quickcheck_narrowing returns some timing information
bulwahn
parents: 43329
diff changeset
   258
      in case with_size 0 [compilation_time] of
8c4b383e5143 quickcheck_narrowing returns some timing information
bulwahn
parents: 43329
diff changeset
   259
           (NONE, exec_times) => (NONE, exec_times)
8c4b383e5143 quickcheck_narrowing returns some timing information
bulwahn
parents: 43329
diff changeset
   260
         | (SOME response, exec_times) =>
43114
b9fca691addd Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents: 43079
diff changeset
   261
           let
b9fca691addd Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents: 43079
diff changeset
   262
             val output_value = the_default "NONE"
b9fca691addd Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents: 43079
diff changeset
   263
               (try (snd o split_last o filter_out (fn s => s = "") o split_lines) response)
b9fca691addd Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents: 43079
diff changeset
   264
               |> translate_string (fn s => if s = "\\" then "\\\\" else s)
b9fca691addd Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents: 43079
diff changeset
   265
             val ml_code = "\nval _ = Context.set_thread_data (SOME (Context.map_proof (" ^ put_ml
b9fca691addd Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents: 43079
diff changeset
   266
               ^ " (fn () => " ^ output_value ^ ")) (ML_Context.the_generic_context ())))";
b9fca691addd Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents: 43079
diff changeset
   267
             val ctxt' = ctxt
b9fca691addd Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents: 43079
diff changeset
   268
               |> put (fn () => error ("Bad evaluation for " ^ quote put_ml))
b9fca691addd Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents: 43079
diff changeset
   269
               |> Context.proof_map (exec false ml_code);
43379
8c4b383e5143 quickcheck_narrowing returns some timing information
bulwahn
parents: 43329
diff changeset
   270
           in (get ctxt' (), exec_times) end     
41930
1e008cc4883a renaming lazysmallcheck ML file to Quickcheck_Narrowing
bulwahn
parents: 41925
diff changeset
   271
      end
43114
b9fca691addd Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents: 43079
diff changeset
   272
  in with_tmp_dir tmp_prefix run end;
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   273
43308
fd6cc1378fec compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
bulwahn
parents: 43240
diff changeset
   274
fun dynamic_value_strict opts cookie thy postproc t =
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   275
  let
43114
b9fca691addd Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents: 43079
diff changeset
   276
    val ctxt = Proof_Context.init_global thy
43308
fd6cc1378fec compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
bulwahn
parents: 43240
diff changeset
   277
    fun evaluator naming program ((_, vs_ty), t) deps = Exn.interruptible_capture (value opts ctxt cookie)
43114
b9fca691addd Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents: 43079
diff changeset
   278
      (Code_Target.evaluator thy target naming program deps (vs_ty, t));    
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   279
  in Exn.release (Code_Thingol.dynamic_value thy (Exn.map_result o postproc) evaluator t) end;
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   280
43047
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
   281
(** counterexample generator **)
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   282
  
41932
e8f113ce8a94 adapting Quickcheck_Narrowing and example file to new names
bulwahn
parents: 41930
diff changeset
   283
structure Counterexample = Proof_Data
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   284
(
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   285
  type T = unit -> term list option
41936
9792a882da9c renaming tester from lazy_exhaustive to narrowing
bulwahn
parents: 41933
diff changeset
   286
  fun init _ () = error "Counterexample"
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   287
)
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   288
43237
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   289
datatype counterexample = Universal_Counterexample of (term * counterexample)
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   290
  | Existential_Counterexample of (term * counterexample) list
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   291
  | Empty_Assignment
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   292
  
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   293
fun map_counterexample f Empty_Assignment = Empty_Assignment
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   294
  | map_counterexample f (Universal_Counterexample (t, c)) =
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   295
      Universal_Counterexample (f t, map_counterexample f c)
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   296
  | map_counterexample f (Existential_Counterexample cs) =
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   297
      Existential_Counterexample (map (fn (t, c) => (f t, map_counterexample f c)) cs)
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   298
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   299
structure Existential_Counterexample = Proof_Data
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   300
(
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   301
  type T = unit -> counterexample option
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   302
  fun init _ () = error "Counterexample"
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   303
)
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   304
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   305
val put_existential_counterexample = Existential_Counterexample.put
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   306
42024
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   307
val put_counterexample = Counterexample.put
42023
1bd4b6d1c482 adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
bulwahn
parents: 42020
diff changeset
   308
43240
da47097bd589 adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents: 43237
diff changeset
   309
fun finitize_functions (xTs, t) =
42024
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   310
  let
43240
da47097bd589 adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents: 43237
diff changeset
   311
    val (names, boundTs) = split_list xTs
42024
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   312
    fun mk_eval_ffun dT rT =
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   313
      Const (@{const_name "Quickcheck_Narrowing.eval_ffun"}, 
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   314
        Type (@{type_name "Quickcheck_Narrowing.ffun"}, [dT, rT]) --> dT --> rT)
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   315
    fun mk_eval_cfun dT rT =
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   316
      Const (@{const_name "Quickcheck_Narrowing.eval_cfun"}, 
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   317
        Type (@{type_name "Quickcheck_Narrowing.cfun"}, [rT]) --> dT --> rT)
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   318
    fun eval_function (T as Type (@{type_name fun}, [dT, rT])) =
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   319
      let
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   320
        val (rt', rT') = eval_function rT
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   321
      in
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   322
        case dT of
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   323
          Type (@{type_name fun}, _) =>
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   324
            (fn t => absdummy (dT, rt' (mk_eval_cfun dT rT' $ incr_boundvars 1 t $ Bound 0)),
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   325
            Type (@{type_name "Quickcheck_Narrowing.cfun"}, [rT']))
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   326
        | _ => (fn t => absdummy (dT, rt' (mk_eval_ffun dT rT' $ incr_boundvars 1 t $ Bound 0)),
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   327
            Type (@{type_name "Quickcheck_Narrowing.ffun"}, [dT, rT']))
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   328
      end
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   329
      | eval_function T = (I, T)
43240
da47097bd589 adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents: 43237
diff changeset
   330
    val (tt, boundTs') = split_list (map eval_function boundTs)
da47097bd589 adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents: 43237
diff changeset
   331
    val t' = subst_bounds (map2 (fn f => fn x => f x) (rev tt) (map_index (Bound o fst) boundTs), t)
42024
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   332
  in
43240
da47097bd589 adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents: 43237
diff changeset
   333
    (names ~~ boundTs', t')
42024
51df23535105 handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
bulwahn
parents: 42023
diff changeset
   334
  end
42023
1bd4b6d1c482 adding minimalistic setup and transformation to handle functions as data to enable naive function generation for Quickcheck_Narrowing
bulwahn
parents: 42020
diff changeset
   335
43114
b9fca691addd Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents: 43079
diff changeset
   336
(** tester **)
43237
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   337
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   338
val rewrs =
43240
da47097bd589 adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents: 43237
diff changeset
   339
    map (swap o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of)
da47097bd589 adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents: 43237
diff changeset
   340
      (@{thms all_simps} @ @{thms ex_simps})
da47097bd589 adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents: 43237
diff changeset
   341
    @ map (HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of)
da47097bd589 adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents: 43237
diff changeset
   342
        [@{thm iff_conv_conj_imp}, @{thm not_ex}, @{thm not_all}]
43237
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   343
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   344
fun make_pnf_term thy t = Pattern.rewrite_term thy rewrs [] t
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   345
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   346
fun strip_quantifiers (Const (@{const_name Ex}, _) $ Abs (x, T, t)) =
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   347
    apfst (cons (@{const_name Ex}, (x, T))) (strip_quantifiers t)
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   348
  | strip_quantifiers (Const (@{const_name All}, _) $ Abs (x, T, t)) =
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   349
    apfst (cons (@{const_name All}, (x, T))) (strip_quantifiers t)
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   350
  | strip_quantifiers t = ([], t)
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   351
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   352
fun contains_existentials t = exists (fn (Q, _) => Q = @{const_name Ex}) (fst (strip_quantifiers t))
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   353
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   354
fun mk_property qs t =
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   355
  let
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   356
    fun enclose (@{const_name Ex}, (x, T)) t =
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   357
        Const (@{const_name Quickcheck_Narrowing.exists}, (T --> @{typ property}) --> @{typ property})
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   358
          $ Abs (x, T, t)
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   359
      | enclose (@{const_name All}, (x, T)) t =
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   360
        Const (@{const_name Quickcheck_Narrowing.all}, (T --> @{typ property}) --> @{typ property})
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   361
          $ Abs (x, T, t)
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   362
  in
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   363
    fold_rev enclose qs (@{term Quickcheck_Narrowing.Property} $
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   364
      (list_comb (t , map Bound (((length qs) - 1) downto 0))))
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   365
  end
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   366
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   367
fun mk_case_term ctxt p ((@{const_name Ex}, (x, T)) :: qs') (Existential_Counterexample cs) =
43317
f9283eb3a4bf fixing code generation test
bulwahn
parents: 43314
diff changeset
   368
    Datatype.make_case ctxt Datatype_Case.Quiet [] (Free (x, T)) (map (fn (t, c) =>
f9283eb3a4bf fixing code generation test
bulwahn
parents: 43314
diff changeset
   369
      (t, mk_case_term ctxt (p - 1) qs' c)) cs)
43237
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   370
  | mk_case_term ctxt p ((@{const_name All}, (x, T)) :: qs') (Universal_Counterexample (t, c)) =
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   371
    if p = 0 then t else mk_case_term ctxt (p - 1) qs' c
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   372
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   373
fun mk_terms ctxt qs result =
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   374
  let
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   375
    val
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   376
      ps = filter (fn (_, (@{const_name All}, _)) => true | _ => false) (map_index I qs)
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   377
    in
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   378
      map (fn (p, (_, (x, T))) => (x, mk_case_term ctxt p qs result)) ps
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   379
    end
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   380
  
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   381
fun test_term ctxt (limit_time, is_interactive) (t, eval_terms) =
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   382
  let
43308
fd6cc1378fec compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
bulwahn
parents: 43240
diff changeset
   383
    val opts = (Config.get ctxt Quickcheck.quiet, Config.get ctxt Quickcheck.size)
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42258
diff changeset
   384
    val thy = Proof_Context.theory_of ctxt
43237
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   385
    val t' = fold_rev (fn (x, T) => fn t => HOLogic.mk_all (x, T, t)) (Term.add_frees t []) t
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   386
    val pnf_t = make_pnf_term thy t'
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   387
  in
43237
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   388
    if Config.get ctxt allow_existentials andalso contains_existentials pnf_t then
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   389
      let
43240
da47097bd589 adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents: 43237
diff changeset
   390
        fun wrap f (qs, t) =
da47097bd589 adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents: 43237
diff changeset
   391
          let val (qs1, qs2) = split_list qs in
da47097bd589 adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents: 43237
diff changeset
   392
          apfst (map2 pair qs1) (f (qs2, t)) end
da47097bd589 adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents: 43237
diff changeset
   393
        val finitize = if Config.get ctxt finite_functions then wrap finitize_functions else I
da47097bd589 adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents: 43237
diff changeset
   394
        val (qs, prop_t) = finitize (strip_quantifiers pnf_t)
da47097bd589 adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents: 43237
diff changeset
   395
        val prop_term = fold_rev (fn (_, (x, T)) => fn t => Abs (x, T, t)) qs prop_t
43237
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   396
        val ((prop_def, _), ctxt') = Local_Theory.define ((Binding.conceal (Binding.name "test_property"), NoSyn),
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   397
          ((Binding.conceal Binding.empty, [Code.add_default_eqn_attrib]), prop_term)) ctxt
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   398
        val (prop_def', thy') = Local_Theory.exit_result_global Morphism.term (prop_def, ctxt') 
43379
8c4b383e5143 quickcheck_narrowing returns some timing information
bulwahn
parents: 43329
diff changeset
   399
        val (result, timings) = dynamic_value_strict (true, opts)
43237
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   400
          (Existential_Counterexample.get, Existential_Counterexample.put,
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   401
            "Narrowing_Generators.put_existential_counterexample")
43379
8c4b383e5143 quickcheck_narrowing returns some timing information
bulwahn
parents: 43329
diff changeset
   402
          thy' (apfst o Option.map o map_counterexample) (mk_property qs prop_def')
43237
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   403
        val result' = Option.map (mk_terms ctxt' (fst (strip_quantifiers pnf_t))) result
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   404
      in
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   405
        Quickcheck.Result {counterexample = result', evaluation_terms = Option.map (K []) result,
43379
8c4b383e5143 quickcheck_narrowing returns some timing information
bulwahn
parents: 43329
diff changeset
   406
          timings = timings, reports = []}
43237
8f5c3c6c2909 adding compilation that allows existentials in Quickcheck_Narrowing
bulwahn
parents: 43114
diff changeset
   407
      end
43240
da47097bd589 adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents: 43237
diff changeset
   408
    else
da47097bd589 adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents: 43237
diff changeset
   409
      let
da47097bd589 adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents: 43237
diff changeset
   410
        val t' = Term.list_abs_free (Term.add_frees t [], t)
da47097bd589 adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents: 43237
diff changeset
   411
        fun wrap f t = list_abs (f (strip_abs t))
da47097bd589 adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents: 43237
diff changeset
   412
        val finitize = if Config.get ctxt finite_functions then wrap finitize_functions else I
da47097bd589 adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents: 43237
diff changeset
   413
        fun ensure_testable t =
da47097bd589 adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents: 43237
diff changeset
   414
          Const (@{const_name Quickcheck_Narrowing.ensure_testable}, fastype_of t --> fastype_of t) $ t
43379
8c4b383e5143 quickcheck_narrowing returns some timing information
bulwahn
parents: 43329
diff changeset
   415
        val (result, timings) = dynamic_value_strict (false, opts)
43240
da47097bd589 adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents: 43237
diff changeset
   416
          (Counterexample.get, Counterexample.put, "Narrowing_Generators.put_counterexample")
43379
8c4b383e5143 quickcheck_narrowing returns some timing information
bulwahn
parents: 43329
diff changeset
   417
          thy (apfst o Option.map o map) (ensure_testable (finitize t'))
43240
da47097bd589 adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents: 43237
diff changeset
   418
      in
da47097bd589 adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents: 43237
diff changeset
   419
        Quickcheck.Result {counterexample = Option.map ((curry (op ~~)) (Term.add_free_names t [])) result,
43379
8c4b383e5143 quickcheck_narrowing returns some timing information
bulwahn
parents: 43329
diff changeset
   420
          evaluation_terms = Option.map (K []) result, timings = timings, reports = []}
43240
da47097bd589 adding finitize_functions and processing of equivalences in existential compilation in quickcheck_narrowing
bulwahn
parents: 43237
diff changeset
   421
      end
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   422
  end;
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   423
43114
b9fca691addd Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents: 43079
diff changeset
   424
fun test_goals ctxt (limit_time, is_interactive) insts goals =
43314
a9090cabca14 adding a nicer error message for quickcheck_narrowing; hiding fact empty_def
bulwahn
parents: 43308
diff changeset
   425
  if (not (getenv "ISABELLE_GHC" = "")) then
a9090cabca14 adding a nicer error message for quickcheck_narrowing; hiding fact empty_def
bulwahn
parents: 43308
diff changeset
   426
    let
a9090cabca14 adding a nicer error message for quickcheck_narrowing; hiding fact empty_def
bulwahn
parents: 43308
diff changeset
   427
      val correct_inst_goals = Quickcheck.instantiate_goals ctxt insts goals
a9090cabca14 adding a nicer error message for quickcheck_narrowing; hiding fact empty_def
bulwahn
parents: 43308
diff changeset
   428
    in
43114
b9fca691addd Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents: 43079
diff changeset
   429
      Quickcheck.collect_results (test_term ctxt (limit_time, is_interactive)) (maps (map snd) correct_inst_goals) []
43314
a9090cabca14 adding a nicer error message for quickcheck_narrowing; hiding fact empty_def
bulwahn
parents: 43308
diff changeset
   430
    end
a9090cabca14 adding a nicer error message for quickcheck_narrowing; hiding fact empty_def
bulwahn
parents: 43308
diff changeset
   431
  else
a9090cabca14 adding a nicer error message for quickcheck_narrowing; hiding fact empty_def
bulwahn
parents: 43308
diff changeset
   432
    (if Config.get ctxt Quickcheck.quiet then () else Output.urgent_message
a9090cabca14 adding a nicer error message for quickcheck_narrowing; hiding fact empty_def
bulwahn
parents: 43308
diff changeset
   433
      ("Environment variable ISABELLE_GHC is not set. To use narrowing-based quickcheck, please set "
a9090cabca14 adding a nicer error message for quickcheck_narrowing; hiding fact empty_def
bulwahn
parents: 43308
diff changeset
   434
        ^ "this variable to your GHC Haskell compiler in your settings file."); [Quickcheck.empty_result])
43114
b9fca691addd Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents: 43079
diff changeset
   435
43047
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
   436
(* setup *)
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   437
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   438
val setup =
43047
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
   439
  Code.datatype_interpretation ensure_partial_term_of
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
   440
  #> Code.datatype_interpretation ensure_partial_term_of_code
26774ccb1c74 automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
bulwahn
parents: 42616
diff changeset
   441
  #> Datatype.interpretation (Quickcheck_Common.ensure_sort_datatype
42258
79cb339d8989 changing ensure_sort_datatype call in narrowing quickcheck (missed in 1491b7209e76)
bulwahn
parents: 42214
diff changeset
   442
    (((@{sort typerep}, @{sort term_of}), @{sort narrowing}), instantiate_narrowing_datatype))
43114
b9fca691addd Quickcheck Narrowing only requires one compilation with GHC now
bulwahn
parents: 43079
diff changeset
   443
  #> Context.theory_map (Quickcheck.add_tester ("narrowing", test_goals))
41905
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   444
    
e2611bc96022 adding files for prototype of lazysmallcheck
bulwahn
parents:
diff changeset
   445
end;