(* Title: Library/normarith.ML
Author: Amine Chaieb, University of Cambridge
Description: A simple decision procedure for linear problems in euclidean space
*)
(* Now the norm procedure for euclidean spaces *)
signature NORM_ARITH =
sig
val norm_arith : Proof.context -> conv
val norm_arith_tac : Proof.context -> int -> tactic
end
structure NormArith : NORM_ARITH =
struct
open Conv Thm Conv2;
val bool_eq = op = : bool *bool -> bool
fun dest_ratconst t = case term_of t of
Const(@{const_name divide}, _)$a$b => Rat.rat_of_quotient(HOLogic.dest_number a |> snd, HOLogic.dest_number b |> snd)
| Const(@{const_name inverse}, _)$a => Rat.rat_of_quotient(1, HOLogic.dest_number a |> snd)
| _ => Rat.rat_of_int (HOLogic.dest_number (term_of t) |> snd)
fun is_ratconst t = can dest_ratconst t
fun augment_norm b t acc = case term_of t of
Const(@{const_name norm}, _) $ _ => insert (eq_pair bool_eq (op aconvc)) (b,dest_arg t) acc
| _ => acc
fun find_normedterms t acc = case term_of t of
@{term "op + :: real => _"}$_$_ =>
find_normedterms (dest_arg1 t) (find_normedterms (dest_arg t) acc)
| @{term "op * :: real => _"}$_$n =>
if not (is_ratconst (dest_arg1 t)) then acc else
augment_norm (dest_ratconst (dest_arg1 t) >=/ Rat.zero)
(dest_arg t) acc
| _ => augment_norm true t acc
val cterm_lincomb_neg = Ctermfunc.mapf Rat.neg
fun cterm_lincomb_cmul c t =
if c =/ Rat.zero then Ctermfunc.undefined else Ctermfunc.mapf (fn x => x */ c) t
fun cterm_lincomb_add l r = Ctermfunc.combine (curry op +/) (fn x => x =/ Rat.zero) l r
fun cterm_lincomb_sub l r = cterm_lincomb_add l (cterm_lincomb_neg r)
fun cterm_lincomb_eq l r = Ctermfunc.is_undefined (cterm_lincomb_sub l r)
val int_lincomb_neg = Intfunc.mapf Rat.neg
fun int_lincomb_cmul c t =
if c =/ Rat.zero then Intfunc.undefined else Intfunc.mapf (fn x => x */ c) t
fun int_lincomb_add l r = Intfunc.combine (curry op +/) (fn x => x =/ Rat.zero) l r
fun int_lincomb_sub l r = int_lincomb_add l (int_lincomb_neg r)
fun int_lincomb_eq l r = Intfunc.is_undefined (int_lincomb_sub l r)
fun vector_lincomb t = case term_of t of
Const(@{const_name plus},Type("fun",[Type("Finite_Cartesian_Product.^",_),_])) $ _ $ _ =>
cterm_lincomb_add (vector_lincomb (dest_arg1 t)) (vector_lincomb (dest_arg t))
| Const(@{const_name minus},Type("fun",[Type("Finite_Cartesian_Product.^",_),_])) $ _ $ _ =>
cterm_lincomb_sub (vector_lincomb (dest_arg1 t)) (vector_lincomb (dest_arg t))
| Const(@{const_name vector_scalar_mult},Type("fun",[Type("Finite_Cartesian_Product.^",_),_]))$_$_ =>
cterm_lincomb_cmul (dest_ratconst (dest_arg1 t)) (vector_lincomb (dest_arg t))
| Const(@{const_name uminus},Type("fun",[Type("Finite_Cartesian_Product.^",_),_]))$_ =>
cterm_lincomb_neg (vector_lincomb (dest_arg t))
| Const(@{const_name vec},_)$_ =>
let
val b = ((snd o HOLogic.dest_number o term_of o dest_arg) t = 0
handle TERM _=> false)
in if b then Ctermfunc.onefunc (t,Rat.one)
else Ctermfunc.undefined
end
| _ => Ctermfunc.onefunc (t,Rat.one)
fun vector_lincombs ts =
fold_rev
(fn t => fn fns => case AList.lookup (op aconvc) fns t of
NONE =>
let val f = vector_lincomb t
in case find_first (fn (_,f') => cterm_lincomb_eq f f') fns of
SOME (_,f') => (t,f') :: fns
| NONE => (t,f) :: fns
end
| SOME _ => fns) ts []
fun replacenegnorms cv t = case term_of t of
@{term "op + :: real => _"}$_$_ => binop_conv (replacenegnorms cv) t
| @{term "op * :: real => _"}$_$_ =>
if dest_ratconst (dest_arg1 t) </ Rat.zero then arg_conv cv t else reflexive t
| _ => reflexive t
fun flip v eq =
if Ctermfunc.defined eq v
then Ctermfunc.update (v, Rat.neg (Ctermfunc.apply eq v)) eq else eq
fun allsubsets s = case s of
[] => [[]]
|(a::t) => let val res = allsubsets t in
map (cons a) res @ res end
fun evaluate env lin =
Intfunc.fold (fn (x,c) => fn s => s +/ c */ (Intfunc.apply env x))
lin Rat.zero
fun solve (vs,eqs) = case (vs,eqs) of
([],[]) => SOME (Intfunc.onefunc (0,Rat.one))
|(_,eq::oeqs) =>
(case filter (member (op =) vs) (Intfunc.dom eq) of (*FIXME use find_first here*)
[] => NONE
| v::_ =>
if Intfunc.defined eq v
then
let
val c = Intfunc.apply eq v
val vdef = int_lincomb_cmul (Rat.neg (Rat.inv c)) eq
fun eliminate eqn = if not (Intfunc.defined eqn v) then eqn
else int_lincomb_add (int_lincomb_cmul (Intfunc.apply eqn v) vdef) eqn
in (case solve (vs \ v,map eliminate oeqs) of
NONE => NONE
| SOME soln => SOME (Intfunc.update (v, evaluate soln (Intfunc.undefine v vdef)) soln))
end
else NONE)
fun combinations k l = if k = 0 then [[]] else
case l of
[] => []
| h::t => map (cons h) (combinations (k - 1) t) @ combinations k t
fun forall2 p l1 l2 = case (l1,l2) of
([],[]) => true
| (h1::t1,h2::t2) => p h1 h2 andalso forall2 p t1 t2
| _ => false;
fun vertices vs eqs =
let
fun vertex cmb = case solve(vs,cmb) of
NONE => NONE
| SOME soln => SOME (map (fn v => Intfunc.tryapplyd soln v Rat.zero) vs)
val rawvs = map_filter vertex (combinations (length vs) eqs)
val unset = filter (forall (fn c => c >=/ Rat.zero)) rawvs
in fold_rev (insert (uncurry (forall2 (curry op =/)))) unset []
end
fun subsumes l m = forall2 (fn x => fn y => Rat.abs x <=/ Rat.abs y) l m
fun subsume todo dun = case todo of
[] => dun
|v::ovs =>
let val dun' = if exists (fn w => subsumes w v) dun then dun
else v::(filter (fn w => not(subsumes v w)) dun)
in subsume ovs dun'
end;
fun match_mp PQ P = P RS PQ;
fun cterm_of_rat x =
let val (a, b) = Rat.quotient_of_rat x
in
if b = 1 then Numeral.mk_cnumber @{ctyp "real"} a
else Thm.capply (Thm.capply @{cterm "op / :: real => _"}
(Numeral.mk_cnumber @{ctyp "real"} a))
(Numeral.mk_cnumber @{ctyp "real"} b)
end;
fun norm_cmul_rule c th = instantiate' [] [SOME (cterm_of_rat c)] (th RS @{thm norm_cmul_rule_thm});
fun norm_add_rule th1 th2 = [th1, th2] MRS @{thm norm_add_rule_thm};
(* I think here the static context should be sufficient!! *)
fun inequality_canon_rule ctxt =
let
(* FIXME : Should be computed statically!! *)
val real_poly_conv =
Normalizer.semiring_normalize_wrapper ctxt
(valOf (NormalizerData.match ctxt @{cterm "(0::real) + 1"}))
in fconv_rule (arg_conv ((rewr_conv @{thm ge_iff_diff_ge_0}) then_conv arg_conv (field_comp_conv then_conv real_poly_conv)))
end;
fun absc cv ct = case term_of ct of
Abs (v,_, _) =>
let val (x,t) = Thm.dest_abs (SOME v) ct
in Thm.abstract_rule ((fst o dest_Free o term_of) x) x (cv t)
end
| _ => all_conv ct;
fun sub_conv cv ct = (comb_conv cv else_conv absc cv) ct;
fun botc1 conv ct =
((sub_conv (botc1 conv)) then_conv (conv else_conv all_conv)) ct;
fun rewrs_conv eqs ct = first_conv (map rewr_conv eqs) ct;
val apply_pth1 = rewr_conv @{thm pth_1};
val apply_pth2 = rewr_conv @{thm pth_2};
val apply_pth3 = rewr_conv @{thm pth_3};
val apply_pth4 = rewrs_conv @{thms pth_4};
val apply_pth5 = rewr_conv @{thm pth_5};
val apply_pth6 = rewr_conv @{thm pth_6};
val apply_pth7 = rewrs_conv @{thms pth_7};
val apply_pth8 = rewr_conv @{thm pth_8} then_conv arg1_conv field_comp_conv then_conv (try_conv (rewr_conv (mk_meta_eq @{thm vector_smult_lzero})));
val apply_pth9 = rewrs_conv @{thms pth_9} then_conv arg1_conv (arg1_conv field_comp_conv);
val apply_ptha = rewr_conv @{thm pth_a};
val apply_pthb = rewrs_conv @{thms pth_b};
val apply_pthc = rewrs_conv @{thms pth_c};
val apply_pthd = try_conv (rewr_conv @{thm pth_d});
fun headvector t = case t of
Const(@{const_name plus}, Type("fun",[Type("Finite_Cartesian_Product.^",_),_]))$
(Const(@{const_name vector_scalar_mult}, _)$l$v)$r => v
| Const(@{const_name vector_scalar_mult}, _)$l$v => v
| _ => error "headvector: non-canonical term"
fun vector_cmul_conv ct =
((apply_pth5 then_conv arg1_conv field_comp_conv) else_conv
(apply_pth6 then_conv binop_conv vector_cmul_conv)) ct
fun vector_add_conv ct = apply_pth7 ct
handle CTERM _ =>
(apply_pth8 ct
handle CTERM _ =>
(case term_of ct of
Const(@{const_name plus},_)$lt$rt =>
let
val l = headvector lt
val r = headvector rt
in (case TermOrd.fast_term_ord (l,r) of
LESS => (apply_pthb then_conv arg_conv vector_add_conv
then_conv apply_pthd) ct
| GREATER => (apply_pthc then_conv arg_conv vector_add_conv
then_conv apply_pthd) ct
| EQUAL => (apply_pth9 then_conv
((apply_ptha then_conv vector_add_conv) else_conv
arg_conv vector_add_conv then_conv apply_pthd)) ct)
end
| _ => reflexive ct))
fun vector_canon_conv ct = case term_of ct of
Const(@{const_name plus},_)$_$_ =>
let
val ((p,l),r) = Thm.dest_comb ct |>> Thm.dest_comb
val lth = vector_canon_conv l
val rth = vector_canon_conv r
val th = Drule.binop_cong_rule p lth rth
in fconv_rule (arg_conv vector_add_conv) th end
| Const(@{const_name vector_scalar_mult}, _)$_$_ =>
let
val (p,r) = Thm.dest_comb ct
val rth = Drule.arg_cong_rule p (vector_canon_conv r)
in fconv_rule (arg_conv (apply_pth4 else_conv vector_cmul_conv)) rth
end
| Const(@{const_name minus},_)$_$_ => (apply_pth2 then_conv vector_canon_conv) ct
| Const(@{const_name uminus},_)$_ => (apply_pth3 then_conv vector_canon_conv) ct
| Const(@{const_name vec},_)$n =>
let val n = Thm.dest_arg ct
in if is_ratconst n andalso not (dest_ratconst n =/ Rat.zero)
then reflexive ct else apply_pth1 ct
end
| _ => apply_pth1 ct
fun norm_canon_conv ct = case term_of ct of
Const(@{const_name norm},_)$_ => arg_conv vector_canon_conv ct
| _ => raise CTERM ("norm_canon_conv", [ct])
fun fold_rev2 f [] [] z = z
| fold_rev2 f (x::xs) (y::ys) z = f x y (fold_rev2 f xs ys z)
| fold_rev2 f _ _ _ = raise UnequalLengths;
fun int_flip v eq =
if Intfunc.defined eq v
then Intfunc.update (v, Rat.neg (Intfunc.apply eq v)) eq else eq;
local
val pth_zero = @{thm "norm_0"}
val tv_n = (hd o tl o dest_ctyp o ctyp_of_term o dest_arg o dest_arg1 o dest_arg o cprop_of)
pth_zero
val concl = dest_arg o cprop_of
fun real_vector_combo_prover ctxt translator (nubs,ges,gts) =
let
(* FIXME: Should be computed statically!!*)
val real_poly_conv =
Normalizer.semiring_normalize_wrapper ctxt
(valOf (NormalizerData.match ctxt @{cterm "(0::real) + 1"}))
val sources = map (dest_arg o dest_arg1 o concl) nubs
val rawdests = fold_rev (find_normedterms o dest_arg o concl) (ges @ gts) []
val _ = if not (forall fst rawdests) then error "real_vector_combo_prover: Sanity check"
else ()
val dests = distinct (op aconvc) (map snd rawdests)
val srcfuns = map vector_lincomb sources
val destfuns = map vector_lincomb dests
val vvs = fold_rev (curry (gen_union op aconvc) o Ctermfunc.dom) (srcfuns @ destfuns) []
val n = length srcfuns
val nvs = 1 upto n
val srccombs = srcfuns ~~ nvs
fun consider d =
let
fun coefficients x =
let
val inp = if Ctermfunc.defined d x then Intfunc.onefunc (0, Rat.neg(Ctermfunc.apply d x))
else Intfunc.undefined
in fold_rev (fn (f,v) => fn g => if Ctermfunc.defined f x then Intfunc.update (v, Ctermfunc.apply f x) g else g) srccombs inp
end
val equations = map coefficients vvs
val inequalities = map (fn n => Intfunc.onefunc (n,Rat.one)) nvs
fun plausiblevertices f =
let
val flippedequations = map (fold_rev int_flip f) equations
val constraints = flippedequations @ inequalities
val rawverts = vertices nvs constraints
fun check_solution v =
let
val f = fold_rev2 (curry Intfunc.update) nvs v (Intfunc.onefunc (0, Rat.one))
in forall (fn e => evaluate f e =/ Rat.zero) flippedequations
end
val goodverts = filter check_solution rawverts
val signfixups = map (fn n => if n mem_int f then ~1 else 1) nvs
in map (map2 (fn s => fn c => Rat.rat_of_int s */ c) signfixups) goodverts
end
val allverts = fold_rev append (map plausiblevertices (allsubsets nvs)) []
in subsume allverts []
end
fun compute_ineq v =
let
val ths = map_filter (fn (v,t) => if v =/ Rat.zero then NONE
else SOME(norm_cmul_rule v t))
(v ~~ nubs)
in inequality_canon_rule ctxt (end_itlist norm_add_rule ths)
end
val ges' = map_filter (try compute_ineq) (fold_rev (append o consider) destfuns []) @
map (inequality_canon_rule ctxt) nubs @ ges
val zerodests = filter
(fn t => null (Ctermfunc.dom (vector_lincomb t))) (map snd rawdests)
in RealArith.real_linear_prover translator
(map (fn t => instantiate ([(tv_n,(hd o tl o dest_ctyp o ctyp_of_term) t)],[]) pth_zero)
zerodests,
map (fconv_rule (once_depth_conv (norm_canon_conv) then_conv
arg_conv (arg_conv real_poly_conv))) ges',
map (fconv_rule (once_depth_conv (norm_canon_conv) then_conv
arg_conv (arg_conv real_poly_conv))) gts)
end
in val real_vector_combo_prover = real_vector_combo_prover
end;
local
val pth = @{thm norm_imp_pos_and_ge}
val norm_mp = match_mp pth
val concl = dest_arg o cprop_of
fun conjunct1 th = th RS @{thm conjunct1}
fun conjunct2 th = th RS @{thm conjunct2}
fun C f x y = f y x
fun real_vector_ineq_prover ctxt translator (ges,gts) =
let
(* val _ = error "real_vector_ineq_prover: pause" *)
val ntms = fold_rev find_normedterms (map (dest_arg o concl) (ges @ gts)) []
val lctab = vector_lincombs (map snd (filter (not o fst) ntms))
val (fxns, ctxt') = Variable.variant_fixes (replicate (length lctab) "x") ctxt
fun mk_norm t = capply (instantiate_cterm' [SOME (ctyp_of_term t)] [] @{cpat "norm :: (?'a :: norm) => real"}) t
fun mk_equals l r = capply (capply (instantiate_cterm' [SOME (ctyp_of_term l)] [] @{cpat "op == :: ?'a =>_"}) l) r
val asl = map2 (fn (t,_) => fn n => assume (mk_equals (mk_norm t) (cterm_of (ProofContext.theory_of ctxt') (Free(n,@{typ real}))))) lctab fxns
val replace_conv = try_conv (rewrs_conv asl)
val replace_rule = fconv_rule (funpow 2 arg_conv (replacenegnorms replace_conv))
val ges' =
fold_rev (fn th => fn ths => conjunct1(norm_mp th)::ths)
asl (map replace_rule ges)
val gts' = map replace_rule gts
val nubs = map (conjunct2 o norm_mp) asl
val th1 = real_vector_combo_prover ctxt' translator (nubs,ges',gts')
val shs = filter (member (fn (t,th) => t aconvc cprop_of th) asl) (#hyps (crep_thm th1))
val th11 = hd (Variable.export ctxt' ctxt [fold implies_intr shs th1])
val cps = map (swap o dest_equals) (cprems_of th11)
val th12 = instantiate ([], cps) th11
val th13 = fold (C implies_elim) (map (reflexive o snd) cps) th12;
in hd (Variable.export ctxt' ctxt [th13])
end
in val real_vector_ineq_prover = real_vector_ineq_prover
end;
local
val rawrule = fconv_rule (arg_conv (rewr_conv @{thm real_eq_0_iff_le_ge_0}))
fun conj_pair th = (th RS @{thm conjunct1}, th RS @{thm conjunct2})
fun simple_cterm_ord t u = TermOrd.term_ord (term_of t, term_of u) = LESS;
(* FIXME: Lookup in the context every time!!! Fix this !!!*)
fun splitequation ctxt th acc =
let
val real_poly_neg_conv = #neg
(Normalizer.semiring_normalizers_ord_wrapper ctxt
(valOf (NormalizerData.match ctxt @{cterm "(0::real) + 1"})) simple_cterm_ord)
val (th1,th2) = conj_pair(rawrule th)
in th1::fconv_rule (arg_conv (arg_conv real_poly_neg_conv)) th2::acc
end
in fun real_vector_prover ctxt translator (eqs,ges,gts) =
real_vector_ineq_prover ctxt translator
(fold_rev (splitequation ctxt) eqs ges,gts)
end;
fun init_conv ctxt =
Simplifier.rewrite (Simplifier.context ctxt
(HOL_basic_ss addsimps ([@{thm vec_0}, @{thm vec_1}, @{thm dist_vector_def}, @{thm diff_0_right}, @{thm right_minus}, @{thm diff_self}, @{thm norm_0}] @ @{thms arithmetic_simps} @ @{thms norm_pths})))
then_conv field_comp_conv
then_conv nnf_conv
fun pure ctxt = RealArith.gen_prover_real_arith ctxt (real_vector_prover ctxt);
fun norm_arith ctxt ct =
let
val ctxt' = Variable.declare_term (term_of ct) ctxt
val th = init_conv ctxt' ct
in equal_elim (Drule.arg_cong_rule @{cterm Trueprop} (symmetric th))
(pure ctxt' (rhs_of th))
end
fun norm_arith_tac ctxt =
clarify_tac HOL_cs THEN'
ObjectLogic.full_atomize_tac THEN'
CSUBGOAL ( fn (p,i) => rtac (norm_arith ctxt (Thm.dest_arg p )) i);
end;