--- a/src/Doc/Implementation/Proof.thy Thu Jul 09 15:52:11 2015 +0200
+++ b/src/Doc/Implementation/Proof.thy Thu Jul 09 17:36:04 2015 +0200
@@ -109,7 +109,7 @@
@{index_ML Variable.polymorphic: "Proof.context -> term list -> term list"} \\
@{index_ML Variable.import: "bool -> thm list -> Proof.context ->
((((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list) * thm list) * Proof.context"} \\
- @{index_ML Variable.focus: "term -> Proof.context ->
+ @{index_ML Variable.focus: "binding list option -> term -> Proof.context ->
((string * (string * typ)) list * term) * Proof.context"} \\
\end{mldecls}
@@ -153,8 +153,8 @@
should be accessible to the user, otherwise newly introduced names
are marked as ``internal'' (\secref{sec:names}).
- \item @{ML Variable.focus}~@{text B} decomposes the outermost @{text
- "\<And>"} prefix of proposition @{text "B"}.
+ \item @{ML Variable.focus}~@{text "bindings B"} decomposes the outermost @{text
+ "\<And>"} prefix of proposition @{text "B"}, using the given name bindings.
\end{description}
\<close>
@@ -394,9 +394,12 @@
Proof.context -> int -> tactic"} \\
@{index_ML Subgoal.FOCUS_PARAMS: "(Subgoal.focus -> tactic) ->
Proof.context -> int -> tactic"} \\
- @{index_ML Subgoal.focus: "Proof.context -> int -> thm -> Subgoal.focus * thm"} \\
- @{index_ML Subgoal.focus_prems: "Proof.context -> int -> thm -> Subgoal.focus * thm"} \\
- @{index_ML Subgoal.focus_params: "Proof.context -> int -> thm -> Subgoal.focus * thm"} \\
+ @{index_ML Subgoal.focus: "Proof.context -> int -> binding list option ->
+ thm -> Subgoal.focus * thm"} \\
+ @{index_ML Subgoal.focus_prems: "Proof.context -> int -> binding list option ->
+ thm -> Subgoal.focus * thm"} \\
+ @{index_ML Subgoal.focus_params: "Proof.context -> int -> binding list option ->
+ thm -> Subgoal.focus * thm"} \\
\end{mldecls}
\begin{mldecls}
@@ -473,7 +476,7 @@
ML_val
\<open>val {goal, context = goal_ctxt, ...} = @{Isar.goal};
val (focus as {params, asms, concl, ...}, goal') =
- Subgoal.focus goal_ctxt 1 goal;
+ Subgoal.focus goal_ctxt 1 (SOME [@{binding x}]) goal;
val [A, B] = #prems focus;
val [(_, x)] = #params focus;\<close>
sorry
--- a/src/Doc/more_antiquote.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/Doc/more_antiquote.ML Thu Jul 09 17:36:04 2015 +0200
@@ -25,7 +25,7 @@
let
val thy = Proof_Context.theory_of ctxt;
val const = Code.check_const thy raw_const;
- val (_, eqngr) = Code_Preproc.obtain true thy [const] [];
+ val (_, eqngr) = Code_Preproc.obtain true ctxt [const] [];
fun holize thm = @{thm meta_eq_to_obj_eq} OF [thm];
val thms = Code_Preproc.cert eqngr const
|> Code.equations_of_cert thy
--- a/src/HOL/Decision_Procs/Polynomial_List.thy Thu Jul 09 15:52:11 2015 +0200
+++ b/src/HOL/Decision_Procs/Polynomial_List.thy Thu Jul 09 17:36:04 2015 +0200
@@ -129,29 +129,38 @@
by (cases a) (simp_all add: add.commute)
qed
-lemma (in comm_semiring_0) padd_assoc: "\<forall>b c. (a +++ b) +++ c = a +++ (b +++ c)"
- apply (induct a)
- apply simp
- apply clarify
- apply (case_tac b)
- apply (simp_all add: ac_simps)
- done
+lemma (in comm_semiring_0) padd_assoc: "(a +++ b) +++ c = a +++ (b +++ c)"
+proof (induct a arbitrary: b c)
+ case Nil
+ then show ?case
+ by simp
+next
+ case Cons
+ then show ?case
+ by (cases b) (simp_all add: ac_simps)
+qed
lemma (in semiring_0) poly_cmult_distr: "a %* (p +++ q) = a %* p +++ a %* q"
- apply (induct p arbitrary: q)
- apply simp
- apply (case_tac q)
- apply (simp_all add: distrib_left)
- done
+proof (induct p arbitrary: q)
+ case Nil
+ then show ?case
+ by simp
+next
+ case Cons
+ then show ?case
+ by (cases q) (simp_all add: distrib_left)
+qed
lemma (in ring_1) pmult_by_x[simp]: "[0, 1] *** t = 0 # t"
- apply (induct t)
- apply simp
- apply (auto simp add: padd_commut)
- apply (case_tac t)
- apply auto
- done
-
+proof (induct t)
+ case Nil
+ then show ?case
+ by simp
+next
+ case (Cons a t)
+ then show ?case
+ by (cases t) (auto simp add: padd_commut)
+qed
text \<open>Properties of evaluation of polynomials.\<close>
@@ -167,18 +176,21 @@
qed
lemma (in comm_semiring_0) poly_cmult: "poly (c %* p) x = c * poly p x"
- apply (induct p)
- apply (case_tac [2] "x = zero")
- apply (auto simp add: distrib_left ac_simps)
- done
+proof (induct p)
+ case Nil
+ then show ?case
+ by simp
+next
+ case Cons
+ then show ?case
+ by (cases "x = zero") (auto simp add: distrib_left ac_simps)
+qed
lemma (in comm_semiring_0) poly_cmult_map: "poly (map (op * c) p) x = c * poly p x"
by (induct p) (auto simp add: distrib_left ac_simps)
lemma (in comm_ring_1) poly_minus: "poly (-- p) x = - (poly p x)"
- apply (simp add: poly_minus_def)
- apply (auto simp add: poly_cmult)
- done
+ by (simp add: poly_minus_def) (auto simp add: poly_cmult)
lemma (in comm_semiring_0) poly_mult: "poly (p1 *** p2) x = poly p1 x * poly p2 x"
proof (induct p1 arbitrary: p2)
@@ -186,7 +198,7 @@
then show ?case
by simp
next
- case (Cons a as p2)
+ case (Cons a as)
then show ?case
by (cases as) (simp_all add: poly_cmult poly_add distrib_right distrib_left ac_simps)
qed
@@ -216,16 +228,16 @@
subsection \<open>Key Property: if @{term "f a = 0"} then @{term "(x - a)"} divides @{term "p(x)"}.\<close>
-lemma (in comm_ring_1) lemma_poly_linear_rem: "\<forall>h. \<exists>q r. h#t = [r] +++ [-a, 1] *** q"
-proof (induct t)
+lemma (in comm_ring_1) lemma_poly_linear_rem: "\<exists>q r. h#t = [r] +++ [-a, 1] *** q"
+proof (induct t arbitrary: h)
case Nil
- have "[h] = [h] +++ [- a, 1] *** []" for h by simp
+ have "[h] = [h] +++ [- a, 1] *** []" by simp
then show ?case by blast
next
case (Cons x xs)
- have "\<exists>q r. h # x # xs = [r] +++ [-a, 1] *** q" for h
+ have "\<exists>q r. h # x # xs = [r] +++ [-a, 1] *** q"
proof -
- from Cons.hyps obtain q r where qr: "x # xs = [r] +++ [- a, 1] *** q"
+ from Cons obtain q r where qr: "x # xs = [r] +++ [- a, 1] *** q"
by blast
have "h # x # xs = [a * r + h] +++ [-a, 1] *** (r # q)"
using qr by (cases q) (simp_all add: algebra_simps)
@@ -237,7 +249,7 @@
lemma (in comm_ring_1) poly_linear_rem: "\<exists>q r. h#t = [r] +++ [-a, 1] *** q"
using lemma_poly_linear_rem [where t = t and a = a] by auto
-lemma (in comm_ring_1) poly_linear_divides: "(poly p a = 0) \<longleftrightarrow> p = [] \<or> (\<exists>q. p = [-a, 1] *** q)"
+lemma (in comm_ring_1) poly_linear_divides: "poly p a = 0 \<longleftrightarrow> p = [] \<or> (\<exists>q. p = [-a, 1] *** q)"
proof (cases p)
case Nil
then show ?thesis by simp
@@ -264,12 +276,12 @@
qed
lemma (in semiring_0) lemma_poly_length_mult[simp]:
- "\<forall>h k a. length (k %* p +++ (h # (a %* p))) = Suc (length p)"
- by (induct p) auto
+ "length (k %* p +++ (h # (a %* p))) = Suc (length p)"
+ by (induct p arbitrary: h k a) auto
lemma (in semiring_0) lemma_poly_length_mult2[simp]:
- "\<forall>h k. length (k %* p +++ (h # p)) = Suc (length p)"
- by (induct p) auto
+ "length (k %* p +++ (h # p)) = Suc (length p)"
+ by (induct p arbitrary: h k) auto
lemma (in ring_1) poly_length_mult[simp]: "length([-a,1] *** q) = Suc (length q)"
by auto
@@ -281,9 +293,9 @@
by (induct p) auto
lemma (in semiring_0) poly_add_length: "length (p1 +++ p2) = max (length p1) (length p2)"
- by (induct p1 arbitrary: p2) (simp_all, arith)
+ by (induct p1 arbitrary: p2) auto
-lemma (in semiring_0) poly_root_mult_length[simp]: "length([a,b] *** p) = Suc (length p)"
+lemma (in semiring_0) poly_root_mult_length[simp]: "length ([a, b] *** p) = Suc (length p)"
by (simp add: poly_add_length)
lemma (in idom) poly_mult_not_eq_poly_Nil[simp]:
@@ -301,15 +313,15 @@
text \<open>A nontrivial polynomial of degree n has no more than n roots.\<close>
lemma (in idom) poly_roots_index_lemma:
- assumes p: "poly p x \<noteq> poly [] x"
- and n: "length p = n"
+ assumes "poly p x \<noteq> poly [] x"
+ and "length p = n"
shows "\<exists>i. \<forall>x. poly p x = 0 \<longrightarrow> (\<exists>m\<le>n. x = i m)"
- using p n
+ using assms
proof (induct n arbitrary: p x)
case 0
then show ?case by simp
next
- case (Suc n p x)
+ case (Suc n)
have False if C: "\<And>i. \<exists>x. poly p x = 0 \<and> (\<forall>m\<le>Suc n. x \<noteq> i m)"
proof -
from Suc.prems have p0: "poly p x \<noteq> 0" "p \<noteq> []"
@@ -325,14 +337,14 @@
using q Suc.prems(2) by simp
from q p0 have qx: "poly q x \<noteq> poly [] x"
by (auto simp add: poly_mult poly_add poly_cmult)
- from Suc.hyps[OF qx lg] obtain i where i: "\<forall>x. poly q x = 0 \<longrightarrow> (\<exists>m\<le>n. x = i m)"
+ from Suc.hyps[OF qx lg] obtain i where i: "\<And>x. poly q x = 0 \<longrightarrow> (\<exists>m\<le>n. x = i m)"
by blast
let ?i = "\<lambda>m. if m = Suc n then a else i m"
from C[of ?i] obtain y where y: "poly p y = 0" "\<forall>m\<le> Suc n. y \<noteq> ?i m"
by blast
from y have "y = a \<or> poly q y = 0"
by (simp only: q poly_mult_eq_zero_disj poly_add) (simp add: algebra_simps)
- with i[rule_format, of y] y(1) y(2) show ?thesis
+ with i[of y] y(1) y(2) show ?thesis
apply auto
apply (erule_tac x = "m" in allE)
apply auto
@@ -343,12 +355,13 @@
lemma (in idom) poly_roots_index_length:
- "poly p x \<noteq> poly [] x \<Longrightarrow> \<exists>i. \<forall>x. (poly p x = 0) \<longrightarrow> (\<exists>n. n \<le> length p \<and> x = i n)"
+ "poly p x \<noteq> poly [] x \<Longrightarrow> \<exists>i. \<forall>x. poly p x = 0 \<longrightarrow> (\<exists>n. n \<le> length p \<and> x = i n)"
by (blast intro: poly_roots_index_lemma)
lemma (in idom) poly_roots_finite_lemma1:
- "poly p x \<noteq> poly [] x \<Longrightarrow> \<exists>N i. \<forall>x. (poly p x = 0) \<longrightarrow> (\<exists>n::nat. n < N \<and> x = i n)"
- apply (drule poly_roots_index_length, safe)
+ "poly p x \<noteq> poly [] x \<Longrightarrow> \<exists>N i. \<forall>x. poly p x = 0 \<longrightarrow> (\<exists>n::nat. n < N \<and> x = i n)"
+ apply (drule poly_roots_index_length)
+ apply safe
apply (rule_tac x = "Suc (length p)" in exI)
apply (rule_tac x = i in exI)
apply (simp add: less_Suc_eq_le)
@@ -358,8 +371,10 @@
assumes "\<forall>x. P x \<longrightarrow> (\<exists>n. n < length j \<and> x = j!n)"
shows "finite {x. P x}"
proof -
- from assms have "{x. P x} \<subseteq> set j" by auto
- then show ?thesis using finite_subset by auto
+ from assms have "{x. P x} \<subseteq> set j"
+ by auto
+ then show ?thesis
+ using finite_subset by auto
qed
lemma (in idom) poly_roots_finite_lemma2:
@@ -379,7 +394,8 @@
assume F: "finite (UNIV :: 'a set)"
have "finite (UNIV :: nat set)"
proof (rule finite_imageD)
- have "of_nat ` UNIV \<subseteq> UNIV" by simp
+ have "of_nat ` UNIV \<subseteq> UNIV"
+ by simp
then show "finite (of_nat ` UNIV :: 'a set)"
using F by (rule finite_subset)
show "inj (of_nat :: nat \<Rightarrow> 'a)"
@@ -499,12 +515,18 @@
qed
lemma (in idom_char_0) poly_zero: "poly p = poly [] \<longleftrightarrow> (\<forall>c \<in> set p. c = 0)"
- apply (induct p)
- apply simp
- apply (rule iffI)
- apply (drule poly_zero_lemma')
- apply auto
- done
+proof (induct p)
+ case Nil
+ then show ?case by simp
+next
+ case Cons
+ show ?case
+ apply (rule iffI)
+ apply (drule poly_zero_lemma')
+ using Cons
+ apply auto
+ done
+qed
lemma (in idom_char_0) poly_0: "\<forall>c \<in> set p. c = 0 \<Longrightarrow> poly p x = 0"
unfolding poly_zero[symmetric] by simp
@@ -594,10 +616,10 @@
text \<open>At last, we can consider the order of a root.\<close>
lemma (in idom_char_0) poly_order_exists_lemma:
- assumes lp: "length p = d"
- and p: "poly p \<noteq> poly []"
+ assumes "length p = d"
+ and "poly p \<noteq> poly []"
shows "\<exists>n q. p = mulexp n [-a, 1] q \<and> poly q a \<noteq> 0"
- using lp p
+ using assms
proof (induct d arbitrary: p)
case 0
then show ?case by simp
@@ -613,15 +635,15 @@
from True[unfolded poly_linear_divides] pN obtain q where q: "p = [-a, 1] *** q"
by blast
from q h True have qh: "length q = n" "poly q \<noteq> poly []"
- apply -
- apply simp
+ apply simp_all
apply (simp only: fun_eq_iff)
apply (rule ccontr)
apply (simp add: fun_eq_iff poly_add poly_cmult)
done
from Suc.hyps[OF qh] obtain m r where mr: "q = mulexp m [-a,1] r" "poly r a \<noteq> 0"
by blast
- from mr q have "p = mulexp (Suc m) [-a,1] r \<and> poly r a \<noteq> 0" by simp
+ from mr q have "p = mulexp (Suc m) [-a,1] r \<and> poly r a \<noteq> 0"
+ by simp
then show ?thesis by blast
next
case False
@@ -680,7 +702,7 @@
case 0
show ?case
proof (rule ccontr)
- assume "\<not> poly (mulexp 0 [- a, 1] q) \<noteq> poly ([- a, 1] %^ Suc 0 *** m)"
+ assume "\<not> ?thesis"
then have "poly q a = 0"
by (simp add: poly_add poly_cmult)
with \<open>poly q a \<noteq> 0\<close> show False
@@ -689,7 +711,7 @@
next
case (Suc n)
show ?case
- by (rule pexp_Suc [THEN ssubst], rule ccontr)
+ by (rule pexp_Suc [THEN ssubst])
(simp add: poly_mult_left_cancel poly_mult_assoc Suc del: pmult_Cons pexp_Suc)
qed
ultimately show False by simp
@@ -906,8 +928,7 @@
lemma (in semiring_0) pnormal_id: "pnormal p \<longleftrightarrow> 0 < length p \<and> last p \<noteq> 0"
using pnormal_last_length pnormal_length pnormal_last_nonzero by blast
-lemma (in idom_char_0) poly_Cons_eq:
- "poly (c # cs) = poly (d # ds) \<longleftrightarrow> c = d \<and> poly cs = poly ds"
+lemma (in idom_char_0) poly_Cons_eq: "poly (c # cs) = poly (d # ds) \<longleftrightarrow> c = d \<and> poly cs = poly ds"
(is "?lhs \<longleftrightarrow> ?rhs")
proof
show ?rhs if ?lhs
@@ -974,11 +995,12 @@
"last ((a %* p) +++ (x # (b %* p))) = (if p = [] then x else b * last p)"
apply (induct p arbitrary: a x b)
apply auto
- apply (rename_tac a p aa x b)
- apply (subgoal_tac "padd (cmult aa p) (times b a # cmult b p) \<noteq> []")
- apply simp
- apply (induct_tac p)
- apply auto
+ subgoal for a p c x b
+ apply (subgoal_tac "padd (cmult c p) (times b a # cmult b p) \<noteq> []")
+ apply simp
+ apply (induct p)
+ apply auto
+ done
done
lemma (in semiring_1) last_linear_mul:
@@ -1097,13 +1119,18 @@
lemma poly_mono:
fixes x :: "'a::linordered_idom"
shows "abs x \<le> k \<Longrightarrow> abs (poly p x) \<le> poly (map abs p) k"
- apply (induct p)
- apply auto
- apply (rename_tac a p)
- apply (rule_tac y = "abs a + abs (x * poly p x)" in order_trans)
- apply (rule abs_triangle_ineq)
- apply (auto intro!: mult_mono simp add: abs_mult)
- done
+proof (induct p)
+ case Nil
+ then show ?case by simp
+next
+ case (Cons a p)
+ then show ?case
+ apply auto
+ apply (rule_tac y = "abs a + abs (x * poly p x)" in order_trans)
+ apply (rule abs_triangle_ineq)
+ apply (auto intro!: mult_mono simp add: abs_mult)
+ done
+qed
lemma (in semiring_0) poly_Sing: "poly [c] x = c"
by simp
--- a/src/HOL/Decision_Procs/Rat_Pair.thy Thu Jul 09 15:52:11 2015 +0200
+++ b/src/HOL/Decision_Procs/Rat_Pair.thy Thu Jul 09 17:36:04 2015 +0200
@@ -328,13 +328,7 @@
proof cases
case 1
then show ?thesis
- apply (cases "a = 0")
- apply (simp_all add: x y Nmul_def INum_def Let_def)
- apply (cases "b = 0")
- apply simp_all
- apply (cases "a' = 0")
- apply simp_all
- done
+ by (auto simp add: x y Nmul_def INum_def)
next
case neq: 2
let ?g = "gcd (a * a') (b * b')"
@@ -347,21 +341,21 @@
qed
qed
-lemma Nneg[simp]: "INum (~\<^sub>N x) = - (INum x ::'a::field)"
+lemma Nneg[simp]: "INum (~\<^sub>N x) = - (INum x :: 'a::field)"
by (simp add: Nneg_def split_def INum_def)
-lemma Nsub[simp]: "INum (x -\<^sub>N y) = INum x - (INum y:: 'a :: {field_char_0,field})"
+lemma Nsub[simp]: "INum (x -\<^sub>N y) = INum x - (INum y:: 'a::{field_char_0,field})"
by (simp add: Nsub_def split_def)
-lemma Ninv[simp]: "INum (Ninv x) = (1::'a :: field) / (INum x)"
+lemma Ninv[simp]: "INum (Ninv x) = (1 :: 'a::field) / INum x"
by (simp add: Ninv_def INum_def split_def)
-lemma Ndiv[simp]: "INum (x \<div>\<^sub>N y) = INum x / (INum y ::'a :: {field_char_0,field})"
+lemma Ndiv[simp]: "INum (x \<div>\<^sub>N y) = INum x / (INum y :: 'a::{field_char_0,field})"
by (simp add: Ndiv_def)
lemma Nlt0_iff[simp]:
assumes nx: "isnormNum x"
- shows "((INum x :: 'a::{field_char_0,linordered_field})< 0) = 0>\<^sub>N x"
+ shows "((INum x :: 'a::{field_char_0,linordered_field}) < 0) = 0>\<^sub>N x"
proof -
obtain a b where x: "x = (a, b)" by (cases x)
show ?thesis
@@ -401,7 +395,7 @@
lemma Ngt0_iff[simp]:
assumes nx: "isnormNum x"
- shows "((INum x :: 'a::{field_char_0,linordered_field})> 0) = 0<\<^sub>N x"
+ shows "((INum x :: 'a::{field_char_0,linordered_field}) > 0) = 0<\<^sub>N x"
proof -
obtain a b where x: "x = (a, b)" by (cases x)
show ?thesis
@@ -421,7 +415,7 @@
lemma Nge0_iff[simp]:
assumes nx: "isnormNum x"
- shows "((INum x :: 'a::{field_char_0,linordered_field}) \<ge> 0) = 0\<le>\<^sub>N x"
+ shows "(INum x :: 'a::{field_char_0,linordered_field}) \<ge> 0 \<longleftrightarrow> 0\<le>\<^sub>N x"
proof -
obtain a b where x: "x = (a, b)" by (cases x)
show ?thesis
@@ -442,12 +436,12 @@
lemma Nlt_iff[simp]:
assumes nx: "isnormNum x"
and ny: "isnormNum y"
- shows "((INum x :: 'a::{field_char_0,linordered_field}) < INum y) = (x <\<^sub>N y)"
+ shows "((INum x :: 'a::{field_char_0,linordered_field}) < INum y) \<longleftrightarrow> x <\<^sub>N y"
proof -
let ?z = "0::'a"
- have "((INum x ::'a) < INum y) = (INum (x -\<^sub>N y) < ?z)"
+ have "((INum x ::'a) < INum y) \<longleftrightarrow> INum (x -\<^sub>N y) < ?z"
using nx ny by simp
- also have "\<dots> = (0>\<^sub>N (x -\<^sub>N y))"
+ also have "\<dots> \<longleftrightarrow> 0>\<^sub>N (x -\<^sub>N y)"
using Nlt0_iff[OF Nsub_normN[OF ny]] by simp
finally show ?thesis
by (simp add: Nlt_def)
@@ -456,11 +450,11 @@
lemma Nle_iff[simp]:
assumes nx: "isnormNum x"
and ny: "isnormNum y"
- shows "((INum x :: 'a::{field_char_0,linordered_field})\<le> INum y) = (x \<le>\<^sub>N y)"
+ shows "((INum x :: 'a::{field_char_0,linordered_field}) \<le> INum y) \<longleftrightarrow> x \<le>\<^sub>N y"
proof -
- have "((INum x ::'a) \<le> INum y) = (INum (x -\<^sub>N y) \<le> (0::'a))"
+ have "((INum x ::'a) \<le> INum y) \<longleftrightarrow> INum (x -\<^sub>N y) \<le> (0::'a)"
using nx ny by simp
- also have "\<dots> = (0\<ge>\<^sub>N (x -\<^sub>N y))"
+ also have "\<dots> \<longleftrightarrow> 0\<ge>\<^sub>N (x -\<^sub>N y)"
using Nle0_iff[OF Nsub_normN[OF ny]] by simp
finally show ?thesis
by (simp add: Nle_def)
@@ -508,7 +502,7 @@
shows "normNum (normNum x) = normNum x"
by simp
-lemma normNum0[simp]: "normNum (0,b) = 0\<^sub>N" "normNum (a,0) = 0\<^sub>N"
+lemma normNum0[simp]: "normNum (0, b) = 0\<^sub>N" "normNum (a, 0) = 0\<^sub>N"
by (simp_all add: normNum_def)
lemma normNum_Nadd:
@@ -520,29 +514,40 @@
assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
shows "normNum x +\<^sub>N y = x +\<^sub>N y"
proof -
- have n: "isnormNum (normNum x +\<^sub>N y)" "isnormNum (x +\<^sub>N y)" by simp_all
- have "INum (normNum x +\<^sub>N y) = INum x + (INum y :: 'a)" by simp
- also have "\<dots> = INum (x +\<^sub>N y)" by simp
- finally show ?thesis using isnormNum_unique[OF n] by simp
+ have n: "isnormNum (normNum x +\<^sub>N y)" "isnormNum (x +\<^sub>N y)"
+ by simp_all
+ have "INum (normNum x +\<^sub>N y) = INum x + (INum y :: 'a)"
+ by simp
+ also have "\<dots> = INum (x +\<^sub>N y)"
+ by simp
+ finally show ?thesis
+ using isnormNum_unique[OF n] by simp
qed
lemma Nadd_normNum2[simp]:
assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
shows "x +\<^sub>N normNum y = x +\<^sub>N y"
proof -
- have n: "isnormNum (x +\<^sub>N normNum y)" "isnormNum (x +\<^sub>N y)" by simp_all
- have "INum (x +\<^sub>N normNum y) = INum x + (INum y :: 'a)" by simp
- also have "\<dots> = INum (x +\<^sub>N y)" by simp
- finally show ?thesis using isnormNum_unique[OF n] by simp
+ have n: "isnormNum (x +\<^sub>N normNum y)" "isnormNum (x +\<^sub>N y)"
+ by simp_all
+ have "INum (x +\<^sub>N normNum y) = INum x + (INum y :: 'a)"
+ by simp
+ also have "\<dots> = INum (x +\<^sub>N y)"
+ by simp
+ finally show ?thesis
+ using isnormNum_unique[OF n] by simp
qed
lemma Nadd_assoc:
assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
shows "x +\<^sub>N y +\<^sub>N z = x +\<^sub>N (y +\<^sub>N z)"
proof -
- have n: "isnormNum (x +\<^sub>N y +\<^sub>N z)" "isnormNum (x +\<^sub>N (y +\<^sub>N z))" by simp_all
- have "INum (x +\<^sub>N y +\<^sub>N z) = (INum (x +\<^sub>N (y +\<^sub>N z)) :: 'a)" by simp
- with isnormNum_unique[OF n] show ?thesis by simp
+ have n: "isnormNum (x +\<^sub>N y +\<^sub>N z)" "isnormNum (x +\<^sub>N (y +\<^sub>N z))"
+ by simp_all
+ have "INum (x +\<^sub>N y +\<^sub>N z) = (INum (x +\<^sub>N (y +\<^sub>N z)) :: 'a)"
+ by simp
+ with isnormNum_unique[OF n] show ?thesis
+ by simp
qed
lemma Nmul_commute: "isnormNum x \<Longrightarrow> isnormNum y \<Longrightarrow> x *\<^sub>N y = y *\<^sub>N x"
@@ -557,8 +562,10 @@
proof -
from nx ny nz have n: "isnormNum (x *\<^sub>N y *\<^sub>N z)" "isnormNum (x *\<^sub>N (y *\<^sub>N z))"
by simp_all
- have "INum (x +\<^sub>N y +\<^sub>N z) = (INum (x +\<^sub>N (y +\<^sub>N z)) :: 'a)" by simp
- with isnormNum_unique[OF n] show ?thesis by simp
+ have "INum (x +\<^sub>N y +\<^sub>N z) = (INum (x +\<^sub>N (y +\<^sub>N z)) :: 'a)"
+ by simp
+ with isnormNum_unique[OF n] show ?thesis
+ by simp
qed
lemma Nsub0:
@@ -568,9 +575,12 @@
shows "x -\<^sub>N y = 0\<^sub>N \<longleftrightarrow> x = y"
proof -
from isnormNum_unique[where 'a = 'a, OF Nsub_normN[OF y], where y="0\<^sub>N"]
- have "(x -\<^sub>N y = 0\<^sub>N) = (INum (x -\<^sub>N y) = (INum 0\<^sub>N :: 'a)) " by simp
- also have "\<dots> = (INum x = (INum y :: 'a))" by simp
- also have "\<dots> = (x = y)" using x y by simp
+ have "x -\<^sub>N y = 0\<^sub>N \<longleftrightarrow> INum (x -\<^sub>N y) = (INum 0\<^sub>N :: 'a)"
+ by simp
+ also have "\<dots> \<longleftrightarrow> INum x = (INum y :: 'a)"
+ by simp
+ also have "\<dots> \<longleftrightarrow> x = y"
+ using x y by simp
finally show ?thesis .
qed
--- a/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy Thu Jul 09 15:52:11 2015 +0200
+++ b/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy Thu Jul 09 17:36:04 2015 +0200
@@ -8,7 +8,7 @@
imports Complex_Main Rat_Pair Polynomial_List
begin
-subsection\<open>Datatype of polynomial expressions\<close>
+subsection \<open>Datatype of polynomial expressions\<close>
datatype poly = C Num | Bound nat | Add poly poly | Sub poly poly
| Mul poly poly| Neg poly| Pw poly nat| CN poly nat poly
@@ -66,7 +66,7 @@
| "decrpoly a = a"
-subsection\<open>Degrees and heads and coefficients\<close>
+subsection \<open>Degrees and heads and coefficients\<close>
fun degree :: "poly \<Rightarrow> nat"
where
@@ -78,7 +78,8 @@
"head (CN c 0 p) = head p"
| "head p = p"
-(* More general notions of degree and head *)
+text \<open>More general notions of degree and head.\<close>
+
fun degreen :: "poly \<Rightarrow> nat \<Rightarrow> nat"
where
"degreen (CN c n p) = (\<lambda>m. if n = m then 1 + degreen p n else 0)"
@@ -110,7 +111,7 @@
| "headconst (C n) = n"
-subsection\<open>Operations for normalization\<close>
+subsection \<open>Operations for normalization\<close>
declare if_cong[fundef_cong del]
declare let_cong[fundef_cong del]
@@ -179,7 +180,7 @@
| "polynate (Pw p n) = polynate p ^\<^sub>p n"
| "polynate (CN c n p) = polynate (Add c (Mul (Bound n) p))"
| "polynate (C c) = C (normNum c)"
-by pat_completeness auto
+ by pat_completeness auto
termination by (relation "measure polysize") auto
fun poly_cmul :: "Num \<Rightarrow> poly \<Rightarrow> poly"
@@ -233,7 +234,7 @@
| "poly_deriv p = 0\<^sub>p"
-subsection\<open>Semantics of the polynomial representation\<close>
+subsection \<open>Semantics of the polynomial representation\<close>
primrec Ipoly :: "'a list \<Rightarrow> poly \<Rightarrow> 'a::{field_char_0,field,power}"
where
@@ -246,8 +247,7 @@
| "Ipoly bs (Pw t n) = Ipoly bs t ^ n"
| "Ipoly bs (CN c n p) = Ipoly bs c + (bs!n) * Ipoly bs p"
-abbreviation Ipoly_syntax :: "poly \<Rightarrow> 'a list \<Rightarrow>'a::{field_char_0,field,power}"
- ("\<lparr>_\<rparr>\<^sub>p\<^bsup>_\<^esup>")
+abbreviation Ipoly_syntax :: "poly \<Rightarrow> 'a list \<Rightarrow>'a::{field_char_0,field,power}" ("\<lparr>_\<rparr>\<^sub>p\<^bsup>_\<^esup>")
where "\<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<equiv> Ipoly bs p"
lemma Ipoly_CInt: "Ipoly bs (C (i, 1)) = of_int i"
@@ -273,14 +273,14 @@
definition isnpoly :: "poly \<Rightarrow> bool"
where "isnpoly p = isnpolyh p 0"
-text\<open>polyadd preserves normal forms\<close>
+text \<open>polyadd preserves normal forms\<close>
lemma polyadd_normh: "isnpolyh p n0 \<Longrightarrow> isnpolyh q n1 \<Longrightarrow> isnpolyh (polyadd p q) (min n0 n1)"
proof (induct p q arbitrary: n0 n1 rule: polyadd.induct)
case (2 ab c' n' p' n0 n1)
from 2 have th1: "isnpolyh (C ab) (Suc n')"
by simp
- from 2(3) have th2: "isnpolyh c' (Suc n')" and nplen1: "n' \<ge> n1"
+ from 2(3) have th2: "isnpolyh c' (Suc n')" and nplen1: "n' \<ge> n1"
by simp_all
with isnpolyh_mono have cp: "isnpolyh c' (Suc n')"
by simp
@@ -314,11 +314,11 @@
by simp
from 4 have n'gen1: "n' \<ge> n1"
by simp
- have "n < n' \<or> n' < n \<or> n = n'"
- by auto
- moreover
- {
- assume eq: "n = n'"
+ consider (eq) "n = n'" | (lt) "n < n'" | (gt) "n > n'"
+ by arith
+ then show ?case
+ proof cases
+ case eq
with "4.hyps"(3)[OF nc nc']
have ncc':"isnpolyh (c +\<^sub>p c') (Suc n)"
by auto
@@ -329,12 +329,10 @@
by simp
have minle: "min n0 n1 \<le> n'"
using ngen0 n'gen1 eq by simp
- from minle npp' ncc'n01 4 eq ngen0 n'gen1 ncc' have ?case
+ from minle npp' ncc'n01 4 eq ngen0 n'gen1 ncc' show ?thesis
by (simp add: Let_def)
- }
- moreover
- {
- assume lt: "n < n'"
+ next
+ case lt
have "min n0 n1 \<le> n0"
by simp
with 4 lt have th1:"min n0 n1 \<le> n"
@@ -347,12 +345,10 @@
by arith
from "4.hyps"(1)[OF th21 th22] have "isnpolyh (polyadd c (CN c' n' p')) (Suc n)"
using th23 by simp
- with 4 lt th1 have ?case
+ with 4 lt th1 show ?thesis
by simp
- }
- moreover
- {
- assume gt: "n' < n"
+ next
+ case gt
then have gt': "n' < n \<and> \<not> n < n'"
by simp
have "min n0 n1 \<le> n1"
@@ -367,10 +363,9 @@
by arith
from "4.hyps"(2)[OF th22 th21] have "isnpolyh (polyadd (CN c n p) c') (Suc n')"
using th23 by simp
- with 4 gt th1 have ?case
+ with 4 gt th1 show ?thesis
by simp
- }
- ultimately show ?case by blast
+ qed
qed auto
lemma polyadd[simp]: "Ipoly bs (polyadd p q) = Ipoly bs p + Ipoly bs q"
@@ -378,9 +373,9 @@
(auto simp add: Let_def field_simps distrib_left[symmetric] simp del: distrib_left_NO_MATCH)
lemma polyadd_norm: "isnpoly p \<Longrightarrow> isnpoly q \<Longrightarrow> isnpoly (polyadd p q)"
- using polyadd_normh[of "p" "0" "q" "0"] isnpoly_def by simp
+ using polyadd_normh[of p 0 q 0] isnpoly_def by simp
-text\<open>The degree of addition and other general lemmas needed for the normal form of polymul\<close>
+text \<open>The degree of addition and other general lemmas needed for the normal form of polymul.\<close>
lemma polyadd_different_degreen:
assumes "isnpolyh p n0"
@@ -391,17 +386,13 @@
using assms
proof (induct p q arbitrary: m n0 n1 rule: polyadd.induct)
case (4 c n p c' n' p' m n0 n1)
- have "n' = n \<or> n < n' \<or> n' < n" by arith
- then show ?case
- proof (elim disjE)
- assume [simp]: "n' = n"
- from 4(4)[of n n m] 4(3)[of "Suc n" "Suc n" m] 4(5-7)
+ show ?case
+ proof (cases "n = n'")
+ case True
+ with 4(4)[of n n m] 4(3)[of "Suc n" "Suc n" m] 4(5-7)
show ?thesis by (auto simp: Let_def)
next
- assume "n < n'"
- with 4 show ?thesis by auto
- next
- assume "n' < n"
+ case False
with 4 show ?thesis by auto
qed
qed auto
@@ -441,13 +432,15 @@
by (cases n) auto
next
case (4 c n p c' n' p' n0 n1 m)
- have "n' = n \<or> n < n' \<or> n' < n" by arith
- then show ?case
- proof (elim disjE)
- assume [simp]: "n' = n"
- from 4(4)[of n n m] 4(3)[of "Suc n" "Suc n" m] 4(5-7)
+ show ?case
+ proof (cases "n = n'")
+ case True
+ with 4(4)[of n n m] 4(3)[of "Suc n" "Suc n" m] 4(5-7)
show ?thesis by (auto simp: Let_def)
- qed simp_all
+ next
+ case False
+ then show ?thesis by simp
+ qed
qed auto
lemma polyadd_eq_const_degreen:
@@ -458,26 +451,16 @@
using assms
proof (induct p q arbitrary: m n0 n1 c rule: polyadd.induct)
case (4 c n p c' n' p' m n0 n1 x)
- {
- assume nn': "n' < n"
- then have ?case using 4 by simp
- }
- moreover
- {
- assume nn': "\<not> n' < n"
- then have "n < n' \<or> n = n'" by arith
- moreover { assume "n < n'" with 4 have ?case by simp }
- moreover
- {
- assume eq: "n = n'"
- then have ?case using 4
- apply (cases "p +\<^sub>p p' = 0\<^sub>p")
- apply (auto simp add: Let_def)
- done
- }
- ultimately have ?case by blast
- }
- ultimately show ?case by blast
+ consider "n = n'" | "n > n' \<or> n < n'" by arith
+ then show ?case
+ proof cases
+ case 1
+ with 4 show ?thesis
+ by (cases "p +\<^sub>p p' = 0\<^sub>p") (auto simp add: Let_def)
+ next
+ case 2
+ with 4 show ?thesis by auto
+ qed
qed simp_all
lemma polymul_properties:
@@ -500,7 +483,7 @@
then show ?case by auto
next
case (3 n0 n1)
- then show ?case using "2.hyps" by auto
+ then show ?case using "2.hyps" by auto
}
next
case (3 c n p c')
@@ -720,7 +703,7 @@
qed
-text\<open>polyneg is a negation and preserves normal forms\<close>
+text \<open>polyneg is a negation and preserves normal forms\<close>
lemma polyneg[simp]: "Ipoly bs (polyneg p) = - Ipoly bs p"
by (induct p rule: polyneg.induct) auto
@@ -738,7 +721,7 @@
using isnpoly_def polyneg_normh by simp
-text\<open>polysub is a substraction and preserves normal forms\<close>
+text \<open>polysub is a substraction and preserves normal forms\<close>
lemma polysub[simp]: "Ipoly bs (polysub p q) = Ipoly bs p - Ipoly bs q"
by (simp add: polysub_def)
@@ -762,7 +745,7 @@
by (induct p q arbitrary: n0 n1 rule:polyadd.induct)
(auto simp: Nsub0[simplified Nsub_def] Let_def)
-text\<open>polypow is a power function and preserves normal forms\<close>
+text \<open>polypow is a power function and preserves normal forms\<close>
lemma polypow[simp]:
"Ipoly bs (polypow n p) = (Ipoly bs p :: 'a::{field_char_0,field}) ^ n"
@@ -830,7 +813,7 @@
shows "isnpoly p \<Longrightarrow> isnpoly (polypow k p)"
by (simp add: polypow_normh isnpoly_def)
-text\<open>Finally the whole normalization\<close>
+text \<open>Finally the whole normalization\<close>
lemma polynate [simp]:
"Ipoly bs (polynate p) = (Ipoly bs p :: 'a ::{field_char_0,field})"
@@ -843,7 +826,7 @@
(simp_all add: polyadd_norm polymul_norm polysub_norm polyneg_norm polypow_norm,
simp_all add: isnpoly_def)
-text\<open>shift1\<close>
+text \<open>shift1\<close>
lemma shift1: "Ipoly bs (shift1 p) = Ipoly bs (Mul (Bound 0) p)"
@@ -1254,7 +1237,7 @@
qed
-text\<open>consequences of unicity on the algorithms for polynomial normalization\<close>
+text \<open>consequences of unicity on the algorithms for polynomial normalization\<close>
lemma polyadd_commute:
assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
@@ -1328,7 +1311,7 @@
unfolding poly_nate_polypoly' by auto
-subsection\<open>heads, degrees and all that\<close>
+subsection \<open>heads, degrees and all that\<close>
lemma degree_eq_degreen0: "degree p = degreen p 0"
by (induct p rule: degree.induct) simp_all
--- a/src/HOL/Eisbach/match_method.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/HOL/Eisbach/match_method.ML Thu Jul 09 17:36:04 2015 +0200
@@ -304,7 +304,7 @@
fun prep_fact_pat ((x, args), pat) ctxt =
let
- val ((params, pat'), ctxt') = Variable.focus pat ctxt;
+ val ((params, pat'), ctxt') = Variable.focus NONE pat ctxt;
val params' = map (Free o snd) params;
val morphism =
@@ -500,10 +500,10 @@
(* Fix schematics in the goal *)
-fun focus_concl ctxt i goal =
+fun focus_concl ctxt i bindings goal =
let
val ({context = ctxt', concl, params, prems, asms, schematics}, goal') =
- Subgoal.focus_params ctxt i goal;
+ Subgoal.focus_params ctxt i bindings goal;
val (inst, ctxt'') = Variable.import_inst true [Thm.term_of concl] ctxt';
val (_, schematic_terms) = Thm.certify_inst ctxt'' inst;
@@ -711,7 +711,7 @@
val (goal_thins, goal) = get_thinned_prems goal;
val ((local_premids, {context = focus_ctxt, params, asms, concl, ...}), focused_goal) =
- focus_cases (K Subgoal.focus_prems) (focus_concl) ctxt 1 goal
+ focus_cases (K Subgoal.focus_prems) focus_concl ctxt 1 NONE goal
|>> augment_focus;
val texts =
--- a/src/HOL/Library/Float.thy Thu Jul 09 15:52:11 2015 +0200
+++ b/src/HOL/Library/Float.thy Thu Jul 09 17:36:04 2015 +0200
@@ -22,6 +22,7 @@
real_of_float_def[code_unfold]: "real \<equiv> real_of_float"
instance ..
+
end
lemma type_definition_float': "type_definition real float_of float"
@@ -34,7 +35,8 @@
declare [[coercion "real :: float \<Rightarrow> real"]]
lemma real_of_float_eq:
- fixes f1 f2 :: float shows "f1 = f2 \<longleftrightarrow> real f1 = real f2"
+ fixes f1 f2 :: float
+ shows "f1 = f2 \<longleftrightarrow> real f1 = real f2"
unfolding real_of_float_def real_of_float_inject ..
lemma float_of_real[simp]: "float_of (real x) = x"
@@ -43,40 +45,63 @@
lemma real_float[simp]: "x \<in> float \<Longrightarrow> real (float_of x) = x"
unfolding real_of_float_def by (rule float_of_inverse)
+
subsection \<open>Real operations preserving the representation as floating point number\<close>
lemma floatI: fixes m e :: int shows "m * 2 powr e = x \<Longrightarrow> x \<in> float"
by (auto simp: float_def)
-lemma zero_float[simp]: "0 \<in> float" by (auto simp: float_def)
-lemma one_float[simp]: "1 \<in> float" by (intro floatI[of 1 0]) simp
-lemma numeral_float[simp]: "numeral i \<in> float" by (intro floatI[of "numeral i" 0]) simp
-lemma neg_numeral_float[simp]: "- numeral i \<in> float" by (intro floatI[of "- numeral i" 0]) simp
-lemma real_of_int_float[simp]: "real (x :: int) \<in> float" by (intro floatI[of x 0]) simp
-lemma real_of_nat_float[simp]: "real (x :: nat) \<in> float" by (intro floatI[of x 0]) simp
-lemma two_powr_int_float[simp]: "2 powr (real (i::int)) \<in> float" by (intro floatI[of 1 i]) simp
-lemma two_powr_nat_float[simp]: "2 powr (real (i::nat)) \<in> float" by (intro floatI[of 1 i]) simp
-lemma two_powr_minus_int_float[simp]: "2 powr - (real (i::int)) \<in> float" by (intro floatI[of 1 "-i"]) simp
-lemma two_powr_minus_nat_float[simp]: "2 powr - (real (i::nat)) \<in> float" by (intro floatI[of 1 "-i"]) simp
-lemma two_powr_numeral_float[simp]: "2 powr numeral i \<in> float" by (intro floatI[of 1 "numeral i"]) simp
-lemma two_powr_neg_numeral_float[simp]: "2 powr - numeral i \<in> float" by (intro floatI[of 1 "- numeral i"]) simp
-lemma two_pow_float[simp]: "2 ^ n \<in> float" by (intro floatI[of 1 "n"]) (simp add: powr_realpow)
-lemma real_of_float_float[simp]: "real (f::float) \<in> float" by (cases f) simp
+lemma zero_float[simp]: "0 \<in> float"
+ by (auto simp: float_def)
+lemma one_float[simp]: "1 \<in> float"
+ by (intro floatI[of 1 0]) simp
+lemma numeral_float[simp]: "numeral i \<in> float"
+ by (intro floatI[of "numeral i" 0]) simp
+lemma neg_numeral_float[simp]: "- numeral i \<in> float"
+ by (intro floatI[of "- numeral i" 0]) simp
+lemma real_of_int_float[simp]: "real (x :: int) \<in> float"
+ by (intro floatI[of x 0]) simp
+lemma real_of_nat_float[simp]: "real (x :: nat) \<in> float"
+ by (intro floatI[of x 0]) simp
+lemma two_powr_int_float[simp]: "2 powr (real (i::int)) \<in> float"
+ by (intro floatI[of 1 i]) simp
+lemma two_powr_nat_float[simp]: "2 powr (real (i::nat)) \<in> float"
+ by (intro floatI[of 1 i]) simp
+lemma two_powr_minus_int_float[simp]: "2 powr - (real (i::int)) \<in> float"
+ by (intro floatI[of 1 "-i"]) simp
+lemma two_powr_minus_nat_float[simp]: "2 powr - (real (i::nat)) \<in> float"
+ by (intro floatI[of 1 "-i"]) simp
+lemma two_powr_numeral_float[simp]: "2 powr numeral i \<in> float"
+ by (intro floatI[of 1 "numeral i"]) simp
+lemma two_powr_neg_numeral_float[simp]: "2 powr - numeral i \<in> float"
+ by (intro floatI[of 1 "- numeral i"]) simp
+lemma two_pow_float[simp]: "2 ^ n \<in> float"
+ by (intro floatI[of 1 "n"]) (simp add: powr_realpow)
+lemma real_of_float_float[simp]: "real (f::float) \<in> float"
+ by (cases f) simp
lemma plus_float[simp]: "r \<in> float \<Longrightarrow> p \<in> float \<Longrightarrow> r + p \<in> float"
unfolding float_def
proof (safe, simp)
- fix e1 m1 e2 m2 :: int
- { fix e1 m1 e2 m2 :: int assume "e1 \<le> e2"
- then have "m1 * 2 powr e1 + m2 * 2 powr e2 = (m1 + m2 * 2 ^ nat (e2 - e1)) * 2 powr e1"
+ have *: "\<exists>(m::int) (e::int). m1 * 2 powr e1 + m2 * 2 powr e2 = m * 2 powr e"
+ if "e1 \<le> e2" for e1 m1 e2 m2 :: int
+ proof -
+ from that have "m1 * 2 powr e1 + m2 * 2 powr e2 = (m1 + m2 * 2 ^ nat (e2 - e1)) * 2 powr e1"
by (simp add: powr_realpow[symmetric] powr_divide2[symmetric] field_simps)
- then have "\<exists>(m::int) (e::int). m1 * 2 powr e1 + m2 * 2 powr e2 = m * 2 powr e"
- by blast }
- note * = this
- show "\<exists>(m::int) (e::int). m1 * 2 powr e1 + m2 * 2 powr e2 = m * 2 powr e"
- proof (cases e1 e2 rule: linorder_le_cases)
- assume "e2 \<le> e1" from *[OF this, of m2 m1] show ?thesis by (simp add: ac_simps)
- qed (rule *)
+ then show ?thesis
+ by blast
+ qed
+ fix e1 m1 e2 m2 :: int
+ consider "e2 \<le> e1" | "e1 \<le> e2" by (rule linorder_le_cases)
+ then show "\<exists>(m::int) (e::int). m1 * 2 powr e1 + m2 * 2 powr e2 = m * 2 powr e"
+ proof cases
+ case 1
+ from *[OF this, of m2 m1] show ?thesis
+ by (simp add: ac_simps)
+ next
+ case 2
+ then show ?thesis by (rule *)
+ qed
qed
lemma uminus_float[simp]: "x \<in> float \<Longrightarrow> -x \<in> float"
@@ -125,7 +150,8 @@
done
lemma div_numeral_Bit0_float[simp]:
- assumes x: "x / numeral n \<in> float" shows "x / (numeral (Num.Bit0 n)) \<in> float"
+ assumes x: "x / numeral n \<in> float"
+ shows "x / (numeral (Num.Bit0 n)) \<in> float"
proof -
have "(x / numeral n) / 2^1 \<in> float"
by (intro x div_power_2_float)
@@ -135,32 +161,38 @@
qed
lemma div_neg_numeral_Bit0_float[simp]:
- assumes x: "x / numeral n \<in> float" shows "x / (- numeral (Num.Bit0 n)) \<in> float"
+ assumes x: "x / numeral n \<in> float"
+ shows "x / (- numeral (Num.Bit0 n)) \<in> float"
proof -
- have "- (x / numeral (Num.Bit0 n)) \<in> float" using x by simp
+ have "- (x / numeral (Num.Bit0 n)) \<in> float"
+ using x by simp
also have "- (x / numeral (Num.Bit0 n)) = x / - numeral (Num.Bit0 n)"
by simp
finally show ?thesis .
qed
-lemma power_float[simp]: assumes "a \<in> float" shows "a ^ b \<in> float"
+lemma power_float[simp]:
+ assumes "a \<in> float"
+ shows "a ^ b \<in> float"
proof -
- from assms obtain m e::int where "a = m * 2 powr e"
+ from assms obtain m e :: int where "a = m * 2 powr e"
by (auto simp: float_def)
- thus ?thesis
+ then show ?thesis
by (auto intro!: floatI[where m="m^b" and e = "e*b"]
simp: power_mult_distrib powr_realpow[symmetric] powr_powr)
qed
-lift_definition Float :: "int \<Rightarrow> int \<Rightarrow> float" is "\<lambda>(m::int) (e::int). m * 2 powr e" by simp
+lift_definition Float :: "int \<Rightarrow> int \<Rightarrow> float" is "\<lambda>(m::int) (e::int). m * 2 powr e"
+ by simp
declare Float.rep_eq[simp]
lemma compute_real_of_float[code]:
"real_of_float (Float m e) = (if e \<ge> 0 then m * 2 ^ nat e else m / 2 ^ (nat (-e)))"
-by (simp add: real_of_float_def[symmetric] powr_int)
+ by (simp add: real_of_float_def[symmetric] powr_int)
code_datatype Float
+
subsection \<open>Arithmetic operations on floating point numbers\<close>
instantiation float :: "{ring_1, linorder, linordered_ring, linordered_idom, numeral, equal}"
@@ -192,16 +224,20 @@
declare less_float.rep_eq[simp]
instance
- proof qed (transfer, fastforce simp add: field_simps intro: mult_left_mono mult_right_mono)+
+ by (standard; transfer; fastforce simp add: field_simps intro: mult_left_mono mult_right_mono)+
+
end
lemma Float_0_eq_0[simp]: "Float 0 e = 0"
by transfer simp
-lemma real_of_float_power[simp]: fixes f::float shows "real (f^n) = real f^n"
+lemma real_of_float_power[simp]:
+ fixes f :: float
+ shows "real (f^n) = real f^n"
by (induct n) simp_all
-lemma fixes x y::float
+lemma
+ fixes x y :: float
shows real_of_float_min: "real (min x y) = min (real x) (real y)"
and real_of_float_max: "real (max x y) = max (real x) (real y)"
by (simp_all add: min_def max_def)
@@ -219,9 +255,9 @@
apply transfer
apply simp
done
- assume "a < b"
- then show "\<exists>c. a < c \<and> c < b"
- apply (intro exI[of _ "(a + b) * Float 1 (- 1)"])
+ show "\<exists>c. a < c \<and> c < b" if "a < b"
+ apply (rule exI[of _ "(a + b) * Float 1 (- 1)"])
+ using that
apply transfer
apply (simp add: powr_minus)
done
@@ -230,11 +266,11 @@
instantiation float :: lattice_ab_group_add
begin
-definition inf_float::"float\<Rightarrow>float\<Rightarrow>float"
-where "inf_float a b = min a b"
+definition inf_float :: "float \<Rightarrow> float \<Rightarrow> float"
+ where "inf_float a b = min a b"
-definition sup_float::"float\<Rightarrow>float\<Rightarrow>float"
-where "sup_float a b = max a b"
+definition sup_float :: "float \<Rightarrow> float \<Rightarrow> float"
+ where "sup_float a b = max a b"
instance
by (standard; transfer; simp add: inf_float_def sup_float_def real_of_float_min real_of_float_max)
@@ -250,20 +286,21 @@
lemma transfer_numeral [transfer_rule]:
"rel_fun (op =) pcr_float (numeral :: _ \<Rightarrow> real) (numeral :: _ \<Rightarrow> float)"
- unfolding rel_fun_def float.pcr_cr_eq cr_float_def by simp
+ by (simp add: rel_fun_def float.pcr_cr_eq cr_float_def)
lemma float_neg_numeral[simp]: "real (- numeral x :: float) = - numeral x"
by simp
lemma transfer_neg_numeral [transfer_rule]:
"rel_fun (op =) pcr_float (- numeral :: _ \<Rightarrow> real) (- numeral :: _ \<Rightarrow> float)"
- unfolding rel_fun_def float.pcr_cr_eq cr_float_def by simp
+ by (simp add: rel_fun_def float.pcr_cr_eq cr_float_def)
lemma
shows float_of_numeral[simp]: "numeral k = float_of (numeral k)"
and float_of_neg_numeral[simp]: "- numeral k = float_of (- numeral k)"
unfolding real_of_float_eq by simp_all
+
subsection \<open>Quickcheck\<close>
instantiation float :: exhaustive
@@ -311,39 +348,51 @@
assumes H: "\<And>n. (\<And>i. \<bar>i\<bar> < \<bar>n\<bar> \<Longrightarrow> P i) \<Longrightarrow> P n"
shows "P j"
proof (induct "nat \<bar>j\<bar>" arbitrary: j rule: less_induct)
- case less show ?case by (rule H[OF less]) simp
+ case less
+ show ?case by (rule H[OF less]) simp
qed
lemma int_cancel_factors:
- fixes n :: int assumes "1 < r" shows "n = 0 \<or> (\<exists>k i. n = k * r ^ i \<and> \<not> r dvd k)"
+ fixes n :: int
+ assumes "1 < r"
+ shows "n = 0 \<or> (\<exists>k i. n = k * r ^ i \<and> \<not> r dvd k)"
proof (induct n rule: int_induct_abs)
case (less n)
- { fix m assume n: "n \<noteq> 0" "n = m * r"
- then have "\<bar>m \<bar> < \<bar>n\<bar>"
+ have "\<exists>k i. n = k * r ^ Suc i \<and> \<not> r dvd k" if "n \<noteq> 0" "n = m * r" for m
+ proof -
+ from that have "\<bar>m \<bar> < \<bar>n\<bar>"
using \<open>1 < r\<close> by (simp add: abs_mult)
- from less[OF this] n have "\<exists>k i. n = k * r ^ Suc i \<and> \<not> r dvd k" by auto }
+ from less[OF this] that show ?thesis by auto
+ qed
then show ?case
by (metis dvd_def monoid_mult_class.mult.right_neutral mult.commute power_0)
qed
lemma mult_powr_eq_mult_powr_iff_asym:
fixes m1 m2 e1 e2 :: int
- assumes m1: "\<not> 2 dvd m1" and "e1 \<le> e2"
+ assumes m1: "\<not> 2 dvd m1"
+ and "e1 \<le> e2"
shows "m1 * 2 powr e1 = m2 * 2 powr e2 \<longleftrightarrow> m1 = m2 \<and> e1 = e2"
+ (is "?lhs \<longleftrightarrow> ?rhs")
proof
- have "m1 \<noteq> 0" using m1 unfolding dvd_def by auto
- assume eq: "m1 * 2 powr e1 = m2 * 2 powr e2"
- with \<open>e1 \<le> e2\<close> have "m1 = m2 * 2 powr nat (e2 - e1)"
- by (simp add: powr_divide2[symmetric] field_simps)
- also have "\<dots> = m2 * 2^nat (e2 - e1)"
- by (simp add: powr_realpow)
- finally have m1_eq: "m1 = m2 * 2^nat (e2 - e1)"
- unfolding real_of_int_inject .
- with m1 have "m1 = m2"
- by (cases "nat (e2 - e1)") (auto simp add: dvd_def)
- then show "m1 = m2 \<and> e1 = e2"
- using eq \<open>m1 \<noteq> 0\<close> by (simp add: powr_inj)
-qed simp
+ show ?rhs if eq: ?lhs
+ proof -
+ have "m1 \<noteq> 0"
+ using m1 unfolding dvd_def by auto
+ from \<open>e1 \<le> e2\<close> eq have "m1 = m2 * 2 powr nat (e2 - e1)"
+ by (simp add: powr_divide2[symmetric] field_simps)
+ also have "\<dots> = m2 * 2^nat (e2 - e1)"
+ by (simp add: powr_realpow)
+ finally have m1_eq: "m1 = m2 * 2^nat (e2 - e1)"
+ unfolding real_of_int_inject .
+ with m1 have "m1 = m2"
+ by (cases "nat (e2 - e1)") (auto simp add: dvd_def)
+ then show ?thesis
+ using eq \<open>m1 \<noteq> 0\<close> by (simp add: powr_inj)
+ qed
+ show ?lhs if ?rhs
+ using that by simp
+qed
lemma mult_powr_eq_mult_powr_iff:
fixes m1 m2 e1 e2 :: int
@@ -356,16 +405,18 @@
assumes x: "x \<in> float"
obtains (zero) "x = 0"
| (powr) m e :: int where "x = m * 2 powr e" "\<not> 2 dvd m" "x \<noteq> 0"
-proof atomize_elim
- { assume "x \<noteq> 0"
- from x obtain m e :: int where x: "x = m * 2 powr e" by (auto simp: float_def)
+proof -
+ {
+ assume "x \<noteq> 0"
+ from x obtain m e :: int where x: "x = m * 2 powr e"
+ by (auto simp: float_def)
with \<open>x \<noteq> 0\<close> int_cancel_factors[of 2 m] obtain k i where "m = k * 2 ^ i" "\<not> 2 dvd k"
by auto
with \<open>\<not> 2 dvd k\<close> x have "\<exists>(m::int) (e::int). x = m * 2 powr e \<and> \<not> (2::int) dvd m"
by (rule_tac exI[of _ "k"], rule_tac exI[of _ "e + int i"])
- (simp add: powr_add powr_realpow) }
- then show "x = 0 \<or> (\<exists>(m::int) (e::int). x = m * 2 powr e \<and> \<not> (2::int) dvd m \<and> x \<noteq> 0)"
- by blast
+ (simp add: powr_add powr_realpow)
+ }
+ with that show thesis by blast
qed
lemma float_normed_cases:
@@ -373,7 +424,8 @@
obtains (zero) "f = 0"
| (powr) m e :: int where "real f = m * 2 powr e" "\<not> 2 dvd m" "f \<noteq> 0"
proof (atomize_elim, induct f)
- case (float_of y) then show ?case
+ case (float_of y)
+ then show ?case
by (cases rule: floatE_normed) (auto simp: zero_float_def)
qed
@@ -389,7 +441,8 @@
shows exponent_0[simp]: "exponent (float_of 0) = 0" (is ?E)
and mantissa_0[simp]: "mantissa (float_of 0) = 0" (is ?M)
proof -
- have "\<And>p::int \<times> int. fst p = 0 \<and> snd p = 0 \<longleftrightarrow> p = (0, 0)" by auto
+ have "\<And>p::int \<times> int. fst p = 0 \<and> snd p = 0 \<longleftrightarrow> p = (0, 0)"
+ by auto
then show ?E ?M
by (auto simp add: mantissa_def exponent_def zero_float_def)
qed
@@ -398,17 +451,20 @@
shows mantissa_exponent: "real f = mantissa f * 2 powr exponent f" (is ?E)
and mantissa_not_dvd: "f \<noteq> (float_of 0) \<Longrightarrow> \<not> 2 dvd mantissa f" (is "_ \<Longrightarrow> ?D")
proof cases
- assume [simp]: "f \<noteq> (float_of 0)"
+ assume [simp]: "f \<noteq> float_of 0"
have "f = mantissa f * 2 powr exponent f \<and> \<not> 2 dvd mantissa f"
proof (cases f rule: float_normed_cases)
+ case zero
+ then show ?thesis by (simp add: zero_float_def)
+ next
case (powr m e)
- then have "\<exists>p::int \<times> int. (f = 0 \<and> fst p = 0 \<and> snd p = 0)
- \<or> (f \<noteq> 0 \<and> real f = real (fst p) * 2 powr real (snd p) \<and> \<not> 2 dvd fst p)"
+ then have "\<exists>p::int \<times> int. (f = 0 \<and> fst p = 0 \<and> snd p = 0) \<or>
+ (f \<noteq> 0 \<and> real f = real (fst p) * 2 powr real (snd p) \<and> \<not> 2 dvd fst p)"
by auto
then show ?thesis
unfolding exponent_def mantissa_def
by (rule someI2_ex) (simp add: zero_float_def)
- qed (simp add: zero_float_def)
+ qed
then show ?E ?D by auto
qed simp
@@ -422,31 +478,33 @@
shows mantissa_float: "mantissa f = m" (is "?M")
and exponent_float: "m \<noteq> 0 \<Longrightarrow> exponent f = e" (is "_ \<Longrightarrow> ?E")
proof cases
- assume "m = 0" with dvd show "mantissa f = m" by auto
+ assume "m = 0"
+ with dvd show "mantissa f = m" by auto
next
assume "m \<noteq> 0"
then have f_not_0: "f \<noteq> float_of 0" by (simp add: f_def)
- from mantissa_exponent[of f]
- have "m * 2 powr e = mantissa f * 2 powr exponent f"
+ from mantissa_exponent[of f] have "m * 2 powr e = mantissa f * 2 powr exponent f"
by (auto simp add: f_def)
then show "?M" "?E"
using mantissa_not_dvd[OF f_not_0] dvd
by (auto simp: mult_powr_eq_mult_powr_iff)
qed
+
subsection \<open>Compute arithmetic operations\<close>
lemma Float_mantissa_exponent: "Float (mantissa f) (exponent f) = f"
unfolding real_of_float_eq mantissa_exponent[of f] by simp
-lemma Float_cases[case_names Float, cases type: float]:
+lemma Float_cases [cases type: float]:
fixes f :: float
obtains (Float) m e :: int where "f = Float m e"
using Float_mantissa_exponent[symmetric]
by (atomize_elim) auto
lemma denormalize_shift:
- assumes f_def: "f \<equiv> Float m e" and not_0: "f \<noteq> float_of 0"
+ assumes f_def: "f \<equiv> Float m e"
+ and not_0: "f \<noteq> float_of 0"
obtains i where "m = mantissa f * 2 ^ i" "e = exponent f - i"
proof
from mantissa_exponent[of f] f_def
@@ -481,87 +539,75 @@
unfolding real_of_int_inject by auto
qed
-lemma compute_float_zero[code_unfold, code]: "0 = Float 0 0"
- by transfer simp
-hide_fact (open) compute_float_zero
+context
+begin
-lemma compute_float_one[code_unfold, code]: "1 = Float 1 0"
+qualified lemma compute_float_zero[code_unfold, code]: "0 = Float 0 0"
by transfer simp
-hide_fact (open) compute_float_one
+
+qualified lemma compute_float_one[code_unfold, code]: "1 = Float 1 0"
+ by transfer simp
lift_definition normfloat :: "float \<Rightarrow> float" is "\<lambda>x. x" .
lemma normloat_id[simp]: "normfloat x = x" by transfer rule
-lemma compute_normfloat[code]: "normfloat (Float m e) =
+qualified lemma compute_normfloat[code]: "normfloat (Float m e) =
(if m mod 2 = 0 \<and> m \<noteq> 0 then normfloat (Float (m div 2) (e + 1))
else if m = 0 then 0 else Float m e)"
by transfer (auto simp add: powr_add zmod_eq_0_iff)
-hide_fact (open) compute_normfloat
-lemma compute_float_numeral[code_abbrev]: "Float (numeral k) 0 = numeral k"
+qualified lemma compute_float_numeral[code_abbrev]: "Float (numeral k) 0 = numeral k"
by transfer simp
-hide_fact (open) compute_float_numeral
-lemma compute_float_neg_numeral[code_abbrev]: "Float (- numeral k) 0 = - numeral k"
+qualified lemma compute_float_neg_numeral[code_abbrev]: "Float (- numeral k) 0 = - numeral k"
by transfer simp
-hide_fact (open) compute_float_neg_numeral
-lemma compute_float_uminus[code]: "- Float m1 e1 = Float (- m1) e1"
+qualified lemma compute_float_uminus[code]: "- Float m1 e1 = Float (- m1) e1"
by transfer simp
-hide_fact (open) compute_float_uminus
-lemma compute_float_times[code]: "Float m1 e1 * Float m2 e2 = Float (m1 * m2) (e1 + e2)"
+qualified lemma compute_float_times[code]: "Float m1 e1 * Float m2 e2 = Float (m1 * m2) (e1 + e2)"
by transfer (simp add: field_simps powr_add)
-hide_fact (open) compute_float_times
-lemma compute_float_plus[code]: "Float m1 e1 + Float m2 e2 =
+qualified lemma compute_float_plus[code]: "Float m1 e1 + Float m2 e2 =
(if m1 = 0 then Float m2 e2 else if m2 = 0 then Float m1 e1 else
if e1 \<le> e2 then Float (m1 + m2 * 2^nat (e2 - e1)) e1
else Float (m2 + m1 * 2^nat (e1 - e2)) e2)"
by transfer (simp add: field_simps powr_realpow[symmetric] powr_divide2[symmetric])
-hide_fact (open) compute_float_plus
-lemma compute_float_minus[code]: fixes f g::float shows "f - g = f + (-g)"
+qualified lemma compute_float_minus[code]: fixes f g::float shows "f - g = f + (-g)"
by simp
-hide_fact (open) compute_float_minus
-lemma compute_float_sgn[code]: "sgn (Float m1 e1) = (if 0 < m1 then 1 else if m1 < 0 then -1 else 0)"
+qualified lemma compute_float_sgn[code]: "sgn (Float m1 e1) = (if 0 < m1 then 1 else if m1 < 0 then -1 else 0)"
by transfer (simp add: sgn_times)
-hide_fact (open) compute_float_sgn
lift_definition is_float_pos :: "float \<Rightarrow> bool" is "op < 0 :: real \<Rightarrow> bool" .
-lemma compute_is_float_pos[code]: "is_float_pos (Float m e) \<longleftrightarrow> 0 < m"
+qualified lemma compute_is_float_pos[code]: "is_float_pos (Float m e) \<longleftrightarrow> 0 < m"
by transfer (auto simp add: zero_less_mult_iff not_le[symmetric, of _ 0])
-hide_fact (open) compute_is_float_pos
-lemma compute_float_less[code]: "a < b \<longleftrightarrow> is_float_pos (b - a)"
+qualified lemma compute_float_less[code]: "a < b \<longleftrightarrow> is_float_pos (b - a)"
by transfer (simp add: field_simps)
-hide_fact (open) compute_float_less
lift_definition is_float_nonneg :: "float \<Rightarrow> bool" is "op \<le> 0 :: real \<Rightarrow> bool" .
-lemma compute_is_float_nonneg[code]: "is_float_nonneg (Float m e) \<longleftrightarrow> 0 \<le> m"
+qualified lemma compute_is_float_nonneg[code]: "is_float_nonneg (Float m e) \<longleftrightarrow> 0 \<le> m"
by transfer (auto simp add: zero_le_mult_iff not_less[symmetric, of _ 0])
-hide_fact (open) compute_is_float_nonneg
-lemma compute_float_le[code]: "a \<le> b \<longleftrightarrow> is_float_nonneg (b - a)"
+qualified lemma compute_float_le[code]: "a \<le> b \<longleftrightarrow> is_float_nonneg (b - a)"
by transfer (simp add: field_simps)
-hide_fact (open) compute_float_le
lift_definition is_float_zero :: "float \<Rightarrow> bool" is "op = 0 :: real \<Rightarrow> bool" .
-lemma compute_is_float_zero[code]: "is_float_zero (Float m e) \<longleftrightarrow> 0 = m"
+qualified lemma compute_is_float_zero[code]: "is_float_zero (Float m e) \<longleftrightarrow> 0 = m"
by transfer (auto simp add: is_float_zero_def)
-hide_fact (open) compute_is_float_zero
-lemma compute_float_abs[code]: "abs (Float m e) = Float (abs m) e"
+qualified lemma compute_float_abs[code]: "abs (Float m e) = Float (abs m) e"
by transfer (simp add: abs_mult)
-hide_fact (open) compute_float_abs
-lemma compute_float_eq[code]: "equal_class.equal f g = is_float_zero (f - g)"
+qualified lemma compute_float_eq[code]: "equal_class.equal f g = is_float_zero (f - g)"
by transfer simp
-hide_fact (open) compute_float_eq
+
+end
subsection \<open>Lemmas for types @{typ real}, @{typ nat}, @{typ int}\<close>
@@ -590,11 +636,11 @@
subsection \<open>Rounding Real Numbers\<close>
-definition round_down :: "int \<Rightarrow> real \<Rightarrow> real" where
- "round_down prec x = floor (x * 2 powr prec) * 2 powr -prec"
+definition round_down :: "int \<Rightarrow> real \<Rightarrow> real"
+ where "round_down prec x = floor (x * 2 powr prec) * 2 powr -prec"
-definition round_up :: "int \<Rightarrow> real \<Rightarrow> real" where
- "round_up prec x = ceiling (x * 2 powr prec) * 2 powr -prec"
+definition round_up :: "int \<Rightarrow> real \<Rightarrow> real"
+ where "round_up prec x = ceiling (x * 2 powr prec) * 2 powr -prec"
lemma round_down_float[simp]: "round_down prec x \<in> float"
unfolding round_down_def
@@ -692,7 +738,7 @@
from neg have "2 powr real p \<le> 2 powr 0"
by (intro powr_mono) auto
also have "\<dots> \<le> \<lfloor>2 powr 0::real\<rfloor>" by simp
- also have "... \<le> \<lfloor>x * 2 powr (real p)\<rfloor>"
+ also have "\<dots> \<le> \<lfloor>x * 2 powr (real p)\<rfloor>"
unfolding real_of_int_le_iff
using x x_le by (intro floor_mono) (simp add: powr_minus_divide field_simps)
finally show ?thesis
@@ -707,9 +753,11 @@
subsection \<open>Rounding Floats\<close>
-definition div_twopow::"int \<Rightarrow> nat \<Rightarrow> int" where [simp]: "div_twopow x n = x div (2 ^ n)"
+definition div_twopow :: "int \<Rightarrow> nat \<Rightarrow> int"
+ where [simp]: "div_twopow x n = x div (2 ^ n)"
-definition mod_twopow::"int \<Rightarrow> nat \<Rightarrow> int" where [simp]: "mod_twopow x n = x mod (2 ^ n)"
+definition mod_twopow :: "int \<Rightarrow> nat \<Rightarrow> int"
+ where [simp]: "mod_twopow x n = x mod (2 ^ n)"
lemma compute_div_twopow[code]:
"div_twopow x n = (if x = 0 \<or> x = -1 \<or> n = 0 then x else div_twopow (x div 2) (n - 1))"
@@ -722,51 +770,54 @@
lift_definition float_up :: "int \<Rightarrow> float \<Rightarrow> float" is round_up by simp
declare float_up.rep_eq[simp]
-lemma round_up_correct:
- shows "round_up e f - f \<in> {0..2 powr -e}"
-unfolding atLeastAtMost_iff
+lemma round_up_correct: "round_up e f - f \<in> {0..2 powr -e}"
+ unfolding atLeastAtMost_iff
proof
- have "round_up e f - f \<le> round_up e f - round_down e f" using round_down by simp
- also have "\<dots> \<le> 2 powr -e" using round_up_diff_round_down by simp
+ have "round_up e f - f \<le> round_up e f - round_down e f"
+ using round_down by simp
+ also have "\<dots> \<le> 2 powr -e"
+ using round_up_diff_round_down by simp
finally show "round_up e f - f \<le> 2 powr - (real e)"
by simp
qed (simp add: algebra_simps round_up)
-lemma float_up_correct:
- shows "real (float_up e f) - real f \<in> {0..2 powr -e}"
+lemma float_up_correct: "real (float_up e f) - real f \<in> {0..2 powr -e}"
by transfer (rule round_up_correct)
lift_definition float_down :: "int \<Rightarrow> float \<Rightarrow> float" is round_down by simp
declare float_down.rep_eq[simp]
-lemma round_down_correct:
- shows "f - (round_down e f) \<in> {0..2 powr -e}"
-unfolding atLeastAtMost_iff
+lemma round_down_correct: "f - (round_down e f) \<in> {0..2 powr -e}"
+ unfolding atLeastAtMost_iff
proof
- have "f - round_down e f \<le> round_up e f - round_down e f" using round_up by simp
- also have "\<dots> \<le> 2 powr -e" using round_up_diff_round_down by simp
+ have "f - round_down e f \<le> round_up e f - round_down e f"
+ using round_up by simp
+ also have "\<dots> \<le> 2 powr -e"
+ using round_up_diff_round_down by simp
finally show "f - round_down e f \<le> 2 powr - (real e)"
by simp
qed (simp add: algebra_simps round_down)
-lemma float_down_correct:
- shows "real f - real (float_down e f) \<in> {0..2 powr -e}"
+lemma float_down_correct: "real f - real (float_down e f) \<in> {0..2 powr -e}"
by transfer (rule round_down_correct)
-lemma compute_float_down[code]:
+context
+begin
+
+qualified lemma compute_float_down[code]:
"float_down p (Float m e) =
(if p + e < 0 then Float (div_twopow m (nat (-(p + e)))) (-p) else Float m e)"
-proof cases
- assume "p + e < 0"
- hence "real ((2::int) ^ nat (-(p + e))) = 2 powr (-(p + e))"
+proof (cases "p + e < 0")
+ case True
+ then have "real ((2::int) ^ nat (-(p + e))) = 2 powr (-(p + e))"
using powr_realpow[of 2 "nat (-(p + e))"] by simp
- also have "... = 1 / 2 powr p / 2 powr e"
+ also have "\<dots> = 1 / 2 powr p / 2 powr e"
unfolding powr_minus_divide real_of_int_minus by (simp add: powr_add)
finally show ?thesis
using \<open>p + e < 0\<close>
by transfer (simp add: ac_simps round_down_def floor_divide_eq_div[symmetric])
next
- assume "\<not> p + e < 0"
+ case False
then have r: "real e + real p = real (nat (e + p))" by simp
have r: "\<lfloor>(m * 2 powr e) * 2 powr real p\<rfloor> = (m * 2 powr e) * 2 powr real p"
by (auto intro: exI[where x="m*2^nat (e+p)"]
@@ -774,7 +825,6 @@
with \<open>\<not> p + e < 0\<close> show ?thesis
by transfer (auto simp add: round_down_def field_simps powr_add powr_minus)
qed
-hide_fact (open) compute_float_down
lemma abs_round_down_le: "\<bar>f - (round_down e f)\<bar> \<le> 2 powr -e"
using round_down_correct[of f e] by simp
@@ -786,75 +836,100 @@
by (auto simp: round_down_def)
lemma ceil_divide_floor_conv:
-assumes "b \<noteq> 0"
-shows "\<lceil>real a / real b\<rceil> = (if b dvd a then a div b else \<lfloor>real a / real b\<rfloor> + 1)"
-proof cases
- assume "\<not> b dvd a"
- hence "a mod b \<noteq> 0" by auto
- hence ne: "real (a mod b) / real b \<noteq> 0" using \<open>b \<noteq> 0\<close> by auto
+ assumes "b \<noteq> 0"
+ shows "\<lceil>real a / real b\<rceil> = (if b dvd a then a div b else \<lfloor>real a / real b\<rfloor> + 1)"
+proof (cases "b dvd a")
+ case True
+ then show ?thesis
+ by (simp add: ceiling_def real_of_int_minus[symmetric] divide_minus_left[symmetric]
+ floor_divide_eq_div dvd_neg_div del: divide_minus_left real_of_int_minus)
+next
+ case False
+ then have "a mod b \<noteq> 0"
+ by auto
+ then have ne: "real (a mod b) / real b \<noteq> 0"
+ using \<open>b \<noteq> 0\<close> by auto
have "\<lceil>real a / real b\<rceil> = \<lfloor>real a / real b\<rfloor> + 1"
- apply (rule ceiling_eq) apply (auto simp: floor_divide_eq_div[symmetric])
+ apply (rule ceiling_eq)
+ apply (auto simp: floor_divide_eq_div[symmetric])
proof -
- have "real \<lfloor>real a / real b\<rfloor> \<le> real a / real b" by simp
+ have "real \<lfloor>real a / real b\<rfloor> \<le> real a / real b"
+ by simp
moreover have "real \<lfloor>real a / real b\<rfloor> \<noteq> real a / real b"
- apply (subst (2) real_of_int_div_aux) unfolding floor_divide_eq_div using ne \<open>b \<noteq> 0\<close> by auto
+ apply (subst (2) real_of_int_div_aux)
+ unfolding floor_divide_eq_div
+ using ne \<open>b \<noteq> 0\<close> apply auto
+ done
ultimately show "real \<lfloor>real a / real b\<rfloor> < real a / real b" by arith
qed
- thus ?thesis using \<open>\<not> b dvd a\<close> by simp
-qed (simp add: ceiling_def real_of_int_minus[symmetric] divide_minus_left[symmetric]
- floor_divide_eq_div dvd_neg_div del: divide_minus_left real_of_int_minus)
+ then show ?thesis
+ using \<open>\<not> b dvd a\<close> by simp
+qed
-lemma compute_float_up[code]:
- "float_up p x = - float_down p (-x)"
+qualified lemma compute_float_up[code]: "float_up p x = - float_down p (-x)"
by transfer (simp add: round_down_uminus_eq)
-hide_fact (open) compute_float_up
+
+end
subsection \<open>Compute bitlen of integers\<close>
-definition bitlen :: "int \<Rightarrow> int" where
- "bitlen a = (if a > 0 then \<lfloor>log 2 a\<rfloor> + 1 else 0)"
+definition bitlen :: "int \<Rightarrow> int"
+ where "bitlen a = (if a > 0 then \<lfloor>log 2 a\<rfloor> + 1 else 0)"
lemma bitlen_nonneg: "0 \<le> bitlen x"
proof -
- {
- assume "0 > x"
- have "-1 = log 2 (inverse 2)" by (subst log_inverse) simp_all
- also have "... < log 2 (-x)" using \<open>0 > x\<close> by auto
- finally have "-1 < log 2 (-x)" .
- } thus "0 \<le> bitlen x" unfolding bitlen_def by (auto intro!: add_nonneg_nonneg)
+ have "-1 < log 2 (-x)" if "0 > x"
+ proof -
+ have "-1 = log 2 (inverse 2)"
+ by (subst log_inverse) simp_all
+ also have "\<dots> < log 2 (-x)"
+ using \<open>0 > x\<close> by auto
+ finally show ?thesis .
+ qed
+ then show ?thesis
+ unfolding bitlen_def by (auto intro!: add_nonneg_nonneg)
qed
lemma bitlen_bounds:
assumes "x > 0"
shows "2 ^ nat (bitlen x - 1) \<le> x \<and> x < 2 ^ nat (bitlen x)"
proof
- have "(2::real) ^ nat \<lfloor>log 2 (real x)\<rfloor> = 2 powr real (floor (log 2 (real x)))"
- using powr_realpow[symmetric, of 2 "nat \<lfloor>log 2 (real x)\<rfloor>"] \<open>x > 0\<close>
- using real_nat_eq_real[of "floor (log 2 (real x))"]
- by simp
- also have "... \<le> 2 powr log 2 (real x)"
- by simp
- also have "... = real x"
- using \<open>0 < x\<close> by simp
- finally have "2 ^ nat \<lfloor>log 2 (real x)\<rfloor> \<le> real x" by simp
- thus "2 ^ nat (bitlen x - 1) \<le> x" using \<open>x > 0\<close>
- by (simp add: bitlen_def)
-next
- have "x \<le> 2 powr (log 2 x)" using \<open>x > 0\<close> by simp
- also have "... < 2 ^ nat (\<lfloor>log 2 (real x)\<rfloor> + 1)"
- apply (simp add: powr_realpow[symmetric])
- using \<open>x > 0\<close> by simp
- finally show "x < 2 ^ nat (bitlen x)" using \<open>x > 0\<close>
- by (simp add: bitlen_def ac_simps)
+ show "2 ^ nat (bitlen x - 1) \<le> x"
+ proof -
+ have "(2::real) ^ nat \<lfloor>log 2 (real x)\<rfloor> = 2 powr real (floor (log 2 (real x)))"
+ using powr_realpow[symmetric, of 2 "nat \<lfloor>log 2 (real x)\<rfloor>"] \<open>x > 0\<close>
+ using real_nat_eq_real[of "floor (log 2 (real x))"]
+ by simp
+ also have "\<dots> \<le> 2 powr log 2 (real x)"
+ by simp
+ also have "\<dots> = real x"
+ using \<open>0 < x\<close> by simp
+ finally have "2 ^ nat \<lfloor>log 2 (real x)\<rfloor> \<le> real x"
+ by simp
+ then show ?thesis
+ using \<open>0 < x\<close> by (simp add: bitlen_def)
+ qed
+ show "x < 2 ^ nat (bitlen x)"
+ proof -
+ have "x \<le> 2 powr (log 2 x)"
+ using \<open>x > 0\<close> by simp
+ also have "\<dots> < 2 ^ nat (\<lfloor>log 2 (real x)\<rfloor> + 1)"
+ apply (simp add: powr_realpow[symmetric])
+ using \<open>x > 0\<close> apply simp
+ done
+ finally show ?thesis
+ using \<open>x > 0\<close> by (simp add: bitlen_def ac_simps)
+ qed
qed
lemma bitlen_pow2[simp]:
assumes "b > 0"
shows "bitlen (b * 2 ^ c) = bitlen b + c"
proof -
- from assms have "b * 2 ^ c > 0" by auto
- thus ?thesis
+ from assms have "b * 2 ^ c > 0"
+ by auto
+ then show ?thesis
using floor_add[of "log 2 b" c] assms
by (auto simp add: log_mult log_nat_power bitlen_def)
qed
@@ -868,9 +943,9 @@
then show ?thesis by (simp add: f_def bitlen_def Float_def)
next
case False
- hence "f \<noteq> float_of 0"
+ then have "f \<noteq> float_of 0"
unfolding real_of_float_eq by (simp add: f_def)
- hence "mantissa f \<noteq> 0"
+ then have "mantissa f \<noteq> 0"
by (simp add: mantissa_noteq_0)
moreover
obtain i where "m = mantissa f * 2 ^ i" "e = exponent f - int i"
@@ -878,22 +953,27 @@
ultimately show ?thesis by (simp add: abs_mult)
qed
-lemma compute_bitlen[code]:
- shows "bitlen x = (if x > 0 then bitlen (x div 2) + 1 else 0)"
+context
+begin
+
+qualified lemma compute_bitlen[code]: "bitlen x = (if x > 0 then bitlen (x div 2) + 1 else 0)"
proof -
{ assume "2 \<le> x"
then have "\<lfloor>log 2 (x div 2)\<rfloor> + 1 = \<lfloor>log 2 (x - x mod 2)\<rfloor>"
by (simp add: log_mult zmod_zdiv_equality')
also have "\<dots> = \<lfloor>log 2 (real x)\<rfloor>"
- proof cases
- assume "x mod 2 = 0" then show ?thesis by simp
+ proof (cases "x mod 2 = 0")
+ case True
+ then show ?thesis by simp
next
+ case False
def n \<equiv> "\<lfloor>log 2 (real x)\<rfloor>"
then have "0 \<le> n"
using \<open>2 \<le> x\<close> by simp
- assume "x mod 2 \<noteq> 0"
- with \<open>2 \<le> x\<close> have "x mod 2 = 1" "\<not> 2 dvd x" by (auto simp add: dvd_eq_mod_eq_0)
- with \<open>2 \<le> x\<close> have "x \<noteq> 2^nat n" by (cases "nat n") auto
+ from \<open>2 \<le> x\<close> False have "x mod 2 = 1" "\<not> 2 dvd x"
+ by (auto simp add: dvd_eq_mod_eq_0)
+ with \<open>2 \<le> x\<close> have "x \<noteq> 2 ^ nat n"
+ by (cases "nat n") auto
moreover
{ have "real (2^nat n :: int) = 2 powr (nat n)"
by (simp add: powr_realpow)
@@ -922,62 +1002,87 @@
unfolding bitlen_def
by (auto simp: pos_imp_zdiv_pos_iff not_le)
qed
-hide_fact (open) compute_bitlen
+
+end
lemma float_gt1_scale: assumes "1 \<le> Float m e"
shows "0 \<le> e + (bitlen m - 1)"
proof -
have "0 < Float m e" using assms by auto
- hence "0 < m" using powr_gt_zero[of 2 e]
+ then have "0 < m" using powr_gt_zero[of 2 e]
apply (auto simp: zero_less_mult_iff)
- using not_le powr_ge_pzero by blast
- hence "m \<noteq> 0" by auto
+ using not_le powr_ge_pzero apply blast
+ done
+ then have "m \<noteq> 0" by auto
show ?thesis
proof (cases "0 \<le> e")
- case True thus ?thesis using \<open>0 < m\<close> by (simp add: bitlen_def)
+ case True
+ then show ?thesis
+ using \<open>0 < m\<close> by (simp add: bitlen_def)
next
+ case False
have "(1::int) < 2" by simp
- case False let ?S = "2^(nat (-e))"
- have "inverse (2 ^ nat (- e)) = 2 powr e" using assms False powr_realpow[of 2 "nat (-e)"]
+ let ?S = "2^(nat (-e))"
+ have "inverse (2 ^ nat (- e)) = 2 powr e"
+ using assms False powr_realpow[of 2 "nat (-e)"]
by (auto simp: powr_minus field_simps)
- hence "1 \<le> real m * inverse ?S" using assms False powr_realpow[of 2 "nat (-e)"]
+ then have "1 \<le> real m * inverse ?S"
+ using assms False powr_realpow[of 2 "nat (-e)"]
by (auto simp: powr_minus)
- hence "1 * ?S \<le> real m * inverse ?S * ?S" by (rule mult_right_mono, auto)
- hence "?S \<le> real m" unfolding mult.assoc by auto
- hence "?S \<le> m" unfolding real_of_int_le_iff[symmetric] by auto
+ then have "1 * ?S \<le> real m * inverse ?S * ?S"
+ by (rule mult_right_mono) auto
+ then have "?S \<le> real m"
+ unfolding mult.assoc by auto
+ then have "?S \<le> m"
+ unfolding real_of_int_le_iff[symmetric] by auto
from this bitlen_bounds[OF \<open>0 < m\<close>, THEN conjunct2]
- have "nat (-e) < (nat (bitlen m))" unfolding power_strict_increasing_iff[OF \<open>1 < 2\<close>, symmetric]
+ have "nat (-e) < (nat (bitlen m))"
+ unfolding power_strict_increasing_iff[OF \<open>1 < 2\<close>, symmetric]
by (rule order_le_less_trans)
- hence "-e < bitlen m" using False by auto
- thus ?thesis by auto
+ then have "-e < bitlen m"
+ using False by auto
+ then show ?thesis
+ by auto
qed
qed
lemma bitlen_div:
assumes "0 < m"
- shows "1 \<le> real m / 2^nat (bitlen m - 1)" and "real m / 2^nat (bitlen m - 1) < 2"
+ shows "1 \<le> real m / 2^nat (bitlen m - 1)"
+ and "real m / 2^nat (bitlen m - 1) < 2"
proof -
let ?B = "2^nat(bitlen m - 1)"
have "?B \<le> m" using bitlen_bounds[OF \<open>0 <m\<close>] ..
- hence "1 * ?B \<le> real m" unfolding real_of_int_le_iff[symmetric] by auto
- thus "1 \<le> real m / ?B" by auto
+ then have "1 * ?B \<le> real m"
+ unfolding real_of_int_le_iff[symmetric] by auto
+ then show "1 \<le> real m / ?B"
+ by auto
- have "m \<noteq> 0" using assms by auto
- have "0 \<le> bitlen m - 1" using \<open>0 < m\<close> by (auto simp: bitlen_def)
+ have "m \<noteq> 0"
+ using assms by auto
+ have "0 \<le> bitlen m - 1"
+ using \<open>0 < m\<close> by (auto simp: bitlen_def)
- have "m < 2^nat(bitlen m)" using bitlen_bounds[OF \<open>0 <m\<close>] ..
- also have "\<dots> = 2^nat(bitlen m - 1 + 1)" using \<open>0 < m\<close> by (auto simp: bitlen_def)
- also have "\<dots> = ?B * 2" unfolding nat_add_distrib[OF \<open>0 \<le> bitlen m - 1\<close> zero_le_one] by auto
- finally have "real m < 2 * ?B" unfolding real_of_int_less_iff[symmetric] by auto
- hence "real m / ?B < 2 * ?B / ?B" by (rule divide_strict_right_mono, auto)
- thus "real m / ?B < 2" by auto
+ have "m < 2^nat(bitlen m)"
+ using bitlen_bounds[OF \<open>0 <m\<close>] ..
+ also have "\<dots> = 2^nat(bitlen m - 1 + 1)"
+ using \<open>0 < m\<close> by (auto simp: bitlen_def)
+ also have "\<dots> = ?B * 2"
+ unfolding nat_add_distrib[OF \<open>0 \<le> bitlen m - 1\<close> zero_le_one] by auto
+ finally have "real m < 2 * ?B"
+ unfolding real_of_int_less_iff[symmetric] by auto
+ then have "real m / ?B < 2 * ?B / ?B"
+ by (rule divide_strict_right_mono) auto
+ then show "real m / ?B < 2"
+ by auto
qed
+
subsection \<open>Truncating Real Numbers\<close>
-definition truncate_down::"nat \<Rightarrow> real \<Rightarrow> real" where
- "truncate_down prec x = round_down (prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor> - 1) x"
+definition truncate_down::"nat \<Rightarrow> real \<Rightarrow> real"
+ where "truncate_down prec x = round_down (prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor> - 1) x"
lemma truncate_down: "truncate_down prec x \<le> x"
using round_down by (simp add: truncate_down_def)
@@ -991,8 +1096,8 @@
lemma truncate_down_float[simp]: "truncate_down p x \<in> float"
by (auto simp: truncate_down_def)
-definition truncate_up::"nat \<Rightarrow> real \<Rightarrow> real" where
- "truncate_up prec x = round_up (prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor> - 1) x"
+definition truncate_up::"nat \<Rightarrow> real \<Rightarrow> real"
+ where "truncate_up prec x = round_up (prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor> - 1) x"
lemma truncate_up: "x \<le> truncate_up prec x"
using round_up by (simp add: truncate_up_def)
@@ -1035,22 +1140,28 @@
by (auto simp: truncate_up_def round_up_def intro!: mult_nonpos_nonneg)
lemma truncate_up_le1:
- assumes "x \<le> 1" "1 \<le> p" shows "truncate_up p x \<le> 1"
+ assumes "x \<le> 1" "1 \<le> p"
+ shows "truncate_up p x \<le> 1"
proof -
- {
- assume "x \<le> 0"
- with truncate_up_nonpos[OF this, of p] have ?thesis by simp
- } moreover {
- assume "x > 0"
- hence le: "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> \<le> 0"
+ consider "x \<le> 0" | "x > 0"
+ by arith
+ then show ?thesis
+ proof cases
+ case 1
+ with truncate_up_nonpos[OF this, of p] show ?thesis
+ by simp
+ next
+ case 2
+ then have le: "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> \<le> 0"
using assms by (auto simp: log_less_iff)
from assms have "1 \<le> int p" by simp
from add_mono[OF this le]
- have ?thesis using assms
- by (simp add: truncate_up_def round_up_le1 add_mono)
- } ultimately show ?thesis by arith
+ show ?thesis
+ using assms by (simp add: truncate_up_def round_up_le1 add_mono)
+ qed
qed
+
subsection \<open>Truncating Floats\<close>
lift_definition float_round_up :: "nat \<Rightarrow> float \<Rightarrow> float" is truncate_up
@@ -1078,25 +1189,30 @@
and minus_float_round_down_eq: "- float_round_down prec x = float_round_up prec (- x)"
by (transfer, simp add: truncate_down_uminus_eq truncate_up_uminus_eq)+
-lemma compute_float_round_down[code]:
+context
+begin
+
+qualified lemma compute_float_round_down[code]:
"float_round_down prec (Float m e) = (let d = bitlen (abs m) - int prec in
if 0 < d then Float (div_twopow m (nat d)) (e + d)
else Float m e)"
using Float.compute_float_down[of "prec - bitlen \<bar>m\<bar> - e" m e, symmetric]
by transfer (simp add: field_simps abs_mult log_mult bitlen_def truncate_down_def
cong del: if_weak_cong)
-hide_fact (open) compute_float_round_down
-lemma compute_float_round_up[code]:
+qualified lemma compute_float_round_up[code]:
"float_round_up prec x = - float_round_down prec (-x)"
by transfer (simp add: truncate_down_uminus_eq)
-hide_fact (open) compute_float_round_up
+
+end
subsection \<open>Approximation of positive rationals\<close>
-lemma div_mult_twopow_eq: fixes a b::nat shows "a div ((2::nat) ^ n) div b = a div (b * 2 ^ n)"
- by (cases "b=0") (simp_all add: div_mult2_eq[symmetric] ac_simps)
+lemma div_mult_twopow_eq:
+ fixes a b :: nat
+ shows "a div ((2::nat) ^ n) div b = a div (b * 2 ^ n)"
+ by (cases "b = 0") (simp_all add: div_mult2_eq[symmetric] ac_simps)
lemma real_div_nat_eq_floor_of_divide:
fixes a b :: nat
@@ -1106,23 +1222,29 @@
definition "rat_precision prec x y = int prec - (bitlen x - bitlen y)"
lift_definition lapprox_posrat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float"
- is "\<lambda>prec (x::nat) (y::nat). round_down (rat_precision prec x y) (x / y)" by simp
+ is "\<lambda>prec (x::nat) (y::nat). round_down (rat_precision prec x y) (x / y)"
+ by simp
-lemma compute_lapprox_posrat[code]:
+context
+begin
+
+qualified lemma compute_lapprox_posrat[code]:
fixes prec x y
shows "lapprox_posrat prec x y =
(let
- l = rat_precision prec x y;
- d = if 0 \<le> l then x * 2^nat l div y else x div 2^nat (- l) div y
+ l = rat_precision prec x y;
+ d = if 0 \<le> l then x * 2^nat l div y else x div 2^nat (- l) div y
in normfloat (Float d (- l)))"
unfolding div_mult_twopow_eq
by transfer
(simp add: round_down_def powr_int real_div_nat_eq_floor_of_divide field_simps Let_def
del: two_powr_minus_int_float)
-hide_fact (open) compute_lapprox_posrat
+
+end
lift_definition rapprox_posrat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float"
- is "\<lambda>prec (x::nat) (y::nat). round_up (rat_precision prec x y) (x / y)" by simp
+ is "\<lambda>prec (x::nat) (y::nat). round_up (rat_precision prec x y) (x / y)" by
+ simp
context
notes divmod_int_mod_div[simp]
@@ -1137,14 +1259,16 @@
(d, m) = divmod_int (fst X) (snd X)
in normfloat (Float (d + (if m = 0 \<or> y = 0 then 0 else 1)) (- l)))"
proof (cases "y = 0")
- assume "y = 0" thus ?thesis by transfer simp
+ assume "y = 0"
+ then show ?thesis by transfer simp
next
assume "y \<noteq> 0"
show ?thesis
proof (cases "0 \<le> l")
- assume "0 \<le> l"
+ case True
def x' \<equiv> "x * 2 ^ nat l"
- have "int x * 2 ^ nat l = x'" by (simp add: x'_def int_mult int_power)
+ have "int x * 2 ^ nat l = x'"
+ by (simp add: x'_def int_mult int_power)
moreover have "real x * 2 powr real l = real x'"
by (simp add: powr_realpow[symmetric] \<open>0 \<le> l\<close> x'_def)
ultimately show ?thesis
@@ -1152,7 +1276,7 @@
l_def[symmetric, THEN meta_eq_to_obj_eq]
by transfer (auto simp add: floor_divide_eq_div [symmetric] round_up_def)
next
- assume "\<not> 0 \<le> l"
+ case False
def y' \<equiv> "y * 2 ^ nat (- l)"
from \<open>y \<noteq> 0\<close> have "y' \<noteq> 0" by (simp add: y'_def)
have "int y * 2 ^ nat (- l) = y'" by (simp add: y'_def int_mult int_power)
@@ -1170,38 +1294,48 @@
end
lemma rat_precision_pos:
- assumes "0 \<le> x" and "0 < y" and "2 * x < y" and "0 < n"
+ assumes "0 \<le> x"
+ and "0 < y"
+ and "2 * x < y"
+ and "0 < n"
shows "rat_precision n (int x) (int y) > 0"
proof -
- { assume "0 < x" hence "log 2 x + 1 = log 2 (2 * x)" by (simp add: log_mult) }
- hence "bitlen (int x) < bitlen (int y)" using assms
+ have "0 < x \<Longrightarrow> log 2 x + 1 = log 2 (2 * x)"
+ by (simp add: log_mult)
+ then have "bitlen (int x) < bitlen (int y)"
+ using assms
by (simp add: bitlen_def del: floor_add_one)
(auto intro!: floor_mono simp add: floor_add_one[symmetric] simp del: floor_add floor_add_one)
- thus ?thesis
- using assms by (auto intro!: pos_add_strict simp add: field_simps rat_precision_def)
+ then show ?thesis
+ using assms
+ by (auto intro!: pos_add_strict simp add: field_simps rat_precision_def)
qed
lemma rapprox_posrat_less1:
- shows "0 \<le> x \<Longrightarrow> 0 < y \<Longrightarrow> 2 * x < y \<Longrightarrow> 0 < n \<Longrightarrow> real (rapprox_posrat n x y) < 1"
+ "0 \<le> x \<Longrightarrow> 0 < y \<Longrightarrow> 2 * x < y \<Longrightarrow> 0 < n \<Longrightarrow> real (rapprox_posrat n x y) < 1"
by transfer (simp add: rat_precision_pos round_up_less1)
lift_definition lapprox_rat :: "nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> float" is
- "\<lambda>prec (x::int) (y::int). round_down (rat_precision prec \<bar>x\<bar> \<bar>y\<bar>) (x / y)" by simp
+ "\<lambda>prec (x::int) (y::int). round_down (rat_precision prec \<bar>x\<bar> \<bar>y\<bar>) (x / y)"
+ by simp
-lemma compute_lapprox_rat[code]:
+context
+begin
+
+qualified lemma compute_lapprox_rat[code]:
"lapprox_rat prec x y =
- (if y = 0 then 0
+ (if y = 0 then 0
else if 0 \<le> x then
- (if 0 < y then lapprox_posrat prec (nat x) (nat y)
+ (if 0 < y then lapprox_posrat prec (nat x) (nat y)
else - (rapprox_posrat prec (nat x) (nat (-y))))
else (if 0 < y
then - (rapprox_posrat prec (nat (-x)) (nat y))
else lapprox_posrat prec (nat (-x)) (nat (-y))))"
by transfer (auto simp: round_up_def round_down_def ceiling_def ac_simps)
-hide_fact (open) compute_lapprox_rat
lift_definition rapprox_rat :: "nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> float" is
- "\<lambda>prec (x::int) (y::int). round_up (rat_precision prec \<bar>x\<bar> \<bar>y\<bar>) (x / y)" by simp
+ "\<lambda>prec (x::int) (y::int). round_up (rat_precision prec \<bar>x\<bar> \<bar>y\<bar>) (x / y)"
+ by simp
lemma "rapprox_rat = rapprox_posrat"
by transfer auto
@@ -1209,10 +1343,12 @@
lemma "lapprox_rat = lapprox_posrat"
by transfer auto
-lemma compute_rapprox_rat[code]:
+qualified lemma compute_rapprox_rat[code]:
"rapprox_rat prec x y = - lapprox_rat prec (-x) y"
by transfer (simp add: round_down_uminus_eq)
-hide_fact (open) compute_rapprox_rat
+
+end
+
subsection \<open>Division\<close>
@@ -1223,47 +1359,58 @@
lift_definition float_divl :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" is real_divl
by (simp add: real_divl_def)
-lemma compute_float_divl[code]:
+context
+begin
+
+qualified lemma compute_float_divl[code]:
"float_divl prec (Float m1 s1) (Float m2 s2) = lapprox_rat prec m1 m2 * Float 1 (s1 - s2)"
-proof cases
+proof (cases "m1 \<noteq> 0 \<and> m2 \<noteq> 0")
+ case True
let ?f1 = "real m1 * 2 powr real s1" and ?f2 = "real m2 * 2 powr real s2"
let ?m = "real m1 / real m2" and ?s = "2 powr real (s1 - s2)"
- assume not_0: "m1 \<noteq> 0 \<and> m2 \<noteq> 0"
- then have eq2: "(int prec + \<lfloor>log 2 \<bar>?f2\<bar>\<rfloor> - \<lfloor>log 2 \<bar>?f1\<bar>\<rfloor>) = rat_precision prec \<bar>m1\<bar> \<bar>m2\<bar> + (s2 - s1)"
+ from True have eq2: "(int prec + \<lfloor>log 2 \<bar>?f2\<bar>\<rfloor> - \<lfloor>log 2 \<bar>?f1\<bar>\<rfloor>) =
+ rat_precision prec \<bar>m1\<bar> \<bar>m2\<bar> + (s2 - s1)"
by (simp add: abs_mult log_mult rat_precision_def bitlen_def)
have eq1: "real m1 * 2 powr real s1 / (real m2 * 2 powr real s2) = ?m * ?s"
by (simp add: field_simps powr_divide2[symmetric])
-
- show ?thesis
- using not_0
+ from True show ?thesis
by (transfer fixing: m1 s1 m2 s2 prec) (unfold eq1 eq2 round_down_shift real_divl_def,
simp add: field_simps)
-qed (transfer, auto simp: real_divl_def)
-hide_fact (open) compute_float_divl
+next
+ case False
+ then show ?thesis by transfer (auto simp: real_divl_def)
+qed
lift_definition float_divr :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" is real_divr
by (simp add: real_divr_def)
-lemma compute_float_divr[code]:
+qualified lemma compute_float_divr[code]:
"float_divr prec x y = - float_divl prec (-x) y"
by transfer (simp add: real_divr_def real_divl_def round_down_uminus_eq)
-hide_fact (open) compute_float_divr
+
+end
subsection \<open>Approximate Power\<close>
-lemma div2_less_self[termination_simp]: fixes n::nat shows "odd n \<Longrightarrow> n div 2 < n"
+lemma div2_less_self[termination_simp]:
+ fixes n :: nat
+ shows "odd n \<Longrightarrow> n div 2 < n"
by (simp add: odd_pos)
-fun power_down :: "nat \<Rightarrow> real \<Rightarrow> nat \<Rightarrow> real" where
+fun power_down :: "nat \<Rightarrow> real \<Rightarrow> nat \<Rightarrow> real"
+where
"power_down p x 0 = 1"
| "power_down p x (Suc n) =
- (if odd n then truncate_down (Suc p) ((power_down p x (Suc n div 2))\<^sup>2) else truncate_down (Suc p) (x * power_down p x n))"
+ (if odd n then truncate_down (Suc p) ((power_down p x (Suc n div 2))\<^sup>2)
+ else truncate_down (Suc p) (x * power_down p x n))"
-fun power_up :: "nat \<Rightarrow> real \<Rightarrow> nat \<Rightarrow> real" where
+fun power_up :: "nat \<Rightarrow> real \<Rightarrow> nat \<Rightarrow> real"
+where
"power_up p x 0 = 1"
| "power_up p x (Suc n) =
- (if odd n then truncate_up p ((power_up p x (Suc n div 2))\<^sup>2) else truncate_up p (x * power_up p x n))"
+ (if odd n then truncate_up p ((power_up p x (Suc n div 2))\<^sup>2)
+ else truncate_up p (x * power_up p x n))"
lift_definition power_up_fl :: "nat \<Rightarrow> float \<Rightarrow> nat \<Rightarrow> float" is power_up
by (induct_tac rule: power_up.induct) simp_all
@@ -1279,11 +1426,13 @@
lemma compute_power_up_fl[code]:
"power_up_fl p x 0 = 1"
"power_up_fl p x (Suc n) =
- (if odd n then float_round_up p ((power_up_fl p x (Suc n div 2))\<^sup>2) else float_round_up p (x * power_up_fl p x n))"
+ (if odd n then float_round_up p ((power_up_fl p x (Suc n div 2))\<^sup>2)
+ else float_round_up p (x * power_up_fl p x n))"
and compute_power_down_fl[code]:
"power_down_fl p x 0 = 1"
"power_down_fl p x (Suc n) =
- (if odd n then float_round_down (Suc p) ((power_down_fl p x (Suc n div 2))\<^sup>2) else float_round_down (Suc p) (x * power_down_fl p x n))"
+ (if odd n then float_round_down (Suc p) ((power_down_fl p x (Suc n div 2))\<^sup>2)
+ else float_round_down (Suc p) (x * power_down_fl p x n))"
unfolding atomize_conj
by transfer simp
@@ -1300,7 +1449,7 @@
case (2 p x n)
{
assume "odd n"
- hence "(power_down p x (Suc n div 2)) ^ 2 \<le> (x ^ (Suc n div 2)) ^ 2"
+ then have "(power_down p x (Suc n div 2)) ^ 2 \<le> (x ^ (Suc n div 2)) ^ 2"
using 2
by (auto intro: power_mono power_down_nonneg simp del: odd_Suc_div_two)
also have "\<dots> = x ^ (Suc n div 2 * 2)"
@@ -1310,7 +1459,8 @@
finally have ?case
using \<open>odd n\<close>
by (auto intro!: truncate_down_le simp del: odd_Suc_div_two)
- } thus ?case
+ }
+ then show ?case
by (auto intro!: truncate_down_le mult_left_mono 2 mult_nonneg_nonneg power_down_nonneg)
qed simp
@@ -1319,9 +1469,9 @@
case (2 p x n)
{
assume "odd n"
- hence "Suc n = Suc n div 2 * 2"
+ then have "Suc n = Suc n div 2 * 2"
using \<open>odd n\<close> even_Suc by presburger
- hence "x ^ Suc n \<le> (x ^ (Suc n div 2))\<^sup>2"
+ then have "x ^ Suc n \<le> (x ^ (Suc n div 2))\<^sup>2"
by (simp add: power_mult[symmetric])
also have "\<dots> \<le> (power_up p x (Suc n div 2))\<^sup>2"
using 2 \<open>odd n\<close>
@@ -1329,7 +1479,8 @@
finally have ?case
using \<open>odd n\<close>
by (auto intro!: truncate_up_le simp del: odd_Suc_div_two )
- } thus ?case
+ }
+ then show ?case
by (auto intro!: truncate_up_le mult_left_mono 2)
qed simp
@@ -1383,8 +1534,7 @@
using truncate_down_uminus_eq[of p "x + y"]
by (auto simp: plus_down_def plus_up_def)
-lemma
- truncate_down_log2_eqI:
+lemma truncate_down_log2_eqI:
assumes "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> = \<lfloor>log 2 \<bar>y\<bar>\<rfloor>"
assumes "\<lfloor>x * 2 powr (p - \<lfloor>log 2 \<bar>x\<bar>\<rfloor> - 1)\<rfloor> = \<lfloor>y * 2 powr (p - \<lfloor>log 2 \<bar>x\<bar>\<rfloor> - 1)\<rfloor>"
shows "truncate_down p x = truncate_down p y"
@@ -1395,40 +1545,43 @@
(metis Float.compute_bitlen add.commute bitlen_def bitlen_nonneg less_add_same_cancel2 not_less
zero_less_one)
-lemma
- sum_neq_zeroI:
- fixes a k::real
+lemma sum_neq_zeroI:
+ fixes a k :: real
shows "abs a \<ge> k \<Longrightarrow> abs b < k \<Longrightarrow> a + b \<noteq> 0"
and "abs a > k \<Longrightarrow> abs b \<le> k \<Longrightarrow> a + b \<noteq> 0"
by auto
-lemma
- abs_real_le_2_powr_bitlen[simp]:
- "\<bar>real m2\<bar> < 2 powr real (bitlen \<bar>m2\<bar>)"
-proof cases
- assume "m2 \<noteq> 0"
- hence "\<bar>m2\<bar> < 2 ^ nat (bitlen \<bar>m2\<bar>)"
+lemma abs_real_le_2_powr_bitlen[simp]: "\<bar>real m2\<bar> < 2 powr real (bitlen \<bar>m2\<bar>)"
+proof (cases "m2 = 0")
+ case True
+ then show ?thesis by simp
+next
+ case False
+ then have "\<bar>m2\<bar> < 2 ^ nat (bitlen \<bar>m2\<bar>)"
using bitlen_bounds[of "\<bar>m2\<bar>"]
by (auto simp: powr_add bitlen_nonneg)
- thus ?thesis
+ then show ?thesis
by (simp add: powr_int bitlen_nonneg real_of_int_less_iff[symmetric])
-qed simp
+qed
lemma floor_sum_times_2_powr_sgn_eq:
- fixes ai p q::int
- and a b::real
+ fixes ai p q :: int
+ and a b :: real
assumes "a * 2 powr p = ai"
- assumes b_le_1: "abs (b * 2 powr (p + 1)) \<le> 1"
- assumes leqp: "q \<le> p"
+ and b_le_1: "abs (b * 2 powr (p + 1)) \<le> 1"
+ and leqp: "q \<le> p"
shows "\<lfloor>(a + b) * 2 powr q\<rfloor> = \<lfloor>(2 * ai + sgn b) * 2 powr (q - p - 1)\<rfloor>"
proof -
- {
- assume "b = 0"
- hence ?thesis
+ consider "b = 0" | "b > 0" | "b < 0" by arith
+ then show ?thesis
+ proof cases
+ case 1
+ then show ?thesis
by (simp add: assms(1)[symmetric] powr_add[symmetric] algebra_simps powr_mult_base)
- } moreover {
- assume "b > 0"
- hence "b * 2 powr p < abs (b * 2 powr (p + 1))" by simp
+ next
+ case 2
+ then have "b * 2 powr p < abs (b * 2 powr (p + 1))"
+ by simp
also note b_le_1
finally have b_less_1: "b * 2 powr real p < 1" .
@@ -1455,13 +1608,12 @@
by (simp del: real_of_int_power add: floor_divide_real_eq_div floor_eq)
finally
have "\<lfloor>(2 * ai + (sgn b)) * 2 powr (real (q - p) - 1)\<rfloor> =
- \<lfloor>real ai / real ((2::int) ^ nat (p - q))\<rfloor>"
- .
- } ultimately have ?thesis by simp
- } moreover {
- assume "\<not> 0 \<le> b"
- hence "0 > b" by simp
- hence floor_eq: "\<lfloor>b * 2 powr (real p + 1)\<rfloor> = -1"
+ \<lfloor>real ai / real ((2::int) ^ nat (p - q))\<rfloor>" .
+ }
+ ultimately show ?thesis by simp
+ next
+ case 3
+ then have floor_eq: "\<lfloor>b * 2 powr (real p + 1)\<rfloor> = -1"
using b_le_1
by (auto simp: floor_eq_iff algebra_simps pos_divide_le_eq[symmetric] abs_if divide_powr_uminus
intro!: mult_neg_pos split: split_if_asm)
@@ -1479,22 +1631,26 @@
del: real_of_int_mult real_of_int_power real_of_int_diff)
also have "\<dots> = \<lfloor>(2 * ai - 1) * 2 powr (q - p - 1)\<rfloor>"
using assms by (simp add: algebra_simps divide_powr_uminus powr_realpow[symmetric])
- finally have ?thesis using \<open>b < 0\<close> by simp
- } ultimately show ?thesis by arith
+ finally show ?thesis
+ using \<open>b < 0\<close> by simp
+ qed
qed
-lemma
- log2_abs_int_add_less_half_sgn_eq:
- fixes ai::int and b::real
- assumes "abs b \<le> 1/2" "ai \<noteq> 0"
+lemma log2_abs_int_add_less_half_sgn_eq:
+ fixes ai :: int
+ and b :: real
+ assumes "abs b \<le> 1/2"
+ and "ai \<noteq> 0"
shows "\<lfloor>log 2 \<bar>real ai + b\<bar>\<rfloor> = \<lfloor>log 2 \<bar>ai + sgn b / 2\<bar>\<rfloor>"
-proof cases
- assume "b = 0" thus ?thesis by simp
+proof (cases "b = 0")
+ case True
+ then show ?thesis by simp
next
- assume "b \<noteq> 0"
+ case False
def k \<equiv> "\<lfloor>log 2 \<bar>ai\<bar>\<rfloor>"
- hence "\<lfloor>log 2 \<bar>ai\<bar>\<rfloor> = k" by simp
- hence k: "2 powr k \<le> \<bar>ai\<bar>" "\<bar>ai\<bar> < 2 powr (k + 1)"
+ then have "\<lfloor>log 2 \<bar>ai\<bar>\<rfloor> = k"
+ by simp
+ then have k: "2 powr k \<le> \<bar>ai\<bar>" "\<bar>ai\<bar> < 2 powr (k + 1)"
by (simp_all add: floor_log_eq_powr_iff \<open>ai \<noteq> 0\<close>)
have "k \<ge> 0"
using assms by (auto simp: k_def)
@@ -1502,7 +1658,7 @@
have r: "0 \<le> r" "r < 2 powr k"
using \<open>k \<ge> 0\<close> k
by (auto simp: r_def k_def algebra_simps powr_add abs_if powr_int)
- hence "r \<le> (2::int) ^ nat k - 1"
+ then have "r \<le> (2::int) ^ nat k - 1"
using \<open>k \<ge> 0\<close> by (auto simp: powr_int)
from this[simplified real_of_int_le_iff[symmetric]] \<open>0 \<le> k\<close>
have r_le: "r \<le> 2 powr k - 1"
@@ -1511,7 +1667,7 @@
have "\<bar>ai\<bar> = 2 powr k + r"
using \<open>k \<ge> 0\<close> by (auto simp: k_def r_def powr_realpow[symmetric])
- have pos: "\<And>b::real. abs b < 1 \<Longrightarrow> 0 < 2 powr k + (r + b)"
+ have pos: "abs b < 1 \<Longrightarrow> 0 < 2 powr k + (r + b)" for b :: real
using \<open>0 \<le> k\<close> \<open>ai \<noteq> 0\<close>
by (auto simp add: r_def powr_realpow[symmetric] abs_if sgn_if algebra_simps
split: split_if_asm)
@@ -1557,8 +1713,12 @@
finally show ?thesis .
qed
-lemma compute_far_float_plus_down:
- fixes m1 e1 m2 e2::int and p::nat
+context
+begin
+
+qualified lemma compute_far_float_plus_down:
+ fixes m1 e1 m2 e2 :: int
+ and p :: nat
defines "k1 \<equiv> p - nat (bitlen \<bar>m1\<bar>)"
assumes H: "bitlen \<bar>m2\<bar> \<le> e1 - e2 - k1 - 2" "m1 \<noteq> 0" "m2 \<noteq> 0" "e1 \<ge> e2"
shows "float_plus_down p (Float m1 e1) (Float m2 e2) =
@@ -1584,7 +1744,7 @@
finally have abs_m2_less_half: "\<bar>?m2\<bar> < 1 / 2"
by simp
- hence "\<bar>real m2\<bar> < 2 powr -(?shift + 1)"
+ then have "\<bar>real m2\<bar> < 2 powr -(?shift + 1)"
unfolding powr_minus_divide by (auto simp: bitlen_def field_simps powr_mult_base abs_mult)
also have "\<dots> \<le> 2 powr real (e1 - e2 - 2)"
by simp
@@ -1593,7 +1753,7 @@
also have "1/4 < \<bar>real m1\<bar> / 2" using \<open>m1 \<noteq> 0\<close> by simp
finally have b_less_half_a: "\<bar>?b\<bar> < 1/2 * \<bar>?a\<bar>"
by (simp add: algebra_simps powr_mult_base abs_mult)
- hence a_half_less_sum: "\<bar>?a\<bar> / 2 < \<bar>?sum\<bar>"
+ then have a_half_less_sum: "\<bar>?a\<bar> / 2 < \<bar>?sum\<bar>"
by (auto simp: field_simps abs_if split: split_if_asm)
from b_less_half_a have "\<bar>?b\<bar> < \<bar>?a\<bar>" "\<bar>?b\<bar> \<le> \<bar>?a\<bar>"
@@ -1602,14 +1762,14 @@
have "\<bar>real (Float m1 e1)\<bar> \<ge> 1/4 * 2 powr real e1"
using \<open>m1 \<noteq> 0\<close>
by (auto simp: powr_add powr_int bitlen_nonneg divide_right_mono abs_mult)
- hence "?sum \<noteq> 0" using b_less_quarter
+ then have "?sum \<noteq> 0" using b_less_quarter
by (rule sum_neq_zeroI)
- hence "?m1 + ?m2 \<noteq> 0"
+ then have "?m1 + ?m2 \<noteq> 0"
unfolding sum_eq by (simp add: abs_mult zero_less_mult_iff)
have "\<bar>real ?m1\<bar> \<ge> 2 ^ Suc k1" "\<bar>?m2'\<bar> < 2 ^ Suc k1"
using \<open>m1 \<noteq> 0\<close> \<open>m2 \<noteq> 0\<close> by (auto simp: sgn_if less_1_mult abs_mult simp del: power.simps)
- hence sum'_nz: "?m1 + ?m2' \<noteq> 0"
+ then have sum'_nz: "?m1 + ?m2' \<noteq> 0"
by (intro sum_neq_zeroI)
have "\<lfloor>log 2 \<bar>real (Float m1 e1) + real (Float m2 e2)\<bar>\<rfloor> = \<lfloor>log 2 \<bar>?m1 + ?m2\<bar>\<rfloor> + ?e"
@@ -1624,13 +1784,13 @@
also
have "\<bar>?m1 + ?m2'\<bar> * 2 powr ?e = \<bar>?m1 * 2 + sgn m2\<bar> * 2 powr (?e - 1)"
by (auto simp: field_simps powr_minus[symmetric] powr_divide2[symmetric] powr_mult_base)
- hence "\<lfloor>log 2 \<bar>?m1 + ?m2'\<bar>\<rfloor> + ?e = \<lfloor>log 2 \<bar>real (Float (?m1 * 2 + sgn m2) (?e - 1))\<bar>\<rfloor>"
+ then have "\<lfloor>log 2 \<bar>?m1 + ?m2'\<bar>\<rfloor> + ?e = \<lfloor>log 2 \<bar>real (Float (?m1 * 2 + sgn m2) (?e - 1))\<bar>\<rfloor>"
using \<open>?m1 + ?m2' \<noteq> 0\<close>
unfolding floor_add[symmetric]
by (simp add: log_add_eq_powr abs_mult_pos)
finally
have "\<lfloor>log 2 \<bar>?sum\<bar>\<rfloor> = \<lfloor>log 2 \<bar>real (Float (?m1*2 + sgn m2) (?e - 1))\<bar>\<rfloor>" .
- hence "plus_down p (Float m1 e1) (Float m2 e2) =
+ then have "plus_down p (Float m1 e1) (Float m2 e2) =
truncate_down p (Float (?m1*2 + sgn m2) (?e - 1))"
unfolding plus_down_def
proof (rule truncate_down_log2_eqI)
@@ -1668,14 +1828,14 @@
finally
show "\<lfloor>(?a + ?b) * 2 powr ?f\<rfloor> = \<lfloor>real (Float (?m1 * 2 + sgn m2) (?e - 1)) * 2 powr ?f\<rfloor>" .
qed
- thus ?thesis
+ then show ?thesis
by transfer (simp add: plus_down_def ac_simps Let_def)
qed
lemma compute_float_plus_down_naive[code]: "float_plus_down p x y = float_round_down p (x + y)"
by transfer (auto simp: plus_down_def)
-lemma compute_float_plus_down[code]:
+qualified lemma compute_float_plus_down[code]:
fixes p::nat and m1 e1 m2 e2::int
shows "float_plus_down p (Float m1 e1) (Float m2 e2) =
(if m1 = 0 then float_round_down p (Float m2 e2)
@@ -1689,53 +1849,66 @@
else float_plus_down p (Float m2 e2) (Float m1 e1)))"
proof -
{
- assume H: "bitlen \<bar>m2\<bar> \<le> e1 - e2 - (p - nat (bitlen \<bar>m1\<bar>)) - 2" "m1 \<noteq> 0" "m2 \<noteq> 0" "e1 \<ge> e2"
- note compute_far_float_plus_down[OF H]
+ assume "bitlen \<bar>m2\<bar> \<le> e1 - e2 - (p - nat (bitlen \<bar>m1\<bar>)) - 2" "m1 \<noteq> 0" "m2 \<noteq> 0" "e1 \<ge> e2"
+ note compute_far_float_plus_down[OF this]
}
- thus ?thesis
+ then show ?thesis
by transfer (simp add: Let_def plus_down_def ac_simps)
qed
-hide_fact (open) compute_far_float_plus_down
-hide_fact (open) compute_float_plus_down
-lemma compute_float_plus_up[code]: "float_plus_up p x y = - float_plus_down p (-x) (-y)"
+qualified lemma compute_float_plus_up[code]: "float_plus_up p x y = - float_plus_down p (-x) (-y)"
using truncate_down_uminus_eq[of p "x + y"]
by transfer (simp add: plus_down_def plus_up_def ac_simps)
-hide_fact (open) compute_float_plus_up
lemma mantissa_zero[simp]: "mantissa 0 = 0"
-by (metis mantissa_0 zero_float.abs_eq)
+ by (metis mantissa_0 zero_float.abs_eq)
+
+end
subsection \<open>Lemmas needed by Approximate\<close>
-lemma Float_num[simp]: shows
- "real (Float 1 0) = 1" and "real (Float 1 1) = 2" and "real (Float 1 2) = 4" and
- "real (Float 1 (- 1)) = 1/2" and "real (Float 1 (- 2)) = 1/4" and "real (Float 1 (- 3)) = 1/8" and
- "real (Float (- 1) 0) = -1" and "real (Float (number_of n) 0) = number_of n"
-using two_powr_int_float[of 2] two_powr_int_float[of "-1"] two_powr_int_float[of "-2"] two_powr_int_float[of "-3"]
-using powr_realpow[of 2 2] powr_realpow[of 2 3]
-using powr_minus[of 2 1] powr_minus[of 2 2] powr_minus[of 2 3]
-by auto
+lemma Float_num[simp]:
+ "real (Float 1 0) = 1"
+ "real (Float 1 1) = 2"
+ "real (Float 1 2) = 4"
+ "real (Float 1 (- 1)) = 1/2"
+ "real (Float 1 (- 2)) = 1/4"
+ "real (Float 1 (- 3)) = 1/8"
+ "real (Float (- 1) 0) = -1"
+ "real (Float (number_of n) 0) = number_of n"
+ using two_powr_int_float[of 2] two_powr_int_float[of "-1"] two_powr_int_float[of "-2"]
+ two_powr_int_float[of "-3"]
+ using powr_realpow[of 2 2] powr_realpow[of 2 3]
+ using powr_minus[of 2 1] powr_minus[of 2 2] powr_minus[of 2 3]
+ by auto
-lemma real_of_Float_int[simp]: "real (Float n 0) = real n" by simp
+lemma real_of_Float_int[simp]: "real (Float n 0) = real n"
+ by simp
-lemma float_zero[simp]: "real (Float 0 e) = 0" by simp
+lemma float_zero[simp]: "real (Float 0 e) = 0"
+ by simp
lemma abs_div_2_less: "a \<noteq> 0 \<Longrightarrow> a \<noteq> -1 \<Longrightarrow> abs((a::int) div 2) < abs a"
-by arith
+ by arith
-lemma lapprox_rat:
- shows "real (lapprox_rat prec x y) \<le> real x / real y"
+lemma lapprox_rat: "real (lapprox_rat prec x y) \<le> real x / real y"
using round_down by (simp add: lapprox_rat_def)
-lemma mult_div_le: fixes a b:: int assumes "b > 0" shows "a \<ge> b * (a div b)"
+lemma mult_div_le:
+ fixes a b :: int
+ assumes "b > 0"
+ shows "a \<ge> b * (a div b)"
proof -
- from zmod_zdiv_equality'[of a b]
- have "a = b * (a div b) + a mod b" by simp
- also have "... \<ge> b * (a div b) + 0" apply (rule add_left_mono) apply (rule pos_mod_sign)
- using assms by simp
- finally show ?thesis by simp
+ from zmod_zdiv_equality'[of a b] have "a = b * (a div b) + a mod b"
+ by simp
+ also have "\<dots> \<ge> b * (a div b) + 0"
+ apply (rule add_left_mono)
+ apply (rule pos_mod_sign)
+ using assms apply simp
+ done
+ finally show ?thesis
+ by simp
qed
lemma lapprox_rat_nonneg:
@@ -1758,12 +1931,10 @@
by transfer (auto intro!: round_up_le1 simp: rat_precision_def)
qed
-lemma rapprox_rat_nonneg_nonpos:
- "0 \<le> x \<Longrightarrow> y \<le> 0 \<Longrightarrow> real (rapprox_rat n x y) \<le> 0"
+lemma rapprox_rat_nonneg_nonpos: "0 \<le> x \<Longrightarrow> y \<le> 0 \<Longrightarrow> real (rapprox_rat n x y) \<le> 0"
by transfer (simp add: round_up_le0 divide_nonneg_nonpos)
-lemma rapprox_rat_nonpos_nonneg:
- "x \<le> 0 \<Longrightarrow> 0 \<le> y \<Longrightarrow> real (rapprox_rat n x y) \<le> 0"
+lemma rapprox_rat_nonpos_nonneg: "x \<le> 0 \<Longrightarrow> 0 \<le> y \<Longrightarrow> real (rapprox_rat n x y) \<le> 0"
by transfer (simp add: round_up_le0 divide_nonpos_nonneg)
lemma real_divl: "real_divl prec x y \<le> x / y"
@@ -1793,17 +1964,32 @@
by (simp add: bitlen_def)
lemma mantissa_eq_zero_iff: "mantissa x = 0 \<longleftrightarrow> x = 0"
+ (is "?lhs \<longleftrightarrow> ?rhs")
proof
- assume "mantissa x = 0" hence z: "0 = real x" using mantissa_exponent by simp
- show "x = 0" by (simp add: zero_float_def z)
-qed (simp add: zero_float_def)
+ show ?rhs if ?lhs
+ proof -
+ from that have z: "0 = real x"
+ using mantissa_exponent by simp
+ show ?thesis
+ by (simp add: zero_float_def z)
+ qed
+ show ?lhs if ?rhs
+ using that by (simp add: zero_float_def)
+qed
lemma float_upper_bound: "x \<le> 2 powr (bitlen \<bar>mantissa x\<bar> + exponent x)"
-proof (cases "x = 0", simp)
- assume "x \<noteq> 0" hence "mantissa x \<noteq> 0" using mantissa_eq_zero_iff by auto
- have "x = mantissa x * 2 powr (exponent x)" by (rule mantissa_exponent)
- also have "mantissa x \<le> \<bar>mantissa x\<bar>" by simp
- also have "... \<le> 2 powr (bitlen \<bar>mantissa x\<bar>)"
+proof (cases "x = 0")
+ case True
+ then show ?thesis by simp
+next
+ case False
+ then have "mantissa x \<noteq> 0"
+ using mantissa_eq_zero_iff by auto
+ have "x = mantissa x * 2 powr (exponent x)"
+ by (rule mantissa_exponent)
+ also have "mantissa x \<le> \<bar>mantissa x\<bar>"
+ by simp
+ also have "\<dots> \<le> 2 powr (bitlen \<bar>mantissa x\<bar>)"
using bitlen_bounds[of "\<bar>mantissa x\<bar>"] bitlen_nonneg \<open>mantissa x \<noteq> 0\<close>
by (auto simp del: real_of_int_abs simp add: powr_int)
finally show ?thesis by (simp add: powr_add)
@@ -1813,22 +1999,28 @@
assumes "0 < x" "x \<le> 1" "prec \<ge> 1"
shows "1 \<le> real_divl prec 1 x"
proof -
- have "log 2 x \<le> real prec + real \<lfloor>log 2 x\<rfloor>" using \<open>prec \<ge> 1\<close> by arith
+ have "log 2 x \<le> real prec + real \<lfloor>log 2 x\<rfloor>"
+ using \<open>prec \<ge> 1\<close> by arith
from this assms show ?thesis
by (simp add: real_divl_def log_divide round_down_ge1)
qed
lemma float_divl_pos_less1_bound:
"0 < real x \<Longrightarrow> real x \<le> 1 \<Longrightarrow> prec \<ge> 1 \<Longrightarrow> 1 \<le> real (float_divl prec 1 x)"
- by (transfer, rule real_divl_pos_less1_bound)
+ by transfer (rule real_divl_pos_less1_bound)
lemma float_divr: "real x / real y \<le> real (float_divr prec x y)"
by transfer (rule real_divr)
-lemma real_divr_pos_less1_lower_bound: assumes "0 < x" and "x \<le> 1" shows "1 \<le> real_divr prec 1 x"
+lemma real_divr_pos_less1_lower_bound:
+ assumes "0 < x"
+ and "x \<le> 1"
+ shows "1 \<le> real_divr prec 1 x"
proof -
- have "1 \<le> 1 / x" using \<open>0 < x\<close> and \<open>x <= 1\<close> by auto
- also have "\<dots> \<le> real_divr prec 1 x" using real_divr[where x=1 and y=x] by auto
+ have "1 \<le> 1 / x"
+ using \<open>0 < x\<close> and \<open>x <= 1\<close> by auto
+ also have "\<dots> \<le> real_divr prec 1 x"
+ using real_divr[where x=1 and y=x] by auto
finally show ?thesis by auto
qed
@@ -1855,19 +2047,21 @@
assumes "0 \<le> x" "x \<le> y"
shows "truncate_up prec x \<le> truncate_up prec y"
proof -
- {
- assume "\<lfloor>log 2 x\<rfloor> = \<lfloor>log 2 y\<rfloor>"
- hence ?thesis
+ consider "\<lfloor>log 2 x\<rfloor> = \<lfloor>log 2 y\<rfloor>" | "\<lfloor>log 2 x\<rfloor> \<noteq> \<lfloor>log 2 y\<rfloor>" "0 < x" | "x \<le> 0"
+ by arith
+ then show ?thesis
+ proof cases
+ case 1
+ then show ?thesis
using assms
by (auto simp: truncate_up_def round_up_def intro!: ceiling_mono)
- } moreover {
- assume "0 < x"
- hence "log 2 x \<le> log 2 y" using assms by auto
- moreover
- assume "\<lfloor>log 2 x\<rfloor> \<noteq> \<lfloor>log 2 y\<rfloor>"
- ultimately have logless: "log 2 x < log 2 y" and flogless: "\<lfloor>log 2 x\<rfloor> < \<lfloor>log 2 y\<rfloor>"
- unfolding atomize_conj
- by (metis floor_less_cancel linorder_cases not_le)
+ next
+ case 2
+ from assms \<open>0 < x\<close> have "log 2 x \<le> log 2 y"
+ by auto
+ with \<open>\<lfloor>log 2 x\<rfloor> \<noteq> \<lfloor>log 2 y\<rfloor>\<close>
+ have logless: "log 2 x < log 2 y" and flogless: "\<lfloor>log 2 x\<rfloor> < \<lfloor>log 2 y\<rfloor>"
+ by (metis floor_less_cancel linorder_cases not_le)+
have "truncate_up prec x =
real \<lceil>x * 2 powr real (int prec - \<lfloor>log 2 x\<rfloor> - 1)\<rceil> * 2 powr - real (int prec - \<lfloor>log 2 x\<rfloor> - 1)"
using assms by (simp add: truncate_up_def round_up_def)
@@ -1876,10 +2070,10 @@
have "x * 2 powr real (int prec - \<lfloor>log 2 x\<rfloor> - 1) \<le> x * (2 powr real prec / (2 powr log 2 x))"
using real_of_int_floor_add_one_ge[of "log 2 x"] assms
by (auto simp add: algebra_simps powr_divide2 intro!: mult_left_mono)
- thus "x * 2 powr real (int prec - \<lfloor>log 2 x\<rfloor> - 1) \<le> real ((2::int) ^ prec)"
+ then show "x * 2 powr real (int prec - \<lfloor>log 2 x\<rfloor> - 1) \<le> real ((2::int) ^ prec)"
using \<open>0 < x\<close> by (simp add: powr_realpow)
qed
- hence "real \<lceil>x * 2 powr real (int prec - \<lfloor>log 2 x\<rfloor> - 1)\<rceil> \<le> 2 powr int prec"
+ then have "real \<lceil>x * 2 powr real (int prec - \<lfloor>log 2 x\<rfloor> - 1)\<rceil> \<le> 2 powr int prec"
by (auto simp: powr_realpow)
also
have "2 powr - real (int prec - \<lfloor>log 2 x\<rfloor> - 1) \<le> 2 powr - real (int prec - \<lfloor>log 2 y\<rfloor>)"
@@ -1896,14 +2090,13 @@
by (simp add: powr_add)
also have "\<dots> \<le> truncate_up prec y"
by (rule truncate_up)
- finally have ?thesis .
- } moreover {
- assume "~ 0 < x"
- hence ?thesis
+ finally show ?thesis .
+ next
+ case 3
+ then show ?thesis
using assms
by (auto intro!: truncate_up_le)
- } ultimately show ?thesis
- by blast
+ qed
qed
lemma truncate_up_switch_sign_mono:
@@ -1931,20 +2124,22 @@
finally have "\<lfloor>x * 2 powr (- real \<lfloor>log 2 x\<rfloor> - 1)\<rfloor> < 1"
unfolding less_ceiling_eq real_of_int_minus real_of_one
by simp
- moreover
- have "0 \<le> \<lfloor>x * 2 powr (- real \<lfloor>log 2 x\<rfloor> - 1)\<rfloor>"
+ moreover have "0 \<le> \<lfloor>x * 2 powr (- real \<lfloor>log 2 x\<rfloor> - 1)\<rfloor>"
using \<open>x > 0\<close> by auto
ultimately have "\<lfloor>x * 2 powr (- real \<lfloor>log 2 x\<rfloor> - 1)\<rfloor> \<in> {0 ..< 1}"
by simp
- also have "\<dots> \<subseteq> {0}" by auto
- finally have "\<lfloor>x * 2 powr (- real \<lfloor>log 2 x\<rfloor> - 1)\<rfloor> = 0" by simp
+ also have "\<dots> \<subseteq> {0}"
+ by auto
+ finally have "\<lfloor>x * 2 powr (- real \<lfloor>log 2 x\<rfloor> - 1)\<rfloor> = 0"
+ by simp
with assms show ?thesis
by (auto simp: truncate_down_def round_down_def)
qed
lemma truncate_down_switch_sign_mono:
- assumes "x \<le> 0" "0 \<le> y"
- assumes "x \<le> y"
+ assumes "x \<le> 0"
+ and "0 \<le> y"
+ and "x \<le> y"
shows "truncate_down prec x \<le> truncate_down prec y"
proof -
note truncate_down_le[OF \<open>x \<le> 0\<close>]
@@ -1956,32 +2151,36 @@
assumes "0 \<le> x" "x \<le> y"
shows "truncate_down prec x \<le> truncate_down prec y"
proof -
- {
- assume "0 < x" "prec = 0"
- with assms have ?thesis
+ consider "0 < x" "prec = 0" | "x \<le> 0" | "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> = \<lfloor>log 2 \<bar>y\<bar>\<rfloor>" |
+ "0 < x" "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> \<noteq> \<lfloor>log 2 \<bar>y\<bar>\<rfloor>" "prec \<noteq> 0"
+ by arith
+ then show ?thesis
+ proof cases
+ case 1
+ with assms show ?thesis
by (simp add: truncate_down_zeroprec_mono)
- } moreover {
- assume "~ 0 < x"
+ next
+ case 2
with assms have "x = 0" "0 \<le> y" by simp_all
- hence ?thesis
+ then show ?thesis
by (auto intro!: truncate_down_nonneg)
- } moreover {
- assume "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> = \<lfloor>log 2 \<bar>y\<bar>\<rfloor>"
- hence ?thesis
+ next
+ case 3
+ then show ?thesis
using assms
by (auto simp: truncate_down_def round_down_def intro!: floor_mono)
- } moreover {
- assume "0 < x"
- hence "log 2 x \<le> log 2 y" "0 < y" "0 \<le> y" using assms by auto
- moreover
- assume "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> \<noteq> \<lfloor>log 2 \<bar>y\<bar>\<rfloor>"
- ultimately have logless: "log 2 x < log 2 y" and flogless: "\<lfloor>log 2 x\<rfloor> < \<lfloor>log 2 y\<rfloor>"
+ next
+ case 4
+ from \<open>0 < x\<close> have "log 2 x \<le> log 2 y" "0 < y" "0 \<le> y"
+ using assms by auto
+ with \<open>\<lfloor>log 2 \<bar>x\<bar>\<rfloor> \<noteq> \<lfloor>log 2 \<bar>y\<bar>\<rfloor>\<close>
+ have logless: "log 2 x < log 2 y" and flogless: "\<lfloor>log 2 x\<rfloor> < \<lfloor>log 2 y\<rfloor>"
unfolding atomize_conj abs_of_pos[OF \<open>0 < x\<close>] abs_of_pos[OF \<open>0 < y\<close>]
by (metis floor_less_cancel linorder_cases not_le)
- assume "prec \<noteq> 0" hence [simp]: "prec \<ge> Suc 0" by auto
+ from \<open>prec \<noteq> 0\<close> have [simp]: "prec \<ge> Suc 0"
+ by auto
have "2 powr (prec - 1) \<le> y * 2 powr real (prec - 1) / (2 powr log 2 y)"
- using \<open>0 < y\<close>
- by simp
+ using \<open>0 < y\<close> by simp
also have "\<dots> \<le> y * 2 powr real prec / (2 powr (real \<lfloor>log 2 y\<rfloor> + 1))"
using \<open>0 \<le> y\<close> \<open>0 \<le> x\<close> assms(2)
by (auto intro!: powr_mono divide_left_mono
@@ -1992,7 +2191,7 @@
finally have "(2 ^ (prec - 1)) \<le> \<lfloor>y * 2 powr real (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor> - 1)\<rfloor>"
using \<open>0 \<le> y\<close>
by (auto simp: powr_divide2[symmetric] le_floor_eq powr_realpow)
- hence "(2 ^ (prec - 1)) * 2 powr - real (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor> - 1) \<le> truncate_down prec y"
+ then have "(2 ^ (prec - 1)) * 2 powr - real (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor> - 1) \<le> truncate_down prec y"
by (auto simp: truncate_down_def round_down_def)
moreover
{
@@ -2006,9 +2205,10 @@
by (auto intro!: floor_mono)
finally have "x \<le> (2 ^ (prec - 1)) * 2 powr - real (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor> - 1)"
by (auto simp: powr_realpow[symmetric] powr_divide2[symmetric] assms real_of_nat_diff)
- } ultimately have ?thesis
+ }
+ ultimately show ?thesis
by (metis dual_order.trans truncate_down)
- } ultimately show ?thesis by blast
+ qed
qed
lemma truncate_down_eq_truncate_up: "truncate_down p x = - truncate_up p (-x)"
@@ -2029,49 +2229,62 @@
lemma Float_le_zero_iff: "Float a b \<le> 0 \<longleftrightarrow> a \<le> 0"
by (auto simp: zero_float_def mult_le_0_iff) (simp add: not_less [symmetric])
-lemma real_of_float_pprt[simp]: fixes a::float shows "real (pprt a) = pprt (real a)"
+lemma real_of_float_pprt[simp]:
+ fixes a :: float
+ shows "real (pprt a) = pprt (real a)"
unfolding pprt_def sup_float_def max_def sup_real_def by auto
-lemma real_of_float_nprt[simp]: fixes a::float shows "real (nprt a) = nprt (real a)"
+lemma real_of_float_nprt[simp]:
+ fixes a :: float
+ shows "real (nprt a) = nprt (real a)"
unfolding nprt_def inf_float_def min_def inf_real_def by auto
+context
+begin
+
lift_definition int_floor_fl :: "float \<Rightarrow> int" is floor .
-lemma compute_int_floor_fl[code]:
+qualified lemma compute_int_floor_fl[code]:
"int_floor_fl (Float m e) = (if 0 \<le> e then m * 2 ^ nat e else m div (2 ^ (nat (-e))))"
by transfer (simp add: powr_int int_of_reals floor_divide_eq_div del: real_of_ints)
-hide_fact (open) compute_int_floor_fl
lift_definition floor_fl :: "float \<Rightarrow> float" is "\<lambda>x. real (floor x)" by simp
-lemma compute_floor_fl[code]:
+qualified lemma compute_floor_fl[code]:
"floor_fl (Float m e) = (if 0 \<le> e then Float m e else Float (m div (2 ^ (nat (-e)))) 0)"
by transfer (simp add: powr_int int_of_reals floor_divide_eq_div del: real_of_ints)
-hide_fact (open) compute_floor_fl
+
+end
-lemma floor_fl: "real (floor_fl x) \<le> real x" by transfer simp
+lemma floor_fl: "real (floor_fl x) \<le> real x"
+ by transfer simp
-lemma int_floor_fl: "real (int_floor_fl x) \<le> real x" by transfer simp
+lemma int_floor_fl: "real (int_floor_fl x) \<le> real x"
+ by transfer simp
lemma floor_pos_exp: "exponent (floor_fl x) \<ge> 0"
proof (cases "floor_fl x = float_of 0")
case True
- then show ?thesis by (simp add: floor_fl_def)
+ then show ?thesis
+ by (simp add: floor_fl_def)
next
case False
- have eq: "floor_fl x = Float \<lfloor>real x\<rfloor> 0" by transfer simp
+ have eq: "floor_fl x = Float \<lfloor>real x\<rfloor> 0"
+ by transfer simp
obtain i where "\<lfloor>real x\<rfloor> = mantissa (floor_fl x) * 2 ^ i" "0 = exponent (floor_fl x) - int i"
by (rule denormalize_shift[OF eq[THEN eq_reflection] False])
- then show ?thesis by simp
+ then show ?thesis
+ by simp
qed
lemma compute_mantissa[code]:
- "mantissa (Float m e) = (if m = 0 then 0 else if 2 dvd m then mantissa (normfloat (Float m e)) else m)"
+ "mantissa (Float m e) =
+ (if m = 0 then 0 else if 2 dvd m then mantissa (normfloat (Float m e)) else m)"
by (auto simp: mantissa_float Float.abs_eq)
lemma compute_exponent[code]:
- "exponent (Float m e) = (if m = 0 then 0 else if 2 dvd m then exponent (normfloat (Float m e)) else e)"
+ "exponent (Float m e) =
+ (if m = 0 then 0 else if 2 dvd m then exponent (normfloat (Float m e)) else e)"
by (auto simp: exponent_float Float.abs_eq)
end
-
--- a/src/HOL/Library/Lattice_Algebras.thy Thu Jul 09 15:52:11 2015 +0200
+++ b/src/HOL/Library/Lattice_Algebras.thy Thu Jul 09 17:36:04 2015 +0200
@@ -145,7 +145,8 @@
lemma nprt_le_zero[simp]: "nprt a \<le> 0"
by (simp add: nprt_def)
-lemma le_eq_neg: "a \<le> - b \<longleftrightarrow> a + b \<le> 0" (is "?l = ?r")
+lemma le_eq_neg: "a \<le> - b \<longleftrightarrow> a + b \<le> 0"
+ (is "?l = ?r")
proof
assume ?l
then show ?r
@@ -177,25 +178,24 @@
lemma nprt_eq_0 [simp, no_atp]: "0 \<le> x \<Longrightarrow> nprt x = 0"
by (simp add: nprt_def inf_absorb2)
-lemma sup_0_imp_0: "sup a (- a) = 0 \<Longrightarrow> a = 0"
+lemma sup_0_imp_0:
+ assumes "sup a (- a) = 0"
+ shows "a = 0"
proof -
- {
- fix a :: 'a
- assume hyp: "sup a (- a) = 0"
- then have "sup a (- a) + a = a"
+ have p: "0 \<le> a" if "sup a (- a) = 0" for a :: 'a
+ proof -
+ from that have "sup a (- a) + a = a"
by simp
then have "sup (a + a) 0 = a"
by (simp add: add_sup_distrib_right)
then have "sup (a + a) 0 \<le> a"
by simp
- then have "0 \<le> a"
+ then show ?thesis
by (blast intro: order_trans inf_sup_ord)
- }
- note p = this
- assume hyp:"sup a (-a) = 0"
- then have hyp2:"sup (-a) (-(-a)) = 0"
+ qed
+ from assms have **: "sup (-a) (-(-a)) = 0"
by (simp add: sup_commute)
- from p[OF hyp] p[OF hyp2] show "a = 0"
+ from p[OF assms] p[OF **] show "a = 0"
by simp
qed
@@ -217,49 +217,50 @@
apply simp
done
-lemma zero_le_double_add_iff_zero_le_single_add [simp]:
- "0 \<le> a + a \<longleftrightarrow> 0 \<le> a"
+lemma zero_le_double_add_iff_zero_le_single_add [simp]: "0 \<le> a + a \<longleftrightarrow> 0 \<le> a"
+ (is "?lhs \<longleftrightarrow> ?rhs")
proof
- assume "0 \<le> a + a"
- then have a: "inf (a + a) 0 = 0"
- by (simp add: inf_commute inf_absorb1)
- have "inf a 0 + inf a 0 = inf (inf (a + a) 0) a" (is "?l=_")
- by (simp add: add_sup_inf_distribs inf_aci)
- then have "?l = 0 + inf a 0"
- by (simp add: a, simp add: inf_commute)
- then have "inf a 0 = 0"
- by (simp only: add_right_cancel)
- then show "0 \<le> a"
- unfolding le_iff_inf by (simp add: inf_commute)
-next
- assume a: "0 \<le> a"
- show "0 \<le> a + a"
- by (simp add: add_mono[OF a a, simplified])
+ show ?rhs if ?lhs
+ proof -
+ from that have a: "inf (a + a) 0 = 0"
+ by (simp add: inf_commute inf_absorb1)
+ have "inf a 0 + inf a 0 = inf (inf (a + a) 0) a" (is "?l=_")
+ by (simp add: add_sup_inf_distribs inf_aci)
+ then have "?l = 0 + inf a 0"
+ by (simp add: a, simp add: inf_commute)
+ then have "inf a 0 = 0"
+ by (simp only: add_right_cancel)
+ then show ?thesis
+ unfolding le_iff_inf by (simp add: inf_commute)
+ qed
+ show ?lhs if ?rhs
+ by (simp add: add_mono[OF that that, simplified])
qed
lemma double_zero [simp]: "a + a = 0 \<longleftrightarrow> a = 0"
+ (is "?lhs \<longleftrightarrow> ?rhs")
proof
- assume assm: "a + a = 0"
- then have "a + a + - a = - a"
- by simp
- then have "a + (a + - a) = - a"
- by (simp only: add.assoc)
- then have a: "- a = a"
- by simp
- show "a = 0"
- apply (rule antisym)
- apply (unfold neg_le_iff_le [symmetric, of a])
- unfolding a
- apply simp
- unfolding zero_le_double_add_iff_zero_le_single_add [symmetric, of a]
- unfolding assm
- unfolding le_less
- apply simp_all
- done
-next
- assume "a = 0"
- then show "a + a = 0"
- by simp
+ show ?rhs if ?lhs
+ proof -
+ from that have "a + a + - a = - a"
+ by simp
+ then have "a + (a + - a) = - a"
+ by (simp only: add.assoc)
+ then have a: "- a = a"
+ by simp
+ show ?thesis
+ apply (rule antisym)
+ apply (unfold neg_le_iff_le [symmetric, of a])
+ unfolding a
+ apply simp
+ unfolding zero_le_double_add_iff_zero_le_single_add [symmetric, of a]
+ unfolding that
+ unfolding le_less
+ apply simp_all
+ done
+ qed
+ show ?lhs if ?rhs
+ using that by simp
qed
lemma zero_less_double_add_iff_zero_less_single_add [simp]: "0 < a + a \<longleftrightarrow> 0 < a"
@@ -281,19 +282,17 @@
done
qed
-lemma double_add_le_zero_iff_single_add_le_zero [simp]:
- "a + a \<le> 0 \<longleftrightarrow> a \<le> 0"
+lemma double_add_le_zero_iff_single_add_le_zero [simp]: "a + a \<le> 0 \<longleftrightarrow> a \<le> 0"
proof -
have "a + a \<le> 0 \<longleftrightarrow> 0 \<le> - (a + a)"
- by (subst le_minus_iff, simp)
+ by (subst le_minus_iff) simp
moreover have "\<dots> \<longleftrightarrow> a \<le> 0"
by (simp only: minus_add_distrib zero_le_double_add_iff_zero_le_single_add) simp
ultimately show ?thesis
by blast
qed
-lemma double_add_less_zero_iff_single_less_zero [simp]:
- "a + a < 0 \<longleftrightarrow> a < 0"
+lemma double_add_less_zero_iff_single_less_zero [simp]: "a + a < 0 \<longleftrightarrow> a < 0"
proof -
have "a + a < 0 \<longleftrightarrow> 0 < - (a + a)"
by (subst less_minus_iff) simp
@@ -316,8 +315,8 @@
lemma minus_le_self_iff: "- a \<le> a \<longleftrightarrow> 0 \<le> a"
proof -
- from add_le_cancel_left [of "uminus a" zero "plus a a"]
have "- a \<le> a \<longleftrightarrow> 0 \<le> a + a"
+ using add_le_cancel_left [of "uminus a" zero "plus a a"]
by (simp add: add.assoc[symmetric])
then show ?thesis
by simp
@@ -370,15 +369,14 @@
subclass ordered_ab_group_add_abs
proof
- have abs_ge_zero [simp]: "\<And>a. 0 \<le> \<bar>a\<bar>"
+ have abs_ge_zero [simp]: "0 \<le> \<bar>a\<bar>" for a
proof -
- fix a b
have a: "a \<le> \<bar>a\<bar>" and b: "- a \<le> \<bar>a\<bar>"
by (auto simp add: abs_lattice)
show "0 \<le> \<bar>a\<bar>"
by (rule add_mono [OF a b, simplified])
qed
- have abs_leI: "\<And>a b. a \<le> b \<Longrightarrow> - a \<le> b \<Longrightarrow> \<bar>a\<bar> \<le> b"
+ have abs_leI: "a \<le> b \<Longrightarrow> - a \<le> b \<Longrightarrow> \<bar>a\<bar> \<le> b" for a b
by (simp add: abs_lattice le_supI)
fix a b
show "0 \<le> \<bar>a\<bar>"
@@ -387,15 +385,12 @@
by (auto simp add: abs_lattice)
show "\<bar>-a\<bar> = \<bar>a\<bar>"
by (simp add: abs_lattice sup_commute)
- {
- assume "a \<le> b"
- then show "- a \<le> b \<Longrightarrow> \<bar>a\<bar> \<le> b"
- by (rule abs_leI)
- }
+ show "- a \<le> b \<Longrightarrow> \<bar>a\<bar> \<le> b" if "a \<le> b"
+ using that by (rule abs_leI)
show "\<bar>a + b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"
proof -
have g: "\<bar>a\<bar> + \<bar>b\<bar> = sup (a + b) (sup (- a - b) (sup (- a + b) (a + (- b))))"
- (is "_=sup ?m ?n")
+ (is "_ = sup ?m ?n")
by (simp add: abs_lattice add_sup_inf_distribs ac_simps)
have a: "a + b \<le> sup ?m ?n"
by simp
@@ -420,25 +415,23 @@
end
lemma sup_eq_if:
- fixes a :: "'a::{lattice_ab_group_add, linorder}"
+ fixes a :: "'a::{lattice_ab_group_add,linorder}"
shows "sup a (- a) = (if a < 0 then - a else a)"
-proof -
- note add_le_cancel_right [of a a "- a", symmetric, simplified]
- moreover note add_le_cancel_right [of "-a" a a, symmetric, simplified]
- then show ?thesis by (auto simp: sup_max max.absorb1 max.absorb2)
-qed
+ using add_le_cancel_right [of a a "- a", symmetric, simplified]
+ and add_le_cancel_right [of "-a" a a, symmetric, simplified]
+ by (auto simp: sup_max max.absorb1 max.absorb2)
lemma abs_if_lattice:
- fixes a :: "'a::{lattice_ab_group_add_abs, linorder}"
+ fixes a :: "'a::{lattice_ab_group_add_abs,linorder}"
shows "\<bar>a\<bar> = (if a < 0 then - a else a)"
by auto
lemma estimate_by_abs:
fixes a b c :: "'a::lattice_ab_group_add_abs"
- shows "a + b \<le> c \<Longrightarrow> a \<le> c + \<bar>b\<bar>"
+ assumes "a + b \<le> c"
+ shows "a \<le> c + \<bar>b\<bar>"
proof -
- assume "a + b \<le> c"
- then have "a \<le> c + (- b)"
+ from assms have "a \<le> c + (- b)"
by (simp add: algebra_simps)
have "- b \<le> \<bar>b\<bar>"
by (rule abs_ge_minus_self)
@@ -464,15 +457,12 @@
let ?y = "pprt a * pprt b + pprt a * nprt b + nprt a * pprt b + nprt a * nprt b"
have a: "\<bar>a\<bar> * \<bar>b\<bar> = ?x"
by (simp only: abs_prts[of a] abs_prts[of b] algebra_simps)
- {
- fix u v :: 'a
- have bh: "u = a \<Longrightarrow> v = b \<Longrightarrow>
- u * v = pprt a * pprt b + pprt a * nprt b +
- nprt a * pprt b + nprt a * nprt b"
- apply (subst prts[of u], subst prts[of v])
- apply (simp add: algebra_simps)
- done
- }
+ have bh: "u = a \<Longrightarrow> v = b \<Longrightarrow>
+ u * v = pprt a * pprt b + pprt a * nprt b +
+ nprt a * pprt b + nprt a * nprt b" for u v :: 'a
+ apply (subst prts[of u], subst prts[of v])
+ apply (simp add: algebra_simps)
+ done
note b = this[OF refl[of a] refl[of b]]
have xy: "- ?x \<le> ?y"
apply simp
@@ -552,9 +542,7 @@
pprt a2 * pprt b2 + pprt a1 * nprt b2 + nprt a2 * pprt b1 + nprt a1 * nprt b1"
proof -
have "a * b = (pprt a + nprt a) * (pprt b + nprt b)"
- apply (subst prts[symmetric])+
- apply simp
- done
+ by (subst prts[symmetric])+ simp
then have "a * b = pprt a * pprt b + pprt a * nprt b + nprt a * pprt b + nprt a * nprt b"
by (simp add: algebra_simps)
moreover have "pprt a * pprt b \<le> pprt a2 * pprt b2"
@@ -587,9 +575,7 @@
by simp
qed
ultimately show ?thesis
- apply -
- apply (rule add_mono | simp)+
- done
+ by - (rule add_mono | simp)+
qed
lemma mult_ge_prts:
@@ -607,7 +593,8 @@
by auto
from mult_le_prts[of "- a2" "- a" "- a1" "b1" b "b2",
OF a1 a2 assms(3) assms(4), simplified nprt_neg pprt_neg]
- have le: "- (a * b) \<le> - nprt a1 * pprt b2 + - nprt a2 * nprt b2 +
+ have le: "- (a * b) \<le>
+ - nprt a1 * pprt b2 + - nprt a2 * nprt b2 +
- pprt a1 * pprt b1 + - pprt a2 * nprt b1"
by simp
then have "- (- nprt a1 * pprt b2 + - nprt a2 * nprt b2 +
--- a/src/HOL/Library/Old_SMT/old_smt_solver.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/HOL/Library/Old_SMT/old_smt_solver.ML Thu Jul 09 17:36:04 2015 +0200
@@ -281,7 +281,7 @@
|> Config.put Old_SMT_Config.oracle false
|> Config.put Old_SMT_Config.filter_only_facts true
- val ({context=ctxt', prems, concl, ...}, _) = Subgoal.focus ctxt i goal
+ val ({context=ctxt', prems, concl, ...}, _) = Subgoal.focus ctxt i NONE goal
fun negate ct = Thm.dest_comb ct ||> Thm.apply cnot |-> Thm.apply
val cprop =
(case try negate (Thm.rhs_of (Old_SMT_Normalize.atomize_conv ctxt' concl)) of
--- a/src/HOL/Library/old_recdef.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/HOL/Library/old_recdef.ML Thu Jul 09 17:36:04 2015 +0200
@@ -219,13 +219,7 @@
rows: int list,
TCs: term list list,
full_pats_TCs: (term * term list) list}
- val wfrec_eqns: theory -> xstring -> thm list -> term list ->
- {WFR: term,
- SV: term list,
- proto_def: term,
- extracta: (thm * term list) list,
- pats: pattern list}
- val mk_induction: theory ->
+ val mk_induction: Proof.context ->
{fconst: term, R: term, SV: term list, pat_TCs_list: (term * term list) list} -> thm
val postprocess: Proof.context -> bool ->
{wf_tac: tactic, terminator: tactic, simplifier: cterm -> thm} ->
@@ -2134,60 +2128,6 @@
end;
-(*---------------------------------------------------------------------------
- * Perform the extraction without making the definition. Definition and
- * extraction commute for the non-nested case. (Deferred recdefs)
- *
- * The purpose of wfrec_eqns is merely to instantiate the recursion theorem
- * and extract termination conditions: no definition is made.
- *---------------------------------------------------------------------------*)
-
-fun wfrec_eqns thy fid tflCongs eqns =
- let val ctxt = Proof_Context.init_global thy
- val {lhs,rhs} = USyntax.dest_eq (hd eqns)
- val (f,args) = USyntax.strip_comb lhs
- val (fname,_) = dest_atom f
- val (SV,_) = front_last args (* SV = schematic variables *)
- val g = list_comb(f,SV)
- val h = Free(fname,type_of g)
- val eqns1 = map (subst_free[(g,h)]) eqns
- val {functional as Abs(x, Ty, _), pats} = mk_functional thy eqns1
- val given_pats = givens pats
- (* val f = Free(x,Ty) *)
- val Type("fun", [f_dty, f_rty]) = Ty
- val _ = if x<>fid then
- raise TFL_ERR "wfrec_eqns"
- ("Expected a definition of " ^
- quote fid ^ " but found one of " ^
- quote x)
- else ()
- val (case_rewrites,context_congs) = extraction_thms thy
- val tych = Thry.typecheck thy
- val WFREC_THM0 = Rules.ISPEC (tych functional) @{thm tfl_wfrec}
- val Const(@{const_name All},_) $ Abs(Rname,Rtype,_) = concl WFREC_THM0
- val R = Free (singleton (Name.variant_list (List.foldr Misc_Legacy.add_term_names [] eqns)) Rname,
- Rtype)
- val WFREC_THM = Rules.ISPECL [tych R, tych g] WFREC_THM0
- val ([proto_def, WFR],_) = USyntax.strip_imp(concl WFREC_THM)
- val _ =
- if !trace then
- writeln ("ORIGINAL PROTO_DEF: " ^
- Syntax.string_of_term_global thy proto_def)
- else ()
- val R1 = USyntax.rand WFR
- val corollary' = Rules.UNDISCH (Rules.UNDISCH WFREC_THM)
- val corollaries = map (fn pat => Rules.SPEC (tych pat) corollary') given_pats
- val corollaries' = map (rewrite_rule ctxt case_rewrites) corollaries
- val extract =
- Rules.CONTEXT_REWRITE_RULE ctxt (f, R1::SV, @{thm cut_apply}, tflCongs @ context_congs)
- in {proto_def = proto_def,
- SV=SV,
- WFR=WFR,
- pats=pats,
- extracta = map extract corollaries'}
- end;
-
-
(*----------------------------------------------------------------------------
*
* INDUCTION THEOREM
@@ -2238,9 +2178,9 @@
*
*---------------------------------------------------------------------------*)
-fun mk_case ty_info usednames thy =
+fun mk_case ctxt ty_info usednames =
let
- val ctxt = Proof_Context.init_global thy
+ val thy = Proof_Context.theory_of ctxt
val divide = ipartition (gvvariant usednames)
val tych = Thry.typecheck thy
fun tych_binding(x,y) = (tych x, tych y)
@@ -2290,8 +2230,8 @@
end;
-fun complete_cases thy =
- let val ctxt = Proof_Context.init_global thy
+fun complete_cases ctxt =
+ let val thy = Proof_Context.theory_of ctxt
val tych = Thry.typecheck thy
val ty_info = Thry.induct_info thy
in fn pats =>
@@ -2310,8 +2250,8 @@
Rules.GEN ctxt (tych a)
(Rules.RIGHT_ASSOC ctxt
(Rules.CHOOSE ctxt (tych v, ex_th0)
- (mk_case ty_info (vname::aname::names)
- thy {path=[v], rows=rows})))
+ (mk_case ctxt ty_info (vname::aname::names)
+ {path=[v], rows=rows})))
end end;
@@ -2414,12 +2354,13 @@
* recursion induction (Rinduct) by proving the antecedent of Sinduct from
* the antecedent of Rinduct.
*---------------------------------------------------------------------------*)
-fun mk_induction thy {fconst, R, SV, pat_TCs_list} =
-let val ctxt = Proof_Context.init_global thy
+fun mk_induction ctxt {fconst, R, SV, pat_TCs_list} =
+let
+ val thy = Proof_Context.theory_of ctxt
val tych = Thry.typecheck thy
val Sinduction = Rules.UNDISCH (Rules.ISPEC (tych R) @{thm tfl_wf_induct})
val (pats,TCsl) = ListPair.unzip pat_TCs_list
- val case_thm = complete_cases thy pats
+ val case_thm = complete_cases ctxt pats
val domain = (type_of o hd) pats
val Pname = singleton (Name.variant_list (List.foldr (Library.foldr Misc_Legacy.add_term_names)
[] (pats::TCsl))) "P"
@@ -2454,7 +2395,6 @@
-
(*---------------------------------------------------------------------------
*
* POST PROCESSING
@@ -2661,7 +2601,7 @@
let
val _ = writeln "Proving induction theorem ..."
val ind =
- Prim.mk_induction (Proof_Context.theory_of ctxt)
+ Prim.mk_induction ctxt
{fconst=f, R=R, SV=[], pat_TCs_list=full_pats_TCs}
val _ = writeln "Postprocessing ...";
val {rules, induction, nested_tcs} =
--- a/src/HOL/Tools/Function/function_context_tree.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/HOL/Tools/Function/function_context_tree.ML Thu Jul 09 17:36:04 2015 +0200
@@ -103,7 +103,7 @@
(* Called on the INSTANTIATED branches of the congruence rule *)
fun mk_branch ctxt t =
let
- val ((params, impl), ctxt') = Variable.focus t ctxt
+ val ((params, impl), ctxt') = Variable.focus NONE t ctxt
val (assms, concl) = Logic.strip_horn impl
in
(ctxt', map #2 params, assms, rhs_of concl)
--- a/src/HOL/Tools/Function/induction_schema.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/HOL/Tools/Function/induction_schema.ML Thu Jul 09 17:36:04 2015 +0200
@@ -53,7 +53,7 @@
fun dest_hhf ctxt t =
let
- val ((params, imp), ctxt') = Variable.focus t ctxt
+ val ((params, imp), ctxt') = Variable.focus NONE t ctxt
in
(ctxt', map #2 params, Logic.strip_imp_prems imp, Logic.strip_imp_concl imp)
end
--- a/src/HOL/Tools/Function/partial_function.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/HOL/Tools/Function/partial_function.ML Thu Jul 09 17:36:04 2015 +0200
@@ -230,7 +230,7 @@
val Setup_Data {fixp, mono, fixp_eq, fixp_induct, fixp_induct_user} = setup_data;
val ((fixes, [(eq_abinding, eqn)]), _) = prep fixes_raw [eqn_raw] lthy;
- val ((_, plain_eqn), args_ctxt) = Variable.focus eqn lthy;
+ val ((_, plain_eqn), args_ctxt) = Variable.focus NONE eqn lthy;
val ((f_binding, fT), mixfix) = the_single fixes;
val fname = Binding.name_of f_binding;
--- a/src/HOL/Tools/Meson/meson.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/HOL/Tools/Meson/meson.ML Thu Jul 09 17:36:04 2015 +0200
@@ -709,7 +709,7 @@
resolve_tac ctxt @{thms ccontr} 1,
preskolem_tac,
Subgoal.FOCUS (fn {context = ctxt', prems = negs, ...} =>
- EVERY1 [skolemize_prems_tac ctxt negs,
+ EVERY1 [skolemize_prems_tac ctxt' negs,
Subgoal.FOCUS (cltac o mkcl o #prems) ctxt']) ctxt 1]) i st
handle THM _ => no_tac st; (*probably from make_meta_clause, not first-order*)
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML Thu Jul 09 17:36:04 2015 +0200
@@ -947,7 +947,7 @@
val U = TFree ("'u", @{sort type})
val y = Free (yname, U)
val f' = absdummy (U --> T') (Bound 0 $ y)
- val th' = Thm.certify_instantiate ctxt
+ val th' = Thm.certify_instantiate ctxt'
([(dest_TVar uninst_T, U --> T'), (dest_TVar uninst_T', T')],
[((fst (dest_Var f), (U --> T') --> T'), f')]) th
val [th'] = Variable.export (Variable.declare_thm th' ctxt') ctxt [th']
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_proof.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_proof.ML Thu Jul 09 17:36:04 2015 +0200
@@ -145,20 +145,20 @@
THEN trace_tac ctxt options "after prove_match:"
THEN (DETERM (TRY
(rtac eval_if_P 1
- THEN (SUBPROOF (fn {prems, ...} =>
+ THEN (SUBPROOF (fn {prems, context = ctxt', ...} =>
(REPEAT_DETERM (rtac @{thm conjI} 1
- THEN (SOLVED (asm_simp_tac (put_simpset HOL_basic_ss' ctxt) 1))))
- THEN trace_tac ctxt options "if condition to be solved:"
- THEN asm_simp_tac (put_simpset HOL_basic_ss' ctxt) 1
+ THEN (SOLVED (asm_simp_tac (put_simpset HOL_basic_ss' ctxt') 1))))
+ THEN trace_tac ctxt' options "if condition to be solved:"
+ THEN asm_simp_tac (put_simpset HOL_basic_ss' ctxt') 1
THEN TRY (
let
val prems' = maps dest_conjunct_prem (take nargs prems)
in
- rewrite_goal_tac ctxt
+ rewrite_goal_tac ctxt'
(map (fn th => th RS @{thm sym} RS @{thm eq_reflection}) prems') 1
end
THEN REPEAT_DETERM (rtac @{thm refl} 1))
- THEN trace_tac ctxt options "after if simp; in SUBPROOF") ctxt 1))))
+ THEN trace_tac ctxt' options "after if simp; in SUBPROOF") ctxt 1))))
THEN trace_tac ctxt options "after if simplification"
end;
--- a/src/HOL/Tools/SMT/smt_solver.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/HOL/Tools/SMT/smt_solver.ML Thu Jul 09 17:36:04 2015 +0200
@@ -252,7 +252,7 @@
let
val ctxt = ctxt0 |> Config.put SMT_Config.timeout (Time.toReal time_limit)
- val ({context = ctxt, prems, concl, ...}, _) = Subgoal.focus ctxt i goal
+ val ({context = ctxt, prems, concl, ...}, _) = Subgoal.focus ctxt i NONE goal
fun negate ct = Thm.dest_comb ct ||> Thm.apply @{cterm Not} |-> Thm.apply
val cprop =
(case try negate (Thm.rhs_of (SMT_Normalize.atomize_conv ctxt concl)) of
@@ -295,7 +295,7 @@
fun tac prove ctxt rules =
CONVERSION (SMT_Normalize.atomize_conv ctxt)
THEN' rtac @{thm ccontr}
- THEN' SUBPROOF (fn {context = ctxt, prems, ...} => resolve (prove ctxt (rules @ prems))) ctxt
+ THEN' SUBPROOF (fn {context = ctxt', prems, ...} => resolve (prove ctxt' (rules @ prems))) ctxt
in
val smt_tac = tac safe_solve
--- a/src/HOL/Tools/cnf.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/HOL/Tools/cnf.ML Thu Jul 09 17:36:04 2015 +0200
@@ -539,9 +539,9 @@
fun cnf_rewrite_tac ctxt i =
(* cut the CNF formulas as new premises *)
- Subgoal.FOCUS (fn {prems, ...} =>
+ Subgoal.FOCUS (fn {prems, context = ctxt', ...} =>
let
- val cnf_thms = map (make_cnf_thm ctxt o HOLogic.dest_Trueprop o Thm.prop_of) prems
+ val cnf_thms = map (make_cnf_thm ctxt' o HOLogic.dest_Trueprop o Thm.prop_of) prems
val cut_thms = map (fn (th, pr) => cnftac_eq_imp OF [th, pr]) (cnf_thms ~~ prems)
in
cut_facts_tac cut_thms 1
@@ -561,9 +561,9 @@
fun cnfx_rewrite_tac ctxt i =
(* cut the CNF formulas as new premises *)
- Subgoal.FOCUS (fn {prems, ...} =>
+ Subgoal.FOCUS (fn {prems, context = ctxt', ...} =>
let
- val cnfx_thms = map (make_cnfx_thm ctxt o HOLogic.dest_Trueprop o Thm.prop_of) prems
+ val cnfx_thms = map (make_cnfx_thm ctxt' o HOLogic.dest_Trueprop o Thm.prop_of) prems
val cut_thms = map (fn (th, pr) => cnftac_eq_imp OF [th, pr]) (cnfx_thms ~~ prems)
in
cut_facts_tac cut_thms 1
--- a/src/HOL/Tools/reification.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/HOL/Tools/reification.ML Thu Jul 09 17:36:04 2015 +0200
@@ -23,11 +23,12 @@
val pure_subst = @{lemma "x == y ==> PROP P y ==> PROP P x" by simp}
-fun lift_conv ctxt conv some_t = Subgoal.FOCUS (fn { concl, ... } =>
+fun lift_conv ctxt conv some_t = Subgoal.FOCUS (fn {context = ctxt', concl, ...} =>
let
- val ct = case some_t
- of NONE => Thm.dest_arg concl
- | SOME t => Thm.cterm_of ctxt t
+ val ct =
+ (case some_t of
+ NONE => Thm.dest_arg concl
+ | SOME t => Thm.cterm_of ctxt' t)
val thm = conv ct;
in
if Thm.is_reflexive thm then no_tac
--- a/src/HOL/Tools/set_comprehension_pointfree.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/HOL/Tools/set_comprehension_pointfree.ML Thu Jul 09 17:36:04 2015 +0200
@@ -441,9 +441,8 @@
THEN' REPEAT_DETERM o eresolve_tac ctxt @{thms exE}
THEN' eresolve_tac ctxt @{thms conjE}
THEN' REPEAT_DETERM o eresolve_tac ctxt @{thms Pair_inject}
- THEN' Subgoal.FOCUS (fn {prems, ...} =>
- (* FIXME inner context!? *)
- simp_tac (put_simpset HOL_basic_ss ctxt addsimps (filter is_eq prems)) 1) ctxt
+ THEN' Subgoal.FOCUS (fn {prems, context = ctxt', ...} =>
+ simp_tac (put_simpset HOL_basic_ss ctxt' addsimps (filter is_eq prems)) 1) ctxt
THEN' TRY o assume_tac ctxt
in
case try mk_term (Thm.term_of ct) of
--- a/src/HOL/ex/Meson_Test.thy Thu Jul 09 15:52:11 2015 +0200
+++ b/src/HOL/ex/Meson_Test.thy Thu Jul 09 17:36:04 2015 +0200
@@ -32,7 +32,7 @@
apply (tactic {* cut_tac xsko25 1 THEN REPEAT (etac exE 1) *})
ML_val {*
val ctxt = @{context};
- val [_, sko25] = #prems (#1 (Subgoal.focus ctxt 1 (#goal @{Isar.goal})));
+ val [_, sko25] = #prems (#1 (Subgoal.focus ctxt 1 NONE (#goal @{Isar.goal})));
val clauses25 = Meson.make_clauses ctxt [sko25]; (*7 clauses*)
val horns25 = Meson.make_horns clauses25; (*16 Horn clauses*)
val go25 :: _ = Meson.gocls clauses25;
@@ -56,7 +56,7 @@
apply (tactic {* cut_tac xsko26 1 THEN REPEAT (etac exE 1) *})
ML_val {*
val ctxt = @{context};
- val [_, sko26] = #prems (#1 (Subgoal.focus ctxt 1 (#goal @{Isar.goal})));
+ val [_, sko26] = #prems (#1 (Subgoal.focus ctxt 1 NONE (#goal @{Isar.goal})));
val clauses26 = Meson.make_clauses ctxt [sko26];
val _ = @{assert} (length clauses26 = 9);
val horns26 = Meson.make_horns clauses26;
@@ -83,7 +83,7 @@
apply (tactic {* cut_tac xsko43 1 THEN REPEAT (etac exE 1) *})
ML_val {*
val ctxt = @{context};
- val [_, sko43] = #prems (#1 (Subgoal.focus ctxt 1 (#goal @{Isar.goal})));
+ val [_, sko43] = #prems (#1 (Subgoal.focus ctxt 1 NONE (#goal @{Isar.goal})));
val clauses43 = Meson.make_clauses ctxt [sko43];
val _ = @{assert} (length clauses43 = 6);
val horns43 = Meson.make_horns clauses43;
--- a/src/Pure/Isar/element.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/Pure/Isar/element.ML Thu Jul 09 17:36:04 2015 +0200
@@ -215,7 +215,7 @@
fun obtain prop ctxt =
let
- val ((ps, prop'), ctxt') = Variable.focus prop ctxt;
+ val ((ps, prop'), ctxt') = Variable.focus NONE prop ctxt;
fun fix (x, T) = (Binding.name (Variable.revert_fixed ctxt' x), SOME T, NoSyn);
val xs = map (fix o #2) ps;
val As = Logic.strip_imp_prems prop';
--- a/src/Pure/Isar/isar_cmd.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/Pure/Isar/isar_cmd.ML Thu Jul 09 17:36:04 2015 +0200
@@ -182,8 +182,6 @@
val local_done_proof = Toplevel.proof Proof.local_done_proof;
val local_skip_proof = Toplevel.proof' Proof.local_skip_proof;
-val skip_local_qed = Toplevel.skip_proof (fn i => if i > 1 then i - 1 else raise Toplevel.UNDEF);
-
(* global endings *)
@@ -194,12 +192,10 @@
val global_skip_proof = Toplevel.end_proof Proof.global_skip_proof;
val global_done_proof = Toplevel.end_proof (K Proof.global_done_proof);
-val skip_global_qed = Toplevel.skip_proof_to_theory (fn n => n = 1);
-
(* common endings *)
-fun qed m = local_qed m o global_qed m o skip_local_qed o skip_global_qed;
+fun qed m = local_qed m o global_qed m;
fun terminal_proof m = local_terminal_proof m o global_terminal_proof m;
val default_proof = local_default_proof o global_default_proof;
val immediate_proof = local_immediate_proof o global_immediate_proof;
--- a/src/Pure/Isar/isar_syn.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/Pure/Isar/isar_syn.ML Thu Jul 09 17:36:04 2015 +0200
@@ -674,9 +674,7 @@
val _ =
Outer_Syntax.command @{command_keyword proof} "backward proof step"
(Scan.option Method.parse >> (fn m =>
- (Option.map Method.report m;
- Toplevel.actual_proof (Proof_Node.applys (Proof.proof_results m)) o
- Toplevel.skip_proof (fn i => i + 1))));
+ (Option.map Method.report m; Toplevel.proofs (Proof.proof_results m))));
(* subgoal focus *)
@@ -715,7 +713,7 @@
Outer_Syntax.command @{command_keyword back} "explicit backtracking of proof command"
(Scan.succeed
(Toplevel.actual_proof (fn prf => (report_back (); Proof_Node.back prf)) o
- Toplevel.skip_proof (fn h => (report_back (); h))));
+ Toplevel.skip_proof report_back));
--- a/src/Pure/Isar/keyword.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/Pure/Isar/keyword.ML Thu Jul 09 17:36:04 2015 +0200
@@ -22,6 +22,7 @@
val qed_global: string
val prf_goal: string
val prf_block: string
+ val next_block: string
val prf_open: string
val prf_close: string
val prf_chain: string
@@ -62,6 +63,8 @@
val is_proof_goal: keywords -> string -> bool
val is_qed: keywords -> string -> bool
val is_qed_global: keywords -> string -> bool
+ val is_proof_open: keywords -> string -> bool
+ val is_proof_close: keywords -> string -> bool
val is_proof_asm: keywords -> string -> bool
val is_improper: keywords -> string -> bool
val is_printed: keywords -> string -> bool
@@ -90,6 +93,7 @@
val qed_global = "qed_global";
val prf_goal = "prf_goal";
val prf_block = "prf_block";
+val next_block = "next_block";
val prf_open = "prf_open";
val prf_close = "prf_close";
val prf_chain = "prf_chain";
@@ -102,9 +106,9 @@
val kinds =
[diag, document_heading, document_body, document_raw, thy_begin, thy_end, thy_load, thy_decl,
- thy_decl_block, thy_goal, qed, qed_script, qed_block, qed_global, prf_goal, prf_block, prf_open,
- prf_close, prf_chain, prf_decl, prf_asm, prf_asm_goal, prf_script, prf_script_goal,
- prf_script_asm_goal];
+ thy_decl_block, thy_goal, qed, qed_script, qed_block, qed_global, prf_goal, prf_block,
+ next_block, prf_open, prf_close, prf_chain, prf_decl, prf_asm, prf_asm_goal, prf_script,
+ prf_script_goal, prf_script_asm_goal];
(* specifications *)
@@ -241,13 +245,13 @@
[thy_load, thy_decl, thy_decl_block, thy_goal];
val is_proof = command_category
- [qed, qed_script, qed_block, qed_global, prf_goal, prf_block, prf_open, prf_close,
- prf_chain, prf_decl, prf_asm, prf_asm_goal, prf_script, prf_script_goal,
+ [qed, qed_script, qed_block, qed_global, prf_goal, prf_block, next_block, prf_open,
+ prf_close, prf_chain, prf_decl, prf_asm, prf_asm_goal, prf_script, prf_script_goal,
prf_script_asm_goal];
val is_proof_body = command_category
- [diag, document_heading, document_body, document_raw, prf_block, prf_open, prf_close,
- prf_chain, prf_decl, prf_asm, prf_asm_goal, prf_script, prf_script_goal,
+ [diag, document_heading, document_body, document_raw, prf_block, next_block, prf_open,
+ prf_close, prf_chain, prf_decl, prf_asm, prf_asm_goal, prf_script, prf_script_goal,
prf_script_asm_goal];
val is_theory_goal = command_category [thy_goal];
@@ -255,6 +259,10 @@
val is_qed = command_category [qed, qed_script, qed_block];
val is_qed_global = command_category [qed_global];
+val is_proof_open =
+ command_category [prf_goal, prf_asm_goal, prf_script_goal, prf_script_asm_goal, prf_open];
+val is_proof_close = command_category [qed, qed_script, qed_block, prf_close];
+
val is_proof_asm = command_category [prf_asm, prf_asm_goal];
val is_improper = command_category [qed_script, prf_script, prf_script_goal, prf_script_asm_goal];
--- a/src/Pure/Isar/keyword.scala Thu Jul 09 15:52:11 2015 +0200
+++ b/src/Pure/Isar/keyword.scala Thu Jul 09 17:36:04 2015 +0200
@@ -29,6 +29,7 @@
val QED_GLOBAL = "qed_global"
val PRF_GOAL = "prf_goal"
val PRF_BLOCK = "prf_block"
+ val NEXT_BLOCK = "next_block"
val PRF_OPEN = "prf_open"
val PRF_CLOSE = "prf_close"
val PRF_CHAIN = "prf_chain"
@@ -63,13 +64,13 @@
val theory_body = Set(THY_LOAD, THY_DECL, THY_DECL_BLOCK, THY_GOAL)
val proof =
- Set(QED, QED_SCRIPT, QED_BLOCK, QED_GLOBAL, PRF_GOAL, PRF_BLOCK, PRF_OPEN, PRF_CLOSE,
- PRF_CHAIN, PRF_DECL, PRF_ASM, PRF_ASM_GOAL, PRF_SCRIPT, PRF_SCRIPT_GOAL,
+ Set(QED, QED_SCRIPT, QED_BLOCK, QED_GLOBAL, PRF_GOAL, PRF_BLOCK, NEXT_BLOCK, PRF_OPEN,
+ PRF_CLOSE, PRF_CHAIN, PRF_DECL, PRF_ASM, PRF_ASM_GOAL, PRF_SCRIPT, PRF_SCRIPT_GOAL,
PRF_SCRIPT_ASM_GOAL)
val proof_body =
- Set(DIAG, DOCUMENT_HEADING, DOCUMENT_BODY, DOCUMENT_RAW, PRF_BLOCK, PRF_OPEN, PRF_CLOSE,
- PRF_CHAIN, PRF_DECL, PRF_ASM, PRF_ASM_GOAL, PRF_SCRIPT, PRF_SCRIPT_GOAL,
+ Set(DIAG, DOCUMENT_HEADING, DOCUMENT_BODY, DOCUMENT_RAW, PRF_BLOCK, NEXT_BLOCK, PRF_OPEN,
+ PRF_CLOSE, PRF_CHAIN, PRF_DECL, PRF_ASM, PRF_ASM_GOAL, PRF_SCRIPT, PRF_SCRIPT_GOAL,
PRF_SCRIPT_ASM_GOAL)
val theory_goal = Set(THY_GOAL)
@@ -77,6 +78,9 @@
val qed = Set(QED, QED_SCRIPT, QED_BLOCK)
val qed_global = Set(QED_GLOBAL)
+ val proof_open = proof_goal + PRF_OPEN
+ val proof_close = qed + PRF_CLOSE
+
/** keyword tables **/
--- a/src/Pure/Isar/obtain.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/Pure/Isar/obtain.ML Thu Jul 09 17:36:04 2015 +0200
@@ -275,7 +275,7 @@
val ((_, [rule']), ctxt') = Variable.import false [closed_rule] ctxt;
val obtain_rule =
Thm.forall_elim (Thm.cterm_of ctxt (Logic.varify_global (Free thesis_var))) rule';
- val ((params, stmt), fix_ctxt) = Variable.focus_cterm (Thm.cprem_of obtain_rule 1) ctxt';
+ val ((params, stmt), fix_ctxt) = Variable.focus_cterm NONE (Thm.cprem_of obtain_rule 1) ctxt';
val (prems, ctxt'') =
Assumption.add_assms (obtain_export fix_ctxt obtain_rule (map #2 params))
(Drule.strip_imp_prems stmt) fix_ctxt;
--- a/src/Pure/Isar/outer_syntax.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/Pure/Isar/outer_syntax.ML Thu Jul 09 17:36:04 2015 +0200
@@ -45,7 +45,7 @@
datatype command_parser =
Parser of (Toplevel.transition -> Toplevel.transition) parser |
- Limited_Parser of
+ Restricted_Parser of
(bool * Position.T) option -> (Toplevel.transition -> Toplevel.transition) parser;
datatype command = Command of
@@ -147,8 +147,8 @@
fun local_theory_command trans command_keyword comment parse =
raw_command command_keyword comment
- (Limited_Parser (fn limited =>
- Parse.opt_target -- parse >> (fn (target, f) => trans limited target f)));
+ (Restricted_Parser (fn restricted =>
+ Parse.opt_target -- parse >> (fn (target, f) => trans restricted target f)));
val local_theory' = local_theory_command Toplevel.local_theory';
val local_theory = local_theory_command Toplevel.local_theory;
@@ -172,15 +172,21 @@
fun parse_command thy =
Scan.ahead (before_command |-- Parse.position Parse.command_) :|-- (fn (name, pos) =>
let
+ val keywords = Thy_Header.get_keywords thy;
val command_tags = Parse.command_ -- Parse.tags;
- val tr0 = Toplevel.empty |> Toplevel.name name |> Toplevel.position pos;
+ val tr0 =
+ Toplevel.empty
+ |> Toplevel.name name
+ |> Toplevel.position pos
+ |> Keyword.is_proof_open keywords name ? Toplevel.skip_proof_open
+ |> Keyword.is_proof_close keywords name ? Toplevel.skip_proof_close;
in
(case lookup_commands thy name of
SOME (Command {command_parser = Parser parse, ...}) =>
Parse.!!! (command_tags |-- parse) >> (fn f => f tr0)
- | SOME (Command {command_parser = Limited_Parser parse, ...}) =>
- before_command :|-- (fn limited =>
- Parse.!!! (command_tags |-- parse limited)) >> (fn f => f tr0)
+ | SOME (Command {command_parser = Restricted_Parser parse, ...}) =>
+ before_command :|-- (fn restricted =>
+ Parse.!!! (command_tags |-- parse restricted)) >> (fn f => f tr0)
| NONE =>
Scan.fail_with (fn _ => fn _ =>
let
--- a/src/Pure/Isar/outer_syntax.scala Thu Jul 09 15:52:11 2015 +0200
+++ b/src/Pure/Isar/outer_syntax.scala Thu Jul 09 17:36:04 2015 +0200
@@ -164,16 +164,13 @@
((structure.span_depth, structure.after_span_depth) /: tokens) {
case ((x, y), tok) =>
if (tok.is_command) {
- if (tok.is_command_kind(keywords, Keyword.theory_goal))
- (2, 1)
- else if (tok.is_command_kind(keywords, Keyword.theory))
- (1, 0)
- else if (tok.is_command_kind(keywords, Keyword.proof_goal) || tok.is_begin_block)
- (y + 2, y + 1)
- else if (tok.is_command_kind(keywords, Keyword.qed) || tok.is_end_block)
- (y + 1, y - 1)
- else if (tok.is_command_kind(keywords, Keyword.qed_global))
- (1, 0)
+ if (tok.is_command_kind(keywords, Keyword.theory_goal)) (2, 1)
+ else if (tok.is_command_kind(keywords, Keyword.theory)) (1, 0)
+ else if (tok.is_command_kind(keywords, Keyword.PRF_BLOCK == _)) (y + 2, y + 1)
+ else if (tok.is_command_kind(keywords, Keyword.QED_BLOCK == _)) (y + 1, y - 3)
+ else if (tok.is_command_kind(keywords, Keyword.proof_open)) (y + 2, y + 1)
+ else if (tok.is_command_kind(keywords, Keyword.proof_close)) (y + 1, y - 1)
+ else if (tok.is_command_kind(keywords, Keyword.qed_global)) (1, 0)
else (x, y)
}
else (x, y)
--- a/src/Pure/Isar/subgoal.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/Pure/Isar/subgoal.ML Thu Jul 09 17:36:04 2015 +0200
@@ -15,10 +15,10 @@
type focus =
{context: Proof.context, params: (string * cterm) list, prems: thm list, asms: cterm list,
concl: cterm, schematics: ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list}
- val focus_params: Proof.context -> int -> thm -> focus * thm
- val focus_params_fixed: Proof.context -> int -> thm -> focus * thm
- val focus_prems: Proof.context -> int -> thm -> focus * thm
- val focus: Proof.context -> int -> thm -> focus * thm
+ val focus_params: Proof.context -> int -> binding list option -> thm -> focus * thm
+ val focus_params_fixed: Proof.context -> int -> binding list option -> thm -> focus * thm
+ val focus_prems: Proof.context -> int -> binding list option -> thm -> focus * thm
+ val focus: Proof.context -> int -> binding list option -> thm -> focus * thm
val retrofit: Proof.context -> Proof.context -> (string * cterm) list -> cterm list ->
int -> thm -> thm -> thm Seq.seq
val FOCUS_PARAMS: (focus -> tactic) -> Proof.context -> int -> tactic
@@ -41,13 +41,13 @@
{context: Proof.context, params: (string * cterm) list, prems: thm list, asms: cterm list,
concl: cterm, schematics: ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list};
-fun gen_focus (do_prems, do_concl) ctxt i raw_st =
+fun gen_focus (do_prems, do_concl) ctxt i bindings raw_st =
let
val st = raw_st
|> Thm.transfer (Proof_Context.theory_of ctxt)
|> Raw_Simplifier.norm_hhf_protect ctxt;
val ((schematic_types, [st']), ctxt1) = Variable.importT [st] ctxt;
- val ((params, goal), ctxt2) = Variable.focus_cterm (Thm.cprem_of st' i) ctxt1;
+ val ((params, goal), ctxt2) = Variable.focus_cterm bindings (Thm.cprem_of st' i) ctxt1;
val (asms, concl) =
if do_prems then (Drule.strip_imp_prems goal, Drule.strip_imp_concl goal)
@@ -150,7 +150,7 @@
fun GEN_FOCUS flags tac ctxt i st =
if Thm.nprems_of st < i then Seq.empty
else
- let val (args as {context = ctxt', params, asms, ...}, st') = gen_focus flags ctxt i st;
+ let val (args as {context = ctxt', params, asms, ...}, st') = gen_focus flags ctxt i NONE st;
in Seq.lifts (retrofit ctxt' ctxt params asms i) (tac args st') st end;
val FOCUS_PARAMS = GEN_FOCUS (false, false);
@@ -165,10 +165,10 @@
local
-fun rename_params ctxt (param_suffix, raw_param_specs) st =
+fun param_bindings ctxt (param_suffix, raw_param_specs) st =
let
val _ = if Thm.no_prems st then error "No subgoals!" else ();
- val (subgoal, goal') = Logic.dest_implies (Thm.prop_of st);
+ val subgoal = #1 (Logic.dest_implies (Thm.prop_of st));
val subgoal_params =
map (apfst (Name.internal o Name.clean)) (Term.strip_all_vars subgoal)
|> Term.variant_frees subgoal |> map #1;
@@ -184,19 +184,16 @@
raw_param_specs |> map
(fn (NONE, _) => NONE
| (SOME x, pos) =>
- let val b = #1 (#1 (Proof_Context.cert_var (Binding.make (x, pos), NONE, NoSyn) ctxt))
- in SOME (Variable.check_name b) end)
+ let
+ val b = #1 (#1 (Proof_Context.cert_var (Binding.make (x, pos), NONE, NoSyn) ctxt));
+ val _ = Variable.check_name b;
+ in SOME b end)
|> param_suffix ? append (replicate (n - m) NONE);
- fun rename_list (opt_x :: xs) (y :: ys) = (the_default y opt_x :: rename_list xs ys)
- | rename_list _ ys = ys;
-
- fun rename_prop (x :: xs) ((c as Const ("Pure.all", _)) $ Abs (_, T, t)) =
- (c $ Abs (x, T, rename_prop xs t))
- | rename_prop [] t = t;
-
- val subgoal' = rename_prop (rename_list param_specs subgoal_params) subgoal;
- in Thm.renamed_prop (Logic.mk_implies (subgoal', goal')) st end;
+ fun bindings (SOME x :: xs) (_ :: ys) = x :: bindings xs ys
+ | bindings (NONE :: xs) (y :: ys) = Binding.name y :: bindings xs ys
+ | bindings _ ys = map Binding.name ys;
+ in bindings param_specs subgoal_params end;
fun gen_subgoal prep_atts raw_result_binding raw_prems_binding param_specs state =
let
@@ -212,8 +209,8 @@
| NONE => ((Binding.empty, []), false));
val (subgoal_focus, _) =
- rename_params ctxt param_specs st
- |> (if do_prems then focus else focus_params_fixed) ctxt 1;
+ (if do_prems then focus else focus_params_fixed) ctxt
+ 1 (SOME (param_bindings ctxt param_specs st)) st;
fun after_qed (ctxt'', [[result]]) =
Proof.end_block #> (fn state' =>
--- a/src/Pure/Isar/token.scala Thu Jul 09 15:52:11 2015 +0200
+++ b/src/Pure/Isar/token.scala Thu Jul 09 17:36:04 2015 +0200
@@ -263,9 +263,6 @@
def is_command_modifier: Boolean =
is_keyword && (source == "public" || source == "private" || source == "qualified")
- def is_begin_block: Boolean = is_command && source == "{"
- def is_end_block: Boolean = is_command && source == "}"
-
def content: String =
if (kind == Token.Kind.STRING) Scan.Parsers.quoted_content("\"", source)
else if (kind == Token.Kind.ALT_STRING) Scan.Parsers.quoted_content("`", source)
--- a/src/Pure/Isar/toplevel.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/Pure/Isar/toplevel.ML Thu Jul 09 17:36:04 2015 +0200
@@ -69,8 +69,9 @@
val proofs: (Proof.state -> Proof.state Seq.result Seq.seq) -> transition -> transition
val proof: (Proof.state -> Proof.state) -> transition -> transition
val actual_proof: (Proof_Node.T -> Proof_Node.T) -> transition -> transition
- val skip_proof: (int -> int) -> transition -> transition
- val skip_proof_to_theory: (int -> bool) -> transition -> transition
+ val skip_proof: (unit -> unit) -> transition -> transition
+ val skip_proof_open: transition -> transition
+ val skip_proof_close: transition -> transition
val exec_id: Document_ID.exec -> transition -> transition
val setmp_thread_position: transition -> ('a -> 'b) -> 'a -> 'b
val add_hook: (transition -> state -> state -> unit) -> unit
@@ -510,16 +511,24 @@
val proofs = proofs' o K;
val proof = proof' o K;
+
+(* skipped proofs *)
+
fun actual_proof f = transaction (fn _ =>
(fn Proof (prf, x) => Proof (f prf, x)
| _ => raise UNDEF));
fun skip_proof f = transaction (fn _ =>
- (fn Skipped_Proof (h, x) => Skipped_Proof (f h, x)
+ (fn skip as Skipped_Proof _ => (f (); skip)
| _ => raise UNDEF));
-fun skip_proof_to_theory pred = transaction (fn _ =>
- (fn Skipped_Proof (d, (gthy, _)) => if pred d then Theory (gthy, NONE) else raise UNDEF
+val skip_proof_open = transaction (fn _ =>
+ (fn Skipped_Proof (d, x) => Skipped_Proof (d + 1, x)
+ | _ => raise UNDEF));
+
+val skip_proof_close = transaction (fn _ =>
+ (fn Skipped_Proof (0, (gthy, _)) => Theory (gthy, NONE)
+ | Skipped_Proof (d, x) => Skipped_Proof (d - 1, x)
| _ => raise UNDEF));
--- a/src/Pure/Pure.thy Thu Jul 09 15:52:11 2015 +0200
+++ b/src/Pure/Pure.thy Thu Jul 09 17:36:04 2015 +0200
@@ -62,7 +62,7 @@
and "case" :: prf_asm % "proof"
and "{" :: prf_open % "proof"
and "}" :: prf_close % "proof"
- and "next" :: prf_block % "proof"
+ and "next" :: next_block % "proof"
and "qed" :: qed_block % "proof"
and "by" ".." "." "sorry" :: "qed" % "proof"
and "done" :: "qed_script" % "proof"
--- a/src/Pure/Tools/rule_insts.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/Pure/Tools/rule_insts.ML Thu Jul 09 17:36:04 2015 +0200
@@ -218,7 +218,7 @@
val ((_, params), ctxt') = ctxt
|> Variable.declare_constraints goal
|> Variable.improper_fixes
- |> Variable.focus_params goal
+ |> Variable.focus_params NONE goal
||> Variable.restore_proper_fixes ctxt;
in (params, ctxt') end;
--- a/src/Pure/goal.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/Pure/goal.ML Thu Jul 09 17:36:04 2015 +0200
@@ -328,7 +328,7 @@
fun assume_rule_tac ctxt = norm_hhf_tac ctxt THEN' CSUBGOAL (fn (goal, i) =>
let
- val ((_, goal'), ctxt') = Variable.focus_cterm goal ctxt;
+ val ((_, goal'), ctxt') = Variable.focus_cterm NONE goal ctxt;
val goal'' = Drule.cterm_rule (singleton (Variable.export ctxt' ctxt)) goal';
val Rs = filter (non_atomic o Thm.term_of) (Drule.strip_imp_prems goal'');
val tacs = Rs |> map (fn R =>
--- a/src/Pure/variable.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/Pure/variable.ML Thu Jul 09 17:36:04 2015 +0200
@@ -78,10 +78,14 @@
((((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list) * thm list) * Proof.context
val tradeT: (Proof.context -> thm list -> thm list) -> Proof.context -> thm list -> thm list
val trade: (Proof.context -> thm list -> thm list) -> Proof.context -> thm list -> thm list
- val focus_params: term -> Proof.context -> (string list * (string * typ) list) * Proof.context
- val focus: term -> Proof.context -> ((string * (string * typ)) list * term) * Proof.context
- val focus_cterm: cterm -> Proof.context -> ((string * cterm) list * cterm) * Proof.context
- val focus_subgoal: int -> thm -> Proof.context -> ((string * cterm) list * cterm) * Proof.context
+ val focus_params: binding list option -> term -> Proof.context ->
+ (string list * (string * typ) list) * Proof.context
+ val focus: binding list option -> term -> Proof.context ->
+ ((string * (string * typ)) list * term) * Proof.context
+ val focus_cterm: binding list option -> cterm -> Proof.context ->
+ ((string * cterm) list * cterm) * Proof.context
+ val focus_subgoal: binding list option -> int -> thm -> Proof.context ->
+ ((string * cterm) list * cterm) * Proof.context
val warn_extra_tfrees: Proof.context -> Proof.context -> unit
val polymorphic_types: Proof.context -> term list -> (indexname * sort) list * term list
val polymorphic: Proof.context -> term list -> term list
@@ -616,18 +620,21 @@
(* focus on outermost parameters: !!x y z. B *)
-fun focus_params t ctxt =
+fun focus_params bindings t ctxt =
let
val (xs, Ts) =
split_list (Term.variant_frees t (Term.strip_all_vars t)); (*as they are printed :-*)
- val (xs', ctxt') = variant_fixes xs ctxt;
+ val (xs', ctxt') =
+ (case bindings of
+ SOME bs => ctxt |> set_body true |> add_fixes_binding bs ||> restore_body ctxt
+ | NONE => ctxt |> variant_fixes xs);
val ps = xs' ~~ Ts;
val ctxt'' = ctxt' |> fold (declare_constraints o Free) ps;
in ((xs, ps), ctxt'') end;
-fun focus t ctxt =
+fun focus bindings t ctxt =
let
- val ((xs, ps), ctxt') = focus_params t ctxt;
+ val ((xs, ps), ctxt') = focus_params bindings t ctxt;
val t' = Term.subst_bounds (rev (map Free ps), Term.strip_all_body t);
in (((xs ~~ ps), t'), ctxt') end;
@@ -635,20 +642,20 @@
Thm.beta_conversion false (Thm.apply (Thm.dest_arg prop) t)
|> Thm.cprop_of |> Thm.dest_arg;
-fun focus_cterm goal ctxt =
+fun focus_cterm bindings goal ctxt =
let
- val ((xs, ps), ctxt') = focus_params (Thm.term_of goal) ctxt;
+ val ((xs, ps), ctxt') = focus_params bindings (Thm.term_of goal) ctxt;
val ps' = map (Thm.cterm_of ctxt' o Free) ps;
val goal' = fold forall_elim_prop ps' goal;
in ((xs ~~ ps', goal'), ctxt') end;
-fun focus_subgoal i st =
+fun focus_subgoal bindings i st =
let
val all_vars = Thm.fold_terms Term.add_vars st [];
in
fold (unbind_term o #1) all_vars #>
fold (declare_constraints o Var) all_vars #>
- focus_cterm (Thm.cprem_of st i)
+ focus_cterm bindings (Thm.cprem_of st i)
end;
--- a/src/Tools/Code/code_preproc.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/Tools/Code/code_preproc.ML Thu Jul 09 17:36:04 2015 +0200
@@ -21,7 +21,7 @@
val sortargs: code_graph -> string -> sort list
val all: code_graph -> string list
val pretty: Proof.context -> code_graph -> Pretty.T
- val obtain: bool -> theory -> string list -> term list -> code_algebra * code_graph
+ val obtain: bool -> Proof.context -> string list -> term list -> code_algebra * code_graph
val dynamic_conv: Proof.context
-> (code_algebra -> code_graph -> term -> conv) -> conv
val dynamic_value: Proof.context -> ((term -> term) -> 'a -> 'b)
@@ -536,15 +536,15 @@
(** retrieval and evaluation interfaces **)
-fun obtain ignore_cache thy consts ts = apsnd snd
- (Wellsorted.change_yield (if ignore_cache then NONE else SOME thy)
- (extend_arities_eqngr (Proof_Context.init_global thy) consts ts));
+fun obtain ignore_cache ctxt consts ts = apsnd snd
+ (Wellsorted.change_yield (if ignore_cache then NONE else SOME (Proof_Context.theory_of ctxt))
+ (extend_arities_eqngr ctxt consts ts));
fun dynamic_evaluator eval ctxt t =
let
val consts = fold_aterms
(fn Const (c, _) => insert (op =) c | _ => I) t [];
- val (algebra, eqngr) = obtain false (Proof_Context.theory_of ctxt) consts [t];
+ val (algebra, eqngr) = obtain false ctxt consts [t];
in eval algebra eqngr t end;
fun dynamic_conv ctxt conv =
@@ -555,12 +555,12 @@
fun static_conv { ctxt, consts } conv =
let
- val (algebra, eqngr) = obtain true (Proof_Context.theory_of ctxt) consts [];
+ val (algebra, eqngr) = obtain true ctxt consts [];
in Sandwich.conversion (value_sandwich ctxt) (conv { algebra = algebra, eqngr = eqngr }) end;
fun static_value { ctxt, lift_postproc, consts } evaluator =
let
- val (algebra, eqngr) = obtain true (Proof_Context.theory_of ctxt) consts [];
+ val (algebra, eqngr) = obtain true ctxt consts [];
in Sandwich.evaluation (value_sandwich ctxt) lift_postproc (evaluator { algebra = algebra, eqngr = eqngr }) end;
--- a/src/Tools/Code/code_thingol.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/Tools/Code/code_thingol.ML Thu Jul 09 17:36:04 2015 +0200
@@ -790,7 +790,8 @@
fun generate_consts ctxt algebra eqngr =
fold_map (ensure_const ctxt algebra eqngr permissive);
in
- invoke_generation permissive thy (Code_Preproc.obtain false thy consts [])
+ invoke_generation permissive thy
+ (Code_Preproc.obtain false (Proof_Context.init_global thy) consts [])
generate_consts consts
|> snd
end;
@@ -918,7 +919,7 @@
fun code_depgr ctxt consts =
let
- val (_, eqngr) = Code_Preproc.obtain true ((Proof_Context.theory_of ctxt)) consts [];
+ val (_, eqngr) = Code_Preproc.obtain true ctxt consts [];
val all_consts = Graph.all_succs eqngr consts;
in Graph.restrict (member (op =) all_consts) eqngr end;
--- a/src/Tools/induct_tacs.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/Tools/induct_tacs.ML Thu Jul 09 17:36:04 2015 +0200
@@ -68,7 +68,7 @@
let
val goal = Thm.cprem_of st i;
val (_, goal_ctxt) = Rule_Insts.goal_context (Thm.term_of goal) ctxt
- and ((_, goal'), _) = Variable.focus_cterm goal ctxt;
+ and ((_, goal'), _) = Variable.focus_cterm NONE goal ctxt;
val (prems, concl) = Logic.strip_horn (Thm.term_of goal');
fun induct_var name =
--- a/src/ZF/Tools/induct_tacs.ML Thu Jul 09 15:52:11 2015 +0200
+++ b/src/ZF/Tools/induct_tacs.ML Thu Jul 09 17:36:04 2015 +0200
@@ -93,7 +93,7 @@
fun exhaust_induct_tac exh ctxt var fixes i state = SUBGOAL (fn _ =>
let
val thy = Proof_Context.theory_of ctxt
- val ({context = ctxt', asms, ...}, _) = Subgoal.focus ctxt i state
+ val ({context = ctxt', asms, ...}, _) = Subgoal.focus ctxt i NONE state
val tn = find_tname ctxt' var (map Thm.term_of asms)
val rule =
if exh then #exhaustion (datatype_info thy tn)