src/HOL/Analysis/normarith.ML
author wenzelm
Fri, 29 Oct 2021 19:17:24 +0200
changeset 74626 9a1f4a7ddf9e
parent 74572 08b4292abe2b
permissions -rw-r--r--
clarified antiquotations;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63627
6ddb43c6b711 rename HOL-Multivariate_Analysis to HOL-Analysis.
hoelzl
parents: 63211
diff changeset
     1
(*  Title:      HOL/Analysis/normarith.ML
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
     2
    Author:     Amine Chaieb, University of Cambridge
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
     3
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
     4
Simple decision procedure for linear problems in Euclidean space.
31118
541d43bee678 Isolated decision procedure for noms and the general arithmetic solver
chaieb
parents: 30868
diff changeset
     5
*)
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
     6
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
     7
signature NORM_ARITH =
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
     8
sig
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
     9
 val norm_arith : Proof.context -> conv
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
    10
 val norm_arith_tac : Proof.context -> int -> tactic
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
    11
end
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
    12
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
    13
structure NormArith : NORM_ARITH =
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
    14
struct
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
    15
32832
4602cb6ae0b5 additional fixes in normarith.ML due to FuncFun and FuncUtil changes
Philipp Meyer
parents: 32645
diff changeset
    16
 open Conv;
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
    17
 val bool_eq = op = : bool *bool -> bool
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
    18
  fun dest_ratconst t = case Thm.term_of t of
74572
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
    19
    \<^Const_>\<open>divide _ for a b\<close> => Rat.make(HOLogic.dest_number a |> snd, HOLogic.dest_number b |> snd)
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
    20
  | \<^Const_>\<open>inverse _ for a\<close> => Rat.make(1, HOLogic.dest_number a |> snd)
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
    21
  | _ => Rat.of_int (HOLogic.dest_number (Thm.term_of t) |> snd)
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
    22
 fun is_ratconst t = can dest_ratconst t
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
    23
 fun augment_norm b t acc = case Thm.term_of t of
74572
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
    24
     \<^Const_>\<open>norm _ for _\<close> => insert (eq_pair bool_eq (op aconvc)) (b,Thm.dest_arg t) acc
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
    25
   | _ => acc
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
    26
 fun find_normedterms t acc = case Thm.term_of t of
74572
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
    27
     \<^Const_>\<open>plus \<^typ>\<open>real\<close> for _ _\<close> =>
32832
4602cb6ae0b5 additional fixes in normarith.ML due to FuncFun and FuncUtil changes
Philipp Meyer
parents: 32645
diff changeset
    28
            find_normedterms (Thm.dest_arg1 t) (find_normedterms (Thm.dest_arg t) acc)
74572
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
    29
   | \<^Const_>\<open>times \<^typ>\<open>real\<close> for _ _\<close> =>
32832
4602cb6ae0b5 additional fixes in normarith.ML due to FuncFun and FuncUtil changes
Philipp Meyer
parents: 32645
diff changeset
    30
            if not (is_ratconst (Thm.dest_arg1 t)) then acc else
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
    31
            augment_norm (dest_ratconst (Thm.dest_arg1 t) >= @0)
32832
4602cb6ae0b5 additional fixes in normarith.ML due to FuncFun and FuncUtil changes
Philipp Meyer
parents: 32645
diff changeset
    32
                      (Thm.dest_arg t) acc
74572
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
    33
   | _ => augment_norm true t acc
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
    34
63211
0bec0d1d9998 more adhoc overloading;
wenzelm
parents: 63205
diff changeset
    35
 val cterm_lincomb_neg = FuncUtil.Ctermfunc.map (K ~)
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
    36
 fun cterm_lincomb_cmul c t =
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
    37
    if c = @0 then FuncUtil.Ctermfunc.empty else FuncUtil.Ctermfunc.map (fn _ => fn x => x * c) t
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
    38
 fun cterm_lincomb_add l r = FuncUtil.Ctermfunc.combine (curry op +) (fn x => x = @0) l r
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
    39
 fun cterm_lincomb_sub l r = cterm_lincomb_add l (cterm_lincomb_neg r)
32832
4602cb6ae0b5 additional fixes in normarith.ML due to FuncFun and FuncUtil changes
Philipp Meyer
parents: 32645
diff changeset
    40
 fun cterm_lincomb_eq l r = FuncUtil.Ctermfunc.is_empty (cterm_lincomb_sub l r)
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
    41
44454
6f28f96a09bf avoid warnings
huffman
parents: 43333
diff changeset
    42
(*
63211
0bec0d1d9998 more adhoc overloading;
wenzelm
parents: 63205
diff changeset
    43
 val int_lincomb_neg = FuncUtil.Intfunc.map (K ~)
44454
6f28f96a09bf avoid warnings
huffman
parents: 43333
diff changeset
    44
*)
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
    45
 fun int_lincomb_cmul c t =
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
    46
    if c = @0 then FuncUtil.Intfunc.empty else FuncUtil.Intfunc.map (fn _ => fn x => x * c) t
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
    47
 fun int_lincomb_add l r = FuncUtil.Intfunc.combine (curry op +) (fn x => x = @0) l r
44454
6f28f96a09bf avoid warnings
huffman
parents: 43333
diff changeset
    48
(*
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
    49
 fun int_lincomb_sub l r = int_lincomb_add l (int_lincomb_neg r)
32832
4602cb6ae0b5 additional fixes in normarith.ML due to FuncFun and FuncUtil changes
Philipp Meyer
parents: 32645
diff changeset
    50
 fun int_lincomb_eq l r = FuncUtil.Intfunc.is_empty (int_lincomb_sub l r)
44454
6f28f96a09bf avoid warnings
huffman
parents: 43333
diff changeset
    51
*)
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
    52
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
    53
fun vector_lincomb t = case Thm.term_of t of
74572
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
    54
   \<^Const_>\<open>plus _ for _ _\<close> =>
32832
4602cb6ae0b5 additional fixes in normarith.ML due to FuncFun and FuncUtil changes
Philipp Meyer
parents: 32645
diff changeset
    55
    cterm_lincomb_add (vector_lincomb (Thm.dest_arg1 t)) (vector_lincomb (Thm.dest_arg t))
74572
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
    56
 | \<^Const_>\<open>minus _ for _ _\<close> =>
32832
4602cb6ae0b5 additional fixes in normarith.ML due to FuncFun and FuncUtil changes
Philipp Meyer
parents: 32645
diff changeset
    57
    cterm_lincomb_sub (vector_lincomb (Thm.dest_arg1 t)) (vector_lincomb (Thm.dest_arg t))
74572
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
    58
 | \<^Const_>\<open>scaleR _ for _ _\<close> =>
32832
4602cb6ae0b5 additional fixes in normarith.ML due to FuncFun and FuncUtil changes
Philipp Meyer
parents: 32645
diff changeset
    59
    cterm_lincomb_cmul (dest_ratconst (Thm.dest_arg1 t)) (vector_lincomb (Thm.dest_arg t))
74572
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
    60
 | \<^Const_>\<open>uminus _ for _\<close> =>
32832
4602cb6ae0b5 additional fixes in normarith.ML due to FuncFun and FuncUtil changes
Philipp Meyer
parents: 32645
diff changeset
    61
     cterm_lincomb_neg (vector_lincomb (Thm.dest_arg t))
31445
c8a474a919a7 generalize norm method to work over class real_normed_vector
huffman
parents: 31344
diff changeset
    62
(* FIXME: how should we handle numerals?
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
    63
 | Const(@ {const_name vec},_)$_ =>
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
    64
   let
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
    65
     val b = ((snd o HOLogic.dest_number o term_of o Thm.dest_arg) t = 0
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
    66
               handle TERM _=> false)
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
    67
   in if b then FuncUtil.Ctermfunc.onefunc (t,@1)
32832
4602cb6ae0b5 additional fixes in normarith.ML due to FuncFun and FuncUtil changes
Philipp Meyer
parents: 32645
diff changeset
    68
      else FuncUtil.Ctermfunc.empty
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
    69
   end
31445
c8a474a919a7 generalize norm method to work over class real_normed_vector
huffman
parents: 31344
diff changeset
    70
*)
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
    71
 | _ => FuncUtil.Ctermfunc.onefunc (t,@1)
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
    72
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
    73
 fun vector_lincombs ts =
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
    74
  fold_rev
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
    75
   (fn t => fn fns => case AList.lookup (op aconvc) fns t of
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
    76
     NONE =>
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
    77
       let val f = vector_lincomb t
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
    78
       in case find_first (fn (_,f') => cterm_lincomb_eq f f') fns of
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
    79
           SOME (_,f') => (t,f') :: fns
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
    80
         | NONE => (t,f) :: fns
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
    81
       end
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
    82
   | SOME _ => fns) ts []
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
    83
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
    84
fun replacenegnorms cv t = case Thm.term_of t of
74572
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
    85
  \<^Const_>\<open>plus \<^typ>\<open>real\<close> for _ _\<close> => binop_conv (replacenegnorms cv) t
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
    86
| \<^Const_>\<open>times \<^typ>\<open>real\<close> for _ _\<close> =>
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
    87
    if dest_ratconst (Thm.dest_arg1 t) < @0 then arg_conv cv t else Thm.reflexive t
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36939
diff changeset
    88
| _ => Thm.reflexive t
44454
6f28f96a09bf avoid warnings
huffman
parents: 43333
diff changeset
    89
(*
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
    90
fun flip v eq =
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
    91
  if FuncUtil.Ctermfunc.defined eq v
63211
0bec0d1d9998 more adhoc overloading;
wenzelm
parents: 63205
diff changeset
    92
  then FuncUtil.Ctermfunc.update (v, ~ (FuncUtil.Ctermfunc.apply eq v)) eq else eq
44454
6f28f96a09bf avoid warnings
huffman
parents: 43333
diff changeset
    93
*)
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
    94
fun allsubsets s = case s of
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
    95
  [] => [[]]
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
    96
|(a::t) => let val res = allsubsets t in
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
    97
               map (cons a) res @ res end
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
    98
fun evaluate env lin =
63198
c583ca33076a ad-hoc overloading for standard operations on type Rat.rat;
wenzelm
parents: 61075
diff changeset
    99
 FuncUtil.Intfunc.fold (fn (x,c) => fn s => s + c * (FuncUtil.Intfunc.apply env x))
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   100
   lin @0
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   101
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   102
fun solve (vs,eqs) = case (vs,eqs) of
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   103
  ([],[]) => SOME (FuncUtil.Intfunc.onefunc (0,@1))
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   104
 |(_,eq::oeqs) =>
32832
4602cb6ae0b5 additional fixes in normarith.ML due to FuncFun and FuncUtil changes
Philipp Meyer
parents: 32645
diff changeset
   105
   (case filter (member (op =) vs) (FuncUtil.Intfunc.dom eq) of (*FIXME use find_first here*)
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   106
     [] => NONE
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   107
    | v::_ =>
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   108
       if FuncUtil.Intfunc.defined eq v
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   109
       then
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   110
        let
32832
4602cb6ae0b5 additional fixes in normarith.ML due to FuncFun and FuncUtil changes
Philipp Meyer
parents: 32645
diff changeset
   111
         val c = FuncUtil.Intfunc.apply eq v
63211
0bec0d1d9998 more adhoc overloading;
wenzelm
parents: 63205
diff changeset
   112
         val vdef = int_lincomb_cmul (~ (Rat.inv c)) eq
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   113
         fun eliminate eqn = if not (FuncUtil.Intfunc.defined eqn v) then eqn
32832
4602cb6ae0b5 additional fixes in normarith.ML due to FuncFun and FuncUtil changes
Philipp Meyer
parents: 32645
diff changeset
   114
                             else int_lincomb_add (int_lincomb_cmul (FuncUtil.Intfunc.apply eqn v) vdef) eqn
33040
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33039
diff changeset
   115
        in (case solve (remove (op =) v vs, map eliminate oeqs) of
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   116
            NONE => NONE
32832
4602cb6ae0b5 additional fixes in normarith.ML due to FuncFun and FuncUtil changes
Philipp Meyer
parents: 32645
diff changeset
   117
          | SOME soln => SOME (FuncUtil.Intfunc.update (v, evaluate soln (FuncUtil.Intfunc.delete_safe v vdef)) soln))
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   118
        end
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   119
       else NONE)
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   120
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   121
fun combinations k l = if k = 0 then [[]] else
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   122
 case l of
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   123
  [] => []
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   124
| h::t => map (cons h) (combinations (k - 1) t) @ combinations k t
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   125
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   126
fun vertices vs eqs =
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   127
 let
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   128
  fun vertex cmb = case solve(vs,cmb) of
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   129
    NONE => NONE
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   130
   | SOME soln => SOME (map (fn v => FuncUtil.Intfunc.tryapplyd soln v @0) vs)
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   131
  val rawvs = map_filter vertex (combinations (length vs) eqs)
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   132
  val unset = filter (forall (fn c => c >= @0)) rawvs
63198
c583ca33076a ad-hoc overloading for standard operations on type Rat.rat;
wenzelm
parents: 61075
diff changeset
   133
 in fold_rev (insert (eq_list op =)) unset []
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   134
 end
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   135
63198
c583ca33076a ad-hoc overloading for standard operations on type Rat.rat;
wenzelm
parents: 61075
diff changeset
   136
val subsumes = eq_list (fn (x, y) => Rat.abs x <= Rat.abs y)
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   137
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   138
fun subsume todo dun = case todo of
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   139
 [] => dun
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   140
|v::ovs =>
40718
4d7211968607 eliminated some clones of eq_list;
wenzelm
parents: 39032
diff changeset
   141
   let val dun' = if exists (fn w => subsumes (w, v)) dun then dun
4d7211968607 eliminated some clones of eq_list;
wenzelm
parents: 39032
diff changeset
   142
                  else v:: filter (fn w => not (subsumes (v, w))) dun
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   143
   in subsume ovs dun'
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   144
   end;
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   145
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   146
fun match_mp PQ P = P RS PQ;
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   147
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   148
fun cterm_of_rat x =
63201
f151704c08e4 tuned signature;
wenzelm
parents: 63198
diff changeset
   149
let val (a, b) = Rat.dest x
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   150
in
74572
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
   151
  if b = 1 then Numeral.mk_cnumber \<^ctyp>\<open>real\<close> a
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
   152
  else
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
   153
    \<^instantiate>\<open>
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
   154
        a = \<open>Numeral.mk_cnumber \<^ctyp>\<open>real\<close> a\<close> and
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
   155
        b = \<open>Numeral.mk_cnumber \<^ctyp>\<open>real\<close> b\<close>
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
   156
      in cterm \<open>a / b\<close> for a b :: real\<close>
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   157
end;
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   158
60801
7664e0916eec tuned signature;
wenzelm
parents: 60754
diff changeset
   159
fun norm_cmul_rule c th = Thm.instantiate' [] [SOME (cterm_of_rat c)] (th RS @{thm norm_cmul_rule_thm});
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   160
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   161
fun norm_add_rule th1 th2 = [th1, th2] MRS @{thm norm_add_rule_thm};
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   162
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   163
  (* I think here the static context should be sufficient!! *)
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   164
fun inequality_canon_rule ctxt =
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   165
 let
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   166
  (* FIXME : Should be computed statically!! *)
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   167
  val real_poly_conv =
36753
5cf4e9128f22 renamed Normalizer to the more specific Semiring_Normalizer
haftmann
parents: 36751
diff changeset
   168
    Semiring_Normalizer.semiring_normalize_wrapper ctxt
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69064
diff changeset
   169
     (the (Semiring_Normalizer.match ctxt \<^cterm>\<open>(0::real) + 1\<close>))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   170
 in
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   171
  fconv_rule (arg_conv ((rewr_conv @{thm ge_iff_diff_ge_0}) then_conv
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   172
    arg_conv (Numeral_Simprocs.field_comp_conv ctxt then_conv real_poly_conv)))
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   173
end;
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   174
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   175
 val apply_pth1 = rewr_conv @{thm pth_1};
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   176
 val apply_pth2 = rewr_conv @{thm pth_2};
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   177
 val apply_pth3 = rewr_conv @{thm pth_3};
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   178
 val apply_pth4 = rewrs_conv @{thms pth_4};
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   179
 val apply_pth5 = rewr_conv @{thm pth_5};
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   180
 val apply_pth6 = rewr_conv @{thm pth_6};
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   181
 val apply_pth7 = rewrs_conv @{thms pth_7};
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   182
 fun apply_pth8 ctxt =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   183
  rewr_conv @{thm pth_8} then_conv
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   184
  arg1_conv (Numeral_Simprocs.field_comp_conv ctxt) then_conv
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   185
  (try_conv (rewr_conv (mk_meta_eq @{thm scaleR_zero_left})));
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   186
 fun apply_pth9 ctxt =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   187
  rewrs_conv @{thms pth_9} then_conv
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   188
  arg1_conv (arg1_conv (Numeral_Simprocs.field_comp_conv ctxt));
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   189
 val apply_ptha = rewr_conv @{thm pth_a};
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   190
 val apply_pthb = rewrs_conv @{thms pth_b};
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   191
 val apply_pthc = rewrs_conv @{thms pth_c};
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   192
 val apply_pthd = try_conv (rewr_conv @{thm pth_d});
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   193
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   194
fun headvector t = case t of
74572
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
   195
   \<^Const_>\<open>plus _ for \<^Const_>\<open>scaleR _ for _ v\<close> _\<close> => v
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
   196
 | \<^Const_>\<open>scaleR _ for _ v\<close> => v
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   197
 | _ => error "headvector: non-canonical term"
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   198
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   199
fun vector_cmul_conv ctxt ct =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   200
   ((apply_pth5 then_conv arg1_conv (Numeral_Simprocs.field_comp_conv ctxt)) else_conv
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   201
    (apply_pth6 then_conv binop_conv (vector_cmul_conv ctxt))) ct
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   202
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   203
fun vector_add_conv ctxt ct = apply_pth7 ct
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   204
 handle CTERM _ =>
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   205
  (apply_pth8 ctxt ct
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   206
   handle CTERM _ =>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   207
    (case Thm.term_of ct of
74572
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
   208
     \<^Const_>\<open>plus _ for lt rt\<close> =>
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   209
      let
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   210
       val l = headvector lt
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   211
       val r = headvector rt
35408
b48ab741683b modernized structure Term_Ord;
wenzelm
parents: 33043
diff changeset
   212
      in (case Term_Ord.fast_term_ord (l,r) of
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   213
         LESS => (apply_pthb then_conv arg_conv (vector_add_conv ctxt)
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   214
                  then_conv apply_pthd) ct
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   215
        | GREATER => (apply_pthc then_conv arg_conv (vector_add_conv ctxt)
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   216
                     then_conv apply_pthd) ct
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   217
        | EQUAL => (apply_pth9 ctxt then_conv
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   218
                ((apply_ptha then_conv (vector_add_conv ctxt)) else_conv
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   219
              arg_conv (vector_add_conv ctxt) then_conv apply_pthd)) ct)
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   220
      end
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36939
diff changeset
   221
     | _ => Thm.reflexive ct))
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   222
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   223
fun vector_canon_conv ctxt ct = case Thm.term_of ct of
74572
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
   224
  \<^Const_>\<open>plus _ for _ _\<close> =>
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   225
  let
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   226
   val ((p,l),r) = Thm.dest_comb ct |>> Thm.dest_comb
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   227
   val lth = vector_canon_conv ctxt l
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   228
   val rth = vector_canon_conv ctxt r
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   229
   val th = Drule.binop_cong_rule p lth rth
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   230
  in fconv_rule (arg_conv (vector_add_conv ctxt)) th end
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   231
74572
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
   232
| \<^Const_>\<open>scaleR _ for _ _\<close> =>
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   233
  let
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   234
   val (p,r) = Thm.dest_comb ct
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   235
   val rth = Drule.arg_cong_rule p (vector_canon_conv ctxt r)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   236
  in fconv_rule (arg_conv (apply_pth4 else_conv (vector_cmul_conv ctxt))) rth
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   237
  end
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   238
74572
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
   239
| \<^Const_>\<open>minus _ for _ _\<close> => (apply_pth2 then_conv (vector_canon_conv ctxt)) ct
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   240
74572
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
   241
| \<^Const_>\<open>uminus _ for _\<close> => (apply_pth3 then_conv (vector_canon_conv ctxt)) ct
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   242
31445
c8a474a919a7 generalize norm method to work over class real_normed_vector
huffman
parents: 31344
diff changeset
   243
(* FIXME
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   244
| Const(@{const_name vec},_)$n =>
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   245
  let val n = Thm.dest_arg ct
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   246
  in if is_ratconst n andalso not (dest_ratconst n =/ @0)
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36939
diff changeset
   247
     then Thm.reflexive ct else apply_pth1 ct
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   248
  end
31445
c8a474a919a7 generalize norm method to work over class real_normed_vector
huffman
parents: 31344
diff changeset
   249
*)
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   250
| _ => apply_pth1 ct
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   251
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   252
fun norm_canon_conv ctxt ct = case Thm.term_of ct of
74572
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
   253
   \<^Const_>\<open>norm _ for _\<close> => arg_conv (vector_canon_conv ctxt) ct
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   254
 | _ => raise CTERM ("norm_canon_conv", [ct])
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   255
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   256
fun int_flip v eq =
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   257
  if FuncUtil.Intfunc.defined eq v
63211
0bec0d1d9998 more adhoc overloading;
wenzelm
parents: 63205
diff changeset
   258
  then FuncUtil.Intfunc.update (v, ~ (FuncUtil.Intfunc.apply eq v)) eq else eq;
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   259
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   260
local
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   261
 val concl = Thm.dest_arg o Thm.cprop_of
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   262
 fun real_vector_combo_prover ctxt translator (nubs,ges,gts) =
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   263
  let
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   264
   (* FIXME: Should be computed statically!!*)
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   265
   val real_poly_conv =
36753
5cf4e9128f22 renamed Normalizer to the more specific Semiring_Normalizer
haftmann
parents: 36751
diff changeset
   266
      Semiring_Normalizer.semiring_normalize_wrapper ctxt
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69064
diff changeset
   267
       (the (Semiring_Normalizer.match ctxt \<^cterm>\<open>(0::real) + 1\<close>))
32832
4602cb6ae0b5 additional fixes in normarith.ML due to FuncFun and FuncUtil changes
Philipp Meyer
parents: 32645
diff changeset
   268
   val sources = map (Thm.dest_arg o Thm.dest_arg1 o concl) nubs
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   269
   val rawdests = fold_rev (find_normedterms o Thm.dest_arg o concl) (ges @ gts) []
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   270
   val _ = if not (forall fst rawdests) then error "real_vector_combo_prover: Sanity check"
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   271
           else ()
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   272
   val dests = distinct (op aconvc) (map snd rawdests)
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   273
   val srcfuns = map vector_lincomb sources
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   274
   val destfuns = map vector_lincomb dests
33042
ddf1f03a9ad9 curried union as canonical list operation
haftmann
parents: 33039
diff changeset
   275
   val vvs = fold_rev (union (op aconvc) o FuncUtil.Ctermfunc.dom) (srcfuns @ destfuns) []
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   276
   val n = length srcfuns
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   277
   val nvs = 1 upto n
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   278
   val srccombs = srcfuns ~~ nvs
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   279
   fun consider d =
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   280
    let
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   281
     fun coefficients x =
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   282
      let
63211
0bec0d1d9998 more adhoc overloading;
wenzelm
parents: 63205
diff changeset
   283
       val inp =
0bec0d1d9998 more adhoc overloading;
wenzelm
parents: 63205
diff changeset
   284
        if FuncUtil.Ctermfunc.defined d x
0bec0d1d9998 more adhoc overloading;
wenzelm
parents: 63205
diff changeset
   285
        then FuncUtil.Intfunc.onefunc (0, ~ (FuncUtil.Ctermfunc.apply d x))
0bec0d1d9998 more adhoc overloading;
wenzelm
parents: 63205
diff changeset
   286
        else FuncUtil.Intfunc.empty
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   287
      in fold_rev (fn (f,v) => fn g => if FuncUtil.Ctermfunc.defined f x then FuncUtil.Intfunc.update (v, FuncUtil.Ctermfunc.apply f x) g else g) srccombs inp
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   288
      end
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   289
     val equations = map coefficients vvs
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   290
     val inequalities = map (fn n => FuncUtil.Intfunc.onefunc (n,@1)) nvs
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   291
     fun plausiblevertices f =
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   292
      let
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   293
       val flippedequations = map (fold_rev int_flip f) equations
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   294
       val constraints = flippedequations @ inequalities
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   295
       val rawverts = vertices nvs constraints
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   296
       fun check_solution v =
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   297
        let
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   298
          val f = fold_rev FuncUtil.Intfunc.update (nvs ~~ v) (FuncUtil.Intfunc.onefunc (0, @1))
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   299
        in forall (fn e => evaluate f e = @0) flippedequations
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   300
        end
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   301
       val goodverts = filter check_solution rawverts
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   302
       val signfixups = map (fn n => if member (op =) f n then ~1 else 1) nvs
63201
f151704c08e4 tuned signature;
wenzelm
parents: 63198
diff changeset
   303
      in map (map2 (fn s => fn c => Rat.of_int s * c) signfixups) goodverts
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   304
      end
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   305
     val allverts = fold_rev append (map plausiblevertices (allsubsets nvs)) []
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   306
    in subsume allverts []
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   307
    end
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   308
   fun compute_ineq v =
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   309
    let
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   310
     val ths = map_filter (fn (v,t) => if v = @0 then NONE
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   311
                                     else SOME(norm_cmul_rule v t))
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   312
                            (v ~~ nubs)
32402
5731300da417 added further conversions and conversionals
boehmes
parents: 31446
diff changeset
   313
     fun end_itlist f xs = split_last xs |> uncurry (fold_rev f)
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   314
    in inequality_canon_rule ctxt (end_itlist norm_add_rule ths)
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   315
    end
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   316
   val ges' = map_filter (try compute_ineq) (fold_rev (append o consider) destfuns []) @
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   317
                 map (inequality_canon_rule ctxt) nubs @ ges
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   318
   val zerodests = filter
32832
4602cb6ae0b5 additional fixes in normarith.ML due to FuncFun and FuncUtil changes
Philipp Meyer
parents: 32645
diff changeset
   319
        (fn t => null (FuncUtil.Ctermfunc.dom (vector_lincomb t))) (map snd rawdests)
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   320
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   321
  in fst (RealArith.real_linear_prover translator
74626
9a1f4a7ddf9e clarified antiquotations;
wenzelm
parents: 74572
diff changeset
   322
        (zerodests |> map (fn t =>
9a1f4a7ddf9e clarified antiquotations;
wenzelm
parents: 74572
diff changeset
   323
          \<^instantiate>\<open>'a = \<open>Thm.ctyp_of_cterm t\<close> in
9a1f4a7ddf9e clarified antiquotations;
wenzelm
parents: 74572
diff changeset
   324
            lemma \<open>norm (0::'a::real_normed_vector) = 0\<close> by simp\<close>),
74544
9864ab4c20ce more accurate treatment of context;
wenzelm
parents: 74282
diff changeset
   325
        map (fconv_rule (try_conv (Conv.top_sweep_conv norm_canon_conv ctxt) then_conv
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   326
                       arg_conv (arg_conv real_poly_conv))) ges',
74544
9864ab4c20ce more accurate treatment of context;
wenzelm
parents: 74282
diff changeset
   327
        map (fconv_rule (try_conv (Conv.top_sweep_conv norm_canon_conv ctxt) then_conv
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   328
                       arg_conv (arg_conv real_poly_conv))) gts))
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   329
  end
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   330
in val real_vector_combo_prover = real_vector_combo_prover
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   331
end;
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   332
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   333
local
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   334
 val pth = @{thm norm_imp_pos_and_ge}
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   335
 val norm_mp = match_mp pth
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   336
 val concl = Thm.dest_arg o Thm.cprop_of
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   337
 fun conjunct1 th = th RS @{thm conjunct1}
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   338
 fun conjunct2 th = th RS @{thm conjunct2}
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   339
fun real_vector_ineq_prover ctxt translator (ges,gts) =
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   340
 let
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   341
(*   val _ = error "real_vector_ineq_prover: pause" *)
32832
4602cb6ae0b5 additional fixes in normarith.ML due to FuncFun and FuncUtil changes
Philipp Meyer
parents: 32645
diff changeset
   342
  val ntms = fold_rev find_normedterms (map (Thm.dest_arg o concl) (ges @ gts)) []
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   343
  val lctab = vector_lincombs (map snd (filter (not o fst) ntms))
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   344
  val (fxns, ctxt') = Variable.variant_fixes (replicate (length lctab) "x") ctxt
61075
f6b0d827240e tuned -- avoid slightly odd @{cpat};
wenzelm
parents: 60949
diff changeset
   345
  fun mk_norm t =
74572
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
   346
    let val T = Thm.ctyp_of_cterm t
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
   347
    in \<^instantiate>\<open>'a = T and t in cterm \<open>norm t\<close>\<close> end
61075
f6b0d827240e tuned -- avoid slightly odd @{cpat};
wenzelm
parents: 60949
diff changeset
   348
  fun mk_equals l r =
74572
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
   349
    let val T = Thm.ctyp_of_cterm l
08b4292abe2b more antiquotations;
wenzelm
parents: 74544
diff changeset
   350
    in \<^instantiate>\<open>'a = T and l and r in cterm \<open>l \<equiv> r\<close>\<close> end
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69064
diff changeset
   351
  val asl = map2 (fn (t,_) => fn n => Thm.assume (mk_equals (mk_norm t) (Thm.cterm_of ctxt' (Free(n,\<^typ>\<open>real\<close>))))) lctab fxns
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   352
  val replace_conv = try_conv (rewrs_conv asl)
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   353
  val replace_rule = fconv_rule (funpow 2 arg_conv (replacenegnorms replace_conv))
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   354
  val ges' =
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   355
       fold_rev (fn th => fn ths => conjunct1(norm_mp th)::ths)
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   356
              asl (map replace_rule ges)
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   357
  val gts' = map replace_rule gts
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   358
  val nubs = map (conjunct2 o norm_mp) asl
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   359
  val th1 = real_vector_combo_prover ctxt' translator (nubs,ges',gts')
60949
ccbf9379e355 added Thm.chyps_of;
wenzelm
parents: 60801
diff changeset
   360
  val shs = filter (member (fn (t,th) => t aconvc Thm.cprop_of th) asl) (Thm.chyps_of th1)
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36939
diff changeset
   361
  val th11 = hd (Variable.export ctxt' ctxt [fold Thm.implies_intr shs th1])
32832
4602cb6ae0b5 additional fixes in normarith.ML due to FuncFun and FuncUtil changes
Philipp Meyer
parents: 32645
diff changeset
   362
  val cps = map (swap o Thm.dest_equals) (cprems_of th11)
74282
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 69597
diff changeset
   363
  val th12 = Drule.instantiate_normalize (TVars.empty, Vars.make (map (apfst (dest_Var o Thm.term_of)) cps)) th11
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36939
diff changeset
   364
  val th13 = fold Thm.elim_implies (map (Thm.reflexive o snd) cps) th12;
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   365
 in hd (Variable.export ctxt' ctxt [th13])
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   366
 end
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   367
in val real_vector_ineq_prover = real_vector_ineq_prover
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   368
end;
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   369
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   370
local
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   371
 val rawrule = fconv_rule (arg_conv (rewr_conv @{thm real_eq_0_iff_le_ge_0}))
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   372
 fun conj_pair th = (th RS @{thm conjunct1}, th RS @{thm conjunct2})
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   373
  (* FIXME: Lookup in the context every time!!! Fix this !!!*)
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   374
 fun splitequation ctxt th acc =
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   375
  let
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   376
   val real_poly_neg_conv = #neg
36753
5cf4e9128f22 renamed Normalizer to the more specific Semiring_Normalizer
haftmann
parents: 36751
diff changeset
   377
       (Semiring_Normalizer.semiring_normalizers_ord_wrapper ctxt
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69064
diff changeset
   378
        (the (Semiring_Normalizer.match ctxt \<^cterm>\<open>(0::real) + 1\<close>)) Thm.term_ord)
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   379
   val (th1,th2) = conj_pair(rawrule th)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   380
  in th1::fconv_rule (arg_conv (arg_conv (real_poly_neg_conv ctxt))) th2::acc
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   381
  end
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   382
in fun real_vector_prover ctxt _ translator (eqs,ges,gts) =
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   383
     (real_vector_ineq_prover ctxt translator
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   384
         (fold_rev (splitequation ctxt) eqs ges,gts), RealArith.Trivial)
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   385
end;
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   386
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   387
  fun init_conv ctxt =
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   388
   Simplifier.rewrite (put_simpset HOL_basic_ss ctxt
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   389
    addsimps ([(*@{thm vec_0}, @{thm vec_1},*) @{thm dist_norm}, @{thm right_minus},
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   390
      @{thm diff_self}, @{thm norm_zero}] @ @{thms arithmetic_simps} @ @{thms norm_pths}))
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   391
   then_conv Numeral_Simprocs.field_comp_conv ctxt
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48112
diff changeset
   392
   then_conv nnf_conv ctxt
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   393
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   394
 fun pure ctxt = fst o RealArith.gen_prover_real_arith ctxt (real_vector_prover ctxt);
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   395
 fun norm_arith ctxt ct =
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   396
  let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   397
   val ctxt' = Variable.declare_term (Thm.term_of ct) ctxt
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   398
   val th = init_conv ctxt' ct
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69064
diff changeset
   399
  in Thm.equal_elim (Drule.arg_cong_rule \<^cterm>\<open>Trueprop\<close> (Thm.symmetric th))
32832
4602cb6ae0b5 additional fixes in normarith.ML due to FuncFun and FuncUtil changes
Philipp Meyer
parents: 32645
diff changeset
   400
                (pure ctxt' (Thm.rhs_of th))
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   401
 end
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   402
36938
278029c8a462 tuned header;
wenzelm
parents: 36937
diff changeset
   403
 fun norm_arith_tac ctxt =
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42361
diff changeset
   404
   clarify_tac (put_claset HOL_cs ctxt) THEN'
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51717
diff changeset
   405
   Object_Logic.full_atomize_tac ctxt THEN'
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   406
   CSUBGOAL ( fn (p,i) => resolve_tac ctxt [norm_arith ctxt (Thm.dest_arg p )] i);
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents:
diff changeset
   407
31445
c8a474a919a7 generalize norm method to work over class real_normed_vector
huffman
parents: 31344
diff changeset
   408
end;