src/HOL/Decision_Procs/approximation.ML
author wenzelm
Sat, 18 Jul 2015 20:54:56 +0200
changeset 60754 02924903a6fd
parent 60642 48dd1cefb4ae
child 62391 1658fc9b2618
permissions -rw-r--r--
prefer tactics with explicit context;
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
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
     7
  val approx: int -> Proof.context -> term -> term
59850
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
     8
  val approximate: Proof.context -> term -> term
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
     9
  val approximation_tac : int -> (string * int) list -> int option -> Proof.context -> int -> tactic
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
    10
end
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
    11
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
    12
structure Approximation: APPROXIMATION =
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
    13
struct
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
    14
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
    15
fun reorder_bounds_tac ctxt prems i =
59850
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    16
  let
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    17
    fun variable_of_bound (Const (@{const_name Trueprop}, _) $
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    18
                           (Const (@{const_name Set.member}, _) $
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    19
                            Free (name, _) $ _)) = name
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    20
      | variable_of_bound (Const (@{const_name Trueprop}, _) $
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    21
                           (Const (@{const_name HOL.eq}, _) $
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    22
                            Free (name, _) $ _)) = name
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    23
      | variable_of_bound t = raise TERM ("variable_of_bound", [t])
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    24
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    25
    val variable_bounds
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    26
      = map (`(variable_of_bound o Thm.prop_of)) prems
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    27
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    28
    fun add_deps (name, bnds)
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    29
      = Graph.add_deps_acyclic (name,
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    30
          remove (op =) name (Term.add_free_names (Thm.prop_of bnds) []))
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    31
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    32
    val order = Graph.empty
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    33
                |> fold Graph.new_node variable_bounds
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    34
                |> fold add_deps variable_bounds
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    35
                |> Graph.strong_conn |> map the_single |> rev
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    36
                |> map_filter (AList.lookup (op =) variable_bounds)
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    37
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
    38
    fun prepend_prem th tac =
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
    39
      tac THEN resolve_tac ctxt [th RSN (2, @{thm mp})] i
59850
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    40
  in
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    41
    fold prepend_prem order all_tac
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    42
  end
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    43
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    44
fun approximation_conv ctxt ct =
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    45
  approximation_oracle (Proof_Context.theory_of ctxt, Thm.term_of ct |> tap (tracing o Syntax.string_of_term ctxt));
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    46
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    47
fun approximate ctxt t =
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    48
  approximation_oracle (Proof_Context.theory_of ctxt, t)
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    49
  |> Thm.prop_of |> Logic.dest_equals |> snd;
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    50
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    51
(* Should be in HOL.thy ? *)
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
    52
fun gen_eval_tac conv ctxt =
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
    53
  CONVERSION
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
    54
    (Object_Logic.judgment_conv ctxt (Conv.params_conv (~1) (K (Conv.concl_conv (~1) conv)) ctxt))
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
    55
  THEN' resolve_tac ctxt [TrueI]
59850
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    56
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    57
val form_equations = @{thms interpret_form_equations};
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    58
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    59
fun rewrite_interpret_form_tac ctxt prec splitting taylor i st = let
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    60
    fun lookup_splitting (Free (name, _))
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    61
      = case AList.lookup (op =) splitting name
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    62
        of SOME s => HOLogic.mk_number @{typ nat} s
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    63
         | NONE => @{term "0 :: nat"}
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    64
    val vs = nth (Thm.prems_of st) (i - 1)
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    65
             |> Logic.strip_imp_concl
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    66
             |> HOLogic.dest_Trueprop
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    67
             |> Term.strip_comb |> snd |> List.last
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    68
             |> HOLogic.dest_list
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    69
    val p = prec
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    70
            |> HOLogic.mk_number @{typ nat}
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    71
            |> Thm.cterm_of ctxt
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    72
  in case taylor
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    73
  of NONE => let
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    74
       val n = vs |> length
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    75
               |> HOLogic.mk_number @{typ nat}
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
       val s = vs
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    78
               |> map lookup_splitting
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    79
               |> HOLogic.mk_list @{typ nat}
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    80
               |> Thm.cterm_of ctxt
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    81
     in
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
    82
       (resolve_tac ctxt [Thm.instantiate ([], [((("n", 0), @{typ nat}), n),
60642
48dd1cefb4ae simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents: 59970
diff changeset
    83
                                   ((("prec", 0), @{typ nat}), p),
48dd1cefb4ae simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents: 59970
diff changeset
    84
                                   ((("ss", 0), @{typ "nat list"}), s)])
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
    85
            @{thm approx_form}] i
59850
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    86
        THEN simp_tac (put_simpset (simpset_of @{context}) ctxt) i) st
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    87
     end
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    88
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    89
   | SOME t =>
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    90
     if length vs <> 1
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    91
     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
    92
     else let
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    93
       val t = t
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    94
            |> HOLogic.mk_number @{typ nat}
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    95
            |> Thm.cterm_of ctxt
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    96
       val s = vs |> map lookup_splitting |> hd
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    97
            |> Thm.cterm_of ctxt
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
    98
     in
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
    99
       resolve_tac ctxt [Thm.instantiate ([], [((("s", 0), @{typ nat}), s),
60642
48dd1cefb4ae simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents: 59970
diff changeset
   100
                                   ((("t", 0), @{typ nat}), t),
48dd1cefb4ae simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents: 59970
diff changeset
   101
                                   ((("prec", 0), @{typ nat}), p)])
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   102
            @{thm approx_tse_form}] i st
59850
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
   103
     end
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
   104
  end
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
   105
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   106
fun calculated_subterms (@{const Trueprop} $ t) = calculated_subterms t
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   107
  | calculated_subterms (@{const HOL.implies} $ _ $ t) = calculated_subterms t
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   108
  | calculated_subterms (@{term "op <= :: real \<Rightarrow> real \<Rightarrow> bool"} $ t1 $ t2) = [t1, t2]
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   109
  | calculated_subterms (@{term "op < :: real \<Rightarrow> real \<Rightarrow> bool"} $ t1 $ t2) = [t1, t2]
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   110
  | calculated_subterms (@{term "op : :: real \<Rightarrow> real set \<Rightarrow> bool"} $ t1 $
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   111
                         (@{term "atLeastAtMost :: real \<Rightarrow> real \<Rightarrow> real set"} $ t2 $ t3)) = [t1, t2, t3]
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   112
  | calculated_subterms t = raise TERM ("calculated_subterms", [t])
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   113
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   114
fun dest_interpret_form (@{const "interpret_form"} $ b $ xs) = (b, xs)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   115
  | dest_interpret_form t = raise TERM ("dest_interpret_form", [t])
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   116
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   117
fun dest_interpret (@{const "interpret_floatarith"} $ b $ xs) = (b, xs)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   118
  | dest_interpret t = raise TERM ("dest_interpret", [t])
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   119
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   120
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   121
fun dest_float (@{const "Float"} $ m $ e) =
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   122
  (snd (HOLogic.dest_number m), snd (HOLogic.dest_number e))
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   123
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   124
fun dest_ivl (Const (@{const_name "Some"}, _) $
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   125
              (Const (@{const_name Pair}, _) $ u $ l)) = SOME (dest_float u, dest_float l)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   126
  | dest_ivl (Const (@{const_name "None"}, _)) = NONE
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   127
  | dest_ivl t = raise TERM ("dest_result", [t])
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   128
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   129
fun mk_approx' prec t = (@{const "approx'"}
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   130
                       $ HOLogic.mk_number @{typ nat} prec
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   131
                       $ t $ @{term "[] :: (float * float) option list"})
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   132
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   133
fun mk_approx_form_eval prec t xs = (@{const "approx_form_eval"}
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   134
                       $ HOLogic.mk_number @{typ nat} prec
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   135
                       $ t $ xs)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   136
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   137
fun float2_float10 prec round_down (m, e) = (
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   138
  let
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   139
    val (m, e) = (if e < 0 then (m,e) else (m * Integer.pow e 2, 0))
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   140
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   141
    fun frac c p 0 digits cnt = (digits, cnt, 0)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   142
      | frac c 0 r digits cnt = (digits, cnt, r)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   143
      | frac c p r digits cnt = (let
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   144
        val (d, r) = Integer.div_mod (r * 10) (Integer.pow (~e) 2)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   145
      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
   146
              (digits * 10 + d) (cnt + 1)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   147
      end)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   148
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   149
    val sgn = Int.sign m
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   150
    val m = abs m
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   151
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   152
    val round_down = (sgn = 1 andalso round_down) orelse
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   153
                     (sgn = ~1 andalso not round_down)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   154
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   155
    val (x, r) = Integer.div_mod m (Integer.pow (~e) 2)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   156
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   157
    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
   158
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   159
    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
   160
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   161
    val digits = if round_down orelse r = 0 then digits else digits + 1
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   162
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   163
  in (sgn * (digits + x * (Integer.pow e10 10)), ~e10)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   164
  end)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   165
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   166
fun mk_result prec (SOME (l, u)) =
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   167
  (let
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   168
    fun mk_float10 rnd x = (let val (m, e) = float2_float10 prec rnd x
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   169
                       in if e = 0 then HOLogic.mk_number @{typ real} m
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   170
                     else if e = 1 then @{term "divide :: real \<Rightarrow> real \<Rightarrow> real"} $
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   171
                                        HOLogic.mk_number @{typ real} m $
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   172
                                        @{term "10"}
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   173
                                   else @{term "divide :: real \<Rightarrow> real \<Rightarrow> real"} $
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   174
                                        HOLogic.mk_number @{typ real} m $
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   175
                                        (@{term "power 10 :: nat \<Rightarrow> real"} $
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   176
                                         HOLogic.mk_number @{typ nat} (~e)) end)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   177
    in @{term "atLeastAtMost :: real \<Rightarrow> real \<Rightarrow> real set"} $ mk_float10 true l $ mk_float10 false u end)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   178
  | mk_result _ NONE = @{term "UNIV :: real set"}
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   179
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   180
fun realify t =
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   181
  let
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   182
    val t = Logic.varify_global t
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   183
    val m = map (fn (name, _) => (name, @{typ real})) (Term.add_tvars t [])
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   184
    val t = Term.subst_TVars m t
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   185
  in t end
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   186
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   187
fun apply_tactic ctxt term tactic =
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   188
  Thm.cterm_of ctxt term
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   189
  |> Goal.init
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   190
  |> SINGLE tactic
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   191
  |> the |> Thm.prems_of |> hd
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   192
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   193
fun prepare_form ctxt term = apply_tactic ctxt term (
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   194
    REPEAT (FIRST' [eresolve_tac ctxt @{thms intervalE},
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   195
      eresolve_tac ctxt @{thms meta_eqE},
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   196
      resolve_tac ctxt @{thms impI}] 1)
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   197
    THEN Subgoal.FOCUS (fn {prems, context = ctxt', ...} => reorder_bounds_tac ctxt' prems 1) ctxt 1
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59174
diff changeset
   198
    THEN DETERM (TRY (filter_prems_tac ctxt (K false) 1)))
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   199
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   200
fun reify_form ctxt term = apply_tactic ctxt term
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   201
   (Reification.tac ctxt form_equations NONE 1)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   202
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   203
fun approx_form prec ctxt t =
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   204
        realify t
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   205
     |> prepare_form ctxt
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   206
     |> (fn arith_term => reify_form ctxt arith_term
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   207
         |> HOLogic.dest_Trueprop |> dest_interpret_form
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   208
         |> (fn (data, xs) =>
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   209
            mk_approx_form_eval prec data (HOLogic.mk_list @{typ "(float * float) option"}
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   210
              (map (fn _ => @{term "None :: (float * float) option"}) (HOLogic.dest_list xs)))
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   211
         |> approximate ctxt
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   212
         |> HOLogic.dest_list
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   213
         |> curry ListPair.zip (HOLogic.dest_list xs @ calculated_subterms arith_term)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   214
         |> map (fn (elem, s) => @{term "op : :: real \<Rightarrow> real set \<Rightarrow> bool"} $ elem $ mk_result prec (dest_ivl s))
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   215
         |> foldr1 HOLogic.mk_conj))
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   216
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   217
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
   218
     |> Thm.cterm_of ctxt
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   219
     |> Reification.conv ctxt form_equations
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   220
     |> Thm.prop_of
56813
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   221
     |> Logic.dest_equals |> snd
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   222
     |> dest_interpret |> fst
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   223
     |> mk_approx' prec
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   224
     |> approximate ctxt
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   225
     |> dest_ivl
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   226
     |> mk_result prec
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   227
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   228
fun approx prec ctxt t =
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   229
  if type_of t = @{typ prop} then approx_form prec ctxt t
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   230
  else if type_of t = @{typ bool} then approx_form prec ctxt (@{const Trueprop} $ t)
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   231
  else approx_arith prec ctxt t
80a5905c1610 separate ML module
haftmann
parents:
diff changeset
   232
56923
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   233
fun approximate_cmd modes raw_t state =
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   234
  let
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   235
    val ctxt = Toplevel.context_of state;
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   236
    val t = Syntax.read_term ctxt raw_t;
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   237
    val t' = approx 30 ctxt t;
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   238
    val ty' = Term.type_of t';
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   239
    val ctxt' = Variable.auto_fixes t' ctxt;
59174
wenzelm
parents: 58893
diff changeset
   240
  in
wenzelm
parents: 58893
diff changeset
   241
    Print_Mode.with_modes modes (fn () =>
56923
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   242
      Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk,
59174
wenzelm
parents: 58893
diff changeset
   243
        Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) ()
wenzelm
parents: 58893
diff changeset
   244
  end |> Pretty.writeln;
56923
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   245
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   246
val opt_modes =
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   247
  Scan.optional (@{keyword "("} |-- Parse.!!! (Scan.repeat1 Parse.xname --| @{keyword ")"})) [];
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   248
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   249
val _ =
59936
b8ffc3dc9e24 @{command_spec} is superseded by @{command_keyword};
wenzelm
parents: 59850
diff changeset
   250
  Outer_Syntax.command @{command_keyword approximate} "print approximation of term"
56923
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   251
    (opt_modes -- Parse.term
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   252
      >> (fn (modes, t) => Toplevel.keep (approximate_cmd modes t)));
c062543d380e prefer separate command for approximation
haftmann
parents: 56813
diff changeset
   253
59850
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
   254
fun approximation_tac prec splitting taylor ctxt i =
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   255
  REPEAT (FIRST' [eresolve_tac ctxt @{thms intervalE},
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   256
                  eresolve_tac ctxt @{thms meta_eqE},
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   257
                  resolve_tac ctxt @{thms impI}] i)
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   258
  THEN Subgoal.FOCUS (fn {prems, context = ctxt', ...} => reorder_bounds_tac ctxt' prems i) ctxt i
59850
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
   259
  THEN DETERM (TRY (filter_prems_tac ctxt (K false) i))
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
   260
  THEN DETERM (Reification.tac ctxt form_equations NONE i)
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
   261
  THEN rewrite_interpret_form_tac ctxt prec splitting taylor i
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
   262
  THEN gen_eval_tac (approximation_conv ctxt) ctxt i
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
   263
    
f339ff48a6ee exposed approximation in ML
eberlm
parents: 59621
diff changeset
   264
end;