src/HOL/Decision_Procs/approximation.ML
author immler
Sun, 03 Nov 2019 19:59:56 -0500
changeset 71037 f630f2e707a6
parent 71034 e0755162093f
child 73019 05e2cab9af8b
permissions -rw-r--r--
refactor Approximation.thy to use more abstract type of intervals
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
     1
(*  Title:      HOL/Decision_Procs/approximation.ML
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
     2
    Author:     Johannes Hoelzl, TU Muenchen
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
     3
*)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
     4
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
     5
signature APPROXIMATION =
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
     6
sig
63929
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
     7
  val reify_form: Proof.context -> term -> term
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
     8
  val approx: int -> Proof.context -> term -> term
59850
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
     9
  val approximate: Proof.context -> term -> term
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    10
  val approximation_tac : int -> (string * int) list -> int option -> Proof.context -> int -> tactic
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
    11
end
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
    12
71037
f630f2e707a6 refactor Approximation.thy to use more abstract type of intervals
immler
parents: 71034
diff changeset
    13
structure Approximation =
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
    14
struct
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
    15
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
    16
fun reorder_bounds_tac ctxt prems i =
59850
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    17
  let
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
    18
    fun variable_of_bound (Const (\<^const_name>\<open>Trueprop\<close>, _) $
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
    19
                           (Const (\<^const_name>\<open>Set.member\<close>, _) $
59850
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    20
                            Free (name, _) $ _)) = name
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
    21
      | variable_of_bound (Const (\<^const_name>\<open>Trueprop\<close>, _) $
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
    22
                           (Const (\<^const_name>\<open>HOL.eq\<close>, _) $
59850
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    23
                            Free (name, _) $ _)) = name
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    24
      | variable_of_bound t = raise TERM ("variable_of_bound", [t])
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    25
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    26
    val variable_bounds
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    27
      = map (`(variable_of_bound o Thm.prop_of)) prems
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    28
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    29
    fun add_deps (name, bnds)
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    30
      = Graph.add_deps_acyclic (name,
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    31
          remove (op =) name (Term.add_free_names (Thm.prop_of bnds) []))
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    32
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    33
    val order = Graph.empty
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    34
                |> fold Graph.new_node variable_bounds
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    35
                |> fold add_deps variable_bounds
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    36
                |> Graph.strong_conn |> map the_single |> rev
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    37
                |> map_filter (AList.lookup (op =) variable_bounds)
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    38
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
    39
    fun prepend_prem th tac =
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
    40
      tac THEN resolve_tac ctxt [th RSN (2, @{thm mp})] i
59850
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    41
  in
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    42
    fold prepend_prem order all_tac
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    43
  end
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    44
71034
e0755162093f replace approximation oracle by less ad-hoc @{computation}s
immler
parents: 70308
diff changeset
    45
fun approximate ctxt t = case fastype_of t
e0755162093f replace approximation oracle by less ad-hoc @{computation}s
immler
parents: 70308
diff changeset
    46
   of \<^typ>\<open>bool\<close> =>
e0755162093f replace approximation oracle by less ad-hoc @{computation}s
immler
parents: 70308
diff changeset
    47
        Approximation_Computation.approx_bool ctxt t
71037
f630f2e707a6 refactor Approximation.thy to use more abstract type of intervals
immler
parents: 71034
diff changeset
    48
    | \<^typ>\<open>(float interval) option\<close> =>
71034
e0755162093f replace approximation oracle by less ad-hoc @{computation}s
immler
parents: 70308
diff changeset
    49
        Approximation_Computation.approx_arith ctxt t
71037
f630f2e707a6 refactor Approximation.thy to use more abstract type of intervals
immler
parents: 71034
diff changeset
    50
    | \<^typ>\<open>(float interval) option list\<close> =>
71034
e0755162093f replace approximation oracle by less ad-hoc @{computation}s
immler
parents: 70308
diff changeset
    51
        Approximation_Computation.approx_form_eval ctxt t
e0755162093f replace approximation oracle by less ad-hoc @{computation}s
immler
parents: 70308
diff changeset
    52
    | _ => error ("Bad term: " ^ Syntax.string_of_term ctxt t);
59850
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    53
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    54
fun rewrite_interpret_form_tac ctxt prec splitting taylor i st = let
63930
867ca0d92ea2 provide more information on error
immler
parents: 63929
diff changeset
    55
    fun lookup_splitting (Free (name, _)) =
867ca0d92ea2 provide more information on error
immler
parents: 63929
diff changeset
    56
        (case AList.lookup (op =) splitting name
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
    57
          of SOME s => HOLogic.mk_number \<^typ>\<open>nat\<close> s
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
    58
           | NONE => \<^term>\<open>0 :: nat\<close>)
63930
867ca0d92ea2 provide more information on error
immler
parents: 63929
diff changeset
    59
      | lookup_splitting t = raise TERM ("lookup_splitting", [t])
59850
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    60
    val vs = nth (Thm.prems_of st) (i - 1)
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    61
             |> Logic.strip_imp_concl
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    62
             |> HOLogic.dest_Trueprop
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    63
             |> Term.strip_comb |> snd |> List.last
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    64
             |> HOLogic.dest_list
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    65
    val p = prec
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
    66
            |> HOLogic.mk_number \<^typ>\<open>nat\<close>
59850
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    67
            |> Thm.cterm_of ctxt
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    68
  in case taylor
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    69
  of NONE => let
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    70
       val n = vs |> length
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
    71
               |> HOLogic.mk_number \<^typ>\<open>nat\<close>
59850
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    72
               |> Thm.cterm_of ctxt
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    73
       val s = vs
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    74
               |> map lookup_splitting
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
    75
               |> HOLogic.mk_list \<^typ>\<open>nat\<close>
59850
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    76
               |> Thm.cterm_of ctxt
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    77
     in
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
    78
       (resolve_tac ctxt [Thm.instantiate ([], [((("n", 0), \<^typ>\<open>nat\<close>), n),
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
    79
                                   ((("prec", 0), \<^typ>\<open>nat\<close>), p),
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
    80
                                   ((("ss", 0), \<^typ>\<open>nat list\<close>), s)])
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
    81
            @{thm approx_form}] i
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
    82
        THEN simp_tac (put_simpset (simpset_of \<^context>) ctxt) i) st
59850
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    83
     end
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    84
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    85
   | SOME t =>
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    86
     if length vs <> 1
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    87
     then raise (TERM ("More than one variable used for taylor series expansion", [Thm.prop_of st]))
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    88
     else let
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    89
       val t = t
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
    90
            |> HOLogic.mk_number \<^typ>\<open>nat\<close>
59850
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    91
            |> Thm.cterm_of ctxt
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    92
       val s = vs |> map lookup_splitting |> hd
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    93
            |> Thm.cterm_of ctxt
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    94
     in
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
    95
       resolve_tac ctxt [Thm.instantiate ([], [((("s", 0), \<^typ>\<open>nat\<close>), s),
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
    96
                                   ((("t", 0), \<^typ>\<open>nat\<close>), t),
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
    97
                                   ((("prec", 0), \<^typ>\<open>nat\<close>), p)])
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
    98
            @{thm approx_tse_form}] i st
59850
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    99
     end
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
   100
  end
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
   101
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   102
fun calculated_subterms (\<^const>\<open>Trueprop\<close> $ t) = calculated_subterms t
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   103
  | calculated_subterms (\<^const>\<open>HOL.implies\<close> $ _ $ t) = calculated_subterms t
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   104
  | calculated_subterms (\<^term>\<open>(\<le>) :: real \<Rightarrow> real \<Rightarrow> bool\<close> $ t1 $ t2) = [t1, t2]
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   105
  | calculated_subterms (\<^term>\<open>(<) :: real \<Rightarrow> real \<Rightarrow> bool\<close> $ t1 $ t2) = [t1, t2]
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   106
  | calculated_subterms (\<^term>\<open>(\<in>) :: real \<Rightarrow> real set \<Rightarrow> bool\<close> $ t1 $
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   107
                         (\<^term>\<open>atLeastAtMost :: real \<Rightarrow> real \<Rightarrow> real set\<close> $ t2 $ t3)) = [t1, t2, t3]
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   108
  | calculated_subterms t = raise TERM ("calculated_subterms", [t])
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   109
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   110
fun dest_interpret_form (\<^const>\<open>interpret_form\<close> $ b $ xs) = (b, xs)
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   111
  | dest_interpret_form t = raise TERM ("dest_interpret_form", [t])
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   112
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   113
fun dest_interpret (\<^const>\<open>interpret_floatarith\<close> $ b $ xs) = (b, xs)
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   114
  | dest_interpret t = raise TERM ("dest_interpret", [t])
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   115
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   116
fun dest_interpret_env (\<^const>\<open>interpret_form\<close> $ _ $ xs) = xs
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   117
  | dest_interpret_env (\<^const>\<open>interpret_floatarith\<close> $ _ $ xs) = xs
63930
867ca0d92ea2 provide more information on error
immler
parents: 63929
diff changeset
   118
  | dest_interpret_env t = raise TERM ("dest_interpret_env", [t])
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   119
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   120
fun dest_float (\<^const>\<open>Float\<close> $ m $ e) = (snd (HOLogic.dest_number m), snd (HOLogic.dest_number e))
63930
867ca0d92ea2 provide more information on error
immler
parents: 63929
diff changeset
   121
  | dest_float t = raise TERM ("dest_float", [t])
867ca0d92ea2 provide more information on error
immler
parents: 63929
diff changeset
   122
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   123
71037
f630f2e707a6 refactor Approximation.thy to use more abstract type of intervals
immler
parents: 71034
diff changeset
   124
fun dest_ivl
f630f2e707a6 refactor Approximation.thy to use more abstract type of intervals
immler
parents: 71034
diff changeset
   125
  (Const (\<^const_name>\<open>Some\<close>, _) $
f630f2e707a6 refactor Approximation.thy to use more abstract type of intervals
immler
parents: 71034
diff changeset
   126
    (Const (\<^const_name>\<open>Interval\<close>, _) $ ((Const (\<^const_name>\<open>Pair\<close>, _) $ u $ l)))) =
f630f2e707a6 refactor Approximation.thy to use more abstract type of intervals
immler
parents: 71034
diff changeset
   127
      SOME (dest_float u, dest_float l)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   128
  | dest_ivl (Const (\<^const_name>\<open>None\<close>, _)) = NONE
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   129
  | dest_ivl t = raise TERM ("dest_result", [t])
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   130
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   131
fun mk_approx' prec t = (\<^const>\<open>approx'\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   132
                       $ HOLogic.mk_number \<^typ>\<open>nat\<close> prec
71037
f630f2e707a6 refactor Approximation.thy to use more abstract type of intervals
immler
parents: 71034
diff changeset
   133
                       $ t $ \<^term>\<open>[] :: (float interval) option list\<close>)
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   134
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   135
fun mk_approx_form_eval prec t xs = (\<^const>\<open>approx_form_eval\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   136
                       $ HOLogic.mk_number \<^typ>\<open>nat\<close> prec
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   137
                       $ t $ xs)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   138
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   139
fun float2_float10 prec round_down (m, e) = (
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   140
  let
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   141
    val (m, e) = (if e < 0 then (m,e) else (m * Integer.pow e 2, 0))
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   142
63930
867ca0d92ea2 provide more information on error
immler
parents: 63929
diff changeset
   143
    fun frac _ _ 0 digits cnt = (digits, cnt, 0)
867ca0d92ea2 provide more information on error
immler
parents: 63929
diff changeset
   144
      | frac _ 0 r digits cnt = (digits, cnt, r)
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   145
      | frac c p r digits cnt = (let
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   146
        val (d, r) = Integer.div_mod (r * 10) (Integer.pow (~e) 2)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   147
      in frac (c orelse d <> 0) (if d <> 0 orelse c then p - 1 else p) r
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   148
              (digits * 10 + d) (cnt + 1)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   149
      end)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   150
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   151
    val sgn = Int.sign m
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   152
    val m = abs m
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   153
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   154
    val round_down = (sgn = 1 andalso round_down) orelse
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   155
                     (sgn = ~1 andalso not round_down)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   156
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   157
    val (x, r) = Integer.div_mod m (Integer.pow (~e) 2)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   158
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   159
    val p = ((if x = 0 then prec else prec - (IntInf.log2 x + 1)) * 3) div 10 + 1
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   160
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   161
    val (digits, e10, r) = if p > 0 then frac (x <> 0) p r 0 0 else (0,0,0)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   162
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   163
    val digits = if round_down orelse r = 0 then digits else digits + 1
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   164
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   165
  in (sgn * (digits + x * (Integer.pow e10 10)), ~e10)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   166
  end)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   167
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   168
fun mk_result prec (SOME (l, u)) =
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   169
  (let
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   170
    fun mk_float10 rnd x = (let val (m, e) = float2_float10 prec rnd x
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   171
                       in if e = 0 then HOLogic.mk_number \<^typ>\<open>real\<close> m
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   172
                     else if e = 1 then \<^term>\<open>divide :: real \<Rightarrow> real \<Rightarrow> real\<close> $
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   173
                                        HOLogic.mk_number \<^typ>\<open>real\<close> m $
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   174
                                        \<^term>\<open>10\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   175
                                   else \<^term>\<open>divide :: real \<Rightarrow> real \<Rightarrow> real\<close> $
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   176
                                        HOLogic.mk_number \<^typ>\<open>real\<close> m $
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   177
                                        (\<^term>\<open>power 10 :: nat \<Rightarrow> real\<close> $
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   178
                                         HOLogic.mk_number \<^typ>\<open>nat\<close> (~e)) end)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   179
    in \<^term>\<open>atLeastAtMost :: real \<Rightarrow> real \<Rightarrow> real set\<close> $ mk_float10 true l $ mk_float10 false u end)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   180
  | mk_result _ NONE = \<^term>\<open>UNIV :: real set\<close>
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   181
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   182
fun realify t =
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   183
  let
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   184
    val t = Logic.varify_global t
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   185
    val m = map (fn (name, _) => (name, \<^typ>\<open>real\<close>)) (Term.add_tvars t [])
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   186
    val t = Term.subst_TVars m t
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   187
  in t end
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   188
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   189
fun apply_tactic ctxt term tactic =
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   190
  Thm.cterm_of ctxt term
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   191
  |> Goal.init
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   192
  |> SINGLE tactic
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   193
  |> the |> Thm.prems_of |> hd
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   194
63929
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
   195
fun preproc_form_conv ctxt =
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
   196
  Simplifier.rewrite
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
   197
   (put_simpset HOL_basic_ss ctxt addsimps
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   198
     (Named_Theorems.get ctxt \<^named_theorems>\<open>approximation_preproc\<close>))
63929
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
   199
63930
867ca0d92ea2 provide more information on error
immler
parents: 63929
diff changeset
   200
fun reify_form_conv ctxt ct =
867ca0d92ea2 provide more information on error
immler
parents: 63929
diff changeset
   201
  let
867ca0d92ea2 provide more information on error
immler
parents: 63929
diff changeset
   202
    val thm =
867ca0d92ea2 provide more information on error
immler
parents: 63929
diff changeset
   203
       Reification.conv ctxt @{thms interpret_form.simps interpret_floatarith.simps} ct
867ca0d92ea2 provide more information on error
immler
parents: 63929
diff changeset
   204
       handle ERROR msg =>
867ca0d92ea2 provide more information on error
immler
parents: 63929
diff changeset
   205
        cat_error ("Reification failed: " ^ msg)
867ca0d92ea2 provide more information on error
immler
parents: 63929
diff changeset
   206
          ("Approximation does not support " ^
867ca0d92ea2 provide more information on error
immler
parents: 63929
diff changeset
   207
            quote (Syntax.string_of_term ctxt (Thm.term_of ct)))
867ca0d92ea2 provide more information on error
immler
parents: 63929
diff changeset
   208
    fun check_env (Free _) = ()
867ca0d92ea2 provide more information on error
immler
parents: 63929
diff changeset
   209
      | check_env (Var _) = ()
867ca0d92ea2 provide more information on error
immler
parents: 63929
diff changeset
   210
      | check_env t =
867ca0d92ea2 provide more information on error
immler
parents: 63929
diff changeset
   211
          cat_error "Term not supported by approximation:" (Syntax.string_of_term ctxt t)
867ca0d92ea2 provide more information on error
immler
parents: 63929
diff changeset
   212
    val _ = Thm.rhs_of thm |> Thm.term_of |> dest_interpret_env |> HOLogic.dest_list |> map check_env
867ca0d92ea2 provide more information on error
immler
parents: 63929
diff changeset
   213
  in thm end
867ca0d92ea2 provide more information on error
immler
parents: 63929
diff changeset
   214
63929
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
   215
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
   216
fun reify_form_tac ctxt i = CONVERSION (Conv.arg_conv (reify_form_conv ctxt)) i
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   217
63929
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
   218
fun prepare_form_tac ctxt i =
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
   219
  REPEAT (FIRST' [eresolve_tac ctxt @{thms intervalE},
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
   220
    eresolve_tac ctxt @{thms meta_eqE},
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
   221
    resolve_tac ctxt @{thms impI}] i)
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
   222
  THEN Subgoal.FOCUS (fn {prems, context = ctxt', ...} => reorder_bounds_tac ctxt' prems i) ctxt i
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
   223
  THEN DETERM (TRY (filter_prems_tac ctxt (K false) i))
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
   224
  THEN CONVERSION (Conv.arg_conv (preproc_form_conv ctxt)) i
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
   225
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
   226
fun prepare_form ctxt term = apply_tactic ctxt term (prepare_form_tac ctxt 1)
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
   227
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
   228
fun apply_reify_form ctxt t = apply_tactic ctxt t (reify_form_tac ctxt 1)
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
   229
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
   230
fun reify_form ctxt t = HOLogic.mk_Trueprop t
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
   231
  |> prepare_form ctxt
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
   232
  |> apply_reify_form ctxt
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
   233
  |> HOLogic.dest_Trueprop
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   234
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   235
fun approx_form prec ctxt t =
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   236
        realify t
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   237
     |> prepare_form ctxt
63929
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
   238
     |> (fn arith_term => apply_reify_form ctxt arith_term
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
   239
         |> HOLogic.dest_Trueprop
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
   240
         |> dest_interpret_form
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   241
         |> (fn (data, xs) =>
71037
f630f2e707a6 refactor Approximation.thy to use more abstract type of intervals
immler
parents: 71034
diff changeset
   242
            mk_approx_form_eval prec data (HOLogic.mk_list \<^typ>\<open>(float interval) option\<close>
f630f2e707a6 refactor Approximation.thy to use more abstract type of intervals
immler
parents: 71034
diff changeset
   243
              (map (fn _ => \<^term>\<open>None :: (float interval) option\<close>) (HOLogic.dest_list xs)))
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   244
         |> approximate ctxt
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   245
         |> HOLogic.dest_list
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   246
         |> curry ListPair.zip (HOLogic.dest_list xs @ calculated_subterms arith_term)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   247
         |> map (fn (elem, s) => \<^term>\<open>(\<in>) :: real \<Rightarrow> real set \<Rightarrow> bool\<close> $ elem $ mk_result prec (dest_ivl s))
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   248
         |> foldr1 HOLogic.mk_conj))
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   249
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   250
fun approx_arith prec ctxt t = realify t
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   251
     |> Thm.cterm_of ctxt
63929
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
   252
     |> (preproc_form_conv ctxt then_conv reify_form_conv ctxt)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   253
     |> Thm.prop_of
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   254
     |> Logic.dest_equals |> snd
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   255
     |> dest_interpret |> fst
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   256
     |> mk_approx' prec
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   257
     |> approximate ctxt
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   258
     |> dest_ivl
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   259
     |> mk_result prec
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   260
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   261
fun approx prec ctxt t =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   262
  if type_of t = \<^typ>\<open>prop\<close> then approx_form prec ctxt t
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   263
  else if type_of t = \<^typ>\<open>bool\<close> then approx_form prec ctxt (\<^const>\<open>Trueprop\<close> $ t)
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   264
  else approx_arith prec ctxt t
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   265
56923
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   266
fun approximate_cmd modes raw_t state =
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   267
  let
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   268
    val ctxt = Toplevel.context_of state;
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   269
    val t = Syntax.read_term ctxt raw_t;
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   270
    val t' = approx 30 ctxt t;
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   271
    val ty' = Term.type_of t';
70308
7f568724d67e clarified signature;
wenzelm
parents: 69597
diff changeset
   272
    val ctxt' = Proof_Context.augment t' ctxt;
59174
wenzelm
parents: 58893
diff changeset
   273
  in
wenzelm
parents: 58893
diff changeset
   274
    Print_Mode.with_modes modes (fn () =>
56923
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   275
      Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk,
59174
wenzelm
parents: 58893
diff changeset
   276
        Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) ()
wenzelm
parents: 58893
diff changeset
   277
  end |> Pretty.writeln;
56923
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   278
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   279
val opt_modes =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   280
  Scan.optional (\<^keyword>\<open>(\<close> |-- Parse.!!! (Scan.repeat1 Parse.name --| \<^keyword>\<open>)\<close>)) [];
56923
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   281
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   282
val _ =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   283
  Outer_Syntax.command \<^command_keyword>\<open>approximate\<close> "print approximation of term"
56923
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   284
    (opt_modes -- Parse.term
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   285
      >> (fn (modes, t) => Toplevel.keep (approximate_cmd modes t)));
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   286
71034
e0755162093f replace approximation oracle by less ad-hoc @{computation}s
immler
parents: 70308
diff changeset
   287
fun approximation_tac prec splitting taylor ctxt =
e0755162093f replace approximation oracle by less ad-hoc @{computation}s
immler
parents: 70308
diff changeset
   288
  prepare_form_tac ctxt
e0755162093f replace approximation oracle by less ad-hoc @{computation}s
immler
parents: 70308
diff changeset
   289
  THEN' reify_form_tac ctxt
e0755162093f replace approximation oracle by less ad-hoc @{computation}s
immler
parents: 70308
diff changeset
   290
  THEN' rewrite_interpret_form_tac ctxt prec splitting taylor
e0755162093f replace approximation oracle by less ad-hoc @{computation}s
immler
parents: 70308
diff changeset
   291
  THEN' CONVERSION (Approximation_Computation.approx_conv ctxt)
e0755162093f replace approximation oracle by less ad-hoc @{computation}s
immler
parents: 70308
diff changeset
   292
  THEN' resolve_tac ctxt [TrueI]
63929
b673e7221b16 approximation: rewrite for reduction to base expressions
immler
parents: 62969
diff changeset
   293
62391
1658fc9b2618 more canonical names
nipkow
parents: 60754
diff changeset
   294
end;