src/HOL/Spec_Check/gen_construction.ML
author wenzelm
Thu, 30 May 2013 21:28:54 +0200
changeset 52254 994055f7db80
parent 52253 afca6a99a361
child 52257 9e97fd77a879
permissions -rw-r--r--
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
     1
(*  Title:      HOL/Spec_Check/gen_construction.ML
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
     2
    Author:     Lukas Bulwahn and Nicolai Schaffroth, TU Muenchen
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
     3
    Author:     Christopher League
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
     4
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
     5
Constructing generators and pretty printing function for complex types.
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
     6
*)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
     7
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
     8
signature GEN_CONSTRUCTION =
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
     9
sig
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    10
  val register : string * (string * string) -> theory -> theory
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    11
  type mltype
52254
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
    12
  val parse_pred : string -> string * mltype
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
    13
  val build_check : Proof.context -> string -> mltype * string -> string
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    14
  (*val safe_check : string -> mltype * string -> string*)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    15
  val string_of_bool : bool -> string
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    16
  val string_of_ref : ('a -> string) -> 'a Unsynchronized.ref -> string
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    17
end;
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    18
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    19
structure Gen_Construction : GEN_CONSTRUCTION =
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    20
struct
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    21
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    22
(* Parsing ML types *)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    23
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    24
datatype mltype = Var | Con of string * mltype list | Tuple of mltype list;
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    25
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    26
(*Split string into tokens for parsing*)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    27
fun split s =
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    28
  let
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    29
    fun split_symbol #"(" = "( "
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    30
      | split_symbol #")" = " )"
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    31
      | split_symbol #"," = " ,"
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    32
      | split_symbol #":" = " :"
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    33
      | split_symbol c = Char.toString c
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    34
    fun is_space c = c = #" "
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    35
  in String.tokens is_space (String.translate split_symbol s) end;
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    36
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    37
(*Accept anything that is not a recognized symbol*)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    38
val scan_name = Scan.one (fn s => not (String.isSubstring s "(),*->;"));
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    39
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    40
(*Turn a type list into a nested Con*)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    41
fun make_con nil = raise Empty
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    42
  | make_con [c] = c
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    43
  | make_con (Con (s, _) :: cl) = Con (s, [make_con cl]);
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    44
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    45
(*Parse a type*)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    46
fun parse_type s = (parse_fun || parse_tuple || parse_type_single) s
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    47
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    48
and parse_type_arg s = (parse_tuple || parse_type_single) s
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    49
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    50
and parse_type_single s = (parse_con || parse_type_basic) s
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    51
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    52
and parse_type_basic s = (parse_var || $$ "(" |-- parse_type --| $$ ")") s
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    53
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    54
and parse_list s =
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    55
  ($$ "(" |-- parse_type -- Scan.repeat1 ($$ "," |-- parse_type) --| $$ ")" >> op::) s
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    56
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    57
and parse_var s = (Scan.one (String.isPrefix "'") >> (fn _ => Var)) s
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    58
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    59
and parse_con s = ((parse_con_nest
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    60
  || parse_type_basic -- parse_con_nest >> (fn (b, Con (t, _) :: tl) => Con (t, [b]) :: tl)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    61
  || parse_list -- parse_con_nest >> (fn (l, Con (t, _) :: tl) => Con (t, l) :: tl))
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    62
  >> (make_con o rev)) s
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    63
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    64
and parse_con_nest s = Scan.unless parse_var (Scan.repeat1 (scan_name >> (fn t => Con (t, nil)))) s
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    65
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    66
and parse_fun s = (parse_type_arg -- $$ "->" -- parse_type >> (fn ((a, f), r) => Con (f, [a, r]))) s
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    67
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    68
and parse_tuple s = (parse_type_single -- Scan.repeat1 ($$ "*" |-- parse_type_single)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    69
  >> (fn (t, tl) => Tuple (t :: tl))) s;
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    70
52253
afca6a99a361 more conventional spelling and grammar;
wenzelm
parents: 52248
diff changeset
    71
(*Parse entire type + name*)
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    72
fun parse_function s =
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    73
  let
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    74
    val p = $$ "val" |-- scan_name --| ($$ "=" -- $$ "fn" -- $$ ":")
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    75
    val (name, ty) = p (split s)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    76
    val stop = Scan.stopper (fn _ => ";") (fn s => s = ";");
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    77
    val (typ, _) = Scan.finite stop parse_type ty
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    78
  in (name, typ) end;
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    79
52253
afca6a99a361 more conventional spelling and grammar;
wenzelm
parents: 52248
diff changeset
    80
(*Create desired output*)
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    81
fun parse_pred s =
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    82
  let
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    83
    val (name, Con ("->", t :: _)) = parse_function s
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    84
  in (name, t) end;
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    85
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    86
(* Construct Generators and Pretty Printers *)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    87
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    88
(*copied from smt_config.ML *)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    89
fun string_of_bool b = if b then "true" else "false"
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    90
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    91
fun string_of_ref f r = f (!r) ^ " ref";
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    92
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    93
val initial_content = Symtab.make
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    94
  [("bool", ("Generator.flip", "Gen_Construction.string_of_bool")),
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    95
  ("option", ("Generator.option (Generator.flip' (1, 3))", "ML_Syntax.print_option")),
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    96
  ("list", ("Generator.list (Generator.flip' (1, 3))", "ML_Syntax.print_list")),
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    97
  ("unit", ("gen_unit", "fn () => \"()\"")),
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    98
  ("int", ("Generator.int", "string_of_int")),
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
    99
  ("real", ("Generator.real", "string_of_real")),
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   100
  ("char", ("Generator.char", "fn c => \"#'\" ^ (Char.toString c) ^ \"'\"")),
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   101
  ("string", ("Generator.string (Generator.range (0, 100), Generator.char)", "ML_Syntax.print_string")),
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   102
  ("->", ("Generator.function_lazy", "fn (_, _) => fn _ => \"fn\"")),
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   103
  ("typ", ("Generator.typ 10", "ML_Syntax.print_typ")),
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   104
  ("term", ("Generator.term 10", "ML_Syntax.print_term"))]
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   105
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   106
structure Data = Theory_Data
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   107
(
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   108
  type T = (string * string) Symtab.table
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   109
  val empty = initial_content
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   110
  val extend = I
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   111
  fun merge data : T = Symtab.merge (K true) data
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   112
)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   113
52254
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   114
fun data_of ctxt tycon =
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   115
  (case Symtab.lookup (Data.get (Proof_Context.theory_of ctxt)) tycon of
52253
afca6a99a361 more conventional spelling and grammar;
wenzelm
parents: 52248
diff changeset
   116
    SOME data => data
afca6a99a361 more conventional spelling and grammar;
wenzelm
parents: 52248
diff changeset
   117
  | NONE => error ("No generator and printer defined for ML type constructor " ^ quote tycon))
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   118
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   119
val generator_of = fst oo data_of
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   120
val printer_of = snd oo data_of
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   121
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   122
fun register (ty, data) = Data.map (Symtab.update (ty, data))
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   123
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   124
(*
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   125
fun remove_gen ty = gen_table := AList.delete (op =) ty (!gen_table);
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   126
*)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   127
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   128
fun combine dict [] = dict
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   129
  | combine dict dicts = enclose "(" ")" dict ^ " " ^ enclose "(" ")" (commas dicts)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   130
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   131
fun compose_generator _ Var = "Generator.int"
52254
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   132
  | compose_generator ctxt (Con (s, types)) =
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   133
      combine (generator_of ctxt s) (map (compose_generator ctxt) types)
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   134
  | compose_generator ctxt (Tuple t) =
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   135
      let
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   136
        fun tuple_body t = space_implode ""
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   137
          (map (fn (ty, n) => "val (x" ^ string_of_int n ^ ", r" ^ string_of_int n ^ ") = " ^
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   138
          compose_generator ctxt ty ^ " r" ^ string_of_int (n - 1) ^ " ") (t ~~ (1 upto (length t))))
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   139
        fun tuple_ret a = commas (map (fn n => "x" ^ string_of_int n) (1 upto a))
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   140
      in
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   141
        "fn r0 => let " ^ tuple_body t ^
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   142
        "in ((" ^ tuple_ret (length t) ^ "), r" ^ string_of_int (length t) ^ ") end"
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   143
      end;
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   144
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   145
fun compose_printer _ Var = "Int.toString"
52254
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   146
  | compose_printer ctxt (Con (s, types)) =
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   147
      combine (printer_of ctxt s) (map (compose_printer ctxt) types)
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   148
  | compose_printer ctxt (Tuple t) =
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   149
      let
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   150
        fun tuple_head a = commas (map (fn n => "x" ^ string_of_int n) (1 upto a))
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   151
        fun tuple_body t = space_implode " ^ \", \" ^ "
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   152
          (map (fn (ty, n) => "(" ^ compose_printer ctxt ty ^ ") x" ^ string_of_int n)
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   153
          (t ~~ (1 upto (length t))))
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   154
      in "fn (" ^ tuple_head (length t) ^ ") => \"(\" ^ " ^ tuple_body t ^ " ^ \")\"" end;
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   155
52253
afca6a99a361 more conventional spelling and grammar;
wenzelm
parents: 52248
diff changeset
   156
(*produce compilable string*)
52254
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   157
fun build_check ctxt name (ty, spec) =
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   158
  "Spec_Check.checkGen (ML_Context.the_local_context ()) ("
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   159
  ^ compose_generator ctxt ty ^ ", SOME (" ^ compose_printer ctxt ty ^ ")) (\""
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   160
  ^ name ^ "\", Property.pred (" ^ spec ^ "));";
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   161
52253
afca6a99a361 more conventional spelling and grammar;
wenzelm
parents: 52248
diff changeset
   162
(*produce compilable string - non-eqtype functions*)
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   163
(*
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   164
fun safe_check name (ty, spec) =
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   165
  let
52254
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   166
    val default =
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   167
      (case AList.lookup (op =) (!gen_table) "->" of
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   168
        NONE => ("gen_function_rand", "fn (_, _) => fn _ => \"fn\"")
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   169
      | SOME entry => entry)
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   170
  in
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   171
   (gen_table :=
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   172
     AList.update (op =) ("->", ("gen_function_safe", "fn (_, _) => fn _ => \"fn\"")) (!gen_table);
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   173
    build_check name (ty, spec) before
52254
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   174
    gen_table := AList.update (op =) ("->", default) (!gen_table))
994055f7db80 simplified context and data management -- plain ctxt: Proof.context is default for most operations;
wenzelm
parents: 52253
diff changeset
   175
  end;
52248
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   176
*)
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   177
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   178
end;
2c893e0c1def added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
bulwahn
parents:
diff changeset
   179