| author | wenzelm | 
| Tue, 25 May 2021 23:58:49 +0200 | |
| changeset 73786 | 18d80cd51823 | 
| parent 69593 | 3dda49e08b9d | 
| permissions | -rw-r--r-- | 
| 53164 
beb4ee344c22
clarified position of Spec_Check for Isabelle/ML -- it is unrelated to Isabelle/HOL;
 wenzelm parents: 
52278diff
changeset | 1 | (* Title: Tools/Spec_Check/spec_check.ML | 
| 52248 
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 | Specification-based testing of ML programs with random values. | 
| 
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 SPEC_CHECK = | 
| 
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 gen_target : int Config.T | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 11 | val gen_max : int Config.T | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 12 | val examples : int Config.T | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 13 | val sort_examples : bool Config.T | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 14 | val show_stats : bool Config.T | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 15 | val column_width : int Config.T | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 16 | val style : string Config.T | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 17 | |
| 52254 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 18 | type output_style = Proof.context -> string -> | 
| 52252 | 19 |     {status : string option * Property.result * (Property.stats  * string option list) -> unit,
 | 
| 20 | finish: Property.stats * string option list -> unit} | |
| 52248 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 21 | |
| 52258 | 22 | val register_style : string -> output_style -> theory -> theory | 
| 52248 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 23 | |
| 52254 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 24 | val checkGen : Proof.context -> | 
| 52252 | 25 |     'a Generator.gen * ('a -> string) option -> string * 'a Property.prop -> unit
 | 
| 52248 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 26 | |
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 27 | val check_property : Proof.context -> string -> unit | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 28 | end; | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 29 | |
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 30 | structure Spec_Check : SPEC_CHECK = | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 31 | struct | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 32 | |
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 33 | (* configurations *) | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 34 | |
| 69593 | 35 | val gen_target = Attrib.setup_config_int \<^binding>\<open>spec_check_gen_target\<close> (K 100) | 
| 36 | val gen_max = Attrib.setup_config_int \<^binding>\<open>spec_check_gen_max\<close> (K 400) | |
| 37 | val examples = Attrib.setup_config_int \<^binding>\<open>spec_check_examples\<close> (K 5) | |
| 52248 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 38 | |
| 69593 | 39 | val sort_examples = Attrib.setup_config_bool \<^binding>\<open>spec_check_sort_examples\<close> (K true) | 
| 40 | val show_stats = Attrib.setup_config_bool \<^binding>\<open>spec_check_show_stats\<close> (K true) | |
| 41 | val column_width = Attrib.setup_config_int \<^binding>\<open>spec_check_column_width\<close> (K 22) | |
| 42 | val style = Attrib.setup_config_string \<^binding>\<open>spec_check_style\<close> (K "Perl") | |
| 52248 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 43 | |
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 44 | type ('a, 'b) reader = 'b -> ('a * 'b) option
 | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 45 | type 'a rep = ('a -> string) option
 | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 46 | |
| 52254 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 47 | type output_style = Proof.context -> string -> | 
| 52252 | 48 |   {status: string option * Property.result * (Property.stats * string option list) -> unit,
 | 
| 49 | finish: Property.stats * string option list -> unit} | |
| 52248 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 50 | |
| 52254 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 51 | fun limit ctxt gen = Generator.limit (Config.get ctxt gen_max) gen | 
| 52248 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 52 | |
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 53 | structure Style = Theory_Data | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 54 | ( | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 55 | type T = output_style Symtab.table | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 56 | val empty = Symtab.empty | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 57 | val extend = I | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 58 | 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 | 59 | ) | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 60 | |
| 52254 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 61 | fun get_style ctxt = | 
| 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 62 | let val name = Config.get ctxt style in | 
| 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 63 | (case Symtab.lookup (Style.get (Proof_Context.theory_of ctxt)) name of | 
| 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 64 | SOME style => style ctxt | 
| 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 65 |     | NONE => error ("No style called " ^ quote name ^ " found"))
 | 
| 52248 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 66 | end | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 67 | |
| 52258 | 68 | fun register_style name style = Style.map (Symtab.update (name, style)) | 
| 52248 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 69 | |
| 52253 | 70 | |
| 52248 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 71 | (* testing functions *) | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 72 | |
| 52254 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 73 | fun cpsCheck ctxt s0 shrink (next, show) (tag, prop) = | 
| 52248 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 74 | let | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 75 | val apply_show = case show of NONE => (fn _ => NONE) | SOME f => SOME o f | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 76 | |
| 52254 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 77 |     val {status, finish} = get_style ctxt tag
 | 
| 52248 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 78 | fun status' (obj, result, (stats, badobjs)) = | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 79 | let | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 80 | val badobjs' = if Property.failure result then obj :: badobjs else badobjs | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 81 | val _ = status (apply_show obj, result, (stats, map apply_show badobjs')) | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 82 | in badobjs' end | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 83 | |
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 84 | fun try_shrink (obj, result, stats') (stats, badobjs) = | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 85 | let | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 86 | fun is_failure s = | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 87 | let val (result, stats') = Property.test prop (s, stats) | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 88 | in if Property.failure result then SOME (s, result, stats') else NONE end | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 89 | in | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 90 | case get_first is_failure (shrink obj) of | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 91 | SOME (obj, result, stats') => try_shrink (obj, result, stats') (stats, badobjs) | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 92 | | NONE => status' (obj, result, (stats', badobjs)) | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 93 | end | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 94 | |
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 95 | fun iter (NONE, (stats, badobjs)) = finish (stats, map apply_show badobjs) | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 96 | | iter (SOME (obj, stream), (stats, badobjs)) = | 
| 52254 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 97 | if #count stats >= Config.get ctxt gen_target then | 
| 52248 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 98 | finish (stats, map apply_show badobjs) | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 99 | else | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 100 | let | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 101 | val (result, stats') = Property.test prop (obj, stats) | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 102 | val badobjs' = if Property.failure result then | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 103 | try_shrink (obj, result, stats') (stats, badobjs) | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 104 | else | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 105 | status' (obj, result, (stats', badobjs)) | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 106 | in iter (next stream, (stats', badobjs')) end | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 107 | in | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 108 | fn stream => iter (next stream, (s0, [])) | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 109 | end | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 110 | |
| 52254 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 111 | fun check' ctxt s0 = cpsCheck ctxt s0 (fn _ => []) | 
| 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 112 | fun check ctxt = check' ctxt Property.stats | 
| 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 113 | fun checks ctxt = cpsCheck ctxt Property.stats | 
| 52248 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 114 | |
| 52254 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 115 | fun checkGen ctxt (gen, show) (tag, prop) = | 
| 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 116 |   check' ctxt {count = 0, tags = [("__GEN", 0)]}
 | 
| 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 117 | (limit ctxt gen, show) (tag, prop) Generator.stream | 
| 52248 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 118 | |
| 52254 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 119 | fun checkGenShrink ctxt shrink (gen, show) (tag, prop) = | 
| 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 120 |   cpsCheck ctxt {count=0, tags=[("__GEN", 0)]} shrink
 | 
| 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 121 | (limit ctxt gen, show) (tag, prop) Generator.stream | 
| 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 122 | |
| 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 123 | fun checkOne ctxt show (tag, prop) obj = | 
| 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 124 | check ctxt (List.getItem, show) (tag, prop) [obj] | 
| 52248 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 125 | |
| 52253 | 126 | (*call the compiler and pass resulting type string to the parser*) | 
| 127 | fun determine_type ctxt s = | |
| 52248 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 128 | let | 
| 52259 
65fb8cec59a5
more direct Context.setmp_thread_data for one-way passing of context;
 wenzelm parents: 
52258diff
changeset | 129 | val return = Unsynchronized.ref "return" | 
| 62494 | 130 | val context : ML_Compiler0.context = | 
| 62495 | 131 |      {name_space = #name_space ML_Env.context,
 | 
| 62716 
d80b9f4990e4
explicit print_depth for the sake of Spec_Check.determine_type;
 wenzelm parents: 
62495diff
changeset | 132 | print_depth = SOME 1000000, | 
| 62495 | 133 | here = #here ML_Env.context, | 
| 52259 
65fb8cec59a5
more direct Context.setmp_thread_data for one-way passing of context;
 wenzelm parents: 
52258diff
changeset | 134 | print = fn r => return := r, | 
| 62495 | 135 | error = #error ML_Env.context} | 
| 52259 
65fb8cec59a5
more direct Context.setmp_thread_data for one-way passing of context;
 wenzelm parents: 
52258diff
changeset | 136 | val _ = | 
| 62889 | 137 | Context.setmp_generic_context (SOME (Context.Proof ctxt)) | 
| 60956 | 138 | (fn () => | 
| 62902 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 wenzelm parents: 
62889diff
changeset | 139 | ML_Compiler0.ML context | 
| 60956 | 140 |             {line = 0, file = "generated code", verbose = true, debug = false} s) ()
 | 
| 52248 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 141 | in | 
| 52259 
65fb8cec59a5
more direct Context.setmp_thread_data for one-way passing of context;
 wenzelm parents: 
52258diff
changeset | 142 | Gen_Construction.parse_pred (! return) | 
| 52248 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 143 | end; | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 144 | |
| 52253 | 145 | (*call the compiler and run the test*) | 
| 146 | fun run_test ctxt s = | |
| 62889 | 147 | Context.setmp_generic_context (SOME (Context.Proof ctxt)) | 
| 60956 | 148 | (fn () => | 
| 62902 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 wenzelm parents: 
62889diff
changeset | 149 | ML_Compiler0.ML ML_Env.context | 
| 60956 | 150 |         {line = 0, file = "generated code", verbose = false, debug = false} s) ();
 | 
| 52248 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 151 | |
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 152 | (*split input into tokens*) | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 153 | fun input_split s = | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 154 | let | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 155 | fun dot c = c = #"." | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 156 | fun space c = c = #" " | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 157 | val (head, code) = Substring.splitl (not o dot) (Substring.full s) | 
| 52254 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 158 | in | 
| 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 159 | (String.tokens space (Substring.string head), | 
| 52252 | 160 | Substring.string (Substring.dropl dot code)) | 
| 161 | end; | |
| 52248 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 162 | |
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 163 | (*create the function from the input*) | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 164 | fun make_fun s = | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 165 | let | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 166 | val scan_param = Scan.one (fn s => s <> ";") | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 167 | fun parameters s = Scan.repeat1 scan_param s | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 168 | val p = $$ "ALL" |-- parameters | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 169 | val (split, code) = input_split s | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 170 | val stop = Scan.stopper (fn _ => ";") (fn s => s = ";"); | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 171 | val (params, _) = Scan.finite stop p split | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 172 |   in "fn (" ^ commas params ^ ") => " ^ code end;
 | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 173 | |
| 52253 | 174 | (*read input and perform the test*) | 
| 175 | fun gen_check_property check ctxt s = | |
| 52248 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 176 | let | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 177 | val func = make_fun s | 
| 52253 | 178 | val (_, ty) = determine_type ctxt func | 
| 52254 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 179 | in run_test ctxt (check ctxt "Check" (ty, func)) end; | 
| 52248 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 180 | |
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 181 | val check_property = gen_check_property Gen_Construction.build_check | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 182 | (*val check_property_safe = gen_check_property Gen_Construction.safe_check*) | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 183 | |
| 52253 | 184 | (*perform test for specification function*) | 
| 185 | fun gen_check_property_f check ctxt s = | |
| 52248 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 186 | let | 
| 52253 | 187 | val (name, ty) = determine_type ctxt s | 
| 52254 
994055f7db80
simplified context and data management -- plain ctxt: Proof.context is default for most operations;
 wenzelm parents: 
52253diff
changeset | 188 | in run_test ctxt (check ctxt name (ty, s)) end; | 
| 52248 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 189 | |
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 190 | val check_property_f = gen_check_property_f Gen_Construction.build_check | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 191 | (*val check_property_safe_f_ = gen_check_property_f Gen_Construction.safe_check*) | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 192 | |
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 193 | end; | 
| 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 194 | |
| 62876 | 195 | fun check_property s = Spec_Check.check_property (Context.the_local_context ()) s; | 
| 52248 
2c893e0c1def
added Spec_Check -- a Quickcheck tool for Isabelle's ML environment;
 bulwahn parents: diff
changeset | 196 |