--- a/src/HOL/Boogie/Tools/boogie_tactics.ML Sat May 15 07:48:24 2010 -0700
+++ b/src/HOL/Boogie/Tools/boogie_tactics.ML Sat May 15 18:29:18 2010 +0200
@@ -25,8 +25,8 @@
val label_eqs = [assert_at_def, block_at_def]
fun unfold_labels_tac ctxt =
- let val unfold = More_Conv.rewrs_conv label_eqs
- in CONVERSION (More_Conv.top_sweep_conv (K unfold) ctxt) end
+ let val unfold = Conv.rewrs_conv label_eqs
+ in CONVERSION (Conv.top_sweep_conv (K unfold) ctxt) end
fun boogie_tac ctxt rules =
unfold_labels_tac ctxt
--- a/src/HOL/HOL.thy Sat May 15 07:48:24 2010 -0700
+++ b/src/HOL/HOL.thy Sat May 15 18:29:18 2010 +0200
@@ -29,7 +29,6 @@
"~~/src/Tools/induct.ML"
("~~/src/Tools/induct_tacs.ML")
("Tools/recfun_codegen.ML")
- "~~/src/Tools/more_conv.ML"
begin
setup {* Intuitionistic.method_setup @{binding iprover} *}
--- a/src/HOL/IsaMakefile Sat May 15 07:48:24 2010 -0700
+++ b/src/HOL/IsaMakefile Sat May 15 18:29:18 2010 +0200
@@ -60,7 +60,7 @@
HOL-Proofs-Extraction \
HOL-Proofs-Lambda \
HOL-SET_Protocol \
- HOL-SMT_Examples \
+ HOL-Word-SMT_Examples \
HOL-Statespace \
HOL-Subst \
TLA-Buffer \
@@ -128,7 +128,6 @@
$(SRC)/Tools/induct.ML \
$(SRC)/Tools/induct_tacs.ML \
$(SRC)/Tools/intuitionistic.ML \
- $(SRC)/Tools/more_conv.ML \
$(SRC)/Tools/nbe.ML \
$(SRC)/Tools/project_rule.ML \
$(SRC)/Tools/quickcheck.ML \
@@ -402,7 +401,7 @@
Library/Efficient_Nat.thy Library/Sum_Of_Squares.thy \
Library/Dlist.thy Library/Sum_Of_Squares/sos_wrapper.ML \
Library/Sum_Of_Squares/sum_of_squares.ML Library/Fset.thy \
- Library/Glbs.thy Library/normarith.ML Library/Executable_Set.thy \
+ Library/Glbs.thy Library/Executable_Set.thy \
Library/Infinite_Set.thy Library/FuncSet.thy \
Library/Permutations.thy Library/Bit.thy Library/FrechetDeriv.thy \
Library/Fraction_Field.thy Library/Fundamental_Theorem_Algebra.thy \
@@ -1088,27 +1087,27 @@
HOL-Multivariate_Analysis: HOL $(OUT)/HOL-Multivariate_Analysis
-$(OUT)/HOL-Multivariate_Analysis: $(OUT)/HOL \
- Multivariate_Analysis/ROOT.ML \
- Multivariate_Analysis/document/root.tex \
- Multivariate_Analysis/Brouwer_Fixpoint.thy \
- Multivariate_Analysis/Convex_Euclidean_Space.thy \
- Multivariate_Analysis/Derivative.thy \
- Multivariate_Analysis/Determinants.thy \
- Multivariate_Analysis/Euclidean_Space.thy \
- Multivariate_Analysis/Fashoda.thy \
- Multivariate_Analysis/Finite_Cartesian_Product.thy \
- Multivariate_Analysis/Integration.thy \
- Multivariate_Analysis/Integration.certs \
- Multivariate_Analysis/L2_Norm.thy \
- Multivariate_Analysis/Multivariate_Analysis.thy \
- Multivariate_Analysis/Operator_Norm.thy \
- Multivariate_Analysis/Path_Connected.thy \
- Multivariate_Analysis/Real_Integration.thy \
- Multivariate_Analysis/Topology_Euclidean_Space.thy \
- Multivariate_Analysis/Vec1.thy Library/Glbs.thy \
- Library/Inner_Product.thy Library/Numeral_Type.thy \
- Library/Convex.thy Library/FrechetDeriv.thy \
+$(OUT)/HOL-Multivariate_Analysis: $(OUT)/HOL \
+ Multivariate_Analysis/Brouwer_Fixpoint.thy \
+ Multivariate_Analysis/Convex_Euclidean_Space.thy \
+ Multivariate_Analysis/Derivative.thy \
+ Multivariate_Analysis/Determinants.thy \
+ Multivariate_Analysis/Euclidean_Space.thy \
+ Multivariate_Analysis/Fashoda.thy \
+ Multivariate_Analysis/Finite_Cartesian_Product.thy \
+ Multivariate_Analysis/Integration.certs \
+ Multivariate_Analysis/Integration.thy \
+ Multivariate_Analysis/L2_Norm.thy \
+ Multivariate_Analysis/Multivariate_Analysis.thy \
+ Multivariate_Analysis/Operator_Norm.thy \
+ Multivariate_Analysis/Path_Connected.thy \
+ Multivariate_Analysis/ROOT.ML \
+ Multivariate_Analysis/Real_Integration.thy \
+ Multivariate_Analysis/Topology_Euclidean_Space.thy \
+ Multivariate_Analysis/document/root.tex \
+ Multivariate_Analysis/normarith.ML Multivariate_Analysis/Vec1.thy \
+ Library/Glbs.thy Library/Inner_Product.thy Library/Numeral_Type.thy \
+ Library/Convex.thy Library/FrechetDeriv.thy \
Library/Product_Vector.thy Library/Product_plus.thy
@cd Multivariate_Analysis; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL HOL-Multivariate_Analysis
@@ -1254,11 +1253,11 @@
@$(ISABELLE_TOOL) usedir $(OUT)/HOL Mirabelle
-## HOL-SMT_Examples
+## HOL-Word-SMT_Examples
-HOL-SMT_Examples: HOL-Word $(LOG)/HOL-SMT_Examples.gz
+HOL-Word-SMT_Examples: HOL-Word $(LOG)/HOL-Word-SMT_Examples.gz
-$(LOG)/HOL-SMT_Examples.gz: $(OUT)/HOL-Word SMT_Examples/ROOT.ML \
+$(LOG)/HOL-Word-SMT_Examples.gz: $(OUT)/HOL-Word SMT_Examples/ROOT.ML \
SMT_Examples/SMT_Examples.thy SMT_Examples/SMT_Examples.certs \
SMT_Examples/SMT_Word_Examples.thy SMT_Examples/SMT_Tests.thy \
SMT_Examples/SMT_Word_Examples.certs SMT_Examples/SMT_Tests.certs
@@ -1346,15 +1345,15 @@
$(LOG)/HOL-Probability.gz $(LOG)/HOL-Prolog.gz \
$(LOG)/HOL-Proofs.gz $(LOG)/HOL-Proofs-Extraction.gz \
$(LOG)/HOL-Proofs-Lambda.gz $(LOG)/HOL-SET_Protocol.gz \
- $(LOG)/HOL-SMT_Examples.gz $(LOG)/HOL-Statespace.gz \
- $(LOG)/HOL-Subst.gz $(LOG)/HOL-UNITY.gz \
- $(LOG)/HOL-Unix.gz $(LOG)/HOL-Word-Examples.gz \
- $(LOG)/HOL-Word.gz $(LOG)/HOL-ZF.gz $(LOG)/HOL-ex.gz \
- $(LOG)/HOL.gz $(LOG)/HOL4.gz $(LOG)/TLA-Buffer.gz \
- $(LOG)/TLA-Inc.gz $(LOG)/TLA-Memory.gz $(LOG)/TLA.gz \
- $(OUT)/HOL $(OUT)/HOL-Algebra $(OUT)/HOL-Base \
- $(OUT)/HOL-Boogie $(OUT)/HOL-Main \
- $(OUT)/HOL-Multivariate_Analysis $(OUT)/HOL-NSA \
- $(OUT)/HOL-Nominal $(OUT)/HOL-Plain \
+ $(LOG)/HOL-Word-SMT_Examples.gz \
+ $(LOG)/HOL-Statespace.gz $(LOG)/HOL-Subst.gz \
+ $(LOG)/HOL-UNITY.gz $(LOG)/HOL-Unix.gz \
+ $(LOG)/HOL-Word-Examples.gz $(LOG)/HOL-Word.gz \
+ $(LOG)/HOL-ZF.gz $(LOG)/HOL-ex.gz $(LOG)/HOL.gz \
+ $(LOG)/HOL4.gz $(LOG)/TLA-Buffer.gz $(LOG)/TLA-Inc.gz \
+ $(LOG)/TLA-Memory.gz $(LOG)/TLA.gz $(OUT)/HOL \
+ $(OUT)/HOL-Algebra $(OUT)/HOL-Base $(OUT)/HOL-Boogie \
+ $(OUT)/HOL-Main $(OUT)/HOL-Multivariate_Analysis \
+ $(OUT)/HOL-NSA $(OUT)/HOL-Nominal $(OUT)/HOL-Plain \
$(OUT)/HOL-Probability $(OUT)/HOL-Proofs \
$(OUT)/HOL-Word $(OUT)/HOL4 $(OUT)/TLA
--- a/src/HOL/Library/normarith.ML Sat May 15 07:48:24 2010 -0700
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,416 +0,0 @@
-(* 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;
- 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,Thm.dest_arg t) acc
- | _ => acc
- fun find_normedterms t acc = case term_of t of
- @{term "op + :: real => _"}$_$_ =>
- find_normedterms (Thm.dest_arg1 t) (find_normedterms (Thm.dest_arg t) acc)
- | @{term "op * :: real => _"}$_$n =>
- if not (is_ratconst (Thm.dest_arg1 t)) then acc else
- augment_norm (dest_ratconst (Thm.dest_arg1 t) >=/ Rat.zero)
- (Thm.dest_arg t) acc
- | _ => augment_norm true t acc
-
- val cterm_lincomb_neg = FuncUtil.Ctermfunc.map Rat.neg
- fun cterm_lincomb_cmul c t =
- if c =/ Rat.zero then FuncUtil.Ctermfunc.empty else FuncUtil.Ctermfunc.map (fn x => x */ c) t
- fun cterm_lincomb_add l r = FuncUtil.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 = FuncUtil.Ctermfunc.is_empty (cterm_lincomb_sub l r)
-
- val int_lincomb_neg = FuncUtil.Intfunc.map Rat.neg
- fun int_lincomb_cmul c t =
- if c =/ Rat.zero then FuncUtil.Intfunc.empty else FuncUtil.Intfunc.map (fn x => x */ c) t
- fun int_lincomb_add l r = FuncUtil.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 = FuncUtil.Intfunc.is_empty (int_lincomb_sub l r)
-
-fun vector_lincomb t = case term_of t of
- Const(@{const_name plus}, _) $ _ $ _ =>
- cterm_lincomb_add (vector_lincomb (Thm.dest_arg1 t)) (vector_lincomb (Thm.dest_arg t))
- | Const(@{const_name minus}, _) $ _ $ _ =>
- cterm_lincomb_sub (vector_lincomb (Thm.dest_arg1 t)) (vector_lincomb (Thm.dest_arg t))
- | Const(@{const_name scaleR}, _)$_$_ =>
- cterm_lincomb_cmul (dest_ratconst (Thm.dest_arg1 t)) (vector_lincomb (Thm.dest_arg t))
- | Const(@{const_name uminus}, _)$_ =>
- cterm_lincomb_neg (vector_lincomb (Thm.dest_arg t))
-(* FIXME: how should we handle numerals?
- | Const(@ {const_name vec},_)$_ =>
- let
- val b = ((snd o HOLogic.dest_number o term_of o Thm.dest_arg) t = 0
- handle TERM _=> false)
- in if b then FuncUtil.Ctermfunc.onefunc (t,Rat.one)
- else FuncUtil.Ctermfunc.empty
- end
-*)
- | _ => FuncUtil.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 (Thm.dest_arg1 t) </ Rat.zero then arg_conv cv t else reflexive t
-| _ => reflexive t
-fun flip v eq =
- if FuncUtil.Ctermfunc.defined eq v
- then FuncUtil.Ctermfunc.update (v, Rat.neg (FuncUtil.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 =
- FuncUtil.Intfunc.fold (fn (x,c) => fn s => s +/ c */ (FuncUtil.Intfunc.apply env x))
- lin Rat.zero
-
-fun solve (vs,eqs) = case (vs,eqs) of
- ([],[]) => SOME (FuncUtil.Intfunc.onefunc (0,Rat.one))
- |(_,eq::oeqs) =>
- (case filter (member (op =) vs) (FuncUtil.Intfunc.dom eq) of (*FIXME use find_first here*)
- [] => NONE
- | v::_ =>
- if FuncUtil.Intfunc.defined eq v
- then
- let
- val c = FuncUtil.Intfunc.apply eq v
- val vdef = int_lincomb_cmul (Rat.neg (Rat.inv c)) eq
- fun eliminate eqn = if not (FuncUtil.Intfunc.defined eqn v) then eqn
- else int_lincomb_add (int_lincomb_cmul (FuncUtil.Intfunc.apply eqn v) vdef) eqn
- in (case solve (remove (op =) v vs, map eliminate oeqs) of
- NONE => NONE
- | SOME soln => SOME (FuncUtil.Intfunc.update (v, evaluate soln (FuncUtil.Intfunc.delete_safe 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 => FuncUtil.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 =
- Semiring_Normalizer.semiring_normalize_wrapper ctxt
- (the (Semiring_Normalizer.match ctxt @{cterm "(0::real) + 1"}))
- in fconv_rule (arg_conv ((rewr_conv @{thm ge_iff_diff_ge_0}) then_conv arg_conv (Numeral_Simprocs.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 Numeral_Simprocs.field_comp_conv then_conv (try_conv (rewr_conv (mk_meta_eq @{thm scaleR_zero_left})));
- val apply_pth9 = rewrs_conv @{thms pth_9} then_conv arg1_conv (arg1_conv Numeral_Simprocs.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}, _)$
- (Const(@{const_name scaleR}, _)$l$v)$r => v
- | Const(@{const_name scaleR}, _)$l$v => v
- | _ => error "headvector: non-canonical term"
-
-fun vector_cmul_conv ct =
- ((apply_pth5 then_conv arg1_conv Numeral_Simprocs.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 Term_Ord.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 scaleR}, _)$_$_ =>
- 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
-
-(* FIXME
-| 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 FuncUtil.Intfunc.defined eq v
- then FuncUtil.Intfunc.update (v, Rat.neg (FuncUtil.Intfunc.apply eq v)) eq else eq;
-
-local
- val pth_zero = @{thm norm_zero}
- val tv_n = (ctyp_of_term o Thm.dest_arg o Thm.dest_arg1 o Thm.dest_arg o cprop_of)
- pth_zero
- val concl = Thm.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 =
- Semiring_Normalizer.semiring_normalize_wrapper ctxt
- (the (Semiring_Normalizer.match ctxt @{cterm "(0::real) + 1"}))
- val sources = map (Thm.dest_arg o Thm.dest_arg1 o concl) nubs
- val rawdests = fold_rev (find_normedterms o Thm.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 (union (op aconvc) o FuncUtil.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 FuncUtil.Ctermfunc.defined d x then FuncUtil.Intfunc.onefunc (0, Rat.neg(FuncUtil.Ctermfunc.apply d x))
- else FuncUtil.Intfunc.empty
- 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
- end
- val equations = map coefficients vvs
- val inequalities = map (fn n => FuncUtil.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 FuncUtil.Intfunc.update) nvs v (FuncUtil.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 member (op =) f n 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)
- fun end_itlist f xs = split_last xs |> uncurry (fold_rev f)
- 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 (FuncUtil.Ctermfunc.dom (vector_lincomb t))) (map snd rawdests)
-
- in fst (RealArith.real_linear_prover translator
- (map (fn t => instantiate ([(tv_n, ctyp_of_term t)],[]) pth_zero)
- zerodests,
- map (fconv_rule (try_conv (More_Conv.top_sweep_conv (K norm_canon_conv) ctxt) then_conv
- arg_conv (arg_conv real_poly_conv))) ges',
- map (fconv_rule (try_conv (More_Conv.top_sweep_conv (K norm_canon_conv) ctxt) 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 = Thm.dest_arg o cprop_of
- fun conjunct1 th = th RS @{thm conjunct1}
- fun conjunct2 th = th RS @{thm conjunct2}
-fun real_vector_ineq_prover ctxt translator (ges,gts) =
- let
-(* val _ = error "real_vector_ineq_prover: pause" *)
- val ntms = fold_rev find_normedterms (map (Thm.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 instantiate_cterm' ty tms = Drule.cterm_rule (Drule.instantiate' ty tms)
- fun mk_norm t = Thm.capply (instantiate_cterm' [SOME (ctyp_of_term t)] [] @{cpat "norm :: (?'a :: real_normed_vector) => real"}) t
- fun mk_equals l r = Thm.capply (Thm.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 Thm.dest_equals) (cprems_of th11)
- val th12 = instantiate ([], cps) th11
- val th13 = fold Thm.elim_implies (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 = Term_Ord.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
- (Semiring_Normalizer.semiring_normalizers_ord_wrapper ctxt
- (the (Semiring_Normalizer.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), RealArith.Trivial)
-end;
-
- fun init_conv ctxt =
- Simplifier.rewrite (Simplifier.context ctxt
- (HOL_basic_ss addsimps ([(*@{thm vec_0}, @{thm vec_1},*) @{thm dist_norm}, @{thm diff_0_right}, @{thm right_minus}, @{thm diff_self}, @{thm norm_zero}] @ @{thms arithmetic_simps} @ @{thms norm_pths})))
- then_conv Numeral_Simprocs.field_comp_conv
- then_conv nnf_conv
-
- fun pure ctxt = fst o 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' (Thm.rhs_of th))
- end
-
- fun norm_arith_tac ctxt =
- clarify_tac HOL_cs THEN'
- Object_Logic.full_atomize_tac THEN'
- CSUBGOAL ( fn (p,i) => rtac (norm_arith ctxt (Thm.dest_arg p )) i);
-
-end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/normarith.ML Sat May 15 18:29:18 2010 +0200
@@ -0,0 +1,402 @@
+(* Title: Library/normarith.ML
+ Author: Amine Chaieb, University of Cambridge
+
+Simple decision procedure for linear problems in Euclidean space.
+*)
+
+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;
+ 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,Thm.dest_arg t) acc
+ | _ => acc
+ fun find_normedterms t acc = case term_of t of
+ @{term "op + :: real => _"}$_$_ =>
+ find_normedterms (Thm.dest_arg1 t) (find_normedterms (Thm.dest_arg t) acc)
+ | @{term "op * :: real => _"}$_$n =>
+ if not (is_ratconst (Thm.dest_arg1 t)) then acc else
+ augment_norm (dest_ratconst (Thm.dest_arg1 t) >=/ Rat.zero)
+ (Thm.dest_arg t) acc
+ | _ => augment_norm true t acc
+
+ val cterm_lincomb_neg = FuncUtil.Ctermfunc.map Rat.neg
+ fun cterm_lincomb_cmul c t =
+ if c =/ Rat.zero then FuncUtil.Ctermfunc.empty else FuncUtil.Ctermfunc.map (fn x => x */ c) t
+ fun cterm_lincomb_add l r = FuncUtil.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 = FuncUtil.Ctermfunc.is_empty (cterm_lincomb_sub l r)
+
+ val int_lincomb_neg = FuncUtil.Intfunc.map Rat.neg
+ fun int_lincomb_cmul c t =
+ if c =/ Rat.zero then FuncUtil.Intfunc.empty else FuncUtil.Intfunc.map (fn x => x */ c) t
+ fun int_lincomb_add l r = FuncUtil.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 = FuncUtil.Intfunc.is_empty (int_lincomb_sub l r)
+
+fun vector_lincomb t = case term_of t of
+ Const(@{const_name plus}, _) $ _ $ _ =>
+ cterm_lincomb_add (vector_lincomb (Thm.dest_arg1 t)) (vector_lincomb (Thm.dest_arg t))
+ | Const(@{const_name minus}, _) $ _ $ _ =>
+ cterm_lincomb_sub (vector_lincomb (Thm.dest_arg1 t)) (vector_lincomb (Thm.dest_arg t))
+ | Const(@{const_name scaleR}, _)$_$_ =>
+ cterm_lincomb_cmul (dest_ratconst (Thm.dest_arg1 t)) (vector_lincomb (Thm.dest_arg t))
+ | Const(@{const_name uminus}, _)$_ =>
+ cterm_lincomb_neg (vector_lincomb (Thm.dest_arg t))
+(* FIXME: how should we handle numerals?
+ | Const(@ {const_name vec},_)$_ =>
+ let
+ val b = ((snd o HOLogic.dest_number o term_of o Thm.dest_arg) t = 0
+ handle TERM _=> false)
+ in if b then FuncUtil.Ctermfunc.onefunc (t,Rat.one)
+ else FuncUtil.Ctermfunc.empty
+ end
+*)
+ | _ => FuncUtil.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 (Thm.dest_arg1 t) </ Rat.zero then arg_conv cv t else reflexive t
+| _ => reflexive t
+fun flip v eq =
+ if FuncUtil.Ctermfunc.defined eq v
+ then FuncUtil.Ctermfunc.update (v, Rat.neg (FuncUtil.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 =
+ FuncUtil.Intfunc.fold (fn (x,c) => fn s => s +/ c */ (FuncUtil.Intfunc.apply env x))
+ lin Rat.zero
+
+fun solve (vs,eqs) = case (vs,eqs) of
+ ([],[]) => SOME (FuncUtil.Intfunc.onefunc (0,Rat.one))
+ |(_,eq::oeqs) =>
+ (case filter (member (op =) vs) (FuncUtil.Intfunc.dom eq) of (*FIXME use find_first here*)
+ [] => NONE
+ | v::_ =>
+ if FuncUtil.Intfunc.defined eq v
+ then
+ let
+ val c = FuncUtil.Intfunc.apply eq v
+ val vdef = int_lincomb_cmul (Rat.neg (Rat.inv c)) eq
+ fun eliminate eqn = if not (FuncUtil.Intfunc.defined eqn v) then eqn
+ else int_lincomb_add (int_lincomb_cmul (FuncUtil.Intfunc.apply eqn v) vdef) eqn
+ in (case solve (remove (op =) v vs, map eliminate oeqs) of
+ NONE => NONE
+ | SOME soln => SOME (FuncUtil.Intfunc.update (v, evaluate soln (FuncUtil.Intfunc.delete_safe 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 => FuncUtil.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 =
+ Semiring_Normalizer.semiring_normalize_wrapper ctxt
+ (the (Semiring_Normalizer.match ctxt @{cterm "(0::real) + 1"}))
+ in fconv_rule (arg_conv ((rewr_conv @{thm ge_iff_diff_ge_0}) then_conv arg_conv (Numeral_Simprocs.field_comp_conv then_conv real_poly_conv)))
+end;
+
+ 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 Numeral_Simprocs.field_comp_conv then_conv (try_conv (rewr_conv (mk_meta_eq @{thm scaleR_zero_left})));
+ val apply_pth9 = rewrs_conv @{thms pth_9} then_conv arg1_conv (arg1_conv Numeral_Simprocs.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}, _)$
+ (Const(@{const_name scaleR}, _)$l$v)$r => v
+ | Const(@{const_name scaleR}, _)$l$v => v
+ | _ => error "headvector: non-canonical term"
+
+fun vector_cmul_conv ct =
+ ((apply_pth5 then_conv arg1_conv Numeral_Simprocs.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 Term_Ord.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 scaleR}, _)$_$_ =>
+ 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
+
+(* FIXME
+| 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 FuncUtil.Intfunc.defined eq v
+ then FuncUtil.Intfunc.update (v, Rat.neg (FuncUtil.Intfunc.apply eq v)) eq else eq;
+
+local
+ val pth_zero = @{thm norm_zero}
+ val tv_n = (ctyp_of_term o Thm.dest_arg o Thm.dest_arg1 o Thm.dest_arg o cprop_of)
+ pth_zero
+ val concl = Thm.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 =
+ Semiring_Normalizer.semiring_normalize_wrapper ctxt
+ (the (Semiring_Normalizer.match ctxt @{cterm "(0::real) + 1"}))
+ val sources = map (Thm.dest_arg o Thm.dest_arg1 o concl) nubs
+ val rawdests = fold_rev (find_normedterms o Thm.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 (union (op aconvc) o FuncUtil.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 FuncUtil.Ctermfunc.defined d x then FuncUtil.Intfunc.onefunc (0, Rat.neg(FuncUtil.Ctermfunc.apply d x))
+ else FuncUtil.Intfunc.empty
+ 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
+ end
+ val equations = map coefficients vvs
+ val inequalities = map (fn n => FuncUtil.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 FuncUtil.Intfunc.update) nvs v (FuncUtil.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 member (op =) f n 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)
+ fun end_itlist f xs = split_last xs |> uncurry (fold_rev f)
+ 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 (FuncUtil.Ctermfunc.dom (vector_lincomb t))) (map snd rawdests)
+
+ in fst (RealArith.real_linear_prover translator
+ (map (fn t => instantiate ([(tv_n, ctyp_of_term t)],[]) pth_zero)
+ zerodests,
+ map (fconv_rule (try_conv (Conv.top_sweep_conv (K norm_canon_conv) ctxt) then_conv
+ arg_conv (arg_conv real_poly_conv))) ges',
+ map (fconv_rule (try_conv (Conv.top_sweep_conv (K norm_canon_conv) ctxt) 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 = Thm.dest_arg o cprop_of
+ fun conjunct1 th = th RS @{thm conjunct1}
+ fun conjunct2 th = th RS @{thm conjunct2}
+fun real_vector_ineq_prover ctxt translator (ges,gts) =
+ let
+(* val _ = error "real_vector_ineq_prover: pause" *)
+ val ntms = fold_rev find_normedterms (map (Thm.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 instantiate_cterm' ty tms = Drule.cterm_rule (Drule.instantiate' ty tms)
+ fun mk_norm t = Thm.capply (instantiate_cterm' [SOME (ctyp_of_term t)] [] @{cpat "norm :: (?'a :: real_normed_vector) => real"}) t
+ fun mk_equals l r = Thm.capply (Thm.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 Thm.dest_equals) (cprems_of th11)
+ val th12 = instantiate ([], cps) th11
+ val th13 = fold Thm.elim_implies (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 = Term_Ord.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
+ (Semiring_Normalizer.semiring_normalizers_ord_wrapper ctxt
+ (the (Semiring_Normalizer.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), RealArith.Trivial)
+end;
+
+ fun init_conv ctxt =
+ Simplifier.rewrite (Simplifier.context ctxt
+ (HOL_basic_ss addsimps ([(*@{thm vec_0}, @{thm vec_1},*) @{thm dist_norm}, @{thm diff_0_right}, @{thm right_minus}, @{thm diff_self}, @{thm norm_zero}] @ @{thms arithmetic_simps} @ @{thms norm_pths})))
+ then_conv Numeral_Simprocs.field_comp_conv
+ then_conv nnf_conv
+
+ fun pure ctxt = fst o 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' (Thm.rhs_of th))
+ end
+
+ fun norm_arith_tac ctxt =
+ clarify_tac HOL_cs THEN'
+ Object_Logic.full_atomize_tac THEN'
+ CSUBGOAL ( fn (p,i) => rtac (norm_arith ctxt (Thm.dest_arg p )) i);
+
+end;
--- a/src/HOL/Tools/Function/function_core.ML Sat May 15 07:48:24 2010 -0700
+++ b/src/HOL/Tools/Function/function_core.ML Sat May 15 18:29:18 2010 +0200
@@ -617,7 +617,7 @@
local open Conv in
val lhs_D = fconv_rule (arg_conv (arg_conv (case_hyp_conv))) x_D
val sih =
- fconv_rule (More_Conv.binder_conv
+ fconv_rule (Conv.binder_conv
(K (arg1_conv (arg_conv (arg_conv case_hyp_conv)))) ctxt) aihyp
end
--- a/src/HOL/Tools/Quotient/quotient_tacs.ML Sat May 15 07:48:24 2010 -0700
+++ b/src/HOL/Tools/Quotient/quotient_tacs.ML Sat May 15 18:29:18 2010 +0200
@@ -490,7 +490,7 @@
end
| _ => Conv.all_conv ctrm
-fun lambda_prs_conv ctxt = More_Conv.top_conv lambda_prs_simple_conv ctxt
+fun lambda_prs_conv ctxt = Conv.top_conv lambda_prs_simple_conv ctxt
fun lambda_prs_tac ctxt = CONVERSION (lambda_prs_conv ctxt)
--- a/src/HOL/Tools/SMT/smt_normalize.ML Sat May 15 07:48:24 2010 -0700
+++ b/src/HOL/Tools/SMT/smt_normalize.ML Sat May 15 18:29:18 2010 +0200
@@ -47,11 +47,11 @@
"distinct [x, y] == (x ~= y)"
by simp_all}
fun distinct_conv _ =
- if_true_conv is_trivial_distinct (More_Conv.rewrs_conv thms)
+ if_true_conv is_trivial_distinct (Conv.rewrs_conv thms)
in
fun trivial_distinct ctxt =
map ((Term.exists_subterm is_trivial_distinct o Thm.prop_of) ??
- Conv.fconv_rule (More_Conv.top_conv distinct_conv ctxt))
+ Conv.fconv_rule (Conv.top_conv distinct_conv ctxt))
end
@@ -67,11 +67,11 @@
"(case P of True => x | False => y) == (if P then x else y)"
"(case P of False => y | True => x) == (if P then x else y)"
by (rule eq_reflection, simp)+}
- val unfold_conv = if_true_conv is_bool_case (More_Conv.rewrs_conv thms)
+ val unfold_conv = if_true_conv is_bool_case (Conv.rewrs_conv thms)
in
fun rewrite_bool_cases ctxt =
map ((Term.exists_subterm is_bool_case o Thm.prop_of) ??
- Conv.fconv_rule (More_Conv.top_conv (K unfold_conv) ctxt))
+ Conv.fconv_rule (Conv.top_conv (K unfold_conv) ctxt))
end
@@ -107,7 +107,7 @@
in
fun normalize_numerals ctxt =
map ((Term.exists_subterm (is_strange_number ctxt) o Thm.prop_of) ??
- Conv.fconv_rule (More_Conv.top_sweep_conv pos_conv ctxt))
+ Conv.fconv_rule (Conv.top_sweep_conv pos_conv ctxt))
end
@@ -193,7 +193,7 @@
local
val eta_conv = eta_expand_conv
- fun keep_conv ctxt = More_Conv.binder_conv norm_conv ctxt
+ fun keep_conv ctxt = Conv.binder_conv (norm_conv o snd) ctxt
and eta_binder_conv ctxt = Conv.arg_conv (eta_conv norm_conv ctxt)
and keep_let_conv ctxt = Conv.combination_conv
(Conv.arg_conv (norm_conv ctxt)) (Conv.abs_conv (norm_conv o snd) ctxt)
@@ -267,7 +267,7 @@
Conv.binop_conv (atomize_conv ctxt) then_conv
Conv.rewr_conv @{thm atomize_eq}
| Const (@{const_name all}, _) $ Abs _ =>
- More_Conv.binder_conv atomize_conv ctxt then_conv
+ Conv.binder_conv (atomize_conv o snd) ctxt then_conv
Conv.rewr_conv @{thm atomize_all}
| _ => Conv.all_conv) ct
--- a/src/HOL/Tools/SMT/smtlib_interface.ML Sat May 15 07:48:24 2010 -0700
+++ b/src/HOL/Tools/SMT/smtlib_interface.ML Sat May 15 18:29:18 2010 +0200
@@ -81,7 +81,7 @@
fun unfold_abs_min_max_defs ctxt thm =
if exists_abs_min_max (Thm.prop_of thm)
- then Conv.fconv_rule (More_Conv.top_conv unfold_def_conv ctxt) thm
+ then Conv.fconv_rule (Conv.top_conv unfold_def_conv ctxt) thm
else thm
--- a/src/HOL/Tools/SMT/z3_proof_reconstruction.ML Sat May 15 07:48:24 2010 -0700
+++ b/src/HOL/Tools/SMT/z3_proof_reconstruction.ML Sat May 15 18:29:18 2010 +0200
@@ -540,7 +540,7 @@
fun elim_unused_conv ctxt =
Conv.params_conv ~1 (K (Conv.arg_conv (Conv.arg1_conv
- (More_Conv.rewrs_conv [elim_all, elim_ex])))) ctxt
+ (Conv.rewrs_conv [elim_all, elim_ex])))) ctxt
fun elim_unused_tac ctxt =
REPEAT_ALL_NEW (
--- a/src/HOL/Tools/SMT/z3_proof_tools.ML Sat May 15 07:48:24 2010 -0700
+++ b/src/HOL/Tools/SMT/z3_proof_tools.ML Sat May 15 18:29:18 2010 +0200
@@ -96,7 +96,7 @@
fun unfold_eqs _ [] = Conv.all_conv
| unfold_eqs ctxt eqs =
- More_Conv.top_sweep_conv (K (More_Conv.rewrs_conv eqs)) ctxt
+ Conv.top_sweep_conv (K (Conv.rewrs_conv eqs)) ctxt
fun match_instantiate f ct thm =
Thm.instantiate (Thm.match (f (Thm.cprop_of thm), ct)) thm
@@ -256,7 +256,7 @@
val set5 = @{lemma "x ~: set (y # ys) == x ~= y & x ~: set ys" by simp}
fun set_conv ct =
- (More_Conv.rewrs_conv [set1, set2, set3, set4] else_conv
+ (Conv.rewrs_conv [set1, set2, set3, set4] else_conv
(Conv.rewr_conv set5 then_conv Conv.arg_conv set_conv)) ct
val dist1 = @{lemma "distinct [] == ~False" by simp}
@@ -267,7 +267,7 @@
fun binop_conv cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2
in
fun unfold_distinct_conv ct =
- (More_Conv.rewrs_conv [dist1, dist2] else_conv
+ (Conv.rewrs_conv [dist1, dist2] else_conv
(Conv.rewr_conv dist3 then_conv binop_conv set_conv unfold_distinct_conv)) ct
end
--- a/src/Pure/axclass.ML Sat May 15 07:48:24 2010 -0700
+++ b/src/Pure/axclass.ML Sat May 15 18:29:18 2010 +0200
@@ -412,8 +412,6 @@
(* primitive rules *)
-val shyps_topped = forall null o #shyps o Thm.rep_thm;
-
fun add_classrel raw_th thy =
let
val th = Thm.strip_shyps (Thm.transfer thy raw_th);
@@ -422,9 +420,8 @@
val rel = Logic.dest_classrel prop handle TERM _ => err ();
val (c1, c2) = cert_classrel thy rel handle TYPE _ => err ();
val th' = th
- |> Drule.instantiate' [SOME (ctyp_of thy (TVar ((Name.aT, 0), [c1])))] []
- |> Thm.unconstrainT;
- val _ = shyps_topped th' orelse raise Fail "add_classrel: nontop shyps after unconstrain";
+ |> Thm.unconstrainT
+ |> Drule.instantiate' [SOME (ctyp_of thy (TVar ((Name.aT, 0), [])))] [];
in
thy
|> Sign.primitive_classrel (c1, c2)
@@ -441,16 +438,15 @@
val args = Name.names Name.context Name.aT Ss;
val T = Type (t, map TFree args);
- val std_vars = map (fn (a, S) => SOME (ctyp_of thy (TVar ((a, 0), S)))) args;
+ val std_vars = map (fn (a, S) => SOME (ctyp_of thy (TVar ((a, 0), [])))) args;
val missing_params = Sign.complete_sort thy [c]
|> maps (these o Option.map #params o try (get_info thy))
|> filter_out (fn (const, _) => can (get_inst_param thy) (const, t))
|> (map o apsnd o map_atyps) (K T);
val th' = th
- |> Drule.instantiate' std_vars []
- |> Thm.unconstrainT;
- val _ = shyps_topped th' orelse raise Fail "add_arity: nontop shyps after unconstrain";
+ |> Thm.unconstrainT
+ |> Drule.instantiate' std_vars [];
in
thy
|> fold (#2 oo declare_overloaded) missing_params
--- a/src/Pure/conv.ML Sat May 15 07:48:24 2010 -0700
+++ b/src/Pure/conv.ML Sat May 15 18:29:18 2010 +0200
@@ -1,5 +1,6 @@
(* Title: Pure/conv.ML
Author: Amine Chaieb, TU Muenchen
+ Author: Sascha Boehme, TU Muenchen
Author: Makarius
Conversions: primitive equality reasoning.
@@ -32,10 +33,16 @@
val arg1_conv: conv -> conv
val fun2_conv: conv -> conv
val binop_conv: conv -> conv
+ val binder_conv: (cterm * Proof.context -> conv) -> Proof.context -> conv
val forall_conv: (cterm * Proof.context -> conv) -> Proof.context -> conv
val implies_conv: conv -> conv -> conv
val implies_concl_conv: conv -> conv
val rewr_conv: thm -> conv
+ val rewrs_conv: thm list -> conv
+ val sub_conv: (Proof.context -> conv) -> Proof.context -> conv
+ val bottom_conv: (Proof.context -> conv) -> Proof.context -> conv
+ val top_conv: (Proof.context -> conv) -> Proof.context -> conv
+ val top_sweep_conv: (Proof.context -> conv) -> Proof.context -> conv
val params_conv: int -> (Proof.context -> conv) -> Proof.context -> conv
val prems_conv: int -> conv -> conv
val concl_conv: int -> conv -> conv
@@ -105,6 +112,29 @@
fun binop_conv cv = combination_conv (arg_conv cv) cv;
+fun binder_conv cv ctxt = arg_conv (abs_conv cv ctxt);
+
+
+(* subterm structure *)
+
+(*cf. SUB_CONV in HOL*)
+fun sub_conv conv ctxt =
+ comb_conv (conv ctxt) else_conv
+ abs_conv (conv o snd) ctxt else_conv
+ all_conv;
+
+(*cf. BOTTOM_CONV in HOL*)
+fun bottom_conv conv ctxt ct =
+ (sub_conv (bottom_conv conv) ctxt then_conv conv ctxt) ct;
+
+(*cf. TOP_CONV in HOL*)
+fun top_conv conv ctxt ct =
+ (conv ctxt then_conv sub_conv (top_conv conv) ctxt) ct;
+
+(*cf. TOP_SWEEP_CONV in HOL*)
+fun top_sweep_conv conv ctxt ct =
+ (conv ctxt else_conv sub_conv (top_sweep_conv conv) ctxt) ct;
+
(* primitive logic *)
@@ -136,6 +166,8 @@
handle Pattern.MATCH => raise CTERM ("rewr_conv", [lhs, ct])
end;
+fun rewrs_conv rules = first_conv (map rewr_conv rules);
+
(* conversions on HHF rules *)
--- a/src/Tools/more_conv.ML Sat May 15 07:48:24 2010 -0700
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,43 +0,0 @@
-(* Title: Tools/more_conv.ML
- Author: Sascha Boehme, TU Muenchen
-
-Further conversions and conversionals.
-*)
-
-signature MORE_CONV =
-sig
- val rewrs_conv: thm list -> conv
-
- val sub_conv: (Proof.context -> conv) -> Proof.context -> conv
- val bottom_conv: (Proof.context -> conv) -> Proof.context -> conv
- val top_conv: (Proof.context -> conv) -> Proof.context -> conv
- val top_sweep_conv: (Proof.context -> conv) -> Proof.context -> conv
-
- val binder_conv: (Proof.context -> conv) -> Proof.context -> conv
-end
-
-structure More_Conv : MORE_CONV =
-struct
-
-fun rewrs_conv eqs = Conv.first_conv (map Conv.rewr_conv eqs)
-
-
-fun sub_conv conv ctxt =
- Conv.comb_conv (conv ctxt) else_conv
- Conv.abs_conv (fn (_, cx) => conv cx) ctxt else_conv
- Conv.all_conv
-
-fun bottom_conv conv ctxt ct =
- (sub_conv (bottom_conv conv) ctxt then_conv conv ctxt) ct
-
-fun top_conv conv ctxt ct =
- (conv ctxt then_conv sub_conv (top_conv conv) ctxt) ct
-
-fun top_sweep_conv conv ctxt ct =
- (conv ctxt else_conv sub_conv (top_sweep_conv conv) ctxt) ct
-
-
-fun binder_conv cv ctxt =
- Conv.arg_conv (Conv.abs_conv (fn (_, cx) => cv cx) ctxt)
-
-end