src/HOL/Decision_Procs/approximation_generator.ML
author wenzelm
Wed, 12 Mar 2025 11:39:00 +0100
changeset 82265 4b875a4c83b0
parent 81984 6c052e21664f
permissions -rw-r--r--
update for release;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58988
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
     1
(*  Title:      HOL/Decision_Procs/approximation_generator.ML
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
     2
    Author:     Fabian Immler, TU Muenchen
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
     3
*)
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
     4
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
     5
signature APPROXIMATION_GENERATOR =
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
     6
sig
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
     7
  val custom_seed: int Config.T
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
     8
  val precision: int Config.T
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
     9
  val epsilon: real Config.T
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    10
  val approximation_generator:
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    11
    Proof.context ->
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    12
    (term * term list) list ->
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    13
    bool -> int list -> (bool * term list) option * Quickcheck.report option
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    14
  val setup: theory -> theory
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    15
end;
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    16
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    17
structure Approximation_Generator : APPROXIMATION_GENERATOR =
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    18
struct
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    19
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 64519
diff changeset
    20
val custom_seed = Attrib.setup_config_int \<^binding>\<open>quickcheck_approximation_custom_seed\<close> (K ~1)
58988
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    21
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 64519
diff changeset
    22
val precision = Attrib.setup_config_int \<^binding>\<open>quickcheck_approximation_precision\<close> (K 30)
58988
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    23
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 64519
diff changeset
    24
val epsilon = Attrib.setup_config_real \<^binding>\<open>quickcheck_approximation_epsilon\<close> (K 0.0)
58988
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    25
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    26
val random_float = @{code "random_class.random::_ \<Rightarrow> _ \<Rightarrow> (float \<times> (unit \<Rightarrow> term)) \<times> _"}
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    27
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    28
fun nat_of_term t =
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    29
  (HOLogic.dest_nat t handle TERM _ => snd (HOLogic.dest_number t)
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    30
    handle TERM _ => raise TERM ("nat_of_term", [t]));
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    31
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    32
fun int_of_term t = snd (HOLogic.dest_number t) handle TERM _ => raise TERM ("int_of_term", [t]);
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    33
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    34
fun real_of_man_exp m e = Real.fromManExp {man = Real.fromInt m, exp = e}
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    35
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 71037
diff changeset
    36
fun mapprox_float \<^Const_>\<open>Float for m e\<close> = real_of_man_exp (int_of_term m) (int_of_term e)
58988
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    37
  | mapprox_float t = Real.fromInt (snd (HOLogic.dest_number t))
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    38
      handle TERM _ => raise TERM ("mapprox_float", [t]);
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    39
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    40
(* TODO: define using compiled terms? *)
81984
6c052e21664f more explicit real operations
haftmann
parents: 74397
diff changeset
    41
fun mapprox_floatarith \<^Const_>\<open>Add for a b\<close> xs = Real.+ (mapprox_floatarith a xs, mapprox_floatarith b xs)
6c052e21664f more explicit real operations
haftmann
parents: 74397
diff changeset
    42
  | mapprox_floatarith \<^Const_>\<open>Minus for a\<close> xs = Real.~ (mapprox_floatarith a xs)
6c052e21664f more explicit real operations
haftmann
parents: 74397
diff changeset
    43
  | mapprox_floatarith \<^Const_>\<open>Mult for a b\<close> xs = Real.* (mapprox_floatarith a xs, mapprox_floatarith b xs)
6c052e21664f more explicit real operations
haftmann
parents: 74397
diff changeset
    44
  | mapprox_floatarith \<^Const_>\<open>Inverse for a\<close> xs = 1.0 / (mapprox_floatarith a xs : real)
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 71037
diff changeset
    45
  | mapprox_floatarith \<^Const_>\<open>Cos for a\<close> xs = Math.cos (mapprox_floatarith a xs)
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 71037
diff changeset
    46
  | mapprox_floatarith \<^Const_>\<open>Arctan for a\<close> xs = Math.atan (mapprox_floatarith a xs)
81984
6c052e21664f more explicit real operations
haftmann
parents: 74397
diff changeset
    47
  | mapprox_floatarith \<^Const_>\<open>Abs for a\<close> xs = Real.abs (mapprox_floatarith a xs : real)
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 71037
diff changeset
    48
  | mapprox_floatarith \<^Const_>\<open>Max for a b\<close> xs =
58988
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    49
      Real.max (mapprox_floatarith a xs, mapprox_floatarith b xs)
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 71037
diff changeset
    50
  | mapprox_floatarith \<^Const_>\<open>Min for a b\<close> xs =
58988
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    51
      Real.min (mapprox_floatarith a xs, mapprox_floatarith b xs)
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 71037
diff changeset
    52
  | mapprox_floatarith \<^Const_>\<open>Pi\<close> _ = Math.pi
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 71037
diff changeset
    53
  | mapprox_floatarith \<^Const_>\<open>Sqrt for a\<close> xs = Math.sqrt (mapprox_floatarith a xs)
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 71037
diff changeset
    54
  | mapprox_floatarith \<^Const_>\<open>Exp for a\<close> xs = Math.exp (mapprox_floatarith a xs)
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 71037
diff changeset
    55
  | mapprox_floatarith \<^Const_>\<open>Powr for a b\<close> xs =
64519
51be997d0698 op powr for quickcheck[approximation] (amending 67792e4a5486)
immler
parents: 63931
diff changeset
    56
      Math.pow (mapprox_floatarith a xs, mapprox_floatarith b xs)
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 71037
diff changeset
    57
  | mapprox_floatarith \<^Const_>\<open>Ln for a\<close> xs = Math.ln (mapprox_floatarith a xs)
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 71037
diff changeset
    58
  | mapprox_floatarith \<^Const_>\<open>Power for a n\<close> xs =
58988
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    59
      Math.pow (mapprox_floatarith a xs, Real.fromInt (nat_of_term n))
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 71037
diff changeset
    60
  | mapprox_floatarith \<^Const_>\<open>Floor for a\<close> xs = Real.fromInt (floor (mapprox_floatarith a xs))
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 71037
diff changeset
    61
  | mapprox_floatarith \<^Const_>\<open>Var for n\<close> xs = nth xs (nat_of_term n)
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 71037
diff changeset
    62
  | mapprox_floatarith \<^Const_>\<open>Num for m\<close> _ = mapprox_float m
58988
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    63
  | mapprox_floatarith t _ = raise TERM ("mapprox_floatarith", [t])
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    64
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    65
fun mapprox_atLeastAtMost eps x a b xs =
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    66
    let
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    67
      val x' = mapprox_floatarith x xs
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    68
    in
81984
6c052e21664f more explicit real operations
haftmann
parents: 74397
diff changeset
    69
      Real.<= (Real.+ (mapprox_floatarith a xs, eps), x') andalso Real.<= (Real.+ (x', eps), mapprox_floatarith b xs)
58988
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    70
    end
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    71
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 71037
diff changeset
    72
fun mapprox_form eps \<^Const_>\<open>Bound for x a b f\<close> xs =
58988
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    73
    (not (mapprox_atLeastAtMost eps x a b xs)) orelse mapprox_form eps f xs
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 71037
diff changeset
    74
| mapprox_form eps \<^Const_>\<open>Assign for x a f\<close> xs =
58988
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    75
    (Real.!= (mapprox_floatarith x xs, mapprox_floatarith a xs)) orelse mapprox_form eps f xs
81984
6c052e21664f more explicit real operations
haftmann
parents: 74397
diff changeset
    76
| mapprox_form eps \<^Const_>\<open>Less for a b\<close> xs = Real.< (Real.+ (mapprox_floatarith a xs, eps), mapprox_floatarith b xs)
6c052e21664f more explicit real operations
haftmann
parents: 74397
diff changeset
    77
| mapprox_form eps \<^Const_>\<open>LessEqual for a b\<close> xs = Real.<= (Real.+ (mapprox_floatarith a xs, eps), mapprox_floatarith b xs)
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 71037
diff changeset
    78
| mapprox_form eps \<^Const_>\<open>AtLeastAtMost for x a b\<close> xs = mapprox_atLeastAtMost eps x a b xs
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 71037
diff changeset
    79
| mapprox_form eps \<^Const_>\<open>Conj for f g\<close> xs = mapprox_form eps f xs andalso mapprox_form eps g xs
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 71037
diff changeset
    80
| mapprox_form eps \<^Const_>\<open>Disj for f g\<close> xs = mapprox_form eps f xs orelse mapprox_form eps g xs
58988
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    81
| mapprox_form _ t _ = raise TERM ("mapprox_form", [t])
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    82
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 71037
diff changeset
    83
fun dest_interpret_form \<^Const_>\<open>interpret_form for b xs\<close> = (b, xs)
58988
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    84
  | dest_interpret_form t = raise TERM ("dest_interpret_form", [t])
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    85
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    86
fun random_float_list size xs seed =
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    87
  fold (K (apsnd (random_float size) #-> (fn c => apfst (fn b => b::c)))) xs ([],seed)
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    88
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    89
fun real_of_Float (@{code Float} (m, e)) =
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    90
    real_of_man_exp (@{code integer_of_int} m) (@{code integer_of_int} e)
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    91
71034
e0755162093f replace approximation oracle by less ad-hoc @{computation}s
immler
parents: 69597
diff changeset
    92
fun term_of_int i = (HOLogic.mk_number @{typ int} (@{code integer_of_int} i))
e0755162093f replace approximation oracle by less ad-hoc @{computation}s
immler
parents: 69597
diff changeset
    93
e0755162093f replace approximation oracle by less ad-hoc @{computation}s
immler
parents: 69597
diff changeset
    94
fun term_of_Float (@{code Float} (m, e)) = @{term Float} $ term_of_int m $ term_of_int e
e0755162093f replace approximation oracle by less ad-hoc @{computation}s
immler
parents: 69597
diff changeset
    95
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 71037
diff changeset
    96
fun is_True \<^Const_>\<open>True\<close> = true
58988
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    97
  | is_True _ = false
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    98
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
    99
val postproc_form_eqs =
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   100
  @{lemma
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 59850
diff changeset
   101
    "real_of_float (Float 0 a) = 0"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 59850
diff changeset
   102
    "real_of_float (Float (numeral m) 0) = numeral m"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 59850
diff changeset
   103
    "real_of_float (Float 1 0) = 1"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 59850
diff changeset
   104
    "real_of_float (Float (- 1) 0) = - 1"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 59850
diff changeset
   105
    "real_of_float (Float 1 (numeral e)) = 2 ^ numeral e"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 59850
diff changeset
   106
    "real_of_float (Float 1 (- numeral e)) = 1 / 2 ^ numeral e"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 59850
diff changeset
   107
    "real_of_float (Float a 1) = a * 2"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 59850
diff changeset
   108
    "real_of_float (Float a (-1)) = a / 2"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 59850
diff changeset
   109
    "real_of_float (Float (- a) b) = - real_of_float (Float a b)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 59850
diff changeset
   110
    "real_of_float (Float (numeral m) (numeral e)) = numeral m * 2 ^ (numeral e)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 59850
diff changeset
   111
    "real_of_float (Float (numeral m) (- numeral e)) = numeral m / 2 ^ (numeral e)"
58988
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   112
    "- (c * d::real) = -c * d"
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   113
    "- (c / d::real) = -c / d"
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   114
    "- (0::real) = 0"
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   115
    "int_of_integer (numeral k) = numeral k"
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   116
    "int_of_integer (- numeral k) = - numeral k"
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   117
    "int_of_integer 0 = 0"
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   118
    "int_of_integer 1 = 1"
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   119
    "int_of_integer (- 1) = - 1"
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   120
    by auto
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   121
  }
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   122
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   123
fun rewrite_with ctxt thms = Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps thms)
59629
0d77c51b5040 clarified context;
wenzelm
parents: 59621
diff changeset
   124
fun conv_term ctxt conv r = Thm.cterm_of ctxt r |> conv |> Thm.prop_of |> Logic.dest_equals |> snd
58988
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   125
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   126
fun approx_random ctxt prec eps frees e xs genuine_only size seed =
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   127
  let
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   128
    val (rs, seed') = random_float_list size xs seed
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   129
    fun mk_approx_form e ts =
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 71037
diff changeset
   130
      \<^Const>\<open>approx_form\<close> $
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 71037
diff changeset
   131
        HOLogic.mk_number \<^Type>\<open>nat\<close> prec $
58988
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   132
        e $
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 71037
diff changeset
   133
        (HOLogic.mk_list \<^typ>\<open>float interval option\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 71037
diff changeset
   134
          (map (fn t => \<^Const>\<open>Some \<^typ>\<open>float interval\<close>\<close> $ \<^Const>\<open>interval_of \<^Type>\<open>float\<close> for t\<close>) ts)) $
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 71037
diff changeset
   135
        \<^Const>\<open>Nil \<^Type>\<open>nat\<close>\<close>
58988
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   136
  in
63931
f17a1c60ac39 approximation: preprocessing for nat/int expressions
immler
parents: 63929
diff changeset
   137
    (if
f17a1c60ac39 approximation: preprocessing for nat/int expressions
immler
parents: 63929
diff changeset
   138
      mapprox_form eps e (map (real_of_Float o fst) rs)
f17a1c60ac39 approximation: preprocessing for nat/int expressions
immler
parents: 63929
diff changeset
   139
      handle
f17a1c60ac39 approximation: preprocessing for nat/int expressions
immler
parents: 63929
diff changeset
   140
        General.Overflow => false
f17a1c60ac39 approximation: preprocessing for nat/int expressions
immler
parents: 63929
diff changeset
   141
      | General.Domain => false
f17a1c60ac39 approximation: preprocessing for nat/int expressions
immler
parents: 63929
diff changeset
   142
      | General.Div => false
f17a1c60ac39 approximation: preprocessing for nat/int expressions
immler
parents: 63929
diff changeset
   143
      | IEEEReal.Unordered => false
58988
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   144
    then
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   145
      let
71034
e0755162093f replace approximation oracle by less ad-hoc @{computation}s
immler
parents: 69597
diff changeset
   146
        val ts = map (fn x => term_of_Float (fst x)) rs
58988
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   147
        val ts' = map
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   148
          (AList.lookup op = (map dest_Free xs ~~ ts)
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   149
            #> the_default Term.dummy
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 71037
diff changeset
   150
            #> curry op $ \<^Const>\<open>real_of_float\<close>
59629
0d77c51b5040 clarified context;
wenzelm
parents: 59621
diff changeset
   151
            #> conv_term ctxt (rewrite_with ctxt postproc_form_eqs))
58988
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   152
          frees
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   153
      in
59850
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59629
diff changeset
   154
        if Approximation.approximate ctxt (mk_approx_form e ts) |> is_True
58988
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   155
        then SOME (true, ts')
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   156
        else (if genuine_only then NONE else SOME (false, ts'))
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   157
      end
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   158
    else NONE, seed')
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   159
  end
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   160
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   161
val preproc_form_eqs =
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   162
  @{lemma
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   163
    "(a::real) \<in> {b .. c} \<longleftrightarrow> b \<le> a \<and> a \<le> c"
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   164
    "a = b \<longleftrightarrow> a \<le> b \<and> b \<le> a"
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   165
    "(p \<longrightarrow> q) \<longleftrightarrow> \<not>p \<or> q"
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   166
    "(p \<longleftrightarrow> q) \<longleftrightarrow> (p \<longrightarrow> q) \<and> (q \<longrightarrow> p)"
59629
0d77c51b5040 clarified context;
wenzelm
parents: 59621
diff changeset
   167
    by auto}
58988
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   168
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   169
fun reify_goal ctxt t =
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   170
  HOLogic.mk_not t
59629
0d77c51b5040 clarified context;
wenzelm
parents: 59621
diff changeset
   171
    |> conv_term ctxt (rewrite_with ctxt preproc_form_eqs)
63929
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 61609
diff changeset
   172
    |> Approximation.reify_form ctxt
58988
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   173
    |> dest_interpret_form
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   174
    ||> HOLogic.dest_list
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   175
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   176
fun approximation_generator_raw ctxt t =
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   177
  let
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   178
    val iterations = Config.get ctxt Quickcheck.iterations
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   179
    val prec = Config.get ctxt precision
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   180
    val eps = Config.get ctxt epsilon
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   181
    val cs = Config.get ctxt custom_seed
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   182
    val seed = (Code_Numeral.natural_of_integer (cs + 1), Code_Numeral.natural_of_integer 1)
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   183
    val run = if cs < 0
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   184
      then (fn f => fn seed => (Random_Engine.run f, seed))
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   185
      else (fn f => fn seed => f seed)
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   186
    val frees = Term.add_frees t []
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   187
    val (e, xs) = reify_goal ctxt t
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   188
    fun single_tester b s =
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   189
      approx_random ctxt prec eps frees e xs b s |> run
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   190
    fun iterate _ _ 0 _ = NONE
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   191
      | iterate genuine_only size j seed =
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   192
        case single_tester genuine_only size seed of
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   193
          (NONE, seed') => iterate genuine_only size (j - 1) seed'
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   194
        | (SOME q, _) => SOME q
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   195
  in
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   196
    fn genuine_only => fn size => (iterate genuine_only size iterations seed, NONE)
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   197
  end
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   198
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   199
fun approximation_generator ctxt [(t, _)] =
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   200
  (fn genuine_only =>
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   201
    fn [_, size] =>
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   202
      approximation_generator_raw ctxt t genuine_only
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   203
        (Code_Numeral.natural_of_integer size))
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   204
  | approximation_generator _ _ =
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   205
      error "Quickcheck-approximation does not support type variables (or finite instantiations)"
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   206
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   207
val test_goals =
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   208
  Quickcheck_Common.generator_test_goal_terms
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   209
    ("approximation", (fn _ => fn _ => false, approximation_generator))
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   210
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 64519
diff changeset
   211
val active = Attrib.setup_config_bool \<^binding>\<open>quickcheck_approximation_active\<close> (K false)
58988
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   212
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   213
val setup = Context.theory_map (Quickcheck.add_tester ("approximation", (active, test_goals)))
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   214
6ebf918128b9 added quickcheck[approximation]
immler
parents:
diff changeset
   215
end