--- a/.hgignore Tue Dec 20 16:17:13 2016 +0100
+++ b/.hgignore Tue Dec 20 16:18:56 2016 +0100
@@ -20,5 +20,8 @@
^doc/.*\.pdf
^doc/.*\.ps
^src/Tools/jEdit/dist/
+^src/Tools/VSCode/out/
+^src/Tools/VSCode/extension/node_modules/
+^src/Tools/VSCode/extension/protocol.log
^Admin/jenkins/ci-extras/target/
^stats/
--- a/NEWS Tue Dec 20 16:17:13 2016 +0100
+++ b/NEWS Tue Dec 20 16:18:56 2016 +0100
@@ -3,9 +3,35 @@
(Note: Isabelle/jEdit shows a tree-view of the NEWS file in Sidekick.)
+
New in this Isabelle version
----------------------------
+*** Prover IDE -- Isabelle/Scala/jEdit ***
+
+* Command-line invocation "isabelle jedit -R -l LOGIC" opens the ROOT
+entry of the specified logic session in the editor, while its parent is
+used for formal checking.
+
+
+*** HOL ***
+
+* Swapped orientation of congruence rules mod_add_left_eq,
+mod_add_right_eq, mod_add_eq, mod_mult_left_eq, mod_mult_right_eq,
+mod_mult_eq, mod_minus_eq, mod_diff_left_eq, mod_diff_right_eq,
+mod_diff_eq. INCOMPATIBILITY.
+
+* Generalized some facts:
+ zminus_zmod ~> mod_minus_eq
+ zdiff_zmod_left ~> mod_diff_left_eq
+ zdiff_zmod_right ~> mod_diff_right_eq
+ zmod_eq_dvd_iff ~> mod_eq_dvd_iff
+INCOMPATIBILITY.
+
+* Named collection mod_simps covers various congruence rules
+concerning mod, replacing former zmod_simps.
+INCOMPATIBILITY.
+
* (Co)datatype package:
- The 'size_gen_o_map' lemma is no longer generated for datatypes
with type class annotations. As a result, the tactic that derives
@@ -13,7 +39,7 @@
* The theorem in Permutations has been renamed:
bij_swap_ompose_bij ~> bij_swap_compose_bij
-
+
New in Isabelle2016-1 (December 2016)
-------------------------------------
--- a/src/Doc/JEdit/JEdit.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/Doc/JEdit/JEdit.thy Tue Dec 20 16:18:56 2016 +0100
@@ -233,6 +233,7 @@
Options are:
-D NAME=X set JVM system property
-J OPTION add JVM runtime option
+ -R open ROOT entry of logic session and use its parent
-b build only
-d DIR include session directory
-f fresh build
@@ -256,6 +257,11 @@
The \<^verbatim>\<open>-n\<close> option bypasses the implicit build process for the selected
session image.
+ Option \<^verbatim>\<open>-R\<close> modifies the meaning of option \<^verbatim>\<open>-l\<close> as follows: the \<^verbatim>\<open>ROOT\<close>
+ entry of the specified session is opened in the editor, while its parent
+ session is used for formal checking. This facilitates maintenance of a
+ broken session, by moving the Prover IDE quickly to relevant source files.
+
The \<^verbatim>\<open>-m\<close> option specifies additional print modes for the prover process.
Note that the system option @{system_option_ref jedit_print_mode} allows to
do the same persistently (e.g.\ via the \<^emph>\<open>Plugin Options\<close> dialog of
--- a/src/HOL/Algebra/Coset.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Algebra/Coset.thy Tue Dec 20 16:18:56 2016 +0100
@@ -15,16 +15,16 @@
where "H #>\<^bsub>G\<^esub> a = (\<Union>h\<in>H. {h \<otimes>\<^bsub>G\<^esub> a})"
definition
- l_coset :: "[_, 'a, 'a set] \<Rightarrow> 'a set" (infixl "<#\<index>" 60)
- where "a <#\<^bsub>G\<^esub> H = (\<Union>h\<in>H. {a \<otimes>\<^bsub>G\<^esub> h})"
+ l_coset :: "[_, 'a, 'a set] \<Rightarrow> 'a set" (infixl "\<subset>#\<index>" 60)
+ where "a \<subset>#\<^bsub>G\<^esub> H = (\<Union>h\<in>H. {a \<otimes>\<^bsub>G\<^esub> h})"
definition
RCOSETS :: "[_, 'a set] \<Rightarrow> ('a set)set" ("rcosets\<index> _" [81] 80)
where "rcosets\<^bsub>G\<^esub> H = (\<Union>a\<in>carrier G. {H #>\<^bsub>G\<^esub> a})"
definition
- set_mult :: "[_, 'a set ,'a set] \<Rightarrow> 'a set" (infixl "<#>\<index>" 60)
- where "H <#>\<^bsub>G\<^esub> K = (\<Union>h\<in>H. \<Union>k\<in>K. {h \<otimes>\<^bsub>G\<^esub> k})"
+ set_mult :: "[_, 'a set ,'a set] \<Rightarrow> 'a set" (infixl "\<subset>#>\<index>" 60)
+ where "H \<subset>#>\<^bsub>G\<^esub> K = (\<Union>h\<in>H. \<Union>k\<in>K. {h \<otimes>\<^bsub>G\<^esub> k})"
definition
SET_INV :: "[_,'a set] \<Rightarrow> 'a set" ("set'_inv\<index> _" [81] 80)
@@ -32,7 +32,7 @@
locale normal = subgroup + group +
- assumes coset_eq: "(\<forall>x \<in> carrier G. H #> x = x <# H)"
+ assumes coset_eq: "(\<forall>x \<in> carrier G. H #> x = x \<subset># H)"
abbreviation
normal_rel :: "['a set, ('a, 'b) monoid_scheme] \<Rightarrow> bool" (infixl "\<lhd>" 60) where
@@ -287,7 +287,7 @@
lemma (in monoid) set_mult_closed:
assumes Acarr: "A \<subseteq> carrier G"
and Bcarr: "B \<subseteq> carrier G"
- shows "A <#> B \<subseteq> carrier G"
+ shows "A \<subset>#> B \<subseteq> carrier G"
apply rule apply (simp add: set_mult_def, clarsimp)
proof -
fix a b
@@ -306,7 +306,7 @@
lemma (in comm_group) mult_subgroups:
assumes subH: "subgroup H G"
and subK: "subgroup K G"
- shows "subgroup (H <#> K) G"
+ shows "subgroup (H \<subset>#> K) G"
apply (rule subgroup.intro)
apply (intro set_mult_closed subgroup.subset[OF subH] subgroup.subset[OF subK])
apply (simp add: set_mult_def) apply clarsimp defer 1
@@ -351,7 +351,7 @@
assumes "group G"
assumes carr: "x \<in> carrier G" "x' \<in> carrier G"
and xixH: "(inv x \<otimes> x') \<in> H"
- shows "x' \<in> x <# H"
+ shows "x' \<in> x \<subset># H"
proof -
interpret group G by fact
from xixH
@@ -375,7 +375,7 @@
have "x \<otimes> h = x'" by simp
from this[symmetric] and hH
- show "x' \<in> x <# H"
+ show "x' \<in> x \<subset># H"
unfolding l_coset_def
by fast
qed
@@ -387,7 +387,7 @@
by (simp add: normal_def subgroup_def)
lemma (in group) normalI:
- "subgroup H G \<Longrightarrow> (\<forall>x \<in> carrier G. H #> x = x <# H) \<Longrightarrow> H \<lhd> G"
+ "subgroup H G \<Longrightarrow> (\<forall>x \<in> carrier G. H #> x = x \<subset># H) \<Longrightarrow> H \<lhd> G"
by (simp add: normal_def normal_axioms_def is_group)
lemma (in normal) inv_op_closed1:
@@ -460,18 +460,18 @@
lemma (in group) lcos_m_assoc:
"[| M \<subseteq> carrier G; g \<in> carrier G; h \<in> carrier G |]
- ==> g <# (h <# M) = (g \<otimes> h) <# M"
+ ==> g \<subset># (h \<subset># M) = (g \<otimes> h) \<subset># M"
by (force simp add: l_coset_def m_assoc)
-lemma (in group) lcos_mult_one: "M \<subseteq> carrier G ==> \<one> <# M = M"
+lemma (in group) lcos_mult_one: "M \<subseteq> carrier G ==> \<one> \<subset># M = M"
by (force simp add: l_coset_def)
lemma (in group) l_coset_subset_G:
- "[| H \<subseteq> carrier G; x \<in> carrier G |] ==> x <# H \<subseteq> carrier G"
+ "[| H \<subseteq> carrier G; x \<in> carrier G |] ==> x \<subset># H \<subseteq> carrier G"
by (auto simp add: l_coset_def subsetD)
lemma (in group) l_coset_swap:
- "\<lbrakk>y \<in> x <# H; x \<in> carrier G; subgroup H G\<rbrakk> \<Longrightarrow> x \<in> y <# H"
+ "\<lbrakk>y \<in> x \<subset># H; x \<in> carrier G; subgroup H G\<rbrakk> \<Longrightarrow> x \<in> y \<subset># H"
proof (simp add: l_coset_def)
assume "\<exists>h\<in>H. y = x \<otimes> h"
and x: "x \<in> carrier G"
@@ -487,13 +487,13 @@
qed
lemma (in group) l_coset_carrier:
- "[| y \<in> x <# H; x \<in> carrier G; subgroup H G |] ==> y \<in> carrier G"
+ "[| y \<in> x \<subset># H; x \<in> carrier G; subgroup H G |] ==> y \<in> carrier G"
by (auto simp add: l_coset_def m_assoc
subgroup.subset [THEN subsetD] subgroup.m_closed)
lemma (in group) l_repr_imp_subset:
- assumes y: "y \<in> x <# H" and x: "x \<in> carrier G" and sb: "subgroup H G"
- shows "y <# H \<subseteq> x <# H"
+ assumes y: "y \<in> x \<subset># H" and x: "x \<in> carrier G" and sb: "subgroup H G"
+ shows "y \<subset># H \<subseteq> x \<subset># H"
proof -
from y
obtain h' where "h' \<in> H" "x \<otimes> h' = y" by (auto simp add: l_coset_def)
@@ -503,20 +503,20 @@
qed
lemma (in group) l_repr_independence:
- assumes y: "y \<in> x <# H" and x: "x \<in> carrier G" and sb: "subgroup H G"
- shows "x <# H = y <# H"
+ assumes y: "y \<in> x \<subset># H" and x: "x \<in> carrier G" and sb: "subgroup H G"
+ shows "x \<subset># H = y \<subset># H"
proof
- show "x <# H \<subseteq> y <# H"
+ show "x \<subset># H \<subseteq> y \<subset># H"
by (rule l_repr_imp_subset,
(blast intro: l_coset_swap l_coset_carrier y x sb)+)
- show "y <# H \<subseteq> x <# H" by (rule l_repr_imp_subset [OF y x sb])
+ show "y \<subset># H \<subseteq> x \<subset># H" by (rule l_repr_imp_subset [OF y x sb])
qed
lemma (in group) setmult_subset_G:
- "\<lbrakk>H \<subseteq> carrier G; K \<subseteq> carrier G\<rbrakk> \<Longrightarrow> H <#> K \<subseteq> carrier G"
+ "\<lbrakk>H \<subseteq> carrier G; K \<subseteq> carrier G\<rbrakk> \<Longrightarrow> H \<subset>#> K \<subseteq> carrier G"
by (auto simp add: set_mult_def subsetD)
-lemma (in group) subgroup_mult_id: "subgroup H G \<Longrightarrow> H <#> H = H"
+lemma (in group) subgroup_mult_id: "subgroup H G \<Longrightarrow> H \<subset>#> H = H"
apply (auto simp add: subgroup.m_closed set_mult_def Sigma_def)
apply (rule_tac x = x in bexI)
apply (rule bexI [of _ "\<one>"])
@@ -549,41 +549,41 @@
qed
-subsubsection \<open>Theorems for \<open><#>\<close> with \<open>#>\<close> or \<open><#\<close>.\<close>
+subsubsection \<open>Theorems for \<open>\<subset>#>\<close> with \<open>#>\<close> or \<open>\<subset>#\<close>.\<close>
lemma (in group) setmult_rcos_assoc:
"\<lbrakk>H \<subseteq> carrier G; K \<subseteq> carrier G; x \<in> carrier G\<rbrakk>
- \<Longrightarrow> H <#> (K #> x) = (H <#> K) #> x"
+ \<Longrightarrow> H \<subset>#> (K #> x) = (H \<subset>#> K) #> x"
by (force simp add: r_coset_def set_mult_def m_assoc)
lemma (in group) rcos_assoc_lcos:
"\<lbrakk>H \<subseteq> carrier G; K \<subseteq> carrier G; x \<in> carrier G\<rbrakk>
- \<Longrightarrow> (H #> x) <#> K = H <#> (x <# K)"
+ \<Longrightarrow> (H #> x) \<subset>#> K = H \<subset>#> (x \<subset># K)"
by (force simp add: r_coset_def l_coset_def set_mult_def m_assoc)
lemma (in normal) rcos_mult_step1:
"\<lbrakk>x \<in> carrier G; y \<in> carrier G\<rbrakk>
- \<Longrightarrow> (H #> x) <#> (H #> y) = (H <#> (x <# H)) #> y"
+ \<Longrightarrow> (H #> x) \<subset>#> (H #> y) = (H \<subset>#> (x \<subset># H)) #> y"
by (simp add: setmult_rcos_assoc subset
r_coset_subset_G l_coset_subset_G rcos_assoc_lcos)
lemma (in normal) rcos_mult_step2:
"\<lbrakk>x \<in> carrier G; y \<in> carrier G\<rbrakk>
- \<Longrightarrow> (H <#> (x <# H)) #> y = (H <#> (H #> x)) #> y"
+ \<Longrightarrow> (H \<subset>#> (x \<subset># H)) #> y = (H \<subset>#> (H #> x)) #> y"
by (insert coset_eq, simp add: normal_def)
lemma (in normal) rcos_mult_step3:
"\<lbrakk>x \<in> carrier G; y \<in> carrier G\<rbrakk>
- \<Longrightarrow> (H <#> (H #> x)) #> y = H #> (x \<otimes> y)"
+ \<Longrightarrow> (H \<subset>#> (H #> x)) #> y = H #> (x \<otimes> y)"
by (simp add: setmult_rcos_assoc coset_mult_assoc
subgroup_mult_id normal.axioms subset normal_axioms)
lemma (in normal) rcos_sum:
"\<lbrakk>x \<in> carrier G; y \<in> carrier G\<rbrakk>
- \<Longrightarrow> (H #> x) <#> (H #> y) = H #> (x \<otimes> y)"
+ \<Longrightarrow> (H #> x) \<subset>#> (H #> y) = H #> (x \<otimes> y)"
by (simp add: rcos_mult_step1 rcos_mult_step2 rcos_mult_step3)
-lemma (in normal) rcosets_mult_eq: "M \<in> rcosets H \<Longrightarrow> H <#> M = M"
+lemma (in normal) rcosets_mult_eq: "M \<in> rcosets H \<Longrightarrow> H \<subset>#> M = M"
\<comment> \<open>generalizes \<open>subgroup_mult_id\<close>\<close>
by (auto simp add: RCOSETS_def subset
setmult_rcos_assoc subgroup_mult_id normal.axioms normal_axioms)
@@ -645,7 +645,7 @@
lemma (in subgroup) l_coset_eq_rcong:
assumes "group G"
assumes a: "a \<in> carrier G"
- shows "a <# H = rcong H `` {a}"
+ shows "a \<subset># H = rcong H `` {a}"
proof -
interpret group G by fact
show ?thesis by (force simp add: r_congruent_def l_coset_def m_assoc [symmetric] a )
@@ -684,7 +684,7 @@
text \<open>The relation is a congruence\<close>
lemma (in normal) congruent_rcong:
- shows "congruent2 (rcong H) (rcong H) (\<lambda>a b. a \<otimes> b <# H)"
+ shows "congruent2 (rcong H) (rcong H) (\<lambda>a b. a \<otimes> b \<subset># H)"
proof (intro congruent2I[of "carrier G" _ "carrier G" _] equiv_rcong is_group)
fix a b c
assume abrcong: "(a, b) \<in> rcong H"
@@ -712,10 +712,10 @@
ultimately
have "(inv (a \<otimes> c)) \<otimes> (b \<otimes> c) \<in> H" by simp
from carr and this
- have "(b \<otimes> c) \<in> (a \<otimes> c) <# H"
+ have "(b \<otimes> c) \<in> (a \<otimes> c) \<subset># H"
by (simp add: lcos_module_rev[OF is_group])
from carr and this and is_subgroup
- show "(a \<otimes> c) <# H = (b \<otimes> c) <# H" by (intro l_repr_independence, simp+)
+ show "(a \<otimes> c) \<subset># H = (b \<otimes> c) \<subset># H" by (intro l_repr_independence, simp+)
next
fix a b c
assume abrcong: "(a, b) \<in> rcong H"
@@ -746,10 +746,10 @@
have "inv (c \<otimes> a) \<otimes> (c \<otimes> b) \<in> H" by simp
from carr and this
- have "(c \<otimes> b) \<in> (c \<otimes> a) <# H"
+ have "(c \<otimes> b) \<in> (c \<otimes> a) \<subset># H"
by (simp add: lcos_module_rev[OF is_group])
from carr and this and is_subgroup
- show "(c \<otimes> a) <# H = (c \<otimes> b) <# H" by (intro l_repr_independence, simp+)
+ show "(c \<otimes> a) \<subset># H = (c \<otimes> b) \<subset># H" by (intro l_repr_independence, simp+)
qed
@@ -835,7 +835,7 @@
where "FactGroup G H = \<lparr>carrier = rcosets\<^bsub>G\<^esub> H, mult = set_mult G, one = H\<rparr>"
lemma (in normal) setmult_closed:
- "\<lbrakk>K1 \<in> rcosets H; K2 \<in> rcosets H\<rbrakk> \<Longrightarrow> K1 <#> K2 \<in> rcosets H"
+ "\<lbrakk>K1 \<in> rcosets H; K2 \<in> rcosets H\<rbrakk> \<Longrightarrow> K1 \<subset>#> K2 \<in> rcosets H"
by (auto simp add: rcos_sum RCOSETS_def)
lemma (in normal) setinv_closed:
@@ -844,7 +844,7 @@
lemma (in normal) rcosets_assoc:
"\<lbrakk>M1 \<in> rcosets H; M2 \<in> rcosets H; M3 \<in> rcosets H\<rbrakk>
- \<Longrightarrow> M1 <#> M2 <#> M3 = M1 <#> (M2 <#> M3)"
+ \<Longrightarrow> M1 \<subset>#> M2 \<subset>#> M3 = M1 \<subset>#> (M2 \<subset>#> M3)"
by (auto simp add: RCOSETS_def rcos_sum m_assoc)
lemma (in subgroup) subgroup_in_rcosets:
@@ -859,7 +859,7 @@
qed
lemma (in normal) rcosets_inv_mult_group_eq:
- "M \<in> rcosets H \<Longrightarrow> set_inv M <#> M = H"
+ "M \<in> rcosets H \<Longrightarrow> set_inv M \<subset>#> M = H"
by (auto simp add: RCOSETS_def rcos_inv rcos_sum subgroup.subset normal.axioms normal_axioms)
theorem (in normal) factorgroup_is_group:
@@ -874,7 +874,7 @@
apply (auto dest: rcosets_inv_mult_group_eq simp add: setinv_closed)
done
-lemma mult_FactGroup [simp]: "X \<otimes>\<^bsub>(G Mod H)\<^esub> X' = X <#>\<^bsub>G\<^esub> X'"
+lemma mult_FactGroup [simp]: "X \<otimes>\<^bsub>(G Mod H)\<^esub> X' = X \<subset>#>\<^bsub>G\<^esub> X'"
by (simp add: FactGroup_def)
lemma (in normal) inv_FactGroup:
@@ -951,11 +951,11 @@
hence all: "\<forall>x\<in>X. h x = h g" "\<forall>x\<in>X'. h x = h g'"
and Xsub: "X \<subseteq> carrier G" and X'sub: "X' \<subseteq> carrier G"
by (force simp add: kernel_def r_coset_def image_def)+
- hence "h ` (X <#> X') = {h g \<otimes>\<^bsub>H\<^esub> h g'}" using X X'
+ hence "h ` (X \<subset>#> X') = {h g \<otimes>\<^bsub>H\<^esub> h g'}" using X X'
by (auto dest!: FactGroup_nonempty intro!: image_eqI
simp add: set_mult_def
subsetD [OF Xsub] subsetD [OF X'sub])
- then show "the_elem (h ` (X <#> X')) = the_elem (h ` X) \<otimes>\<^bsub>H\<^esub> the_elem (h ` X')"
+ then show "the_elem (h ` (X \<subset>#> X')) = the_elem (h ` X) \<otimes>\<^bsub>H\<^esub> the_elem (h ` X')"
by (auto simp add: all FactGroup_nonempty X X' the_elem_image_unique)
qed
--- a/src/HOL/Algebra/Divisibility.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Algebra/Divisibility.thy Tue Dec 20 16:18:56 2016 +0100
@@ -1918,7 +1918,7 @@
and afs: "wfactors G as a"
and bfs: "wfactors G bs b"
and carr: "a \<in> carrier G" "b \<in> carrier G" "set as \<subseteq> carrier G" "set bs \<subseteq> carrier G"
- shows "fmset G as \<le># fmset G bs"
+ shows "fmset G as \<subseteq># fmset G bs"
using ab
proof (elim dividesE)
fix c
@@ -1935,7 +1935,7 @@
qed
lemma (in comm_monoid_cancel) fmsubset_divides:
- assumes msubset: "fmset G as \<le># fmset G bs"
+ assumes msubset: "fmset G as \<subseteq># fmset G bs"
and afs: "wfactors G as a"
and bfs: "wfactors G bs b"
and acarr: "a \<in> carrier G"
@@ -1988,7 +1988,7 @@
and "b \<in> carrier G"
and "set as \<subseteq> carrier G"
and "set bs \<subseteq> carrier G"
- shows "a divides b = (fmset G as \<le># fmset G bs)"
+ shows "a divides b = (fmset G as \<subseteq># fmset G bs)"
using assms
by (blast intro: divides_fmsubset fmsubset_divides)
@@ -1996,7 +1996,7 @@
text \<open>Proper factors on multisets\<close>
lemma (in factorial_monoid) fmset_properfactor:
- assumes asubb: "fmset G as \<le># fmset G bs"
+ assumes asubb: "fmset G as \<subseteq># fmset G bs"
and anb: "fmset G as \<noteq> fmset G bs"
and "wfactors G as a"
and "wfactors G bs b"
@@ -2009,7 +2009,7 @@
apply (rule fmsubset_divides[of as bs], fact+)
proof
assume "b divides a"
- then have "fmset G bs \<le># fmset G as"
+ then have "fmset G bs \<subseteq># fmset G as"
by (rule divides_fmsubset) fact+
with asubb have "fmset G as = fmset G bs"
by (rule subset_mset.antisym)
@@ -2024,7 +2024,7 @@
and "b \<in> carrier G"
and "set as \<subseteq> carrier G"
and "set bs \<subseteq> carrier G"
- shows "fmset G as \<le># fmset G bs \<and> fmset G as \<noteq> fmset G bs"
+ shows "fmset G as \<subseteq># fmset G bs \<and> fmset G as \<noteq> fmset G bs"
using pf
apply (elim properfactorE)
apply rule
@@ -2334,11 +2334,11 @@
have "c gcdof a b"
proof (simp add: isgcd_def, safe)
from csmset
- have "fmset G cs \<le># fmset G as"
+ have "fmset G cs \<subseteq># fmset G as"
by (simp add: multiset_inter_def subset_mset_def)
then show "c divides a" by (rule fmsubset_divides) fact+
next
- from csmset have "fmset G cs \<le># fmset G bs"
+ from csmset have "fmset G cs \<subseteq># fmset G bs"
by (simp add: multiset_inter_def subseteq_mset_def, force)
then show "c divides b"
by (rule fmsubset_divides) fact+
@@ -2350,14 +2350,14 @@
by blast
assume "y divides a"
- then have ya: "fmset G ys \<le># fmset G as"
+ then have ya: "fmset G ys \<subseteq># fmset G as"
by (rule divides_fmsubset) fact+
assume "y divides b"
- then have yb: "fmset G ys \<le># fmset G bs"
+ then have yb: "fmset G ys \<subseteq># fmset G bs"
by (rule divides_fmsubset) fact+
- from ya yb csmset have "fmset G ys \<le># fmset G cs"
+ from ya yb csmset have "fmset G ys \<subseteq># fmset G cs"
by (simp add: subset_mset_def)
then show "y divides c"
by (rule fmsubset_divides) fact+
@@ -2420,12 +2420,12 @@
have "c lcmof a b"
proof (simp add: islcm_def, safe)
- from csmset have "fmset G as \<le># fmset G cs"
+ from csmset have "fmset G as \<subseteq># fmset G cs"
by (simp add: subseteq_mset_def, force)
then show "a divides c"
by (rule fmsubset_divides) fact+
next
- from csmset have "fmset G bs \<le># fmset G cs"
+ from csmset have "fmset G bs \<subseteq># fmset G cs"
by (simp add: subset_mset_def)
then show "b divides c"
by (rule fmsubset_divides) fact+
@@ -2437,14 +2437,14 @@
by blast
assume "a divides y"
- then have ya: "fmset G as \<le># fmset G ys"
+ then have ya: "fmset G as \<subseteq># fmset G ys"
by (rule divides_fmsubset) fact+
assume "b divides y"
- then have yb: "fmset G bs \<le># fmset G ys"
+ then have yb: "fmset G bs \<subseteq># fmset G ys"
by (rule divides_fmsubset) fact+
- from ya yb csmset have "fmset G cs \<le># fmset G ys"
+ from ya yb csmset have "fmset G cs \<subseteq># fmset G ys"
apply (simp add: subseteq_mset_def, clarify)
apply (case_tac "count (fmset G as) a < count (fmset G bs) a")
apply simp
--- a/src/HOL/Analysis/Complex_Transcendental.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Analysis/Complex_Transcendental.thy Tue Dec 20 16:18:56 2016 +0100
@@ -3206,7 +3206,7 @@
using of_int_eq_iff apply fastforce
by (metis of_int_add of_int_mult of_int_of_nat_eq)
also have "... \<longleftrightarrow> int j mod int n = int k mod int n"
- by (auto simp: zmod_eq_dvd_iff dvd_def algebra_simps)
+ by (auto simp: mod_eq_dvd_iff dvd_def algebra_simps)
also have "... \<longleftrightarrow> j mod n = k mod n"
by (metis of_nat_eq_iff zmod_int)
finally have "(\<exists>z. \<i> * (of_nat j * (of_real pi * 2)) =
--- a/src/HOL/Code_Numeral.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Code_Numeral.thy Tue Dec 20 16:18:56 2016 +0100
@@ -168,21 +168,9 @@
"integer_of_num (Num.Bit0 Num.One) = 2"
by (transfer, simp)+
-instantiation integer :: "{ring_div, equal, linordered_idom}"
+instantiation integer :: "{linordered_idom, equal}"
begin
-lift_definition divide_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"
- is "divide :: int \<Rightarrow> int \<Rightarrow> int"
- .
-
-declare divide_integer.rep_eq [simp]
-
-lift_definition modulo_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"
- is "modulo :: int \<Rightarrow> int \<Rightarrow> int"
- .
-
-declare modulo_integer.rep_eq [simp]
-
lift_definition abs_integer :: "integer \<Rightarrow> integer"
is "abs :: int \<Rightarrow> int"
.
@@ -199,6 +187,7 @@
is "less_eq :: int \<Rightarrow> int \<Rightarrow> bool"
.
+
lift_definition less_integer :: "integer \<Rightarrow> integer \<Rightarrow> bool"
is "less :: int \<Rightarrow> int \<Rightarrow> bool"
.
@@ -207,8 +196,8 @@
is "HOL.equal :: int \<Rightarrow> int \<Rightarrow> bool"
.
-instance proof
-qed (transfer, simp add: algebra_simps equal less_le_not_le [symmetric] mult_strict_right_mono linear)+
+instance
+ by standard (transfer, simp add: algebra_simps equal less_le_not_le [symmetric] mult_strict_right_mono linear)+
end
@@ -236,6 +225,38 @@
"of_nat (nat_of_integer k) = max 0 k"
by transfer auto
+instantiation integer :: "{ring_div, normalization_semidom}"
+begin
+
+lift_definition normalize_integer :: "integer \<Rightarrow> integer"
+ is "normalize :: int \<Rightarrow> int"
+ .
+
+declare normalize_integer.rep_eq [simp]
+
+lift_definition unit_factor_integer :: "integer \<Rightarrow> integer"
+ is "unit_factor :: int \<Rightarrow> int"
+ .
+
+declare unit_factor_integer.rep_eq [simp]
+
+lift_definition divide_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"
+ is "divide :: int \<Rightarrow> int \<Rightarrow> int"
+ .
+
+declare divide_integer.rep_eq [simp]
+
+lift_definition modulo_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"
+ is "modulo :: int \<Rightarrow> int \<Rightarrow> int"
+ .
+
+declare modulo_integer.rep_eq [simp]
+
+instance
+ by standard (transfer, simp add: mult_sgn_abs sgn_mult)+
+
+end
+
instantiation integer :: semiring_numeral_div
begin
@@ -389,6 +410,14 @@
"Neg m * Neg n = Pos (m * n)"
by simp_all
+lemma normalize_integer_code [code]:
+ "normalize = (abs :: integer \<Rightarrow> integer)"
+ by transfer simp
+
+lemma unit_factor_integer_code [code]:
+ "unit_factor = (sgn :: integer \<Rightarrow> integer)"
+ by transfer simp
+
definition divmod_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer \<times> integer"
where
"divmod_integer k l = (k div l, k mod l)"
@@ -760,21 +789,9 @@
"nat_of_natural (numeral k) = numeral k"
by transfer rule
-instantiation natural :: "{semiring_div, equal, linordered_semiring}"
+instantiation natural :: "{linordered_semiring, equal}"
begin
-lift_definition divide_natural :: "natural \<Rightarrow> natural \<Rightarrow> natural"
- is "divide :: nat \<Rightarrow> nat \<Rightarrow> nat"
- .
-
-declare divide_natural.rep_eq [simp]
-
-lift_definition modulo_natural :: "natural \<Rightarrow> natural \<Rightarrow> natural"
- is "modulo :: nat \<Rightarrow> nat \<Rightarrow> nat"
- .
-
-declare modulo_natural.rep_eq [simp]
-
lift_definition less_eq_natural :: "natural \<Rightarrow> natural \<Rightarrow> bool"
is "less_eq :: nat \<Rightarrow> nat \<Rightarrow> bool"
.
@@ -812,6 +829,38 @@
"nat_of_natural (max k l) = max (nat_of_natural k) (nat_of_natural l)"
by transfer rule
+instantiation natural :: "{semiring_div, normalization_semidom}"
+begin
+
+lift_definition normalize_natural :: "natural \<Rightarrow> natural"
+ is "normalize :: nat \<Rightarrow> nat"
+ .
+
+declare normalize_natural.rep_eq [simp]
+
+lift_definition unit_factor_natural :: "natural \<Rightarrow> natural"
+ is "unit_factor :: nat \<Rightarrow> nat"
+ .
+
+declare unit_factor_natural.rep_eq [simp]
+
+lift_definition divide_natural :: "natural \<Rightarrow> natural \<Rightarrow> natural"
+ is "divide :: nat \<Rightarrow> nat \<Rightarrow> nat"
+ .
+
+declare divide_natural.rep_eq [simp]
+
+lift_definition modulo_natural :: "natural \<Rightarrow> natural \<Rightarrow> natural"
+ is "modulo :: nat \<Rightarrow> nat \<Rightarrow> nat"
+ .
+
+declare modulo_natural.rep_eq [simp]
+
+instance
+ by standard (transfer, auto simp add: algebra_simps unit_factor_nat_def gr0_conv_Suc)+
+
+end
+
lift_definition natural_of_integer :: "integer \<Rightarrow> natural"
is "nat :: int \<Rightarrow> nat"
.
@@ -945,7 +994,32 @@
lemma [code abstract]:
"integer_of_natural (m * n) = integer_of_natural m * integer_of_natural n"
- by transfer (simp add: of_nat_mult)
+ by transfer simp
+
+lemma [code]:
+ "normalize n = n" for n :: natural
+ by transfer simp
+
+lemma [code]:
+ "unit_factor n = of_bool (n \<noteq> 0)" for n :: natural
+proof (cases "n = 0")
+ case True
+ then show ?thesis
+ by simp
+next
+ case False
+ then have "unit_factor n = 1"
+ proof transfer
+ fix n :: nat
+ assume "n \<noteq> 0"
+ then obtain m where "n = Suc m"
+ by (cases n) auto
+ then show "unit_factor n = 1"
+ by simp
+ qed
+ with False show ?thesis
+ by simp
+qed
lemma [code abstract]:
"integer_of_natural (m div n) = integer_of_natural m div integer_of_natural n"
--- a/src/HOL/Decision_Procs/cooper_tac.ML Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Decision_Procs/cooper_tac.ML Tue Dec 20 16:18:56 2016 +0100
@@ -45,8 +45,8 @@
(* Some simpsets for dealing with mod div abs and nat*)
val mod_div_simpset =
put_simpset HOL_basic_ss ctxt
- addsimps @{thms refl mod_add_eq [symmetric] mod_add_left_eq [symmetric]
- mod_add_right_eq [symmetric]
+ addsimps @{thms refl mod_add_eq mod_add_left_eq
+ mod_add_right_eq
div_add1_eq [symmetric] zdiv_zadd1_eq [symmetric]
mod_self
div_by_0 mod_by_0 div_0 mod_0
--- a/src/HOL/Decision_Procs/mir_tac.ML Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Decision_Procs/mir_tac.ML Tue Dec 20 16:18:56 2016 +0100
@@ -31,8 +31,6 @@
@{thm uminus_add_conv_diff [symmetric]}, @{thm "minus_divide_left"}]
val comp_ths = distinct Thm.eq_thm (ths @ comp_arith @ @{thms simp_thms});
-val mod_add_eq = @{thm "mod_add_eq"} RS sym;
-
fun prepare_for_mir q fm =
let
val ps = Logic.strip_params fm
@@ -71,7 +69,7 @@
val (t,np,nh) = prepare_for_mir q g
(* Some simpsets for dealing with mod div abs and nat*)
val mod_div_simpset = put_simpset HOL_basic_ss ctxt
- addsimps [refl, mod_add_eq,
+ addsimps [refl, @{thm mod_add_eq},
@{thm mod_self},
@{thm div_0}, @{thm mod_0},
@{thm div_by_1}, @{thm mod_by_1}, @{thm div_by_Suc_0}, @{thm mod_by_Suc_0},
--- a/src/HOL/Divides.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Divides.thy Tue Dec 20 16:18:56 2016 +0100
@@ -3,37 +3,95 @@
Copyright 1999 University of Cambridge
*)
-section \<open>The division operators div and mod\<close>
+section \<open>Quotient and remainder\<close>
theory Divides
imports Parity
begin
-subsection \<open>Abstract division in commutative semirings.\<close>
-
-class semiring_div = semidom + semiring_modulo +
- assumes div_by_0: "a div 0 = 0"
- and div_0: "0 div a = 0"
- and div_mult_self1 [simp]: "b \<noteq> 0 \<Longrightarrow> (a + c * b) div b = c + a div b"
+subsection \<open>Quotient and remainder in integral domains\<close>
+
+class semidom_modulo = algebraic_semidom + semiring_modulo
+begin
+
+lemma mod_0 [simp]: "0 mod a = 0"
+ using div_mult_mod_eq [of 0 a] by simp
+
+lemma mod_by_0 [simp]: "a mod 0 = a"
+ using div_mult_mod_eq [of a 0] by simp
+
+lemma mod_by_1 [simp]:
+ "a mod 1 = 0"
+proof -
+ from div_mult_mod_eq [of a one] div_by_1 have "a + a mod 1 = a" by simp
+ then have "a + a mod 1 = a + 0" by simp
+ then show ?thesis by (rule add_left_imp_eq)
+qed
+
+lemma mod_self [simp]:
+ "a mod a = 0"
+ using div_mult_mod_eq [of a a] by simp
+
+lemma dvd_imp_mod_0 [simp]:
+ assumes "a dvd b"
+ shows "b mod a = 0"
+ using assms minus_div_mult_eq_mod [of b a] by simp
+
+lemma mod_0_imp_dvd:
+ assumes "a mod b = 0"
+ shows "b dvd a"
+proof -
+ have "b dvd ((a div b) * b)" by simp
+ also have "(a div b) * b = a"
+ using div_mult_mod_eq [of a b] by (simp add: assms)
+ finally show ?thesis .
+qed
+
+lemma mod_eq_0_iff_dvd:
+ "a mod b = 0 \<longleftrightarrow> b dvd a"
+ by (auto intro: mod_0_imp_dvd)
+
+lemma dvd_eq_mod_eq_0 [nitpick_unfold, code]:
+ "a dvd b \<longleftrightarrow> b mod a = 0"
+ by (simp add: mod_eq_0_iff_dvd)
+
+lemma dvd_mod_iff:
+ assumes "c dvd b"
+ shows "c dvd a mod b \<longleftrightarrow> c dvd a"
+proof -
+ from assms have "(c dvd a mod b) \<longleftrightarrow> (c dvd ((a div b) * b + a mod b))"
+ by (simp add: dvd_add_right_iff)
+ also have "(a div b) * b + a mod b = a"
+ using div_mult_mod_eq [of a b] by simp
+ finally show ?thesis .
+qed
+
+lemma dvd_mod_imp_dvd:
+ assumes "c dvd a mod b" and "c dvd b"
+ shows "c dvd a"
+ using assms dvd_mod_iff [of c b a] by simp
+
+end
+
+class idom_modulo = idom + semidom_modulo
+begin
+
+subclass idom_divide ..
+
+lemma div_diff [simp]:
+ "c dvd a \<Longrightarrow> c dvd b \<Longrightarrow> (a - b) div c = a div c - b div c"
+ using div_add [of _ _ "- b"] by (simp add: dvd_neg_div)
+
+end
+
+
+subsection \<open>Quotient and remainder in integral domains with additional properties\<close>
+
+class semiring_div = semidom_modulo +
+ assumes div_mult_self1 [simp]: "b \<noteq> 0 \<Longrightarrow> (a + c * b) div b = c + a div b"
and div_mult_mult1 [simp]: "c \<noteq> 0 \<Longrightarrow> (c * a) div (c * b) = a div b"
begin
-subclass algebraic_semidom
-proof
- fix b a
- assume "b \<noteq> 0"
- then show "a * b div b = a"
- using div_mult_self1 [of b 0 a] by (simp add: ac_simps div_0)
-qed (simp add: div_by_0)
-
-text \<open>@{const divide} and @{const modulo}\<close>
-
-lemma mod_by_0 [simp]: "a mod 0 = a"
- using div_mult_mod_eq [of a zero] by simp
-
-lemma mod_0 [simp]: "0 mod a = 0"
- using div_mult_mod_eq [of zero a] div_0 by simp
-
lemma div_mult_self2 [simp]:
assumes "b \<noteq> 0"
shows "(a + b * c) div b = c + a div b"
@@ -86,18 +144,6 @@
"a * b mod b = 0"
using mod_mult_self1 [of 0 a b] by simp
-lemma mod_by_1 [simp]:
- "a mod 1 = 0"
-proof -
- from div_mult_mod_eq [of a one] div_by_1 have "a + a mod 1 = a" by simp
- then have "a + a mod 1 = a + 0" by simp
- then show ?thesis by (rule add_left_imp_eq)
-qed
-
-lemma mod_self [simp]:
- "a mod a = 0"
- using mod_mult_self2_is_0 [of 1] by simp
-
lemma div_add_self1:
assumes "b \<noteq> 0"
shows "(b + a) div b = a div b + 1"
@@ -116,31 +162,6 @@
"(a + b) mod b = a mod b"
using mod_mult_self1 [of a 1 b] by simp
-lemma dvd_imp_mod_0 [simp]:
- assumes "a dvd b"
- shows "b mod a = 0"
-proof -
- from assms obtain c where "b = a * c" ..
- then have "b mod a = a * c mod a" by simp
- then show "b mod a = 0" by simp
-qed
-
-lemma mod_eq_0_iff_dvd:
- "a mod b = 0 \<longleftrightarrow> b dvd a"
-proof
- assume "b dvd a"
- then show "a mod b = 0" by simp
-next
- assume "a mod b = 0"
- with div_mult_mod_eq [of a b] have "a div b * b = a" by simp
- then have "a = b * (a div b)" by (simp add: ac_simps)
- then show "b dvd a" ..
-qed
-
-lemma dvd_eq_mod_eq_0 [nitpick_unfold, code]:
- "a dvd b \<longleftrightarrow> b mod a = 0"
- by (simp add: mod_eq_0_iff_dvd)
-
lemma mod_div_trivial [simp]:
"a mod b div b = 0"
proof (cases "b = 0")
@@ -168,106 +189,6 @@
finally show ?thesis .
qed
-lemma dvd_mod_imp_dvd:
- assumes "k dvd m mod n" and "k dvd n"
- shows "k dvd m"
-proof -
- from assms have "k dvd (m div n) * n + m mod n"
- by (simp only: dvd_add dvd_mult)
- then show ?thesis by (simp add: div_mult_mod_eq)
-qed
-
-text \<open>Addition respects modular equivalence.\<close>
-
-lemma mod_add_left_eq: \<comment> \<open>FIXME reorient\<close>
- "(a + b) mod c = (a mod c + b) mod c"
-proof -
- have "(a + b) mod c = (a div c * c + a mod c + b) mod c"
- by (simp only: div_mult_mod_eq)
- also have "\<dots> = (a mod c + b + a div c * c) mod c"
- by (simp only: ac_simps)
- also have "\<dots> = (a mod c + b) mod c"
- by (rule mod_mult_self1)
- finally show ?thesis .
-qed
-
-lemma mod_add_right_eq: \<comment> \<open>FIXME reorient\<close>
- "(a + b) mod c = (a + b mod c) mod c"
-proof -
- have "(a + b) mod c = (a + (b div c * c + b mod c)) mod c"
- by (simp only: div_mult_mod_eq)
- also have "\<dots> = (a + b mod c + b div c * c) mod c"
- by (simp only: ac_simps)
- also have "\<dots> = (a + b mod c) mod c"
- by (rule mod_mult_self1)
- finally show ?thesis .
-qed
-
-lemma mod_add_eq: \<comment> \<open>FIXME reorient\<close>
- "(a + b) mod c = (a mod c + b mod c) mod c"
-by (rule trans [OF mod_add_left_eq mod_add_right_eq])
-
-lemma mod_add_cong:
- assumes "a mod c = a' mod c"
- assumes "b mod c = b' mod c"
- shows "(a + b) mod c = (a' + b') mod c"
-proof -
- have "(a mod c + b mod c) mod c = (a' mod c + b' mod c) mod c"
- unfolding assms ..
- thus ?thesis
- by (simp only: mod_add_eq [symmetric])
-qed
-
-text \<open>Multiplication respects modular equivalence.\<close>
-
-lemma mod_mult_left_eq: \<comment> \<open>FIXME reorient\<close>
- "(a * b) mod c = ((a mod c) * b) mod c"
-proof -
- have "(a * b) mod c = ((a div c * c + a mod c) * b) mod c"
- by (simp only: div_mult_mod_eq)
- also have "\<dots> = (a mod c * b + a div c * b * c) mod c"
- by (simp only: algebra_simps)
- also have "\<dots> = (a mod c * b) mod c"
- by (rule mod_mult_self1)
- finally show ?thesis .
-qed
-
-lemma mod_mult_right_eq: \<comment> \<open>FIXME reorient\<close>
- "(a * b) mod c = (a * (b mod c)) mod c"
-proof -
- have "(a * b) mod c = (a * (b div c * c + b mod c)) mod c"
- by (simp only: div_mult_mod_eq)
- also have "\<dots> = (a * (b mod c) + a * (b div c) * c) mod c"
- by (simp only: algebra_simps)
- also have "\<dots> = (a * (b mod c)) mod c"
- by (rule mod_mult_self1)
- finally show ?thesis .
-qed
-
-lemma mod_mult_eq: \<comment> \<open>FIXME reorient\<close>
- "(a * b) mod c = ((a mod c) * (b mod c)) mod c"
-by (rule trans [OF mod_mult_left_eq mod_mult_right_eq])
-
-lemma mod_mult_cong:
- assumes "a mod c = a' mod c"
- assumes "b mod c = b' mod c"
- shows "(a * b) mod c = (a' * b') mod c"
-proof -
- have "(a mod c * (b mod c)) mod c = (a' mod c * (b' mod c)) mod c"
- unfolding assms ..
- thus ?thesis
- by (simp only: mod_mult_eq [symmetric])
-qed
-
-text \<open>Exponentiation respects modular equivalence.\<close>
-
-lemma power_mod: "(a mod b) ^ n mod b = a ^ n mod b"
-apply (induct n, simp_all)
-apply (rule mod_mult_right_eq [THEN trans])
-apply (simp (no_asm_simp))
-apply (rule mod_mult_eq [symmetric])
-done
-
lemma mod_mod_cancel:
assumes "c dvd b"
shows "a mod b mod c = a mod c"
@@ -319,30 +240,120 @@
lemma dvd_mod: "k dvd m \<Longrightarrow> k dvd n \<Longrightarrow> k dvd (m mod n)"
unfolding dvd_def by (auto simp add: mod_mult_mult1)
-lemma dvd_mod_iff: "k dvd n \<Longrightarrow> k dvd (m mod n) \<longleftrightarrow> k dvd m"
-by (blast intro: dvd_mod_imp_dvd dvd_mod)
-
-lemma div_div_eq_right:
- assumes "c dvd b" "b dvd a"
- shows "a div (b div c) = a div b * c"
+named_theorems mod_simps
+
+text \<open>Addition respects modular equivalence.\<close>
+
+lemma mod_add_left_eq [mod_simps]:
+ "(a mod c + b) mod c = (a + b) mod c"
+proof -
+ have "(a + b) mod c = (a div c * c + a mod c + b) mod c"
+ by (simp only: div_mult_mod_eq)
+ also have "\<dots> = (a mod c + b + a div c * c) mod c"
+ by (simp only: ac_simps)
+ also have "\<dots> = (a mod c + b) mod c"
+ by (rule mod_mult_self1)
+ finally show ?thesis
+ by (rule sym)
+qed
+
+lemma mod_add_right_eq [mod_simps]:
+ "(a + b mod c) mod c = (a + b) mod c"
+ using mod_add_left_eq [of b c a] by (simp add: ac_simps)
+
+lemma mod_add_eq:
+ "(a mod c + b mod c) mod c = (a + b) mod c"
+ by (simp add: mod_add_left_eq mod_add_right_eq)
+
+lemma mod_sum_eq [mod_simps]:
+ "(\<Sum>i\<in>A. f i mod a) mod a = sum f A mod a"
+proof (induct A rule: infinite_finite_induct)
+ case (insert i A)
+ then have "(\<Sum>i\<in>insert i A. f i mod a) mod a
+ = (f i mod a + (\<Sum>i\<in>A. f i mod a)) mod a"
+ by simp
+ also have "\<dots> = (f i + (\<Sum>i\<in>A. f i mod a) mod a) mod a"
+ by (simp add: mod_simps)
+ also have "\<dots> = (f i + (\<Sum>i\<in>A. f i) mod a) mod a"
+ by (simp add: insert.hyps)
+ finally show ?case
+ by (simp add: insert.hyps mod_simps)
+qed simp_all
+
+lemma mod_add_cong:
+ assumes "a mod c = a' mod c"
+ assumes "b mod c = b' mod c"
+ shows "(a + b) mod c = (a' + b') mod c"
+proof -
+ have "(a mod c + b mod c) mod c = (a' mod c + b' mod c) mod c"
+ unfolding assms ..
+ then show ?thesis
+ by (simp add: mod_add_eq)
+qed
+
+text \<open>Multiplication respects modular equivalence.\<close>
+
+lemma mod_mult_left_eq [mod_simps]:
+ "((a mod c) * b) mod c = (a * b) mod c"
proof -
- from assms have "a div b * c = (a * c) div b"
- by (subst dvd_div_mult) simp_all
- also from assms have "\<dots> = (a * c) div ((b div c) * c)" by simp
- also have "a * c div (b div c * c) = a div (b div c)"
- by (cases "c = 0") simp_all
- finally show ?thesis ..
+ have "(a * b) mod c = ((a div c * c + a mod c) * b) mod c"
+ by (simp only: div_mult_mod_eq)
+ also have "\<dots> = (a mod c * b + a div c * b * c) mod c"
+ by (simp only: algebra_simps)
+ also have "\<dots> = (a mod c * b) mod c"
+ by (rule mod_mult_self1)
+ finally show ?thesis
+ by (rule sym)
qed
-lemma div_div_div_same:
- assumes "d dvd a" "d dvd b" "b dvd a"
- shows "(a div d) div (b div d) = a div b"
- using assms by (subst dvd_div_mult2_eq [symmetric]) simp_all
-
-lemma cancel_div_mod_rules:
- "((a div b) * b + a mod b) + c = a + c"
- "(b * (a div b) + a mod b) + c = a + c"
- by (simp_all add: div_mult_mod_eq mult_div_mod_eq)
+lemma mod_mult_right_eq [mod_simps]:
+ "(a * (b mod c)) mod c = (a * b) mod c"
+ using mod_mult_left_eq [of b c a] by (simp add: ac_simps)
+
+lemma mod_mult_eq:
+ "((a mod c) * (b mod c)) mod c = (a * b) mod c"
+ by (simp add: mod_mult_left_eq mod_mult_right_eq)
+
+lemma mod_prod_eq [mod_simps]:
+ "(\<Prod>i\<in>A. f i mod a) mod a = prod f A mod a"
+proof (induct A rule: infinite_finite_induct)
+ case (insert i A)
+ then have "(\<Prod>i\<in>insert i A. f i mod a) mod a
+ = (f i mod a * (\<Prod>i\<in>A. f i mod a)) mod a"
+ by simp
+ also have "\<dots> = (f i * ((\<Prod>i\<in>A. f i mod a) mod a)) mod a"
+ by (simp add: mod_simps)
+ also have "\<dots> = (f i * ((\<Prod>i\<in>A. f i) mod a)) mod a"
+ by (simp add: insert.hyps)
+ finally show ?case
+ by (simp add: insert.hyps mod_simps)
+qed simp_all
+
+lemma mod_mult_cong:
+ assumes "a mod c = a' mod c"
+ assumes "b mod c = b' mod c"
+ shows "(a * b) mod c = (a' * b') mod c"
+proof -
+ have "(a mod c * (b mod c)) mod c = (a' mod c * (b' mod c)) mod c"
+ unfolding assms ..
+ then show ?thesis
+ by (simp add: mod_mult_eq)
+qed
+
+text \<open>Exponentiation respects modular equivalence.\<close>
+
+lemma power_mod [mod_simps]:
+ "((a mod b) ^ n) mod b = (a ^ n) mod b"
+proof (induct n)
+ case 0
+ then show ?case by simp
+next
+ case (Suc n)
+ have "(a mod b) ^ Suc n mod b = (a mod b) * ((a mod b) ^ n mod b) mod b"
+ by (simp add: mod_mult_right_eq)
+ with Suc show ?case
+ by (simp add: mod_mult_left_eq mod_mult_right_eq)
+qed
end
@@ -351,9 +362,28 @@
subclass idom_divide ..
+lemma div_minus_minus [simp]: "(- a) div (- b) = a div b"
+ using div_mult_mult1 [of "- 1" a b] by simp
+
+lemma mod_minus_minus [simp]: "(- a) mod (- b) = - (a mod b)"
+ using mod_mult_mult1 [of "- 1" a b] by simp
+
+lemma div_minus_right: "a div (- b) = (- a) div b"
+ using div_minus_minus [of "- a" b] by simp
+
+lemma mod_minus_right: "a mod (- b) = - ((- a) mod b)"
+ using mod_minus_minus [of "- a" b] by simp
+
+lemma div_minus1_right [simp]: "a div (- 1) = - a"
+ using div_minus_right [of a 1] by simp
+
+lemma mod_minus1_right [simp]: "a mod (- 1) = 0"
+ using mod_minus_right [of a 1] by simp
+
text \<open>Negation respects modular equivalence.\<close>
-lemma mod_minus_eq: "(- a) mod b = (- (a mod b)) mod b"
+lemma mod_minus_eq [mod_simps]:
+ "(- (a mod b)) mod b = (- a) mod b"
proof -
have "(- a) mod b = (- (a div b * b + a mod b)) mod b"
by (simp only: div_mult_mod_eq)
@@ -361,7 +391,8 @@
by (simp add: ac_simps)
also have "\<dots> = (- (a mod b)) mod b"
by (rule mod_mult_self1)
- finally show ?thesis .
+ finally show ?thesis
+ by (rule sym)
qed
lemma mod_minus_cong:
@@ -370,73 +401,37 @@
proof -
have "(- (a mod b)) mod b = (- (a' mod b)) mod b"
unfolding assms ..
- thus ?thesis
- by (simp only: mod_minus_eq [symmetric])
+ then show ?thesis
+ by (simp add: mod_minus_eq)
qed
text \<open>Subtraction respects modular equivalence.\<close>
-lemma mod_diff_left_eq:
- "(a - b) mod c = (a mod c - b) mod c"
- using mod_add_cong [of a c "a mod c" "- b" "- b"] by simp
-
-lemma mod_diff_right_eq:
- "(a - b) mod c = (a - b mod c) mod c"
- using mod_add_cong [of a c a "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b] by simp
+lemma mod_diff_left_eq [mod_simps]:
+ "(a mod c - b) mod c = (a - b) mod c"
+ using mod_add_cong [of a c "a mod c" "- b" "- b"]
+ by simp
+
+lemma mod_diff_right_eq [mod_simps]:
+ "(a - b mod c) mod c = (a - b) mod c"
+ using mod_add_cong [of a c a "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b]
+ by simp
lemma mod_diff_eq:
- "(a - b) mod c = (a mod c - b mod c) mod c"
- using mod_add_cong [of a c "a mod c" "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b] by simp
+ "(a mod c - b mod c) mod c = (a - b) mod c"
+ using mod_add_cong [of a c "a mod c" "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b]
+ by simp
lemma mod_diff_cong:
assumes "a mod c = a' mod c"
assumes "b mod c = b' mod c"
shows "(a - b) mod c = (a' - b') mod c"
- using assms mod_add_cong [of a c a' "- b" "- b'"] mod_minus_cong [of b c "b'"] by simp
-
-lemma dvd_neg_div: "y dvd x \<Longrightarrow> -x div y = - (x div y)"
-apply (case_tac "y = 0") apply simp
-apply (auto simp add: dvd_def)
-apply (subgoal_tac "-(y * k) = y * - k")
- apply (simp only:)
- apply (erule nonzero_mult_div_cancel_left)
-apply simp
-done
-
-lemma dvd_div_neg: "y dvd x \<Longrightarrow> x div -y = - (x div y)"
-apply (case_tac "y = 0") apply simp
-apply (auto simp add: dvd_def)
-apply (subgoal_tac "y * k = -y * -k")
- apply (erule ssubst, rule nonzero_mult_div_cancel_left)
- apply simp
-apply simp
-done
-
-lemma div_diff [simp]:
- "z dvd x \<Longrightarrow> z dvd y \<Longrightarrow> (x - y) div z = x div z - y div z"
- using div_add [of _ _ "- y"] by (simp add: dvd_neg_div)
-
-lemma div_minus_minus [simp]: "(-a) div (-b) = a div b"
- using div_mult_mult1 [of "- 1" a b]
- unfolding neg_equal_0_iff_equal by simp
-
-lemma mod_minus_minus [simp]: "(-a) mod (-b) = - (a mod b)"
- using mod_mult_mult1 [of "- 1" a b] by simp
-
-lemma div_minus_right: "a div (-b) = (-a) div b"
- using div_minus_minus [of "-a" b] by simp
-
-lemma mod_minus_right: "a mod (-b) = - ((-a) mod b)"
- using mod_minus_minus [of "-a" b] by simp
-
-lemma div_minus1_right [simp]: "a div (-1) = -a"
- using div_minus_right [of a 1] by simp
-
-lemma mod_minus1_right [simp]: "a mod (-1) = 0"
- using mod_minus_right [of a 1] by simp
+ using assms mod_add_cong [of a c a' "- b" "- b'"] mod_minus_cong [of b c "b'"]
+ by simp
lemma minus_mod_self2 [simp]:
"(a - b) mod b = a mod b"
+ using mod_diff_right_eq [of a b b]
by (simp add: mod_diff_right_eq)
lemma minus_mod_self1 [simp]:
@@ -446,7 +441,7 @@
end
-subsubsection \<open>Parity and division\<close>
+subsection \<open>Parity\<close>
class semiring_div_parity = semiring_div + comm_semiring_1_cancel + numeral +
assumes parity: "a mod 2 = 0 \<or> a mod 2 = 1"
@@ -490,18 +485,21 @@
assume "a mod 2 = 1"
moreover assume "b mod 2 = 1"
ultimately show "(a + b) mod 2 = 0"
- using mod_add_eq [of a b 2] by simp
+ using mod_add_eq [of a 2 b] by simp
next
fix a b
assume "(a * b) mod 2 = 0"
+ then have "(a mod 2) * (b mod 2) mod 2 = 0"
+ by (simp add: mod_mult_eq)
then have "(a mod 2) * (b mod 2) = 0"
- by (cases "a mod 2 = 0") (simp_all add: mod_mult_eq [of a b 2])
+ by (cases "a mod 2 = 0") simp_all
then show "a mod 2 = 0 \<or> b mod 2 = 0"
by (rule divisors_zero)
next
fix a
assume "a mod 2 = 1"
- then have "a = a div 2 * 2 + 1" using div_mult_mod_eq [of a 2] by simp
+ then have "a = a div 2 * 2 + 1"
+ using div_mult_mod_eq [of a 2] by simp
then show "\<exists>b. a = b + 1" ..
qed
@@ -532,7 +530,7 @@
end
-subsection \<open>Generic numeral division with a pragmatic type class\<close>
+subsection \<open>Numeral division with a pragmatic type class\<close>
text \<open>
The following type class contains everything necessary to formulate
@@ -826,8 +824,10 @@
from m n have "Suc m = q' * n + Suc r'" by simp
with True show ?thesis by blast
next
- case False then have "n \<le> Suc r'" by auto
- moreover from n have "Suc r' \<le> n" by auto
+ case False then have "n \<le> Suc r'"
+ by (simp add: not_less)
+ moreover from n have "Suc r' \<le> n"
+ by (simp add: Suc_le_eq)
ultimately have "n = Suc r'" by auto
with m have "Suc m = Suc q' * n + 0" by simp
with \<open>n \<noteq> 0\<close> show ?thesis by blast
@@ -855,7 +855,7 @@
apply (auto simp add: add_mult_distrib)
done
from \<open>n \<noteq> 0\<close> assms have *: "fst qr = fst qr'"
- by (auto simp add: divmod_nat_rel_def intro: order_antisym dest: aux sym)
+ by (auto simp add: divmod_nat_rel_def intro: order_antisym dest: aux sym split: if_splits)
with assms have "snd qr = snd qr'"
by (simp add: divmod_nat_rel_def)
with * show ?thesis by (cases qr, cases qr') simp
@@ -884,10 +884,10 @@
using assms by (auto intro: divmod_nat_rel_unique divmod_nat_rel_divmod_nat)
qualified lemma divmod_nat_zero: "divmod_nat m 0 = (0, m)"
- by (simp add: Divides.divmod_nat_unique divmod_nat_rel_def)
+ by (simp add: divmod_nat_unique divmod_nat_rel_def)
qualified lemma divmod_nat_zero_left: "divmod_nat 0 n = (0, 0)"
- by (simp add: Divides.divmod_nat_unique divmod_nat_rel_def)
+ by (simp add: divmod_nat_unique divmod_nat_rel_def)
qualified lemma divmod_nat_base: "m < n \<Longrightarrow> divmod_nat m n = (0, m)"
by (simp add: divmod_nat_unique divmod_nat_rel_def)
@@ -899,19 +899,31 @@
have "divmod_nat_rel (m - n) n (divmod_nat (m - n) n)"
by (fact divmod_nat_rel_divmod_nat)
then show "divmod_nat_rel m n (apfst Suc (divmod_nat (m - n) n))"
- unfolding divmod_nat_rel_def using assms by auto
+ unfolding divmod_nat_rel_def using assms
+ by (auto split: if_splits simp add: algebra_simps)
qed
end
-
-instantiation nat :: semiring_div
+
+instantiation nat :: "{semidom_modulo, normalization_semidom}"
begin
-definition divide_nat where
- div_nat_def: "m div n = fst (Divides.divmod_nat m n)"
-
-definition modulo_nat where
- mod_nat_def: "m mod n = snd (Divides.divmod_nat m n)"
+definition normalize_nat :: "nat \<Rightarrow> nat"
+ where [simp]: "normalize = (id :: nat \<Rightarrow> nat)"
+
+definition unit_factor_nat :: "nat \<Rightarrow> nat"
+ where "unit_factor n = (if n = 0 then 0 else 1 :: nat)"
+
+lemma unit_factor_simps [simp]:
+ "unit_factor 0 = (0::nat)"
+ "unit_factor (Suc n) = 1"
+ by (simp_all add: unit_factor_nat_def)
+
+definition divide_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"
+ where div_nat_def: "m div n = fst (Divides.divmod_nat m n)"
+
+definition modulo_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"
+ where mod_nat_def: "m mod n = snd (Divides.divmod_nat m n)"
lemma fst_divmod_nat [simp]:
"fst (Divides.divmod_nat m n) = m div n"
@@ -928,15 +940,18 @@
lemma div_nat_unique:
assumes "divmod_nat_rel m n (q, r)"
shows "m div n = q"
- using assms by (auto dest!: Divides.divmod_nat_unique simp add: prod_eq_iff)
+ using assms
+ by (auto dest!: Divides.divmod_nat_unique simp add: prod_eq_iff)
lemma mod_nat_unique:
assumes "divmod_nat_rel m n (q, r)"
shows "m mod n = r"
- using assms by (auto dest!: Divides.divmod_nat_unique simp add: prod_eq_iff)
+ using assms
+ by (auto dest!: Divides.divmod_nat_unique simp add: prod_eq_iff)
lemma divmod_nat_rel: "divmod_nat_rel m n (m div n, m mod n)"
- using Divides.divmod_nat_rel_divmod_nat by (simp add: divmod_nat_div_mod)
+ using Divides.divmod_nat_rel_divmod_nat
+ by (simp add: divmod_nat_div_mod)
text \<open>The ''recursion'' equations for @{const divide} and @{const modulo}\<close>
@@ -964,11 +979,40 @@
shows "m mod n = (m - n) mod n"
using assms Divides.divmod_nat_step by (cases "n = 0") (simp_all add: prod_eq_iff)
+lemma mod_less_divisor [simp]:
+ fixes m n :: nat
+ assumes "n > 0"
+ shows "m mod n < n"
+ using assms divmod_nat_rel [of m n] unfolding divmod_nat_rel_def
+ by (auto split: if_splits)
+
+lemma mod_le_divisor [simp]:
+ fixes m n :: nat
+ assumes "n > 0"
+ shows "m mod n \<le> n"
+proof (rule less_imp_le)
+ from assms show "m mod n < n"
+ by simp
+qed
+
instance proof
fix m n :: nat
show "m div n * n + m mod n = m"
using divmod_nat_rel [of m n] by (simp add: divmod_nat_rel_def)
next
+ fix n :: nat show "n div 0 = 0"
+ by (simp add: div_nat_def Divides.divmod_nat_zero)
+next
+ fix m n :: nat
+ assume "n \<noteq> 0"
+ then show "m * n div n = m"
+ by (auto simp add: divmod_nat_rel_def intro: div_nat_unique [of _ _ _ 0])
+qed (simp_all add: unit_factor_nat_def)
+
+end
+
+instance nat :: semiring_div
+proof
fix m n q :: nat
assume "n \<noteq> 0"
then show "(q + m * n) div n = m + q div n"
@@ -976,48 +1020,33 @@
next
fix m n q :: nat
assume "m \<noteq> 0"
- hence "\<And>a b. divmod_nat_rel n q (a, b) \<Longrightarrow> divmod_nat_rel (m * n) (m * q) (a, m * b)"
- unfolding divmod_nat_rel_def
- by (auto split: if_split_asm, simp_all add: algebra_simps)
- moreover from divmod_nat_rel have "divmod_nat_rel n q (n div q, n mod q)" .
- ultimately have "divmod_nat_rel (m * n) (m * q) (n div q, m * (n mod q))" .
- thus "(m * n) div (m * q) = n div q" by (rule div_nat_unique)
-next
- fix n :: nat show "n div 0 = 0"
- by (simp add: div_nat_def Divides.divmod_nat_zero)
-next
- fix n :: nat show "0 div n = 0"
- by (simp add: div_nat_def Divides.divmod_nat_zero_left)
+ then have "divmod_nat_rel (m * n) (m * q) (n div q, m * (n mod q))"
+ using div_mult_mod_eq [of n q]
+ by (auto simp add: divmod_nat_rel_def algebra_simps distrib_left [symmetric] simp del: distrib_left)
+ then show "(m * n) div (m * q) = n div q"
+ by (rule div_nat_unique)
qed
-end
-
-instantiation nat :: normalization_semidom
-begin
-
-definition normalize_nat
- where [simp]: "normalize = (id :: nat \<Rightarrow> nat)"
-
-definition unit_factor_nat
- where "unit_factor n = (if n = 0 then 0 else 1 :: nat)"
-
-lemma unit_factor_simps [simp]:
- "unit_factor 0 = (0::nat)"
- "unit_factor (Suc n) = 1"
- by (simp_all add: unit_factor_nat_def)
-
-instance
- by standard (simp_all add: unit_factor_nat_def)
-
-end
-
-lemma divmod_nat_if [code]:
- "Divides.divmod_nat m n = (if n = 0 \<or> m < n then (0, m) else
- let (q, r) = Divides.divmod_nat (m - n) n in (Suc q, r))"
- by (simp add: prod_eq_iff case_prod_beta not_less le_div_geq le_mod_geq)
+lemma div_by_Suc_0 [simp]:
+ "m div Suc 0 = m"
+ using div_by_1 [of m] by simp
+
+lemma mod_by_Suc_0 [simp]:
+ "m mod Suc 0 = 0"
+ using mod_by_1 [of m] by simp
+
+lemma mod_greater_zero_iff_not_dvd:
+ fixes m n :: nat
+ shows "m mod n > 0 \<longleftrightarrow> \<not> n dvd m"
+ by (simp add: dvd_eq_mod_eq_0)
text \<open>Simproc for cancelling @{const divide} and @{const modulo}\<close>
+lemma (in semiring_modulo) cancel_div_mod_rules:
+ "((a div b) * b + a mod b) + c = a + c"
+ "(b * (a div b) + a mod b) + c = a + c"
+ by (simp_all add: div_mult_mod_eq mult_div_mod_eq)
+
ML_file "~~/src/Provers/Arith/cancel_div_mod.ML"
ML \<open>
@@ -1048,7 +1077,31 @@
)
\<close>
-simproc_setup cancel_div_mod_nat ("(m::nat) + n") = \<open>K Cancel_Div_Mod_Nat.proc\<close>
+simproc_setup cancel_div_mod_nat ("(m::nat) + n") =
+ \<open>K Cancel_Div_Mod_Nat.proc\<close>
+
+lemma divmod_nat_if [code]:
+ "Divides.divmod_nat m n = (if n = 0 \<or> m < n then (0, m) else
+ let (q, r) = Divides.divmod_nat (m - n) n in (Suc q, r))"
+ by (simp add: prod_eq_iff case_prod_beta not_less le_div_geq le_mod_geq)
+
+lemma mod_Suc_eq [mod_simps]:
+ "Suc (m mod n) mod n = Suc m mod n"
+proof -
+ have "(m mod n + 1) mod n = (m + 1) mod n"
+ by (simp only: mod_simps)
+ then show ?thesis
+ by simp
+qed
+
+lemma mod_Suc_Suc_eq [mod_simps]:
+ "Suc (Suc (m mod n)) mod n = Suc (Suc m) mod n"
+proof -
+ have "(m mod n + 2) mod n = (m + 2) mod n"
+ by (simp only: mod_simps)
+ then show ?thesis
+ by simp
+qed
subsubsection \<open>Quotient\<close>
@@ -1077,16 +1130,11 @@
qed
lemma div_eq_0_iff: "(a div b::nat) = 0 \<longleftrightarrow> a < b \<or> b = 0"
- by (metis div_less div_positive div_by_0 gr0I less_numeral_extra(3) not_less)
+ by auto (metis div_positive less_numeral_extra(3) not_less)
+
subsubsection \<open>Remainder\<close>
-lemma mod_less_divisor [simp]:
- fixes m n :: nat
- assumes "n > 0"
- shows "m mod n < (n::nat)"
- using assms divmod_nat_rel [of m n] unfolding divmod_nat_rel_def by auto
-
lemma mod_Suc_le_divisor [simp]:
"m mod Suc n \<le> n"
using mod_less_divisor [of "Suc n" m] by arith
@@ -1105,13 +1153,6 @@
lemma mod_if: "m mod (n::nat) = (if m < n then m else (m - n) mod n)"
by (simp add: le_mod_geq)
-lemma mod_by_Suc_0 [simp]: "m mod Suc 0 = 0"
-by (induct m) (simp_all add: mod_geq)
-
-lemma mod_le_divisor[simp]: "0 < n \<Longrightarrow> m mod n \<le> (n::nat)"
- apply (drule mod_less_divisor [where m = m])
- apply simp
- done
subsubsection \<open>Quotient and Remainder\<close>
@@ -1180,25 +1221,16 @@
subsubsection \<open>Further Facts about Quotient and Remainder\<close>
-lemma div_by_Suc_0 [simp]:
- "m div Suc 0 = m"
- using div_by_1 [of m] by simp
-
-(* Monotonicity of div in first argument *)
-lemma div_le_mono [rule_format (no_asm)]:
- "\<forall>m::nat. m \<le> n --> (m div k) \<le> (n div k)"
-apply (case_tac "k=0", simp)
-apply (induct "n" rule: nat_less_induct, clarify)
-apply (case_tac "n<k")
-(* 1 case n<k *)
-apply simp
-(* 2 case n >= k *)
-apply (case_tac "m<k")
-(* 2.1 case m<k *)
-apply simp
-(* 2.2 case m>=k *)
-apply (simp add: div_geq diff_le_mono)
-done
+lemma div_le_mono:
+ fixes m n k :: nat
+ assumes "m \<le> n"
+ shows "m div k \<le> n div k"
+proof -
+ from assms obtain q where "n = m + q"
+ by (auto simp add: le_iff_add)
+ then show ?thesis
+ by (simp add: div_add1_eq [of m q k])
+qed
(* Antimonotonicity of div in second argument *)
lemma div_le_mono2: "!!m::nat. [| 0<m; m\<le>n |] ==> (k div n) \<le> (k div m)"
@@ -1519,11 +1551,6 @@
declare Suc_times_mod_eq [of "numeral w", simp] for w
-lemma mod_greater_zero_iff_not_dvd:
- fixes m n :: nat
- shows "m mod n > 0 \<longleftrightarrow> \<not> n dvd m"
- by (simp add: dvd_eq_mod_eq_0)
-
lemma Suc_div_le_mono [simp]: "n div k \<le> (Suc n) div k"
by (simp add: div_le_mono)
@@ -1643,6 +1670,9 @@
subsection \<open>Division on @{typ int}\<close>
+context
+begin
+
definition divmod_int_rel :: "int \<Rightarrow> int \<Rightarrow> int \<times> int \<Rightarrow> bool" \<comment> \<open>definition of quotient and remainder\<close>
where "divmod_int_rel a b = (\<lambda>(q, r). a = b * q + r \<and>
(if 0 < b then 0 \<le> r \<and> r < b else if b < 0 then b < r \<and> r \<le> 0 else q = 0))"
@@ -1678,10 +1708,18 @@
apply (blast intro: unique_quotient)
done
-instantiation int :: modulo
+end
+
+instantiation int :: "{idom_modulo, normalization_semidom}"
begin
-definition divide_int
+definition normalize_int :: "int \<Rightarrow> int"
+ where [simp]: "normalize = (abs :: int \<Rightarrow> int)"
+
+definition unit_factor_int :: "int \<Rightarrow> int"
+ where [simp]: "unit_factor = (sgn :: int \<Rightarrow> int)"
+
+definition divide_int :: "int \<Rightarrow> int \<Rightarrow> int"
where "k div l = (if l = 0 \<or> k = 0 then 0
else if k > 0 \<and> l > 0 \<or> k < 0 \<and> l < 0
then int (nat \<bar>k\<bar> div nat \<bar>l\<bar>)
@@ -1689,32 +1727,35 @@
if l dvd k then - int (nat \<bar>k\<bar> div nat \<bar>l\<bar>)
else - int (Suc (nat \<bar>k\<bar> div nat \<bar>l\<bar>)))"
-definition modulo_int
+definition modulo_int :: "int \<Rightarrow> int \<Rightarrow> int"
where "k mod l = (if l = 0 then k else if l dvd k then 0
else if k > 0 \<and> l > 0 \<or> k < 0 \<and> l < 0
then sgn l * int (nat \<bar>k\<bar> mod nat \<bar>l\<bar>)
else sgn l * (\<bar>l\<bar> - int (nat \<bar>k\<bar> mod nat \<bar>l\<bar>)))"
-instance ..
-
-end
-
lemma divmod_int_rel:
"divmod_int_rel k l (k div l, k mod l)"
- unfolding divmod_int_rel_def divide_int_def modulo_int_def
- apply (cases k rule: int_cases3)
- apply (simp add: mod_greater_zero_iff_not_dvd not_le algebra_simps)
- apply (cases l rule: int_cases3)
- apply (simp add: mod_greater_zero_iff_not_dvd not_le algebra_simps)
- apply (simp_all del: of_nat_add of_nat_mult add: mod_greater_zero_iff_not_dvd not_le algebra_simps int_dvd_iff of_nat_add [symmetric] of_nat_mult [symmetric])
- apply (cases l rule: int_cases3)
- apply (simp_all del: of_nat_add of_nat_mult add: not_le algebra_simps int_dvd_iff of_nat_add [symmetric] of_nat_mult [symmetric])
- done
-
-instantiation int :: ring_div
-begin
-
-subsubsection \<open>Uniqueness and Monotonicity of Quotients and Remainders\<close>
+proof (cases k rule: int_cases3)
+ case zero
+ then show ?thesis
+ by (simp add: divmod_int_rel_def divide_int_def modulo_int_def)
+next
+ case (pos n)
+ then show ?thesis
+ using div_mult_mod_eq [of n]
+ by (cases l rule: int_cases3)
+ (auto simp del: of_nat_mult of_nat_add
+ simp add: mod_greater_zero_iff_not_dvd of_nat_mult [symmetric] of_nat_add [symmetric] algebra_simps
+ divmod_int_rel_def divide_int_def modulo_int_def int_dvd_iff)
+next
+ case (neg n)
+ then show ?thesis
+ using div_mult_mod_eq [of n]
+ by (cases l rule: int_cases3)
+ (auto simp del: of_nat_mult of_nat_add
+ simp add: mod_greater_zero_iff_not_dvd of_nat_mult [symmetric] of_nat_add [symmetric] algebra_simps
+ divmod_int_rel_def divide_int_def modulo_int_def int_dvd_iff)
+qed
lemma divmod_int_unique:
assumes "divmod_int_rel k l (q, r)"
@@ -1722,42 +1763,21 @@
using assms divmod_int_rel [of k l]
using unique_quotient [of k l] unique_remainder [of k l]
by auto
-
-instance
-proof
- fix a b :: int
- show "a div b * b + a mod b = a"
- using divmod_int_rel [of a b]
- unfolding divmod_int_rel_def by (simp add: mult.commute)
-next
- fix a b c :: int
- assume "b \<noteq> 0"
- hence "divmod_int_rel (a + c * b) b (c + a div b, a mod b)"
- using divmod_int_rel [of a b]
- unfolding divmod_int_rel_def by (auto simp: algebra_simps)
- thus "(a + c * b) div b = c + a div b"
- by (rule div_int_unique)
+
+instance proof
+ fix k l :: int
+ show "k div l * l + k mod l = k"
+ using divmod_int_rel [of k l]
+ unfolding divmod_int_rel_def by (simp add: ac_simps)
next
- fix a b c :: int
- assume c: "c \<noteq> 0"
- have "\<And>q r. divmod_int_rel a b (q, r)
- \<Longrightarrow> divmod_int_rel (c * a) (c * b) (q, c * r)"
- unfolding divmod_int_rel_def
- by (rule linorder_cases [of 0 b])
- (use c in \<open>auto simp: algebra_simps
- mult_less_0_iff zero_less_mult_iff mult_strict_right_mono
- mult_strict_right_mono_neg zero_le_mult_iff mult_le_0_iff\<close>)
- hence "divmod_int_rel (c * a) (c * b) (a div b, c * (a mod b))"
- using divmod_int_rel [of a b] .
- thus "(c * a) div (c * b) = a div b"
- by (rule div_int_unique)
-next
- fix a :: int show "a div 0 = 0"
+ fix k :: int show "k div 0 = 0"
by (rule div_int_unique, simp add: divmod_int_rel_def)
next
- fix a :: int show "0 div a = 0"
- by (rule div_int_unique, auto simp add: divmod_int_rel_def)
-qed
+ fix k l :: int
+ assume "l \<noteq> 0"
+ then show "k * l div l = k"
+ by (auto simp add: divmod_int_rel_def ac_simps intro: div_int_unique [of _ _ _ 0])
+qed (simp_all add: sgn_mult mult_sgn_abs abs_sgn_eq)
end
@@ -1765,36 +1785,30 @@
"is_unit (k::int) \<longleftrightarrow> k = 1 \<or> k = - 1"
by auto
-instantiation int :: normalization_semidom
-begin
-
-definition normalize_int
- where [simp]: "normalize = (abs :: int \<Rightarrow> int)"
-
-definition unit_factor_int
- where [simp]: "unit_factor = (sgn :: int \<Rightarrow> int)"
-
-instance
+instance int :: ring_div
proof
- fix k :: int
- assume "k \<noteq> 0"
- then have "\<bar>sgn k\<bar> = 1"
- by (cases "0::int" k rule: linorder_cases) simp_all
- then show "is_unit (unit_factor k)"
- by simp
-qed (simp_all add: sgn_mult mult_sgn_abs)
-
-end
-
-text\<open>Basic laws about division and remainder\<close>
-
-lemma zdiv_int: "int (a div b) = int a div int b"
- by (simp add: divide_int_def)
-
-lemma zmod_int: "int (a mod b) = int a mod int b"
- by (simp add: modulo_int_def int_dvd_iff)
-
-text \<open>Tool setup\<close>
+ fix k l s :: int
+ assume "l \<noteq> 0"
+ then have "divmod_int_rel (k + s * l) l (s + k div l, k mod l)"
+ using divmod_int_rel [of k l]
+ unfolding divmod_int_rel_def by (auto simp: algebra_simps)
+ then show "(k + s * l) div l = s + k div l"
+ by (rule div_int_unique)
+next
+ fix k l s :: int
+ assume "s \<noteq> 0"
+ have "\<And>q r. divmod_int_rel k l (q, r)
+ \<Longrightarrow> divmod_int_rel (s * k) (s * l) (q, s * r)"
+ unfolding divmod_int_rel_def
+ by (rule linorder_cases [of 0 l])
+ (use \<open>s \<noteq> 0\<close> in \<open>auto simp: algebra_simps
+ mult_less_0_iff zero_less_mult_iff mult_strict_right_mono
+ mult_strict_right_mono_neg zero_le_mult_iff mult_le_0_iff\<close>)
+ then have "divmod_int_rel (s * k) (s * l) (k div l, s * (k mod l))"
+ using divmod_int_rel [of k l] .
+ then show "(s * k) div (s * l) = k div l"
+ by (rule div_int_unique)
+qed
ML \<open>
structure Cancel_Div_Mod_Int = Cancel_Div_Mod
@@ -1807,12 +1821,22 @@
val div_mod_eqs = map mk_meta_eq @{thms cancel_div_mod_rules};
- val prove_eq_sums = Arith_Data.prove_conv2 all_tac
- (Arith_Data.simp_all_tac @{thms diff_conv_add_uminus add_0_left add_0_right ac_simps})
+ val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac
+ @{thms diff_conv_add_uminus add_0_left add_0_right ac_simps})
)
\<close>
-simproc_setup cancel_div_mod_int ("(k::int) + l") = \<open>K Cancel_Div_Mod_Int.proc\<close>
+simproc_setup cancel_div_mod_int ("(k::int) + l") =
+ \<open>K Cancel_Div_Mod_Int.proc\<close>
+
+
+text\<open>Basic laws about division and remainder\<close>
+
+lemma zdiv_int: "int (a div b) = int a div int b"
+ by (simp add: divide_int_def)
+
+lemma zmod_int: "int (a mod b) = int a mod int b"
+ by (simp add: modulo_int_def int_dvd_iff)
lemma pos_mod_conj: "(0::int) < b \<Longrightarrow> 0 \<le> a mod b \<and> a mod b < b"
using divmod_int_rel [of a b]
@@ -1890,7 +1914,12 @@
lemma zmod_zminus1_not_zero:
fixes k l :: int
shows "- k mod l \<noteq> 0 \<Longrightarrow> k mod l \<noteq> 0"
- unfolding zmod_zminus1_eq_if by auto
+ by (simp add: mod_eq_0_iff_dvd)
+
+lemma zmod_zminus2_not_zero:
+ fixes k l :: int
+ shows "k mod - l \<noteq> 0 \<Longrightarrow> k mod l \<noteq> 0"
+ by (simp add: mod_eq_0_iff_dvd)
lemma zdiv_zminus2_eq_if:
"b \<noteq> (0::int)
@@ -1902,11 +1931,6 @@
"a mod (-b::int) = (if a mod b = 0 then 0 else (a mod b) - b)"
by (simp add: zmod_zminus1_eq_if mod_minus_right)
-lemma zmod_zminus2_not_zero:
- fixes k l :: int
- shows "k mod - l \<noteq> 0 \<Longrightarrow> k mod l \<noteq> 0"
- unfolding zmod_zminus2_eq_if by auto
-
subsubsection \<open>Monotonicity in the First Argument (Dividend)\<close>
@@ -2121,7 +2145,7 @@
P(n div k :: int)(n mod k) = (\<forall>i j. 0\<le>j & j<k & n = k*i + j --> P i j)"
apply (rule iffI, clarify)
apply (erule_tac P="P x y" for x y in rev_mp)
- apply (subst mod_add_eq)
+ apply (subst mod_add_eq [symmetric])
apply (subst zdiv_zadd1_eq)
apply (simp add: div_pos_pos_trivial mod_pos_pos_trivial)
txt\<open>converse direction\<close>
@@ -2134,7 +2158,7 @@
P(n div k :: int)(n mod k) = (\<forall>i j. k<j & j\<le>0 & n = k*i + j --> P i j)"
apply (rule iffI, clarify)
apply (erule_tac P="P x y" for x y in rev_mp)
- apply (subst mod_add_eq)
+ apply (subst mod_add_eq [symmetric])
apply (subst zdiv_zadd1_eq)
apply (simp add: div_neg_neg_trivial mod_neg_neg_trivial)
txt\<open>converse direction\<close>
@@ -2454,24 +2478,6 @@
apply simp_all
done
-text \<open>by Brian Huffman\<close>
-lemma zminus_zmod: "- ((x::int) mod m) mod m = - x mod m"
-by (rule mod_minus_eq [symmetric])
-
-lemma zdiff_zmod_left: "(x mod m - y) mod m = (x - y) mod (m::int)"
-by (rule mod_diff_left_eq [symmetric])
-
-lemma zdiff_zmod_right: "(x - y mod m) mod m = (x - y) mod (m::int)"
-by (rule mod_diff_right_eq [symmetric])
-
-lemmas zmod_simps =
- mod_add_left_eq [symmetric]
- mod_add_right_eq [symmetric]
- mod_mult_right_eq[symmetric]
- mod_mult_left_eq [symmetric]
- power_mod
- zminus_zmod zdiff_zmod_left zdiff_zmod_right
-
text \<open>Distributive laws for function \<open>nat\<close>.\<close>
lemma nat_div_distrib: "0 \<le> x \<Longrightarrow> nat (x div y) = nat x div nat y"
@@ -2531,28 +2537,29 @@
apply (rule Divides.div_less_dividend, simp_all)
done
-lemma zmod_eq_dvd_iff: "(x::int) mod n = y mod n \<longleftrightarrow> n dvd x - y"
+lemma (in ring_div) mod_eq_dvd_iff:
+ "a mod c = b mod c \<longleftrightarrow> c dvd a - b" (is "?P \<longleftrightarrow> ?Q")
proof
- assume H: "x mod n = y mod n"
- hence "x mod n - y mod n = 0" by simp
- hence "(x mod n - y mod n) mod n = 0" by simp
- hence "(x - y) mod n = 0" by (simp add: mod_diff_eq[symmetric])
- thus "n dvd x - y" by (simp add: dvd_eq_mod_eq_0)
+ assume ?P
+ then have "(a mod c - b mod c) mod c = 0"
+ by simp
+ then show ?Q
+ by (simp add: dvd_eq_mod_eq_0 mod_simps)
next
- assume H: "n dvd x - y"
- then obtain k where k: "x-y = n*k" unfolding dvd_def by blast
- hence "x = n*k + y" by simp
- hence "x mod n = (n*k + y) mod n" by simp
- thus "x mod n = y mod n" by (simp add: mod_add_left_eq)
+ assume ?Q
+ then obtain d where d: "a - b = c * d" ..
+ then have "a = c * d + b"
+ by (simp add: algebra_simps)
+ then show ?P by simp
qed
-lemma nat_mod_eq_lemma: assumes xyn: "(x::nat) mod n = y mod n" and xy:"y \<le> x"
+lemma nat_mod_eq_lemma: assumes xyn: "(x::nat) mod n = y mod n" and xy:"y \<le> x"
shows "\<exists>q. x = y + n * q"
proof-
from xy have th: "int x - int y = int (x - y)" by simp
from xyn have "int x mod int n = int y mod int n"
by (simp add: zmod_int [symmetric])
- hence "int n dvd int x - int y" by (simp only: zmod_eq_dvd_iff[symmetric])
+ hence "int n dvd int x - int y" by (simp only: mod_eq_dvd_iff [symmetric])
hence "n dvd x - y" by (simp add: th zdvd_int)
then show ?thesis using xy unfolding dvd_def apply clarsimp apply (rule_tac x="k" in exI) by arith
qed
@@ -2666,6 +2673,4 @@
declare minus_div_mult_eq_mod [symmetric, nitpick_unfold]
-hide_fact (open) div_0 div_by_0
-
end
--- a/src/HOL/Enum.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Enum.thy Tue Dec 20 16:18:56 2016 +0100
@@ -683,7 +683,7 @@
instance finite_2 :: complete_linorder ..
-instantiation finite_2 :: "{field, ring_div, idom_abs_sgn}" begin
+instantiation finite_2 :: "{field, idom_abs_sgn}" begin
definition [simp]: "0 = a\<^sub>1"
definition [simp]: "1 = a\<^sub>2"
definition "x + y = (case (x, y) of (a\<^sub>1, a\<^sub>1) \<Rightarrow> a\<^sub>1 | (a\<^sub>2, a\<^sub>2) \<Rightarrow> a\<^sub>1 | _ \<Rightarrow> a\<^sub>2)"
@@ -693,19 +693,33 @@
definition "inverse = (\<lambda>x :: finite_2. x)"
definition "divide = (op * :: finite_2 \<Rightarrow> _)"
definition "abs = (\<lambda>x :: finite_2. x)"
-definition "x mod y = (case (x, y) of (a\<^sub>2, a\<^sub>1) \<Rightarrow> a\<^sub>2 | _ \<Rightarrow> a\<^sub>1)"
definition "sgn = (\<lambda>x :: finite_2. x)"
instance
-by intro_classes
- (simp_all add: plus_finite_2_def uminus_finite_2_def minus_finite_2_def times_finite_2_def
- inverse_finite_2_def divide_finite_2_def abs_finite_2_def modulo_finite_2_def sgn_finite_2_def
- split: finite_2.splits)
+ by standard
+ (simp_all add: plus_finite_2_def uminus_finite_2_def minus_finite_2_def times_finite_2_def
+ inverse_finite_2_def divide_finite_2_def abs_finite_2_def sgn_finite_2_def
+ split: finite_2.splits)
end
lemma two_finite_2 [simp]:
"2 = a\<^sub>1"
by (simp add: numeral.simps plus_finite_2_def)
-
+
+lemma dvd_finite_2_unfold:
+ "x dvd y \<longleftrightarrow> x = a\<^sub>2 \<or> y = a\<^sub>1"
+ by (auto simp add: dvd_def times_finite_2_def split: finite_2.splits)
+
+instantiation finite_2 :: "{ring_div, normalization_semidom}" begin
+definition [simp]: "normalize = (id :: finite_2 \<Rightarrow> _)"
+definition [simp]: "unit_factor = (id :: finite_2 \<Rightarrow> _)"
+definition "x mod y = (case (x, y) of (a\<^sub>2, a\<^sub>1) \<Rightarrow> a\<^sub>2 | _ \<Rightarrow> a\<^sub>1)"
+instance
+ by standard
+ (simp_all add: dvd_finite_2_unfold times_finite_2_def
+ divide_finite_2_def modulo_finite_2_def split: finite_2.splits)
+end
+
+
hide_const (open) a\<^sub>1 a\<^sub>2
datatype (plugins only: code "quickcheck" extraction) finite_3 =
@@ -736,6 +750,12 @@
end
+lemma finite_3_not_eq_unfold:
+ "x \<noteq> a\<^sub>1 \<longleftrightarrow> x \<in> {a\<^sub>2, a\<^sub>3}"
+ "x \<noteq> a\<^sub>2 \<longleftrightarrow> x \<in> {a\<^sub>1, a\<^sub>3}"
+ "x \<noteq> a\<^sub>3 \<longleftrightarrow> x \<in> {a\<^sub>1, a\<^sub>2}"
+ by (cases x; simp)+
+
instantiation finite_3 :: linorder
begin
@@ -806,7 +826,7 @@
instance finite_3 :: complete_linorder ..
-instantiation finite_3 :: "{field, ring_div, idom_abs_sgn}" begin
+instantiation finite_3 :: "{field, idom_abs_sgn}" begin
definition [simp]: "0 = a\<^sub>1"
definition [simp]: "1 = a\<^sub>2"
definition
@@ -820,14 +840,33 @@
definition "inverse = (\<lambda>x :: finite_3. x)"
definition "x div y = x * inverse (y :: finite_3)"
definition "abs = (\<lambda>x. case x of a\<^sub>3 \<Rightarrow> a\<^sub>2 | _ \<Rightarrow> x)"
-definition "x mod y = (case (x, y) of (a\<^sub>2, a\<^sub>1) \<Rightarrow> a\<^sub>2 | (a\<^sub>3, a\<^sub>1) \<Rightarrow> a\<^sub>3 | _ \<Rightarrow> a\<^sub>1)"
definition "sgn = (\<lambda>x :: finite_3. x)"
instance
-by intro_classes
- (simp_all add: plus_finite_3_def uminus_finite_3_def minus_finite_3_def times_finite_3_def
- inverse_finite_3_def divide_finite_3_def abs_finite_3_def modulo_finite_3_def sgn_finite_3_def
- less_finite_3_def
- split: finite_3.splits)
+ by standard
+ (simp_all add: plus_finite_3_def uminus_finite_3_def minus_finite_3_def times_finite_3_def
+ inverse_finite_3_def divide_finite_3_def abs_finite_3_def sgn_finite_3_def
+ less_finite_3_def
+ split: finite_3.splits)
+end
+
+lemma two_finite_3 [simp]:
+ "2 = a\<^sub>3"
+ by (simp add: numeral.simps plus_finite_3_def)
+
+lemma dvd_finite_3_unfold:
+ "x dvd y \<longleftrightarrow> x = a\<^sub>2 \<or> x = a\<^sub>3 \<or> y = a\<^sub>1"
+ by (cases x) (auto simp add: dvd_def times_finite_3_def split: finite_3.splits)
+
+instantiation finite_3 :: "{ring_div, normalization_semidom}" begin
+definition "normalize x = (case x of a\<^sub>3 \<Rightarrow> a\<^sub>2 | _ \<Rightarrow> x)"
+definition [simp]: "unit_factor = (id :: finite_3 \<Rightarrow> _)"
+definition "x mod y = (case (x, y) of (a\<^sub>2, a\<^sub>1) \<Rightarrow> a\<^sub>2 | (a\<^sub>3, a\<^sub>1) \<Rightarrow> a\<^sub>3 | _ \<Rightarrow> a\<^sub>1)"
+instance
+ by standard
+ (auto simp add: finite_3_not_eq_unfold plus_finite_3_def
+ dvd_finite_3_unfold times_finite_3_def inverse_finite_3_def
+ normalize_finite_3_def divide_finite_3_def modulo_finite_3_def
+ split: finite_3.splits)
end
--- a/src/HOL/Fields.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Fields.thy Tue Dec 20 16:18:56 2016 +0100
@@ -506,6 +506,21 @@
"y \<noteq> 0 \<Longrightarrow> z + x / y = (x + z * y) / y"
by (simp add: add_divide_distrib add.commute)
+lemma dvd_field_iff:
+ "a dvd b \<longleftrightarrow> (a = 0 \<longrightarrow> b = 0)"
+proof (cases "a = 0")
+ case True
+ then show ?thesis
+ by simp
+next
+ case False
+ then have "b = a * (b / a)"
+ by (simp add: field_simps)
+ then have "a dvd b" ..
+ with False show ?thesis
+ by simp
+qed
+
end
class field_char_0 = field + ring_char_0
--- a/src/HOL/Fun_Def.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Fun_Def.thy Tue Dec 20 16:18:56 2016 +0100
@@ -278,6 +278,16 @@
done
+subsection \<open>Yet another induction principle on the natural numbers\<close>
+
+lemma nat_descend_induct [case_names base descend]:
+ fixes P :: "nat \<Rightarrow> bool"
+ assumes H1: "\<And>k. k > n \<Longrightarrow> P k"
+ assumes H2: "\<And>k. k \<le> n \<Longrightarrow> (\<And>i. i > k \<Longrightarrow> P i) \<Longrightarrow> P k"
+ shows "P m"
+ using assms by induction_schema (force intro!: wf_measure [of "\<lambda>k. Suc n - k"])+
+
+
subsection \<open>Tool setup\<close>
ML_file "Tools/Function/termination.ML"
--- a/src/HOL/GCD.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/GCD.thy Tue Dec 20 16:18:56 2016 +0100
@@ -639,7 +639,6 @@
using dvd_times_left_cancel_iff [of "gcd a b" _ 1] by simp
qed
-
lemma divides_mult:
assumes "a dvd c" and nr: "b dvd c" and "coprime a b"
shows "a * b dvd c"
@@ -695,6 +694,10 @@
using coprime_rmult[of d a b] coprime_lmult[of d a b] coprime_mult[of d a b]
by blast
+lemma coprime_mul_eq':
+ "coprime (a * b) d \<longleftrightarrow> coprime a d \<and> coprime b d"
+ using coprime_mul_eq [of d a b] by (simp add: gcd.commute)
+
lemma gcd_coprime:
assumes c: "gcd a b \<noteq> 0"
and a: "a = a' * gcd a b"
@@ -958,6 +961,24 @@
ultimately show ?thesis by (rule that)
qed
+lemma coprime_crossproduct':
+ fixes a b c d
+ assumes "b \<noteq> 0"
+ assumes unit_factors: "unit_factor b = unit_factor d"
+ assumes coprime: "coprime a b" "coprime c d"
+ shows "a * d = b * c \<longleftrightarrow> a = c \<and> b = d"
+proof safe
+ assume eq: "a * d = b * c"
+ hence "normalize a * normalize d = normalize c * normalize b"
+ by (simp only: normalize_mult [symmetric] mult_ac)
+ with coprime have "normalize b = normalize d"
+ by (subst (asm) coprime_crossproduct) simp_all
+ from this and unit_factors show "b = d"
+ by (rule normalize_unit_factor_eqI)
+ from eq have "a * d = c * d" by (simp only: \<open>b = d\<close> mult_ac)
+ with \<open>b \<noteq> 0\<close> \<open>b = d\<close> show "a = c" by simp
+qed (simp_all add: mult_ac)
+
end
class ring_gcd = comm_ring_1 + semiring_gcd
--- a/src/HOL/Groebner_Basis.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Groebner_Basis.thy Tue Dec 20 16:18:56 2016 +0100
@@ -72,7 +72,7 @@
declare zmod_eq_0_iff[algebra]
declare dvd_0_left_iff[algebra]
declare zdvd1_eq[algebra]
-declare zmod_eq_dvd_iff[algebra]
+declare mod_eq_dvd_iff[algebra]
declare nat_mod_eq_iff[algebra]
context semiring_parity
--- a/src/HOL/Hilbert_Choice.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Hilbert_Choice.thy Tue Dec 20 16:18:56 2016 +0100
@@ -657,6 +657,12 @@
for x :: nat
unfolding Greatest_def by (rule GreatestM_nat_le) auto
+lemma GreatestI_ex: "\<exists>k::nat. P k \<Longrightarrow> \<forall>y. P y \<longrightarrow> y < b \<Longrightarrow> P (GREATEST x. P x)"
+ apply (erule exE)
+ apply (rule GreatestI)
+ apply assumption+
+ done
+
subsection \<open>An aside: bounded accessible part\<close>
--- a/src/HOL/IMP/Abs_Int1_parity.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/IMP/Abs_Int1_parity.thy Tue Dec 20 16:18:56 2016 +0100
@@ -112,7 +112,8 @@
case 3 show ?case by auto
next
case (4 _ a1 _ a2) thus ?case
- by (induction a1 a2 rule: plus_parity.induct) (auto simp add:mod_add_eq)
+ by (induction a1 a2 rule: plus_parity.induct)
+ (auto simp add: mod_add_eq [symmetric])
qed
text{* In case 4 we needed to refer to particular variables.
--- a/src/HOL/Library/Code_Test.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Library/Code_Test.thy Tue Dec 20 16:18:56 2016 +0100
@@ -1,7 +1,7 @@
(* Title: HOL/Library/Code_Test.thy
- Author: Andreas Lochbihler, ETH Zurich
+ Author: Andreas Lochbihler, ETH Zürich
-Test infrastructure for the code generator
+Test infrastructure for the code generator.
*)
theory Code_Test
@@ -100,7 +100,7 @@
"yxml_string_of_xml_tree (xml.Elem name atts ts) rest =
yot_append xml.XY (
yot_append (yot_literal name) (
- foldr (\<lambda>(a, x) rest.
+ foldr (\<lambda>(a, x) rest.
yot_append xml.Y (
yot_append (yot_literal a) (
yot_append (yot_literal (STR ''='')) (
--- a/src/HOL/Library/DAList_Multiset.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Library/DAList_Multiset.thy Tue Dec 20 16:18:56 2016 +0100
@@ -228,17 +228,17 @@
by (rule multiset_eqI) (simp add: count_of_filter DAList.filter.rep_eq)
-lemma mset_eq [code]: "HOL.equal (m1::'a::equal multiset) m2 \<longleftrightarrow> m1 \<le># m2 \<and> m2 \<le># m1"
+lemma mset_eq [code]: "HOL.equal (m1::'a::equal multiset) m2 \<longleftrightarrow> m1 \<subseteq># m2 \<and> m2 \<subseteq># m1"
by (metis equal_multiset_def subset_mset.eq_iff)
text \<open>By default the code for \<open><\<close> is @{prop"xs < ys \<longleftrightarrow> xs \<le> ys \<and> \<not> xs = ys"}.
With equality implemented by \<open>\<le>\<close>, this leads to three calls of \<open>\<le>\<close>.
Here is a more efficient version:\<close>
-lemma mset_less[code]: "xs <# (ys :: 'a multiset) \<longleftrightarrow> xs \<le># ys \<and> \<not> ys \<le># xs"
+lemma mset_less[code]: "xs \<subset># (ys :: 'a multiset) \<longleftrightarrow> xs \<subseteq># ys \<and> \<not> ys \<subseteq># xs"
by (rule subset_mset.less_le_not_le)
lemma mset_less_eq_Bag0:
- "Bag xs \<le># A \<longleftrightarrow> (\<forall>(x, n) \<in> set (DAList.impl_of xs). count_of (DAList.impl_of xs) x \<le> count A x)"
+ "Bag xs \<subseteq># A \<longleftrightarrow> (\<forall>(x, n) \<in> set (DAList.impl_of xs). count_of (DAList.impl_of xs) x \<le> count A x)"
(is "?lhs \<longleftrightarrow> ?rhs")
proof
assume ?lhs
@@ -255,7 +255,7 @@
qed
lemma mset_less_eq_Bag [code]:
- "Bag xs \<le># (A :: 'a multiset) \<longleftrightarrow> (\<forall>(x, n) \<in> set (DAList.impl_of xs). n \<le> count A x)"
+ "Bag xs \<subseteq># (A :: 'a multiset) \<longleftrightarrow> (\<forall>(x, n) \<in> set (DAList.impl_of xs). n \<le> count A x)"
proof -
{
fix x n
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Field_as_Ring.thy Tue Dec 20 16:18:56 2016 +0100
@@ -0,0 +1,108 @@
+(* Title: HOL/Library/Field_as_Ring.thy
+ Author: Manuel Eberl
+*)
+
+theory Field_as_Ring
+imports
+ Complex_Main
+ "~~/src/HOL/Number_Theory/Euclidean_Algorithm"
+begin
+
+context field
+begin
+
+subclass idom_divide ..
+
+definition normalize_field :: "'a \<Rightarrow> 'a"
+ where [simp]: "normalize_field x = (if x = 0 then 0 else 1)"
+definition unit_factor_field :: "'a \<Rightarrow> 'a"
+ where [simp]: "unit_factor_field x = x"
+definition euclidean_size_field :: "'a \<Rightarrow> nat"
+ where [simp]: "euclidean_size_field x = (if x = 0 then 0 else 1)"
+definition mod_field :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
+ where [simp]: "mod_field x y = (if y = 0 then x else 0)"
+
+end
+
+instantiation real :: euclidean_ring
+begin
+
+definition [simp]: "normalize_real = (normalize_field :: real \<Rightarrow> _)"
+definition [simp]: "unit_factor_real = (unit_factor_field :: real \<Rightarrow> _)"
+definition [simp]: "euclidean_size_real = (euclidean_size_field :: real \<Rightarrow> _)"
+definition [simp]: "modulo_real = (mod_field :: real \<Rightarrow> _)"
+
+instance by standard (simp_all add: dvd_field_iff divide_simps)
+end
+
+instantiation real :: euclidean_ring_gcd
+begin
+
+definition gcd_real :: "real \<Rightarrow> real \<Rightarrow> real" where
+ "gcd_real = gcd_eucl"
+definition lcm_real :: "real \<Rightarrow> real \<Rightarrow> real" where
+ "lcm_real = lcm_eucl"
+definition Gcd_real :: "real set \<Rightarrow> real" where
+ "Gcd_real = Gcd_eucl"
+definition Lcm_real :: "real set \<Rightarrow> real" where
+ "Lcm_real = Lcm_eucl"
+
+instance by standard (simp_all add: gcd_real_def lcm_real_def Gcd_real_def Lcm_real_def)
+
+end
+
+instantiation rat :: euclidean_ring
+begin
+
+definition [simp]: "normalize_rat = (normalize_field :: rat \<Rightarrow> _)"
+definition [simp]: "unit_factor_rat = (unit_factor_field :: rat \<Rightarrow> _)"
+definition [simp]: "euclidean_size_rat = (euclidean_size_field :: rat \<Rightarrow> _)"
+definition [simp]: "modulo_rat = (mod_field :: rat \<Rightarrow> _)"
+
+instance by standard (simp_all add: dvd_field_iff divide_simps)
+end
+
+instantiation rat :: euclidean_ring_gcd
+begin
+
+definition gcd_rat :: "rat \<Rightarrow> rat \<Rightarrow> rat" where
+ "gcd_rat = gcd_eucl"
+definition lcm_rat :: "rat \<Rightarrow> rat \<Rightarrow> rat" where
+ "lcm_rat = lcm_eucl"
+definition Gcd_rat :: "rat set \<Rightarrow> rat" where
+ "Gcd_rat = Gcd_eucl"
+definition Lcm_rat :: "rat set \<Rightarrow> rat" where
+ "Lcm_rat = Lcm_eucl"
+
+instance by standard (simp_all add: gcd_rat_def lcm_rat_def Gcd_rat_def Lcm_rat_def)
+
+end
+
+instantiation complex :: euclidean_ring
+begin
+
+definition [simp]: "normalize_complex = (normalize_field :: complex \<Rightarrow> _)"
+definition [simp]: "unit_factor_complex = (unit_factor_field :: complex \<Rightarrow> _)"
+definition [simp]: "euclidean_size_complex = (euclidean_size_field :: complex \<Rightarrow> _)"
+definition [simp]: "modulo_complex = (mod_field :: complex \<Rightarrow> _)"
+
+instance by standard (simp_all add: dvd_field_iff divide_simps)
+end
+
+instantiation complex :: euclidean_ring_gcd
+begin
+
+definition gcd_complex :: "complex \<Rightarrow> complex \<Rightarrow> complex" where
+ "gcd_complex = gcd_eucl"
+definition lcm_complex :: "complex \<Rightarrow> complex \<Rightarrow> complex" where
+ "lcm_complex = lcm_eucl"
+definition Gcd_complex :: "complex set \<Rightarrow> complex" where
+ "Gcd_complex = Gcd_eucl"
+definition Lcm_complex :: "complex set \<Rightarrow> complex" where
+ "Lcm_complex = Lcm_eucl"
+
+instance by standard (simp_all add: gcd_complex_def lcm_complex_def Gcd_complex_def Lcm_complex_def)
+
+end
+
+end
--- a/src/HOL/Library/Formal_Power_Series.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Library/Formal_Power_Series.thy Tue Dec 20 16:18:56 2016 +0100
@@ -1157,7 +1157,40 @@
lemma fps_unit_dvd [simp]: "(f $ 0 :: 'a :: field) \<noteq> 0 \<Longrightarrow> f dvd g"
by (rule dvd_trans, subst fps_is_unit_iff) simp_all
-
+instantiation fps :: (field) normalization_semidom
+begin
+
+definition fps_unit_factor_def [simp]:
+ "unit_factor f = fps_shift (subdegree f) f"
+
+definition fps_normalize_def [simp]:
+ "normalize f = (if f = 0 then 0 else X ^ subdegree f)"
+
+instance proof
+ fix f :: "'a fps"
+ show "unit_factor f * normalize f = f"
+ by (simp add: fps_shift_times_X_power)
+next
+ fix f g :: "'a fps"
+ show "unit_factor (f * g) = unit_factor f * unit_factor g"
+ proof (cases "f = 0 \<or> g = 0")
+ assume "\<not>(f = 0 \<or> g = 0)"
+ thus "unit_factor (f * g) = unit_factor f * unit_factor g"
+ unfolding fps_unit_factor_def
+ by (auto simp: fps_shift_fps_shift fps_shift_mult fps_shift_mult_right)
+ qed auto
+next
+ fix f g :: "'a fps"
+ assume "g \<noteq> 0"
+ then have "f * (fps_shift (subdegree g) g * inverse (fps_shift (subdegree g) g)) = f"
+ by (metis add_cancel_right_left fps_shift_nth inverse_mult_eq_1 mult.commute mult_cancel_left2 nth_subdegree_nonzero)
+ then have "fps_shift (subdegree g) (g * (f * inverse (fps_shift (subdegree g) g))) = f"
+ by (simp add: fps_shift_mult_right mult.commute)
+ with \<open>g \<noteq> 0\<close> show "f * g / g = f"
+ by (simp add: fps_divide_def Let_def ac_simps)
+qed (auto simp add: fps_divide_def Let_def)
+
+end
instantiation fps :: (field) ring_div
begin
@@ -1291,7 +1324,7 @@
also have "fps_shift n (f * inverse h') = f div h"
by (simp add: fps_divide_def Let_def dfs)
finally show "(f + g * h) div h = g + f div h" by simp
-qed (auto simp: fps_divide_def fps_mod_def Let_def)
+qed
end
end
@@ -1365,36 +1398,6 @@
fps_numeral_divide_divide fps_numeral_mult_divide inverse_fps_numeral neg_numeral_fps_const
-
-instantiation fps :: (field) normalization_semidom
-begin
-
-definition fps_unit_factor_def [simp]:
- "unit_factor f = fps_shift (subdegree f) f"
-
-definition fps_normalize_def [simp]:
- "normalize f = (if f = 0 then 0 else X ^ subdegree f)"
-
-instance proof
- fix f :: "'a fps"
- show "unit_factor f * normalize f = f"
- by (simp add: fps_shift_times_X_power)
-next
- fix f g :: "'a fps"
- show "unit_factor (f * g) = unit_factor f * unit_factor g"
- proof (cases "f = 0 \<or> g = 0")
- assume "\<not>(f = 0 \<or> g = 0)"
- thus "unit_factor (f * g) = unit_factor f * unit_factor g"
- unfolding fps_unit_factor_def
- by (auto simp: fps_shift_fps_shift fps_shift_mult fps_shift_mult_right)
- qed auto
-qed auto
-
-end
-
-instance fps :: (field) algebraic_semidom ..
-
-
subsection \<open>Formal power series form a Euclidean ring\<close>
instantiation fps :: (field) euclidean_ring
--- a/src/HOL/Library/Library.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Library/Library.thy Tue Dec 20 16:18:56 2016 +0100
@@ -49,12 +49,14 @@
More_List
Multiset_Order
Multiset_Permutations
+ Nonpos_Ints
Numeral_Type
Omega_Words_Fun
OptionalSugar
Option_ord
Order_Continuity
Parallel
+ Periodic_Fun
Perm
Permutation
Permutations
@@ -72,6 +74,7 @@
Quotient_Type
Ramsey
Reflection
+ Rewrite
Saturated
Set_Algebras
State_Monad
--- a/src/HOL/Library/Multiset.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Library/Multiset.thy Tue Dec 20 16:18:56 2016 +0100
@@ -526,9 +526,11 @@
interpretation subset_mset: ordered_ab_semigroup_add_imp_le "op +" "op -" "op \<subseteq>#" "op \<subset>#"
by standard (auto simp add: subset_mset_def subseteq_mset_def multiset_eq_iff intro: order_trans antisym)
-
-interpretation subset_mset: ordered_ab_semigroup_monoid_add_imp_le "op +" 0 "op -" "op \<le>#" "op <#"
+ \<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close>
+
+interpretation subset_mset: ordered_ab_semigroup_monoid_add_imp_le "op +" 0 "op -" "op \<subseteq>#" "op \<subset>#"
by standard
+ \<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close>
lemma mset_subset_eqI:
"(\<And>a. count A a \<le> count B a) \<Longrightarrow> A \<subseteq># B"
@@ -545,8 +547,9 @@
apply (auto intro: multiset_eq_iff [THEN iffD2])
done
-interpretation subset_mset: ordered_cancel_comm_monoid_diff "op +" 0 "op \<le>#" "op <#" "op -"
+interpretation subset_mset: ordered_cancel_comm_monoid_diff "op +" 0 "op \<subseteq>#" "op \<subset>#" "op -"
by standard (simp, fact mset_subset_eq_exists_conv)
+ \<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close>
declare subset_mset.add_diff_assoc[simp] subset_mset.add_diff_assoc2[simp]
@@ -625,8 +628,8 @@
lemma mset_subset_of_empty[simp]: "A \<subset># {#} \<longleftrightarrow> False"
by (simp only: subset_mset.not_less_zero)
-lemma empty_subset_add_mset[simp]: "{#} <# add_mset x M"
-by(auto intro: subset_mset.gr_zeroI)
+lemma empty_subset_add_mset[simp]: "{#} \<subset># add_mset x M"
+ by (auto intro: subset_mset.gr_zeroI)
lemma empty_le: "{#} \<subseteq># A"
by (fact subset_mset.zero_le)
@@ -684,8 +687,7 @@
by arith
show "class.semilattice_inf op \<inter># op \<subseteq># op \<subset>#"
by standard (auto simp add: multiset_inter_def subseteq_mset_def)
-qed
- \<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close>
+qed \<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close>
definition sup_subset_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset"(infixl "\<union>#" 70)
where "sup_subset_mset A B = A + (B - A)" \<comment> \<open>FIXME irregular fact name\<close>
@@ -696,12 +698,12 @@
by arith
show "class.semilattice_sup op \<union># op \<subseteq># op \<subset>#"
by standard (auto simp add: sup_subset_mset_def subseteq_mset_def)
-qed
- \<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close>
+qed \<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close>
interpretation subset_mset: bounded_lattice_bot "op \<inter>#" "op \<subseteq>#" "op \<subset>#"
"op \<union>#" "{#}"
by standard auto
+ \<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close>
subsubsection \<open>Additional intersection facts\<close>
@@ -885,11 +887,6 @@
by (auto simp: multiset_eq_iff max_def)
-subsubsection \<open>Subset is an order\<close>
-
-interpretation subset_mset: order "op \<subseteq>#" "op \<subset>#" by unfold_locales
-
-
subsection \<open>Replicate and repeat operations\<close>
definition replicate_mset :: "nat \<Rightarrow> 'a \<Rightarrow> 'a multiset" where
@@ -1161,7 +1158,7 @@
by (intro cSup_least) (auto intro: mset_subset_eq_count ge)
finally show "count (Sup A) x \<le> count X x" .
qed
-qed
+qed \<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close>
lemma set_mset_Inf:
assumes "A \<noteq> {}"
@@ -1239,7 +1236,7 @@
fix A B C :: "'a multiset"
show "A \<union># (B \<inter># C) = A \<union># B \<inter># (A \<union># C)"
by (intro multiset_eqI) simp_all
-qed
+qed \<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close>
subsubsection \<open>Filter (with comprehension syntax)\<close>
@@ -1741,6 +1738,10 @@
"(\<forall>x y. (x, y) \<in># M \<longrightarrow> f x y = g x y) \<Longrightarrow> {#f x y. (x, y) \<in># M#} = {#g x y. (x, y) \<in># M#}"
by (metis image_mset_cong split_cong)
+lemma image_mset_const_eq:
+ "{#c. a \<in># M#} = replicate_mset (size M) c"
+ by (induct M) simp_all
+
subsection \<open>Further conversions\<close>
@@ -2313,6 +2314,9 @@
translations
"\<Prod>i \<in># A. b" \<rightleftharpoons> "CONST prod_mset (CONST image_mset (\<lambda>i. b) A)"
+lemma prod_mset_constant [simp]: "(\<Prod>_\<in>#A. c) = c ^ size A"
+ by (simp add: image_mset_const_eq)
+
lemma (in comm_monoid_mult) prod_mset_subset_imp_dvd:
assumes "A \<subseteq># B"
shows "prod_mset A dvd prod_mset B"
--- a/src/HOL/Library/Multiset_Order.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Library/Multiset_Order.thy Tue Dec 20 16:18:56 2016 +0100
@@ -49,7 +49,7 @@
definition less_multiset\<^sub>D\<^sub>M where
"less_multiset\<^sub>D\<^sub>M M N \<longleftrightarrow>
- (\<exists>X Y. X \<noteq> {#} \<and> X \<le># N \<and> M = (N - X) + Y \<and> (\<forall>k. k \<in># Y \<longrightarrow> (\<exists>a. a \<in># X \<and> k < a)))"
+ (\<exists>X Y. X \<noteq> {#} \<and> X \<subseteq># N \<and> M = (N - X) + Y \<and> (\<forall>k. k \<in># Y \<longrightarrow> (\<exists>a. a \<in># X \<and> k < a)))"
text \<open>The Huet--Oppen ordering:\<close>
@@ -123,11 +123,11 @@
proof -
assume "less_multiset\<^sub>D\<^sub>M M N"
then obtain X Y where
- "X \<noteq> {#}" and "X \<le># N" and "M = N - X + Y" and "\<forall>k. k \<in># Y \<longrightarrow> (\<exists>a. a \<in># X \<and> k < a)"
+ "X \<noteq> {#}" and "X \<subseteq># N" and "M = N - X + Y" and "\<forall>k. k \<in># Y \<longrightarrow> (\<exists>a. a \<in># X \<and> k < a)"
unfolding less_multiset\<^sub>D\<^sub>M_def by blast
then have "(N - X + Y, N - X + X) \<in> mult {(x, y). x < y}"
by (intro one_step_implies_mult) (auto simp: Bex_def trans_def)
- with \<open>M = N - X + Y\<close> \<open>X \<le># N\<close> show "(M, N) \<in> mult {(x, y). x < y}"
+ with \<open>M = N - X + Y\<close> \<open>X \<subseteq># N\<close> show "(M, N) \<in> mult {(x, y). x < y}"
by (metis subset_mset.diff_add)
qed
@@ -140,7 +140,7 @@
define X where "X = N - M"
define Y where "Y = M - N"
from z show "X \<noteq> {#}" unfolding X_def by (auto simp: multiset_eq_iff not_less_eq_eq Suc_le_eq)
- from z show "X \<le># N" unfolding X_def by auto
+ from z show "X \<subseteq># N" unfolding X_def by auto
show "M = (N - X) + Y" unfolding X_def Y_def multiset_eq_iff count_union count_diff by force
show "\<forall>k. k \<in># Y \<longrightarrow> (\<exists>a. a \<in># X \<and> k < a)"
proof (intro allI impI)
@@ -175,7 +175,7 @@
lemmas less_multiset\<^sub>H\<^sub>O = mult\<^sub>H\<^sub>O[folded less_multiset_def]
lemma subset_eq_imp_le_multiset:
- shows "M \<le># N \<Longrightarrow> M \<le> N"
+ shows "M \<subseteq># N \<Longrightarrow> M \<le> N"
unfolding less_eq_multiset_def less_multiset\<^sub>H\<^sub>O
by (simp add: less_le_not_le subseteq_mset_def)
@@ -201,7 +201,7 @@
lemma le_multiset_empty_right[simp]: "\<not> M < {#}"
using subset_mset.le_zero_eq less_multiset\<^sub>D\<^sub>M by blast
-lemma union_le_diff_plus: "P \<le># M \<Longrightarrow> N < P \<Longrightarrow> M - P + N < M"
+lemma union_le_diff_plus: "P \<subseteq># M \<Longrightarrow> N < P \<Longrightarrow> M - P + N < M"
by (drule subset_mset.diff_add[symmetric]) (metis union_le_mono2)
instantiation multiset :: (preorder) ordered_ab_semigroup_monoid_add_imp_le
--- a/src/HOL/Library/Normalized_Fraction.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Library/Normalized_Fraction.thy Tue Dec 20 16:18:56 2016 +0100
@@ -1,3 +1,7 @@
+(* Title: HOL/Library/Normalized_Fraction.thy
+ Author: Manuel Eberl
+*)
+
theory Normalized_Fraction
imports
Main
@@ -5,75 +9,6 @@
"~~/src/HOL/Library/Fraction_Field"
begin
-lemma dvd_neg_div': "y dvd (x :: 'a :: idom_divide) \<Longrightarrow> -x div y = - (x div y)"
-apply (case_tac "y = 0") apply simp
-apply (auto simp add: dvd_def)
-apply (subgoal_tac "-(y * k) = y * - k")
-apply (simp only:)
-apply (erule nonzero_mult_div_cancel_left)
-apply simp
-done
-
-(* TODO Move *)
-lemma (in semiring_gcd) coprime_mul_eq': "coprime (a * b) d \<longleftrightarrow> coprime a d \<and> coprime b d"
- using coprime_mul_eq[of d a b] by (simp add: gcd.commute)
-
-lemma dvd_div_eq_0_iff:
- assumes "b dvd (a :: 'a :: semidom_divide)"
- shows "a div b = 0 \<longleftrightarrow> a = 0"
- using assms by (elim dvdE, cases "b = 0") simp_all
-
-lemma dvd_div_eq_0_iff':
- assumes "b dvd (a :: 'a :: semiring_div)"
- shows "a div b = 0 \<longleftrightarrow> a = 0"
- using assms by (elim dvdE, cases "b = 0") simp_all
-
-lemma unit_div_eq_0_iff:
- assumes "is_unit (b :: 'a :: {algebraic_semidom,semidom_divide})"
- shows "a div b = 0 \<longleftrightarrow> a = 0"
- by (rule dvd_div_eq_0_iff) (insert assms, auto)
-
-lemma unit_div_eq_0_iff':
- assumes "is_unit (b :: 'a :: semiring_div)"
- shows "a div b = 0 \<longleftrightarrow> a = 0"
- by (rule dvd_div_eq_0_iff) (insert assms, auto)
-
-lemma dvd_div_eq_cancel:
- "a div c = b div c \<Longrightarrow> (c :: 'a :: semiring_div) dvd a \<Longrightarrow> c dvd b \<Longrightarrow> a = b"
- by (elim dvdE, cases "c = 0") simp_all
-
-lemma dvd_div_eq_iff:
- "(c :: 'a :: semiring_div) dvd a \<Longrightarrow> c dvd b \<Longrightarrow> a div c = b div c \<longleftrightarrow> a = b"
- by (elim dvdE, cases "c = 0") simp_all
-
-lemma normalize_imp_eq:
- "normalize a = normalize b \<Longrightarrow> unit_factor a = unit_factor b \<Longrightarrow> a = b"
- by (cases "a = 0 \<or> b = 0")
- (auto simp add: div_unit_factor [symmetric] unit_div_cancel simp del: div_unit_factor)
-
-lemma coprime_crossproduct':
- fixes a b c d :: "'a :: semiring_gcd"
- assumes nz: "b \<noteq> 0"
- assumes unit_factors: "unit_factor b = unit_factor d"
- assumes coprime: "coprime a b" "coprime c d"
- shows "a * d = b * c \<longleftrightarrow> a = c \<and> b = d"
-proof safe
- assume eq: "a * d = b * c"
- hence "normalize a * normalize d = normalize c * normalize b"
- by (simp only: normalize_mult [symmetric] mult_ac)
- with coprime have "normalize b = normalize d"
- by (subst (asm) coprime_crossproduct) simp_all
- from this and unit_factors show "b = d" by (rule normalize_imp_eq)
- from eq have "a * d = c * d" by (simp only: \<open>b = d\<close> mult_ac)
- with nz \<open>b = d\<close> show "a = c" by simp
-qed (simp_all add: mult_ac)
-
-
-lemma div_mult_unit2: "is_unit c \<Longrightarrow> b dvd a \<Longrightarrow> a div (b * c) = a div b div c"
- by (subst dvd_div_mult2_eq) (simp_all add: mult_unit_dvd_iff)
-(* END TODO *)
-
-
definition quot_to_fract :: "'a :: {idom} \<times> 'a \<Rightarrow> 'a fract" where
"quot_to_fract = (\<lambda>(a,b). Fraction_Field.Fract a b)"
@@ -249,7 +184,7 @@
lemma quot_of_fract_uminus:
"quot_of_fract (-x) = (let (a,b) = quot_of_fract x in (-a, b))"
- by transfer (auto simp: case_prod_unfold Let_def normalize_quot_def dvd_neg_div' mult_unit_dvd_iff)
+ by transfer (auto simp: case_prod_unfold Let_def normalize_quot_def dvd_neg_div mult_unit_dvd_iff)
lemma quot_of_fract_diff:
"quot_of_fract (x - y) =
--- a/src/HOL/Library/Numeral_Type.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Library/Numeral_Type.thy Tue Dec 20 16:18:56 2016 +0100
@@ -133,7 +133,7 @@
lemma comm_ring_1: "OFCLASS('a, comm_ring_1_class)"
apply (intro_classes, unfold definitions)
-apply (simp_all add: Rep_simps zmod_simps field_simps)
+apply (simp_all add: Rep_simps mod_simps field_simps)
done
end
@@ -147,12 +147,12 @@
lemma of_nat_eq: "of_nat k = Abs (int k mod n)"
apply (induct k)
apply (simp add: zero_def)
-apply (simp add: Rep_simps add_def one_def zmod_simps ac_simps)
+apply (simp add: Rep_simps add_def one_def mod_simps ac_simps)
done
lemma of_int_eq: "of_int z = Abs (z mod n)"
apply (cases z rule: int_diff_cases)
-apply (simp add: Rep_simps of_nat_eq diff_def zmod_simps)
+apply (simp add: Rep_simps of_nat_eq diff_def mod_simps)
done
lemma Rep_numeral:
--- a/src/HOL/Library/Omega_Words_Fun.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Library/Omega_Words_Fun.thy Tue Dec 20 16:18:56 2016 +0100
@@ -626,7 +626,7 @@
by (auto simp add: set_conv_nth)
\<comment> "the following bound is terrible, but it simplifies the proof"
from nempty k have "\<forall>m. w\<^sup>\<omega> ((Suc m)*(length w) + k) = a"
- by (simp add: mod_add_left_eq)
+ by (simp add: mod_add_left_eq [symmetric])
moreover
\<comment> "why is the following so hard to prove??"
have "\<forall>m. m < (Suc m)*(length w) + k"
--- a/src/HOL/Library/Permutation.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Library/Permutation.thy Tue Dec 20 16:18:56 2016 +0100
@@ -134,7 +134,7 @@
apply simp
done
-proposition mset_le_perm_append: "mset xs \<le># mset ys \<longleftrightarrow> (\<exists>zs. xs @ zs <~~> ys)"
+proposition mset_le_perm_append: "mset xs \<subseteq># mset ys \<longleftrightarrow> (\<exists>zs. xs @ zs <~~> ys)"
apply (auto simp: mset_eq_perm[THEN sym] mset_subset_eq_exists_conv)
apply (insert surj_mset)
apply (drule surjD)
--- a/src/HOL/Library/Polynomial.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Library/Polynomial.thy Tue Dec 20 16:18:56 2016 +0100
@@ -877,7 +877,7 @@
by (induct n, simp add: monom_0, simp add: monom_Suc)
lemma smult_Poly: "smult c (Poly xs) = Poly (map (op * c) xs)"
- by (auto simp add: poly_eq_iff coeff_Poly_eq nth_default_def)
+ by (auto simp add: poly_eq_iff nth_default_def)
lemma degree_smult_eq [simp]:
fixes a :: "'a::{comm_semiring_0,semiring_no_zero_divisors}"
@@ -1064,6 +1064,111 @@
by (rule le_trans[OF degree_mult_le], insert insert, auto)
qed simp
+
+subsection \<open>Mapping polynomials\<close>
+
+definition map_poly
+ :: "('a :: zero \<Rightarrow> 'b :: zero) \<Rightarrow> 'a poly \<Rightarrow> 'b poly" where
+ "map_poly f p = Poly (map f (coeffs p))"
+
+lemma map_poly_0 [simp]: "map_poly f 0 = 0"
+ by (simp add: map_poly_def)
+
+lemma map_poly_1: "map_poly f 1 = [:f 1:]"
+ by (simp add: map_poly_def)
+
+lemma map_poly_1' [simp]: "f 1 = 1 \<Longrightarrow> map_poly f 1 = 1"
+ by (simp add: map_poly_def one_poly_def)
+
+lemma coeff_map_poly:
+ assumes "f 0 = 0"
+ shows "coeff (map_poly f p) n = f (coeff p n)"
+ by (auto simp: map_poly_def nth_default_def coeffs_def assms
+ not_less Suc_le_eq coeff_eq_0 simp del: upt_Suc)
+
+lemma coeffs_map_poly [code abstract]:
+ "coeffs (map_poly f p) = strip_while (op = 0) (map f (coeffs p))"
+ by (simp add: map_poly_def)
+
+lemma set_coeffs_map_poly:
+ "(\<And>x. f x = 0 \<longleftrightarrow> x = 0) \<Longrightarrow> set (coeffs (map_poly f p)) = f ` set (coeffs p)"
+ by (cases "p = 0") (auto simp: coeffs_map_poly last_map last_coeffs_not_0)
+
+lemma coeffs_map_poly':
+ assumes "(\<And>x. x \<noteq> 0 \<Longrightarrow> f x \<noteq> 0)"
+ shows "coeffs (map_poly f p) = map f (coeffs p)"
+ by (cases "p = 0") (auto simp: coeffs_map_poly last_map last_coeffs_not_0 assms
+ intro!: strip_while_not_last split: if_splits)
+
+lemma degree_map_poly:
+ assumes "\<And>x. x \<noteq> 0 \<Longrightarrow> f x \<noteq> 0"
+ shows "degree (map_poly f p) = degree p"
+ by (simp add: degree_eq_length_coeffs coeffs_map_poly' assms)
+
+lemma map_poly_eq_0_iff:
+ assumes "f 0 = 0" "\<And>x. x \<in> set (coeffs p) \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> f x \<noteq> 0"
+ shows "map_poly f p = 0 \<longleftrightarrow> p = 0"
+proof -
+ {
+ fix n :: nat
+ have "coeff (map_poly f p) n = f (coeff p n)" by (simp add: coeff_map_poly assms)
+ also have "\<dots> = 0 \<longleftrightarrow> coeff p n = 0"
+ proof (cases "n < length (coeffs p)")
+ case True
+ hence "coeff p n \<in> set (coeffs p)" by (auto simp: coeffs_def simp del: upt_Suc)
+ with assms show "f (coeff p n) = 0 \<longleftrightarrow> coeff p n = 0" by auto
+ qed (auto simp: assms length_coeffs nth_default_coeffs_eq [symmetric] nth_default_def)
+ finally have "(coeff (map_poly f p) n = 0) = (coeff p n = 0)" .
+ }
+ thus ?thesis by (auto simp: poly_eq_iff)
+qed
+
+lemma map_poly_smult:
+ assumes "f 0 = 0""\<And>c x. f (c * x) = f c * f x"
+ shows "map_poly f (smult c p) = smult (f c) (map_poly f p)"
+ by (intro poly_eqI) (simp_all add: assms coeff_map_poly)
+
+lemma map_poly_pCons:
+ assumes "f 0 = 0"
+ shows "map_poly f (pCons c p) = pCons (f c) (map_poly f p)"
+ by (intro poly_eqI) (simp_all add: assms coeff_map_poly coeff_pCons split: nat.splits)
+
+lemma map_poly_map_poly:
+ assumes "f 0 = 0" "g 0 = 0"
+ shows "map_poly f (map_poly g p) = map_poly (f \<circ> g) p"
+ by (intro poly_eqI) (simp add: coeff_map_poly assms)
+
+lemma map_poly_id [simp]: "map_poly id p = p"
+ by (simp add: map_poly_def)
+
+lemma map_poly_id' [simp]: "map_poly (\<lambda>x. x) p = p"
+ by (simp add: map_poly_def)
+
+lemma map_poly_cong:
+ assumes "(\<And>x. x \<in> set (coeffs p) \<Longrightarrow> f x = g x)"
+ shows "map_poly f p = map_poly g p"
+proof -
+ from assms have "map f (coeffs p) = map g (coeffs p)" by (intro map_cong) simp_all
+ thus ?thesis by (simp only: coeffs_eq_iff coeffs_map_poly)
+qed
+
+lemma map_poly_monom: "f 0 = 0 \<Longrightarrow> map_poly f (monom c n) = monom (f c) n"
+ by (intro poly_eqI) (simp_all add: coeff_map_poly)
+
+lemma map_poly_idI:
+ assumes "\<And>x. x \<in> set (coeffs p) \<Longrightarrow> f x = x"
+ shows "map_poly f p = p"
+ using map_poly_cong[OF assms, of _ id] by simp
+
+lemma map_poly_idI':
+ assumes "\<And>x. x \<in> set (coeffs p) \<Longrightarrow> f x = x"
+ shows "p = map_poly f p"
+ using map_poly_cong[OF assms, of _ id] by simp
+
+lemma smult_conv_map_poly: "smult c p = map_poly (\<lambda>x. c * x) p"
+ by (intro poly_eqI) (simp_all add: coeff_map_poly)
+
+
subsection \<open>Conversions from natural numbers\<close>
lemma of_nat_poly: "of_nat n = [:of_nat n :: 'a :: comm_semiring_1:]"
@@ -1086,6 +1191,7 @@
lemma numeral_poly: "numeral n = [:numeral n:]"
by (subst of_nat_numeral [symmetric], subst of_nat_poly) simp
+
subsection \<open>Lemmas about divisibility\<close>
lemma dvd_smult: "p dvd q \<Longrightarrow> p dvd smult a q"
@@ -1137,6 +1243,11 @@
apply (simp add: coeff_mult_degree_sum)
done
+lemma degree_mult_eq_0:
+ fixes p q:: "'a :: {comm_semiring_0,semiring_no_zero_divisors} poly"
+ shows "degree (p * q) = 0 \<longleftrightarrow> p = 0 \<or> q = 0 \<or> (p \<noteq> 0 \<and> q \<noteq> 0 \<and> degree p = 0 \<and> degree q = 0)"
+ by (auto simp add: degree_mult_eq)
+
lemma degree_mult_right_le:
fixes p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
assumes "q \<noteq> 0"
@@ -1290,6 +1401,75 @@
text \<open>TODO: Simplification rules for comparisons\<close>
+subsection \<open>Leading coefficient\<close>
+
+definition lead_coeff:: "'a::zero poly \<Rightarrow> 'a" where
+ "lead_coeff p= coeff p (degree p)"
+
+lemma lead_coeff_pCons[simp]:
+ "p\<noteq>0 \<Longrightarrow>lead_coeff (pCons a p) = lead_coeff p"
+ "p=0 \<Longrightarrow> lead_coeff (pCons a p) = a"
+unfolding lead_coeff_def by auto
+
+lemma lead_coeff_0[simp]:"lead_coeff 0 =0"
+ unfolding lead_coeff_def by auto
+
+lemma coeff_0_prod_list: "coeff (prod_list xs) 0 = prod_list (map (\<lambda>p. coeff p 0) xs)"
+ by (induction xs) (simp_all add: coeff_mult)
+
+lemma coeff_0_power: "coeff (p ^ n) 0 = coeff p 0 ^ n"
+ by (induction n) (simp_all add: coeff_mult)
+
+lemma lead_coeff_mult:
+ fixes p q::"'a :: {comm_semiring_0,semiring_no_zero_divisors} poly"
+ shows "lead_coeff (p * q) = lead_coeff p * lead_coeff q"
+by (unfold lead_coeff_def,cases "p=0 \<or> q=0",auto simp add:coeff_mult_degree_sum degree_mult_eq)
+
+lemma lead_coeff_add_le:
+ assumes "degree p < degree q"
+ shows "lead_coeff (p+q) = lead_coeff q"
+using assms unfolding lead_coeff_def
+by (metis coeff_add coeff_eq_0 monoid_add_class.add.left_neutral degree_add_eq_right)
+
+lemma lead_coeff_minus:
+ "lead_coeff (-p) = - lead_coeff p"
+by (metis coeff_minus degree_minus lead_coeff_def)
+
+lemma lead_coeff_smult:
+ "lead_coeff (smult c p :: 'a :: {comm_semiring_0,semiring_no_zero_divisors} poly) = c * lead_coeff p"
+proof -
+ have "smult c p = [:c:] * p" by simp
+ also have "lead_coeff \<dots> = c * lead_coeff p"
+ by (subst lead_coeff_mult) simp_all
+ finally show ?thesis .
+qed
+
+lemma lead_coeff_eq_zero_iff [simp]: "lead_coeff p = 0 \<longleftrightarrow> p = 0"
+ by (simp add: lead_coeff_def)
+
+lemma lead_coeff_1 [simp]: "lead_coeff 1 = 1"
+ by (simp add: lead_coeff_def)
+
+lemma lead_coeff_of_nat [simp]:
+ "lead_coeff (of_nat n) = (of_nat n :: 'a :: {comm_semiring_1,semiring_char_0})"
+ by (induction n) (simp_all add: lead_coeff_def of_nat_poly)
+
+lemma lead_coeff_numeral [simp]:
+ "lead_coeff (numeral n) = numeral n"
+ unfolding lead_coeff_def
+ by (subst of_nat_numeral [symmetric], subst of_nat_poly) simp
+
+lemma lead_coeff_power:
+ "lead_coeff (p ^ n :: 'a :: {comm_semiring_1,semiring_no_zero_divisors} poly) = lead_coeff p ^ n"
+ by (induction n) (simp_all add: lead_coeff_mult)
+
+lemma lead_coeff_nonzero: "p \<noteq> 0 \<Longrightarrow> lead_coeff p \<noteq> 0"
+ by (simp add: lead_coeff_def)
+
+lemma lead_coeff_monom [simp]: "lead_coeff (monom c n) = c"
+ by (cases "c = 0") (simp_all add: lead_coeff_def degree_monom_eq)
+
+
subsection \<open>Synthetic division and polynomial roots\<close>
text \<open>
@@ -1555,7 +1735,7 @@
-subsection\<open>Pseudo-Division and Division of Polynomials\<close>
+subsection \<open>Pseudo-Division and Division of Polynomials\<close>
text\<open>This part is by René Thiemann and Akihisa Yamada.\<close>
@@ -1838,15 +2018,172 @@
lemma divide_poly_0: "f div 0 = (0 :: 'a poly)"
by (simp add: divide_poly_def Let_def divide_poly_main_0)
-instance by (standard, auto simp: divide_poly divide_poly_0)
+instance
+ by standard (auto simp: divide_poly divide_poly_0)
+
end
-
instance poly :: (idom_divide) algebraic_semidom ..
-
-
-subsubsection\<open>Division in Field Polynomials\<close>
+lemma div_const_poly_conv_map_poly:
+ assumes "[:c:] dvd p"
+ shows "p div [:c:] = map_poly (\<lambda>x. x div c) p"
+proof (cases "c = 0")
+ case False
+ from assms obtain q where p: "p = [:c:] * q" by (erule dvdE)
+ moreover {
+ have "smult c q = [:c:] * q" by simp
+ also have "\<dots> div [:c:] = q" by (rule nonzero_mult_div_cancel_left) (insert False, auto)
+ finally have "smult c q div [:c:] = q" .
+ }
+ ultimately show ?thesis by (intro poly_eqI) (auto simp: coeff_map_poly False)
+qed (auto intro!: poly_eqI simp: coeff_map_poly)
+
+lemma is_unit_monom_0:
+ fixes a :: "'a::field"
+ assumes "a \<noteq> 0"
+ shows "is_unit (monom a 0)"
+proof
+ from assms show "1 = monom a 0 * monom (inverse a) 0"
+ by (simp add: mult_monom)
+qed
+
+lemma is_unit_triv:
+ fixes a :: "'a::field"
+ assumes "a \<noteq> 0"
+ shows "is_unit [:a:]"
+ using assms by (simp add: is_unit_monom_0 monom_0 [symmetric])
+
+lemma is_unit_iff_degree:
+ assumes "p \<noteq> (0 :: _ :: field poly)"
+ shows "is_unit p \<longleftrightarrow> degree p = 0" (is "?P \<longleftrightarrow> ?Q")
+proof
+ assume ?Q
+ then obtain a where "p = [:a:]" by (rule degree_eq_zeroE)
+ with assms show ?P by (simp add: is_unit_triv)
+next
+ assume ?P
+ then obtain q where "q \<noteq> 0" "p * q = 1" ..
+ then have "degree (p * q) = degree 1"
+ by simp
+ with \<open>p \<noteq> 0\<close> \<open>q \<noteq> 0\<close> have "degree p + degree q = 0"
+ by (simp add: degree_mult_eq)
+ then show ?Q by simp
+qed
+
+lemma is_unit_pCons_iff:
+ "is_unit (pCons (a::_::field) p) \<longleftrightarrow> p = 0 \<and> a \<noteq> 0"
+ by (cases "p = 0") (auto simp add: is_unit_triv is_unit_iff_degree)
+
+lemma is_unit_monom_trival:
+ fixes p :: "'a::field poly"
+ assumes "is_unit p"
+ shows "monom (coeff p (degree p)) 0 = p"
+ using assms by (cases p) (simp_all add: monom_0 is_unit_pCons_iff)
+
+lemma is_unit_const_poly_iff:
+ "[:c :: 'a :: {comm_semiring_1,semiring_no_zero_divisors}:] dvd 1 \<longleftrightarrow> c dvd 1"
+ by (auto simp: one_poly_def)
+
+lemma is_unit_polyE:
+ fixes p :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly"
+ assumes "p dvd 1" obtains c where "p = [:c:]" "c dvd 1"
+proof -
+ from assms obtain q where "1 = p * q"
+ by (rule dvdE)
+ then have "p \<noteq> 0" and "q \<noteq> 0"
+ by auto
+ from \<open>1 = p * q\<close> have "degree 1 = degree (p * q)"
+ by simp
+ also from \<open>p \<noteq> 0\<close> and \<open>q \<noteq> 0\<close> have "\<dots> = degree p + degree q"
+ by (simp add: degree_mult_eq)
+ finally have "degree p = 0" by simp
+ with degree_eq_zeroE obtain c where c: "p = [:c:]" .
+ moreover with \<open>p dvd 1\<close> have "c dvd 1"
+ by (simp add: is_unit_const_poly_iff)
+ ultimately show thesis
+ by (rule that)
+qed
+
+lemma is_unit_polyE':
+ assumes "is_unit (p::_::field poly)"
+ obtains a where "p = monom a 0" and "a \<noteq> 0"
+proof -
+ obtain a q where "p = pCons a q" by (cases p)
+ with assms have "p = [:a:]" and "a \<noteq> 0"
+ by (simp_all add: is_unit_pCons_iff)
+ with that show thesis by (simp add: monom_0)
+qed
+
+lemma is_unit_poly_iff:
+ fixes p :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly"
+ shows "p dvd 1 \<longleftrightarrow> (\<exists>c. p = [:c:] \<and> c dvd 1)"
+ by (auto elim: is_unit_polyE simp add: is_unit_const_poly_iff)
+
+instantiation poly :: ("{normalization_semidom, idom_divide}") normalization_semidom
+begin
+
+definition unit_factor_poly :: "'a poly \<Rightarrow> 'a poly"
+ where "unit_factor_poly p = monom (unit_factor (lead_coeff p)) 0"
+
+definition normalize_poly :: "'a poly \<Rightarrow> 'a poly"
+ where "normalize_poly p = map_poly (\<lambda>x. x div unit_factor (lead_coeff p)) p"
+
+instance proof
+ fix p :: "'a poly"
+ show "unit_factor p * normalize p = p"
+ by (cases "p = 0")
+ (simp_all add: unit_factor_poly_def normalize_poly_def monom_0
+ smult_conv_map_poly map_poly_map_poly o_def)
+next
+ fix p :: "'a poly"
+ assume "is_unit p"
+ then obtain c where p: "p = [:c:]" "is_unit c"
+ by (auto simp: is_unit_poly_iff)
+ thus "normalize p = 1"
+ by (simp add: normalize_poly_def map_poly_pCons is_unit_normalize one_poly_def)
+next
+ fix p :: "'a poly" assume "p \<noteq> 0"
+ thus "is_unit (unit_factor p)"
+ by (simp add: unit_factor_poly_def monom_0 is_unit_poly_iff)
+qed (simp_all add: normalize_poly_def unit_factor_poly_def monom_0 lead_coeff_mult unit_factor_mult)
+
+end
+
+lemma normalize_poly_eq_div:
+ "normalize p = p div [:unit_factor (lead_coeff p):]"
+proof (cases "p = 0")
+ case False
+ thus ?thesis
+ by (subst div_const_poly_conv_map_poly)
+ (auto simp: normalize_poly_def const_poly_dvd_iff lead_coeff_def )
+qed (auto simp: normalize_poly_def)
+
+lemma unit_factor_pCons:
+ "unit_factor (pCons a p) = (if p = 0 then monom (unit_factor a) 0 else unit_factor p)"
+ by (simp add: unit_factor_poly_def)
+
+lemma normalize_monom [simp]:
+ "normalize (monom a n) = monom (normalize a) n"
+ by (simp add: map_poly_monom normalize_poly_def)
+
+lemma unit_factor_monom [simp]:
+ "unit_factor (monom a n) = monom (unit_factor a) 0"
+ by (simp add: unit_factor_poly_def )
+
+lemma normalize_const_poly: "normalize [:c:] = [:normalize c:]"
+ by (simp add: normalize_poly_def map_poly_pCons)
+
+lemma normalize_smult: "normalize (smult c p) = smult (normalize c) (normalize p)"
+proof -
+ have "smult c p = [:c:] * p" by simp
+ also have "normalize \<dots> = smult (normalize c) (normalize p)"
+ by (subst normalize_mult) (simp add: normalize_const_poly)
+ finally show ?thesis .
+qed
+
+
+subsubsection \<open>Division in Field Polynomials\<close>
text\<open>
This part connects the above result to the division of field polynomials.
@@ -1937,12 +2274,6 @@
from pdivmod_rel[of x y,unfolded pdivmod_rel_def]
show "x div y * y + x mod y = x" by auto
next
- fix x :: "'a poly"
- show "x div 0 = 0" by simp
-next
- fix y :: "'a poly"
- show "0 div y = 0" by simp
-next
fix x y z :: "'a poly"
assume "y \<noteq> 0"
hence "pdivmod_rel (x + z * y) y (z + x div y) (x mod y)"
@@ -1978,58 +2309,6 @@
end
-lemma is_unit_monom_0:
- fixes a :: "'a::field"
- assumes "a \<noteq> 0"
- shows "is_unit (monom a 0)"
-proof
- from assms show "1 = monom a 0 * monom (inverse a) 0"
- by (simp add: mult_monom)
-qed
-
-lemma is_unit_triv:
- fixes a :: "'a::field"
- assumes "a \<noteq> 0"
- shows "is_unit [:a:]"
- using assms by (simp add: is_unit_monom_0 monom_0 [symmetric])
-
-lemma is_unit_iff_degree:
- assumes "p \<noteq> (0 :: _ :: field poly)"
- shows "is_unit p \<longleftrightarrow> degree p = 0" (is "?P \<longleftrightarrow> ?Q")
-proof
- assume ?Q
- then obtain a where "p = [:a:]" by (rule degree_eq_zeroE)
- with assms show ?P by (simp add: is_unit_triv)
-next
- assume ?P
- then obtain q where "q \<noteq> 0" "p * q = 1" ..
- then have "degree (p * q) = degree 1"
- by simp
- with \<open>p \<noteq> 0\<close> \<open>q \<noteq> 0\<close> have "degree p + degree q = 0"
- by (simp add: degree_mult_eq)
- then show ?Q by simp
-qed
-
-lemma is_unit_pCons_iff:
- "is_unit (pCons (a::_::field) p) \<longleftrightarrow> p = 0 \<and> a \<noteq> 0"
- by (cases "p = 0") (auto simp add: is_unit_triv is_unit_iff_degree)
-
-lemma is_unit_monom_trival:
- fixes p :: "'a::field poly"
- assumes "is_unit p"
- shows "monom (coeff p (degree p)) 0 = p"
- using assms by (cases p) (simp_all add: monom_0 is_unit_pCons_iff)
-
-lemma is_unit_polyE:
- assumes "is_unit (p::_::field poly)"
- obtains a where "p = monom a 0" and "a \<noteq> 0"
-proof -
- obtain a q where "p = pCons a q" by (cases p)
- with assms have "p = [:a:]" and "a \<noteq> 0"
- by (simp_all add: is_unit_pCons_iff)
- with that show thesis by (simp add: monom_0)
-qed
-
lemma degree_mod_less:
"y \<noteq> 0 \<Longrightarrow> x mod y = 0 \<or> degree (x mod y) < degree y"
using pdivmod_rel [of x y]
@@ -2860,18 +3139,11 @@
by (cases "finite A", induction rule: finite_induct)
(simp_all add: pcompose_1 pcompose_mult)
-
-(* The remainder of this section and the next were contributed by Wenda Li *)
-
-lemma degree_mult_eq_0:
- fixes p q:: "'a :: {comm_semiring_0,semiring_no_zero_divisors} poly"
- shows "degree (p*q) = 0 \<longleftrightarrow> p=0 \<or> q=0 \<or> (p\<noteq>0 \<and> q\<noteq>0 \<and> degree p =0 \<and> degree q =0)"
-by (auto simp add:degree_mult_eq)
-
-lemma pcompose_const[simp]:"pcompose [:a:] q = [:a:]" by (subst pcompose_pCons,simp)
+lemma pcompose_const [simp]: "pcompose [:a:] q = [:a:]"
+ by (subst pcompose_pCons) simp
lemma pcompose_0': "pcompose p 0 = [:coeff p 0:]"
- by (induct p) (auto simp add:pcompose_pCons)
+ by (induct p) (auto simp add: pcompose_pCons)
lemma degree_pcompose:
fixes p q:: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
@@ -2932,53 +3204,6 @@
thus ?thesis using \<open>p=[:a:]\<close> by simp
qed
-
-subsection \<open>Leading coefficient\<close>
-
-definition lead_coeff:: "'a::zero poly \<Rightarrow> 'a" where
- "lead_coeff p= coeff p (degree p)"
-
-lemma lead_coeff_pCons[simp]:
- "p\<noteq>0 \<Longrightarrow>lead_coeff (pCons a p) = lead_coeff p"
- "p=0 \<Longrightarrow> lead_coeff (pCons a p) = a"
-unfolding lead_coeff_def by auto
-
-lemma lead_coeff_0[simp]:"lead_coeff 0 =0"
- unfolding lead_coeff_def by auto
-
-lemma coeff_0_prod_list: "coeff (prod_list xs) 0 = prod_list (map (\<lambda>p. coeff p 0) xs)"
- by (induction xs) (simp_all add: coeff_mult)
-
-lemma coeff_0_power: "coeff (p ^ n) 0 = coeff p 0 ^ n"
- by (induction n) (simp_all add: coeff_mult)
-
-lemma lead_coeff_mult:
- fixes p q::"'a :: {comm_semiring_0,semiring_no_zero_divisors} poly"
- shows "lead_coeff (p * q) = lead_coeff p * lead_coeff q"
-by (unfold lead_coeff_def,cases "p=0 \<or> q=0",auto simp add:coeff_mult_degree_sum degree_mult_eq)
-
-lemma lead_coeff_add_le:
- assumes "degree p < degree q"
- shows "lead_coeff (p+q) = lead_coeff q"
-using assms unfolding lead_coeff_def
-by (metis coeff_add coeff_eq_0 monoid_add_class.add.left_neutral degree_add_eq_right)
-
-lemma lead_coeff_minus:
- "lead_coeff (-p) = - lead_coeff p"
-by (metis coeff_minus degree_minus lead_coeff_def)
-
-lemma lead_coeff_smult:
- "lead_coeff (smult c p :: 'a :: {comm_semiring_0,semiring_no_zero_divisors} poly) = c * lead_coeff p"
-proof -
- have "smult c p = [:c:] * p" by simp
- also have "lead_coeff \<dots> = c * lead_coeff p"
- by (subst lead_coeff_mult) simp_all
- finally show ?thesis .
-qed
-
-lemma lead_coeff_eq_zero_iff [simp]: "lead_coeff p = 0 \<longleftrightarrow> p = 0"
- by (simp add: lead_coeff_def)
-
lemma lead_coeff_comp:
fixes p q:: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly"
assumes "degree q > 0"
@@ -3009,25 +3234,6 @@
ultimately show ?case by blast
qed
-lemma lead_coeff_1 [simp]: "lead_coeff 1 = 1"
- by (simp add: lead_coeff_def)
-
-lemma lead_coeff_of_nat [simp]:
- "lead_coeff (of_nat n) = (of_nat n :: 'a :: {comm_semiring_1,semiring_char_0})"
- by (induction n) (simp_all add: lead_coeff_def of_nat_poly)
-
-lemma lead_coeff_numeral [simp]:
- "lead_coeff (numeral n) = numeral n"
- unfolding lead_coeff_def
- by (subst of_nat_numeral [symmetric], subst of_nat_poly) simp
-
-lemma lead_coeff_power:
- "lead_coeff (p ^ n :: 'a :: {comm_semiring_1,semiring_no_zero_divisors} poly) = lead_coeff p ^ n"
- by (induction n) (simp_all add: lead_coeff_mult)
-
-lemma lead_coeff_nonzero: "p \<noteq> 0 \<Longrightarrow> lead_coeff p \<noteq> 0"
- by (simp add: lead_coeff_def)
-
subsection \<open>Shifting polynomials\<close>
--- a/src/HOL/Library/Polynomial_Factorial.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Library/Polynomial_Factorial.thy Tue Dec 20 16:18:56 2016 +0100
@@ -9,144 +9,84 @@
theory Polynomial_Factorial
imports
Complex_Main
- "~~/src/HOL/Number_Theory/Euclidean_Algorithm"
"~~/src/HOL/Library/Polynomial"
"~~/src/HOL/Library/Normalized_Fraction"
-begin
-
-subsection \<open>Prelude\<close>
-
-lemma prod_mset_mult:
- "prod_mset (image_mset (\<lambda>x. f x * g x) A) = prod_mset (image_mset f A) * prod_mset (image_mset g A)"
- by (induction A) (simp_all add: mult_ac)
-
-lemma prod_mset_const: "prod_mset (image_mset (\<lambda>_. c) A) = c ^ size A"
- by (induction A) (simp_all add: mult_ac)
-
-lemma dvd_field_iff: "x dvd y \<longleftrightarrow> (x = 0 \<longrightarrow> y = (0::'a::field))"
-proof safe
- assume "x \<noteq> 0"
- hence "y = x * (y / x)" by (simp add: field_simps)
- thus "x dvd y" by (rule dvdI)
-qed auto
-
-lemma nat_descend_induct [case_names base descend]:
- assumes "\<And>k::nat. k > n \<Longrightarrow> P k"
- assumes "\<And>k::nat. k \<le> n \<Longrightarrow> (\<And>i. i > k \<Longrightarrow> P i) \<Longrightarrow> P k"
- shows "P m"
- using assms by induction_schema (force intro!: wf_measure[of "\<lambda>k. Suc n - k"])+
-
-lemma GreatestI_ex: "\<exists>k::nat. P k \<Longrightarrow> \<forall>y. P y \<longrightarrow> y < b \<Longrightarrow> P (GREATEST x. P x)"
- by (metis GreatestI)
-
-
-context field
-begin
-
-subclass idom_divide ..
-
-end
-
-context field
-begin
-
-definition normalize_field :: "'a \<Rightarrow> 'a"
- where [simp]: "normalize_field x = (if x = 0 then 0 else 1)"
-definition unit_factor_field :: "'a \<Rightarrow> 'a"
- where [simp]: "unit_factor_field x = x"
-definition euclidean_size_field :: "'a \<Rightarrow> nat"
- where [simp]: "euclidean_size_field x = (if x = 0 then 0 else 1)"
-definition mod_field :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
- where [simp]: "mod_field x y = (if y = 0 then x else 0)"
-
-end
-
-instantiation real :: euclidean_ring
-begin
-
-definition [simp]: "normalize_real = (normalize_field :: real \<Rightarrow> _)"
-definition [simp]: "unit_factor_real = (unit_factor_field :: real \<Rightarrow> _)"
-definition [simp]: "euclidean_size_real = (euclidean_size_field :: real \<Rightarrow> _)"
-definition [simp]: "modulo_real = (mod_field :: real \<Rightarrow> _)"
-
-instance by standard (simp_all add: dvd_field_iff divide_simps)
-end
-
-instantiation real :: euclidean_ring_gcd
+ "~~/src/HOL/Library/Field_as_Ring"
begin
-definition gcd_real :: "real \<Rightarrow> real \<Rightarrow> real" where
- "gcd_real = gcd_eucl"
-definition lcm_real :: "real \<Rightarrow> real \<Rightarrow> real" where
- "lcm_real = lcm_eucl"
-definition Gcd_real :: "real set \<Rightarrow> real" where
- "Gcd_real = Gcd_eucl"
-definition Lcm_real :: "real set \<Rightarrow> real" where
- "Lcm_real = Lcm_eucl"
+subsection \<open>Various facts about polynomials\<close>
-instance by standard (simp_all add: gcd_real_def lcm_real_def Gcd_real_def Lcm_real_def)
-
-end
+lemma prod_mset_const_poly: "prod_mset (image_mset (\<lambda>x. [:f x:]) A) = [:prod_mset (image_mset f A):]"
+ by (induction A) (simp_all add: one_poly_def mult_ac)
-instantiation rat :: euclidean_ring
-begin
-
-definition [simp]: "normalize_rat = (normalize_field :: rat \<Rightarrow> _)"
-definition [simp]: "unit_factor_rat = (unit_factor_field :: rat \<Rightarrow> _)"
-definition [simp]: "euclidean_size_rat = (euclidean_size_field :: rat \<Rightarrow> _)"
-definition [simp]: "modulo_rat = (mod_field :: rat \<Rightarrow> _)"
-
-instance by standard (simp_all add: dvd_field_iff divide_simps)
-end
-
-instantiation rat :: euclidean_ring_gcd
-begin
+lemma is_unit_smult_iff: "smult c p dvd 1 \<longleftrightarrow> c dvd 1 \<and> p dvd 1"
+proof -
+ have "smult c p = [:c:] * p" by simp
+ also have "\<dots> dvd 1 \<longleftrightarrow> c dvd 1 \<and> p dvd 1"
+ proof safe
+ assume A: "[:c:] * p dvd 1"
+ thus "p dvd 1" by (rule dvd_mult_right)
+ from A obtain q where B: "1 = [:c:] * p * q" by (erule dvdE)
+ have "c dvd c * (coeff p 0 * coeff q 0)" by simp
+ also have "\<dots> = coeff ([:c:] * p * q) 0" by (simp add: mult.assoc coeff_mult)
+ also note B [symmetric]
+ finally show "c dvd 1" by simp
+ next
+ assume "c dvd 1" "p dvd 1"
+ from \<open>c dvd 1\<close> obtain d where "1 = c * d" by (erule dvdE)
+ hence "1 = [:c:] * [:d:]" by (simp add: one_poly_def mult_ac)
+ hence "[:c:] dvd 1" by (rule dvdI)
+ from mult_dvd_mono[OF this \<open>p dvd 1\<close>] show "[:c:] * p dvd 1" by simp
+ qed
+ finally show ?thesis .
+qed
-definition gcd_rat :: "rat \<Rightarrow> rat \<Rightarrow> rat" where
- "gcd_rat = gcd_eucl"
-definition lcm_rat :: "rat \<Rightarrow> rat \<Rightarrow> rat" where
- "lcm_rat = lcm_eucl"
-definition Gcd_rat :: "rat set \<Rightarrow> rat" where
- "Gcd_rat = Gcd_eucl"
-definition Lcm_rat :: "rat set \<Rightarrow> rat" where
- "Lcm_rat = Lcm_eucl"
-
-instance by standard (simp_all add: gcd_rat_def lcm_rat_def Gcd_rat_def Lcm_rat_def)
-
-end
-
-instantiation complex :: euclidean_ring
-begin
+lemma degree_mod_less': "b \<noteq> 0 \<Longrightarrow> a mod b \<noteq> 0 \<Longrightarrow> degree (a mod b) < degree b"
+ using degree_mod_less[of b a] by auto
+
+lemma smult_eq_iff:
+ assumes "(b :: 'a :: field) \<noteq> 0"
+ shows "smult a p = smult b q \<longleftrightarrow> smult (a / b) p = q"
+proof
+ assume "smult a p = smult b q"
+ also from assms have "smult (inverse b) \<dots> = q" by simp
+ finally show "smult (a / b) p = q" by (simp add: field_simps)
+qed (insert assms, auto)
-definition [simp]: "normalize_complex = (normalize_field :: complex \<Rightarrow> _)"
-definition [simp]: "unit_factor_complex = (unit_factor_field :: complex \<Rightarrow> _)"
-definition [simp]: "euclidean_size_complex = (euclidean_size_field :: complex \<Rightarrow> _)"
-definition [simp]: "modulo_complex = (mod_field :: complex \<Rightarrow> _)"
-
-instance by standard (simp_all add: dvd_field_iff divide_simps)
-end
-
-instantiation complex :: euclidean_ring_gcd
-begin
-
-definition gcd_complex :: "complex \<Rightarrow> complex \<Rightarrow> complex" where
- "gcd_complex = gcd_eucl"
-definition lcm_complex :: "complex \<Rightarrow> complex \<Rightarrow> complex" where
- "lcm_complex = lcm_eucl"
-definition Gcd_complex :: "complex set \<Rightarrow> complex" where
- "Gcd_complex = Gcd_eucl"
-definition Lcm_complex :: "complex set \<Rightarrow> complex" where
- "Lcm_complex = Lcm_eucl"
-
-instance by standard (simp_all add: gcd_complex_def lcm_complex_def Gcd_complex_def Lcm_complex_def)
-
-end
-
+lemma irreducible_const_poly_iff:
+ fixes c :: "'a :: {comm_semiring_1,semiring_no_zero_divisors}"
+ shows "irreducible [:c:] \<longleftrightarrow> irreducible c"
+proof
+ assume A: "irreducible c"
+ show "irreducible [:c:]"
+ proof (rule irreducibleI)
+ fix a b assume ab: "[:c:] = a * b"
+ hence "degree [:c:] = degree (a * b)" by (simp only: )
+ also from A ab have "a \<noteq> 0" "b \<noteq> 0" by auto
+ hence "degree (a * b) = degree a + degree b" by (simp add: degree_mult_eq)
+ finally have "degree a = 0" "degree b = 0" by auto
+ then obtain a' b' where ab': "a = [:a':]" "b = [:b':]" by (auto elim!: degree_eq_zeroE)
+ from ab have "coeff [:c:] 0 = coeff (a * b) 0" by (simp only: )
+ hence "c = a' * b'" by (simp add: ab' mult_ac)
+ from A and this have "a' dvd 1 \<or> b' dvd 1" by (rule irreducibleD)
+ with ab' show "a dvd 1 \<or> b dvd 1" by (auto simp: one_poly_def)
+ qed (insert A, auto simp: irreducible_def is_unit_poly_iff)
+next
+ assume A: "irreducible [:c:]"
+ show "irreducible c"
+ proof (rule irreducibleI)
+ fix a b assume ab: "c = a * b"
+ hence "[:c:] = [:a:] * [:b:]" by (simp add: mult_ac)
+ from A and this have "[:a:] dvd 1 \<or> [:b:] dvd 1" by (rule irreducibleD)
+ thus "a dvd 1 \<or> b dvd 1" by (simp add: one_poly_def)
+ qed (insert A, auto simp: irreducible_def one_poly_def)
+qed
subsection \<open>Lifting elements into the field of fractions\<close>
definition to_fract :: "'a :: idom \<Rightarrow> 'a fract" where "to_fract x = Fract x 1"
+ -- \<open>FIXME: name \<open>of_idom\<close>, abbreviation\<close>
lemma to_fract_0 [simp]: "to_fract 0 = 0"
by (simp add: to_fract_def eq_fract Zero_fract_def)
@@ -219,285 +159,6 @@
lemma normalize_snd_quot_of_fract: "normalize (snd (quot_of_fract x)) = snd (quot_of_fract x)"
by (intro unit_factor_1_imp_normalized unit_factor_snd_quot_of_fract)
-
-subsection \<open>Mapping polynomials\<close>
-
-definition map_poly
- :: "('a :: zero \<Rightarrow> 'b :: zero) \<Rightarrow> 'a poly \<Rightarrow> 'b poly" where
- "map_poly f p = Poly (map f (coeffs p))"
-
-lemma map_poly_0 [simp]: "map_poly f 0 = 0"
- by (simp add: map_poly_def)
-
-lemma map_poly_1: "map_poly f 1 = [:f 1:]"
- by (simp add: map_poly_def)
-
-lemma map_poly_1' [simp]: "f 1 = 1 \<Longrightarrow> map_poly f 1 = 1"
- by (simp add: map_poly_def one_poly_def)
-
-lemma coeff_map_poly:
- assumes "f 0 = 0"
- shows "coeff (map_poly f p) n = f (coeff p n)"
- by (auto simp: map_poly_def nth_default_def coeffs_def assms
- not_less Suc_le_eq coeff_eq_0 simp del: upt_Suc)
-
-lemma coeffs_map_poly [code abstract]:
- "coeffs (map_poly f p) = strip_while (op = 0) (map f (coeffs p))"
- by (simp add: map_poly_def)
-
-lemma set_coeffs_map_poly:
- "(\<And>x. f x = 0 \<longleftrightarrow> x = 0) \<Longrightarrow> set (coeffs (map_poly f p)) = f ` set (coeffs p)"
- by (cases "p = 0") (auto simp: coeffs_map_poly last_map last_coeffs_not_0)
-
-lemma coeffs_map_poly':
- assumes "(\<And>x. x \<noteq> 0 \<Longrightarrow> f x \<noteq> 0)"
- shows "coeffs (map_poly f p) = map f (coeffs p)"
- by (cases "p = 0") (auto simp: coeffs_map_poly last_map last_coeffs_not_0 assms
- intro!: strip_while_not_last split: if_splits)
-
-lemma degree_map_poly:
- assumes "\<And>x. x \<noteq> 0 \<Longrightarrow> f x \<noteq> 0"
- shows "degree (map_poly f p) = degree p"
- by (simp add: degree_eq_length_coeffs coeffs_map_poly' assms)
-
-lemma map_poly_eq_0_iff:
- assumes "f 0 = 0" "\<And>x. x \<in> set (coeffs p) \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> f x \<noteq> 0"
- shows "map_poly f p = 0 \<longleftrightarrow> p = 0"
-proof -
- {
- fix n :: nat
- have "coeff (map_poly f p) n = f (coeff p n)" by (simp add: coeff_map_poly assms)
- also have "\<dots> = 0 \<longleftrightarrow> coeff p n = 0"
- proof (cases "n < length (coeffs p)")
- case True
- hence "coeff p n \<in> set (coeffs p)" by (auto simp: coeffs_def simp del: upt_Suc)
- with assms show "f (coeff p n) = 0 \<longleftrightarrow> coeff p n = 0" by auto
- qed (auto simp: assms length_coeffs nth_default_coeffs_eq [symmetric] nth_default_def)
- finally have "(coeff (map_poly f p) n = 0) = (coeff p n = 0)" .
- }
- thus ?thesis by (auto simp: poly_eq_iff)
-qed
-
-lemma map_poly_smult:
- assumes "f 0 = 0""\<And>c x. f (c * x) = f c * f x"
- shows "map_poly f (smult c p) = smult (f c) (map_poly f p)"
- by (intro poly_eqI) (simp_all add: assms coeff_map_poly)
-
-lemma map_poly_pCons:
- assumes "f 0 = 0"
- shows "map_poly f (pCons c p) = pCons (f c) (map_poly f p)"
- by (intro poly_eqI) (simp_all add: assms coeff_map_poly coeff_pCons split: nat.splits)
-
-lemma map_poly_map_poly:
- assumes "f 0 = 0" "g 0 = 0"
- shows "map_poly f (map_poly g p) = map_poly (f \<circ> g) p"
- by (intro poly_eqI) (simp add: coeff_map_poly assms)
-
-lemma map_poly_id [simp]: "map_poly id p = p"
- by (simp add: map_poly_def)
-
-lemma map_poly_id' [simp]: "map_poly (\<lambda>x. x) p = p"
- by (simp add: map_poly_def)
-
-lemma map_poly_cong:
- assumes "(\<And>x. x \<in> set (coeffs p) \<Longrightarrow> f x = g x)"
- shows "map_poly f p = map_poly g p"
-proof -
- from assms have "map f (coeffs p) = map g (coeffs p)" by (intro map_cong) simp_all
- thus ?thesis by (simp only: coeffs_eq_iff coeffs_map_poly)
-qed
-
-lemma map_poly_monom: "f 0 = 0 \<Longrightarrow> map_poly f (monom c n) = monom (f c) n"
- by (intro poly_eqI) (simp_all add: coeff_map_poly)
-
-lemma map_poly_idI:
- assumes "\<And>x. x \<in> set (coeffs p) \<Longrightarrow> f x = x"
- shows "map_poly f p = p"
- using map_poly_cong[OF assms, of _ id] by simp
-
-lemma map_poly_idI':
- assumes "\<And>x. x \<in> set (coeffs p) \<Longrightarrow> f x = x"
- shows "p = map_poly f p"
- using map_poly_cong[OF assms, of _ id] by simp
-
-lemma smult_conv_map_poly: "smult c p = map_poly (\<lambda>x. c * x) p"
- by (intro poly_eqI) (simp_all add: coeff_map_poly)
-
-lemma div_const_poly_conv_map_poly:
- assumes "[:c:] dvd p"
- shows "p div [:c:] = map_poly (\<lambda>x. x div c) p"
-proof (cases "c = 0")
- case False
- from assms obtain q where p: "p = [:c:] * q" by (erule dvdE)
- moreover {
- have "smult c q = [:c:] * q" by simp
- also have "\<dots> div [:c:] = q" by (rule nonzero_mult_div_cancel_left) (insert False, auto)
- finally have "smult c q div [:c:] = q" .
- }
- ultimately show ?thesis by (intro poly_eqI) (auto simp: coeff_map_poly False)
-qed (auto intro!: poly_eqI simp: coeff_map_poly)
-
-
-
-subsection \<open>Various facts about polynomials\<close>
-
-lemma prod_mset_const_poly: "prod_mset (image_mset (\<lambda>x. [:f x:]) A) = [:prod_mset (image_mset f A):]"
- by (induction A) (simp_all add: one_poly_def mult_ac)
-
-lemma degree_mod_less': "b \<noteq> 0 \<Longrightarrow> a mod b \<noteq> 0 \<Longrightarrow> degree (a mod b) < degree b"
- using degree_mod_less[of b a] by auto
-
-lemma is_unit_const_poly_iff:
- "[:c :: 'a :: {comm_semiring_1,semiring_no_zero_divisors}:] dvd 1 \<longleftrightarrow> c dvd 1"
- by (auto simp: one_poly_def)
-
-lemma is_unit_poly_iff:
- fixes p :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly"
- shows "p dvd 1 \<longleftrightarrow> (\<exists>c. p = [:c:] \<and> c dvd 1)"
-proof safe
- assume "p dvd 1"
- then obtain q where pq: "1 = p * q" by (erule dvdE)
- hence "degree 1 = degree (p * q)" by simp
- also from pq have "\<dots> = degree p + degree q" by (intro degree_mult_eq) auto
- finally have "degree p = 0" by simp
- from degree_eq_zeroE[OF this] obtain c where c: "p = [:c:]" .
- with \<open>p dvd 1\<close> show "\<exists>c. p = [:c:] \<and> c dvd 1"
- by (auto simp: is_unit_const_poly_iff)
-qed (auto simp: is_unit_const_poly_iff)
-
-lemma is_unit_polyE:
- fixes p :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly"
- assumes "p dvd 1" obtains c where "p = [:c:]" "c dvd 1"
- using assms by (subst (asm) is_unit_poly_iff) blast
-
-lemma smult_eq_iff:
- assumes "(b :: 'a :: field) \<noteq> 0"
- shows "smult a p = smult b q \<longleftrightarrow> smult (a / b) p = q"
-proof
- assume "smult a p = smult b q"
- also from assms have "smult (inverse b) \<dots> = q" by simp
- finally show "smult (a / b) p = q" by (simp add: field_simps)
-qed (insert assms, auto)
-
-lemma irreducible_const_poly_iff:
- fixes c :: "'a :: {comm_semiring_1,semiring_no_zero_divisors}"
- shows "irreducible [:c:] \<longleftrightarrow> irreducible c"
-proof
- assume A: "irreducible c"
- show "irreducible [:c:]"
- proof (rule irreducibleI)
- fix a b assume ab: "[:c:] = a * b"
- hence "degree [:c:] = degree (a * b)" by (simp only: )
- also from A ab have "a \<noteq> 0" "b \<noteq> 0" by auto
- hence "degree (a * b) = degree a + degree b" by (simp add: degree_mult_eq)
- finally have "degree a = 0" "degree b = 0" by auto
- then obtain a' b' where ab': "a = [:a':]" "b = [:b':]" by (auto elim!: degree_eq_zeroE)
- from ab have "coeff [:c:] 0 = coeff (a * b) 0" by (simp only: )
- hence "c = a' * b'" by (simp add: ab' mult_ac)
- from A and this have "a' dvd 1 \<or> b' dvd 1" by (rule irreducibleD)
- with ab' show "a dvd 1 \<or> b dvd 1" by (auto simp: one_poly_def)
- qed (insert A, auto simp: irreducible_def is_unit_poly_iff)
-next
- assume A: "irreducible [:c:]"
- show "irreducible c"
- proof (rule irreducibleI)
- fix a b assume ab: "c = a * b"
- hence "[:c:] = [:a:] * [:b:]" by (simp add: mult_ac)
- from A and this have "[:a:] dvd 1 \<or> [:b:] dvd 1" by (rule irreducibleD)
- thus "a dvd 1 \<or> b dvd 1" by (simp add: one_poly_def)
- qed (insert A, auto simp: irreducible_def one_poly_def)
-qed
-
-lemma lead_coeff_monom [simp]: "lead_coeff (monom c n) = c"
- by (cases "c = 0") (simp_all add: lead_coeff_def degree_monom_eq)
-
-
-subsection \<open>Normalisation of polynomials\<close>
-
-instantiation poly :: ("{normalization_semidom,idom_divide}") normalization_semidom
-begin
-
-definition unit_factor_poly :: "'a poly \<Rightarrow> 'a poly"
- where "unit_factor_poly p = monom (unit_factor (lead_coeff p)) 0"
-
-definition normalize_poly :: "'a poly \<Rightarrow> 'a poly"
- where "normalize_poly p = map_poly (\<lambda>x. x div unit_factor (lead_coeff p)) p"
-
-lemma normalize_poly_altdef:
- "normalize p = p div [:unit_factor (lead_coeff p):]"
-proof (cases "p = 0")
- case False
- thus ?thesis
- by (subst div_const_poly_conv_map_poly)
- (auto simp: normalize_poly_def const_poly_dvd_iff lead_coeff_def )
-qed (auto simp: normalize_poly_def)
-
-instance
-proof
- fix p :: "'a poly"
- show "unit_factor p * normalize p = p"
- by (cases "p = 0")
- (simp_all add: unit_factor_poly_def normalize_poly_def monom_0
- smult_conv_map_poly map_poly_map_poly o_def)
-next
- fix p :: "'a poly"
- assume "is_unit p"
- then obtain c where p: "p = [:c:]" "is_unit c" by (auto simp: is_unit_poly_iff)
- thus "normalize p = 1"
- by (simp add: normalize_poly_def map_poly_pCons is_unit_normalize one_poly_def)
-next
- fix p :: "'a poly" assume "p \<noteq> 0"
- thus "is_unit (unit_factor p)"
- by (simp add: unit_factor_poly_def monom_0 is_unit_poly_iff)
-qed (simp_all add: normalize_poly_def unit_factor_poly_def monom_0 lead_coeff_mult unit_factor_mult)
-
-end
-
-lemma unit_factor_pCons:
- "unit_factor (pCons a p) = (if p = 0 then monom (unit_factor a) 0 else unit_factor p)"
- by (simp add: unit_factor_poly_def)
-
-lemma normalize_monom [simp]:
- "normalize (monom a n) = monom (normalize a) n"
- by (simp add: map_poly_monom normalize_poly_def)
-
-lemma unit_factor_monom [simp]:
- "unit_factor (monom a n) = monom (unit_factor a) 0"
- by (simp add: unit_factor_poly_def )
-
-lemma normalize_const_poly: "normalize [:c:] = [:normalize c:]"
- by (simp add: normalize_poly_def map_poly_pCons)
-
-lemma normalize_smult: "normalize (smult c p) = smult (normalize c) (normalize p)"
-proof -
- have "smult c p = [:c:] * p" by simp
- also have "normalize \<dots> = smult (normalize c) (normalize p)"
- by (subst normalize_mult) (simp add: normalize_const_poly)
- finally show ?thesis .
-qed
-
-lemma is_unit_smult_iff: "smult c p dvd 1 \<longleftrightarrow> c dvd 1 \<and> p dvd 1"
-proof -
- have "smult c p = [:c:] * p" by simp
- also have "\<dots> dvd 1 \<longleftrightarrow> c dvd 1 \<and> p dvd 1"
- proof safe
- assume A: "[:c:] * p dvd 1"
- thus "p dvd 1" by (rule dvd_mult_right)
- from A obtain q where B: "1 = [:c:] * p * q" by (erule dvdE)
- have "c dvd c * (coeff p 0 * coeff q 0)" by simp
- also have "\<dots> = coeff ([:c:] * p * q) 0" by (simp add: mult.assoc coeff_mult)
- also note B [symmetric]
- finally show "c dvd 1" by simp
- next
- assume "c dvd 1" "p dvd 1"
- from \<open>c dvd 1\<close> obtain d where "1 = c * d" by (erule dvdE)
- hence "1 = [:c:] * [:d:]" by (simp add: one_poly_def mult_ac)
- hence "[:c:] dvd 1" by (rule dvdI)
- from mult_dvd_mono[OF this \<open>p dvd 1\<close>] show "[:c:] * p dvd 1" by simp
- qed
- finally show ?thesis .
-qed
-
subsection \<open>Content and primitive part of a polynomial\<close>
@@ -1243,7 +904,7 @@
end
-
+
subsection \<open>Prime factorisation of polynomials\<close>
context
@@ -1264,7 +925,8 @@
by (simp add: e_def content_prod_mset multiset.map_comp o_def)
also have "image_mset (\<lambda>x. content (primitive_part_fract x)) ?P = image_mset (\<lambda>_. 1) ?P"
by (intro image_mset_cong content_primitive_part_fract) auto
- finally have content_e: "content e = 1" by (simp add: prod_mset_const)
+ finally have content_e: "content e = 1"
+ by simp
have "fract_poly p = unit_factor_field_poly (fract_poly p) *
normalize_field_poly (fract_poly p)" by simp
@@ -1277,7 +939,7 @@
image_mset (\<lambda>x. [:fract_content x:] * fract_poly (primitive_part_fract x)) ?P"
by (intro image_mset_cong) (auto simp: content_times_primitive_part_fract)
also have "prod_mset \<dots> = smult c (fract_poly e)"
- by (subst prod_mset_mult) (simp_all add: prod_mset_fract_poly prod_mset_const_poly c_def e_def)
+ by (subst prod_mset.distrib) (simp_all add: prod_mset_fract_poly prod_mset_const_poly c_def e_def)
also have "[:to_fract (lead_coeff p):] * \<dots> = smult c' (fract_poly e)"
by (simp add: c'_def)
finally have eq: "fract_poly p = smult c' (fract_poly e)" .
@@ -1466,20 +1128,22 @@
smult (gcd (content p) (content q))
(gcd_poly_code_aux (primitive_part p) (primitive_part q)))"
+lemma gcd_poly_code [code]: "gcd p q = gcd_poly_code p q"
+ by (simp add: gcd_poly_code_def gcd_poly_code_aux_correct gcd_poly_decompose [symmetric])
+
lemma lcm_poly_code [code]:
fixes p q :: "'a :: factorial_ring_gcd poly"
shows "lcm p q = normalize (p * q) div gcd p q"
- by (rule lcm_gcd)
-
-lemma gcd_poly_code [code]: "gcd p q = gcd_poly_code p q"
- by (simp add: gcd_poly_code_def gcd_poly_code_aux_correct gcd_poly_decompose [symmetric])
+ by (fact lcm_gcd)
declare Gcd_set
[where ?'a = "'a :: factorial_ring_gcd poly", code]
declare Lcm_set
[where ?'a = "'a :: factorial_ring_gcd poly", code]
+
+text \<open>Example:
+ @{lemma "Lcm {[:1, 2, 3:], [:2, 3, 4:]} = [:[:2:], [:7:], [:16:], [:17:], [:12 :: int:]:]" by eval}
+\<close>
-value [code] "Lcm {[:1,2,3:], [:2,3,4::int poly:]}"
-
end
--- a/src/HOL/Library/code_test.ML Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Library/code_test.ML Tue Dec 20 16:18:56 2016 +0100
@@ -1,59 +1,44 @@
(* Title: HOL/Library/code_test.ML
- Author: Andreas Lochbihler, ETH Zurich
+ Author: Andreas Lochbihler, ETH Zürich
-Test infrastructure for the code generator
+Test infrastructure for the code generator.
*)
-signature CODE_TEST = sig
- val add_driver : string * ((Proof.context -> (string * string) list * string -> Path.T -> string) * string) -> theory -> theory
- val get_driver : theory -> string -> ((Proof.context -> (string * string) list * string -> Path.T -> string) * string) option
- val overlord : bool Config.T
- val successN : string
- val failureN : string
- val start_markerN : string
- val end_markerN : string
- val test_terms : Proof.context -> term list -> string -> unit
- val test_targets : Proof.context -> term list -> string list -> unit list
- val test_code_cmd : string list -> string list -> Toplevel.state -> unit
-
- val eval_term : string -> Proof.context -> term -> term
-
- val gen_driver :
+signature CODE_TEST =
+sig
+ val add_driver:
+ string * ((Proof.context -> (string * string) list * string -> Path.T -> string) * string) ->
+ theory -> theory
+ val overlord: bool Config.T
+ val successN: string
+ val failureN: string
+ val start_markerN: string
+ val end_markerN: string
+ val test_terms: Proof.context -> term list -> string -> unit
+ val test_targets: Proof.context -> term list -> string list -> unit
+ val test_code_cmd: string list -> string list -> Proof.context -> unit
+ val eval_term: string -> Proof.context -> term -> term
+ val evaluate:
(theory -> Path.T -> string list -> string ->
- {files : (Path.T * string) list,
- compile_cmd : string option, run_cmd : string, mk_code_file : string -> Path.T})
- -> string -> string -> string
- -> theory -> (string * string) list * string -> Path.T -> string
-
- val ISABELLE_POLYML : string
- val polymlN : string
- val evaluate_in_polyml : Proof.context -> (string * string) list * string -> Path.T -> string
-
- val mltonN : string
- val ISABELLE_MLTON : string
- val evaluate_in_mlton : Proof.context -> (string * string) list * string -> Path.T -> string
-
- val smlnjN : string
- val ISABELLE_SMLNJ : string
- val evaluate_in_smlnj : Proof.context -> (string * string) list * string -> Path.T -> string
-
- val ocamlN : string
- val ISABELLE_OCAMLC : string
- val evaluate_in_ocaml : Proof.context -> (string * string) list * string -> Path.T -> string
-
- val ghcN : string
- val ISABELLE_GHC : string
- val ghc_options : string Config.T
- val evaluate_in_ghc : Proof.context -> (string * string) list * string -> Path.T -> string
-
- val scalaN : string
- val ISABELLE_SCALA : string
- val evaluate_in_scala : Proof.context -> (string * string) list * string -> Path.T -> string
+ {files: (Path.T * string) list,
+ compile_cmd: string option,
+ run_cmd: string,
+ mk_code_file: string -> Path.T}) -> (string * string) option -> string -> theory ->
+ (string * string) list * string -> Path.T -> string
+ val evaluate_in_polyml: Proof.context -> (string * string) list * string -> Path.T -> string
+ val evaluate_in_mlton: Proof.context -> (string * string) list * string -> Path.T -> string
+ val evaluate_in_smlnj: Proof.context -> (string * string) list * string -> Path.T -> string
+ val evaluate_in_ocaml: Proof.context -> (string * string) list * string -> Path.T -> string
+ val ghc_options: string Config.T
+ val evaluate_in_ghc: Proof.context -> (string * string) list * string -> Path.T -> string
+ val evaluate_in_scala: Proof.context -> (string * string) list * string -> Path.T -> string
end
-structure Code_Test : CODE_TEST = struct
+structure Code_Test: CODE_TEST =
+struct
(* convert a list of terms into nested tuples and back *)
+
fun mk_tuples [] = @{term "()"}
| mk_tuples [t] = t
| mk_tuples (t :: ts) = HOLogic.mk_prod (t, mk_tuples ts)
@@ -62,47 +47,46 @@
| dest_tuples t = [t]
-fun map_option _ NONE = NONE
- | map_option f (SOME x) = SOME (f x)
-
fun last_field sep str =
let
- val n = size sep;
- val len = size str;
+ val n = size sep
+ val len = size str
fun find i =
if i < 0 then NONE
else if String.substring (str, i, n) = sep then SOME i
- else find (i - 1);
+ else find (i - 1)
in
(case find (len - n) of
NONE => NONE
| SOME i => SOME (String.substring (str, 0, i), String.extract (str, i + n, NONE)))
- end;
+ end
fun split_first_last start stop s =
- case first_field start s
- of NONE => NONE
- | SOME (initial, rest) =>
- case last_field stop rest
- of NONE => NONE
- | SOME (middle, tail) => SOME (initial, middle, tail);
+ (case first_field start s of
+ NONE => NONE
+ | SOME (initial, rest) =>
+ (case last_field stop rest of
+ NONE => NONE
+ | SOME (middle, tail) => SOME (initial, middle, tail)))
-(* Data slot for drivers *)
+
+(* data slot for drivers *)
structure Drivers = Theory_Data
(
- type T = (string * ((Proof.context -> (string * string) list * string -> Path.T -> string) * string)) list;
- val empty = [];
- val extend = I;
- fun merge data : T = AList.merge (op =) (K true) data;
+ type T =
+ (string * ((Proof.context -> (string * string) list * string -> Path.T -> string) * string)) list
+ val empty = []
+ val extend = I
+ fun merge data : T = AList.merge (op =) (K true) data
)
-val add_driver = Drivers.map o AList.update (op =);
-val get_driver = AList.lookup (op =) o Drivers.get;
+val add_driver = Drivers.map o AList.update (op =)
+val get_driver = AList.lookup (op =) o Drivers.get
(*
Test drivers must produce output of the following format:
-
+
The start of the relevant data is marked with start_markerN,
its end with end_markerN.
@@ -112,7 +96,8 @@
There must not be any additional whitespace in between.
*)
-(* Parsing of results *)
+
+(* parsing of results *)
val successN = "True"
val failureN = "False"
@@ -121,7 +106,7 @@
fun parse_line line =
if String.isPrefix successN line then (true, NONE)
- else if String.isPrefix failureN line then (false,
+ else if String.isPrefix failureN line then (false,
if size line > size failureN then
String.extract (line, size failureN, NONE)
|> YXML.parse_body
@@ -132,20 +117,21 @@
else raise Fail ("Cannot parse result of evaluation:\n" ^ line)
fun parse_result target out =
- case split_first_last start_markerN end_markerN out
- of NONE => error ("Evaluation failed for " ^ target ^ "!\nBash output:\n" ^ out)
- | SOME (_, middle, _) => middle |> trim_line |> split_lines |> map parse_line
+ (case split_first_last start_markerN end_markerN out of
+ NONE => error ("Evaluation failed for " ^ target ^ "!\nBash output:\n" ^ out)
+ | SOME (_, middle, _) => middle |> trim_line |> split_lines |> map parse_line)
-(* Pretty printing of test results *)
+
+(* pretty printing of test results *)
fun pretty_eval _ NONE _ = []
- | pretty_eval ctxt (SOME evals) ts =
- [Pretty.fbrk,
- Pretty.big_list "Evaluated terms"
- (map (fn (t, eval) => Pretty.block
- [Syntax.pretty_term ctxt t, Pretty.brk 1, Pretty.str "=", Pretty.brk 1,
- Syntax.pretty_term ctxt eval])
- (ts ~~ evals))]
+ | pretty_eval ctxt (SOME evals) ts =
+ [Pretty.fbrk,
+ Pretty.big_list "Evaluated terms"
+ (map (fn (t, eval) => Pretty.block
+ [Syntax.pretty_term ctxt t, Pretty.brk 1, Pretty.str "=", Pretty.brk 1,
+ Syntax.pretty_term ctxt eval])
+ (ts ~~ evals))]
fun pretty_failure ctxt target (((_, evals), query), eval_ts) =
Pretty.block (Pretty.text ("Test in " ^ target ^ " failed for")
@@ -155,60 +141,61 @@
fun pretty_failures ctxt target failures =
Pretty.blk (0, Pretty.fbreaks (map (pretty_failure ctxt target) failures))
-(* Driver invocation *)
-val overlord = Attrib.setup_config_bool @{binding "code_test_overlord"} (K false);
+(* driver invocation *)
+
+val overlord = Attrib.setup_config_bool @{binding "code_test_overlord"} (K false)
fun with_overlord_dir name f =
let
- val path = Path.append (Path.explode "$ISABELLE_HOME_USER") (Path.basic (name ^ serial_string ()))
- val _ = Isabelle_System.mkdirs path;
- in
- Exn.release (Exn.capture f path)
- end;
+ val path =
+ Path.append (Path.explode "$ISABELLE_HOME_USER") (Path.basic (name ^ serial_string ()))
+ val _ = Isabelle_System.mkdirs path
+ in Exn.release (Exn.capture f path) end
fun dynamic_value_strict ctxt t compiler =
let
val thy = Proof_Context.theory_of ctxt
- val (driver, target) = case get_driver thy compiler
- of NONE => error ("No driver for target " ^ compiler)
- | SOME f => f;
+ val (driver, target) =
+ (case get_driver thy compiler of
+ NONE => error ("No driver for target " ^ compiler)
+ | SOME f => f)
val debug = Config.get (Proof_Context.init_global thy) overlord
val with_dir = if debug then with_overlord_dir else Isabelle_System.with_tmp_dir
fun evaluate f = with_dir "Code_Test" (driver ctxt f) |> parse_result compiler
fun evaluator program _ vs_ty deps =
- Exn.interruptible_capture evaluate (Code_Target.computation_text ctxt target program deps true vs_ty);
- fun postproc f = map (apsnd (map_option (map f)))
- in
- Exn.release (Code_Thingol.dynamic_value ctxt (Exn.map_res o postproc) evaluator t)
- end;
+ Exn.interruptible_capture evaluate
+ (Code_Target.computation_text ctxt target program deps true vs_ty)
+ fun postproc f = map (apsnd (Option.map (map f)))
+ in Exn.release (Code_Thingol.dynamic_value ctxt (Exn.map_res o postproc) evaluator t) end
-(* Term preprocessing *)
+
+(* term preprocessing *)
fun add_eval (Const (@{const_name Trueprop}, _) $ t) = add_eval t
| add_eval (Const (@{const_name "HOL.eq"}, _) $ lhs $ rhs) = (fn acc =>
- acc
- |> add_eval rhs
- |> add_eval lhs
- |> cons rhs
- |> cons lhs)
+ acc
+ |> add_eval rhs
+ |> add_eval lhs
+ |> cons rhs
+ |> cons lhs)
| add_eval (Const (@{const_name "Not"}, _) $ t) = add_eval t
| add_eval (Const (@{const_name "Orderings.ord_class.less_eq"}, _) $ lhs $ rhs) = (fn acc =>
- lhs :: rhs :: acc)
+ lhs :: rhs :: acc)
| add_eval (Const (@{const_name "Orderings.ord_class.less"}, _) $ lhs $ rhs) = (fn acc =>
- lhs :: rhs :: acc)
+ lhs :: rhs :: acc)
| add_eval _ = I
fun mk_term_of [] = @{term "None :: (unit \<Rightarrow> yxml_of_term) option"}
| mk_term_of ts =
- let
- val tuple = mk_tuples ts
- val T = fastype_of tuple
- in
- @{term "Some :: (unit \<Rightarrow> yxml_of_term) \<Rightarrow> (unit \<Rightarrow> yxml_of_term) option"} $
- (absdummy @{typ unit} (@{const yxml_string_of_term} $
- (Const (@{const_name Code_Evaluation.term_of}, T --> @{typ term}) $ tuple)))
- end
+ let
+ val tuple = mk_tuples ts
+ val T = fastype_of tuple
+ in
+ @{term "Some :: (unit \<Rightarrow> yxml_of_term) \<Rightarrow> (unit \<Rightarrow> yxml_of_term) option"} $
+ (absdummy @{typ unit} (@{const yxml_string_of_term} $
+ (Const (@{const_name Code_Evaluation.term_of}, T --> @{typ term}) $ tuple)))
+ end
fun test_terms ctxt ts target =
let
@@ -216,109 +203,121 @@
fun term_of t = Sign.of_sort thy (fastype_of t, @{sort term_of})
- fun ensure_bool t = case fastype_of t of @{typ bool} => ()
- | _ => error ("Test case not of type bool: " ^ Pretty.string_of (Syntax.pretty_term ctxt t))
+ fun ensure_bool t =
+ (case fastype_of t of
+ @{typ bool} => ()
+ | _ =>
+ error (Pretty.string_of
+ (Pretty.block [Pretty.str "Test case not of type bool:", Pretty.brk 1,
+ Syntax.pretty_term ctxt t])))
- val _ = map ensure_bool ts
+ val _ = List.app ensure_bool ts
val evals = map (fn t => filter term_of (add_eval t [])) ts
val eval = map mk_term_of evals
- val T = HOLogic.mk_prodT (@{typ bool}, Type (@{type_name option}, [@{typ unit} --> @{typ yxml_of_term}]))
- val t = HOLogic.mk_list T (map HOLogic.mk_prod (ts ~~ eval))
+ val t =
+ HOLogic.mk_list @{typ "bool \<times> (unit \<Rightarrow> yxml_of_term) option"}
+ (map HOLogic.mk_prod (ts ~~ eval))
- val result = dynamic_value_strict ctxt t target;
+ val result = dynamic_value_strict ctxt t target
val failed =
filter_out (fst o fst o fst) (result ~~ ts ~~ evals)
- handle ListPair.UnequalLengths =>
+ handle ListPair.UnequalLengths =>
error ("Evaluation failed!\nWrong number of test results: " ^ Int.toString (length result))
- val _ = case failed of [] => ()
- | _ => error (Pretty.string_of (pretty_failures ctxt target failed))
in
- ()
+ (case failed of [] =>
+ ()
+ | _ => error (Pretty.string_of (pretty_failures ctxt target failed)))
end
-fun test_targets ctxt = map o test_terms ctxt
+fun test_targets ctxt = List.app o test_terms ctxt
-fun test_code_cmd raw_ts targets state =
+fun pretty_free ctxt = Syntax.pretty_term ctxt o Free
+
+fun test_code_cmd raw_ts targets ctxt =
let
- val ctxt = Toplevel.context_of state;
- val ts = Syntax.read_terms ctxt raw_ts;
- val frees = fold Term.add_free_names ts []
- val _ = if frees = [] then () else
- error ("Terms contain free variables: " ^
- Pretty.string_of (Pretty.block (Pretty.commas (map Pretty.str frees))))
- in
- test_targets ctxt ts targets; ()
- end
+ val ts = Syntax.read_terms ctxt raw_ts
+ val frees = fold Term.add_frees ts []
+ val _ =
+ if null frees then ()
+ else error (Pretty.string_of
+ (Pretty.block (Pretty.str "Terms contain free variables:" :: Pretty.brk 1 ::
+ Pretty.commas (map (pretty_free ctxt) frees))))
+ in test_targets ctxt ts targets end
fun eval_term target ctxt t =
let
- val frees = Term.add_free_names t []
- val _ = if frees = [] then () else
- error ("Term contains free variables: " ^
- Pretty.string_of (Pretty.block (Pretty.commas (map Pretty.str frees))))
+ val frees = Term.add_frees t []
+ val _ =
+ if null frees then ()
+ else
+ error (Pretty.string_of
+ (Pretty.block (Pretty.str "Term contains free variables:" :: Pretty.brk 1 ::
+ Pretty.commas (map (pretty_free ctxt) frees))))
- val thy = Proof_Context.theory_of ctxt
+ val T = fastype_of t
+ val _ =
+ if Sign.of_sort (Proof_Context.theory_of ctxt) (T, @{sort term_of}) then ()
+ else error ("Type " ^ Syntax.string_of_typ ctxt T ^
+ " of term not of sort " ^ Syntax.string_of_sort ctxt @{sort term_of})
- val T_t = fastype_of t
- val _ = if Sign.of_sort thy (T_t, @{sort term_of}) then () else error
- ("Type " ^ Pretty.string_of (Syntax.pretty_typ ctxt T_t) ^
- " of term not of sort " ^ Pretty.string_of (Syntax.pretty_sort ctxt @{sort term_of}))
+ val t' =
+ HOLogic.mk_list @{typ "bool \<times> (unit \<Rightarrow> yxml_of_term) option"}
+ [HOLogic.mk_prod (@{term "False"}, mk_term_of [t])]
- val T = HOLogic.mk_prodT (@{typ bool}, Type (@{type_name option}, [@{typ unit} --> @{typ yxml_of_term}]))
- val t' = HOLogic.mk_list T [HOLogic.mk_prod (@{term "False"}, mk_term_of [t])]
+ val result = dynamic_value_strict ctxt t' target
+ in (case result of [(_, SOME [t])] => t | _ => error "Evaluation failed") end
- val result = dynamic_value_strict ctxt t' target;
- in
- case result of [(_, SOME [t])] => t | _ => error "Evaluation failed"
- end
+
+(* generic driver *)
-(* Generic driver *)
-
-fun gen_driver mk_driver env_var env_var_dest compilerN ctxt (code_files, value_name) =
+fun evaluate mk_driver opt_env_var compilerN ctxt (code_files, value_name) =
let
- val compiler = getenv env_var
- val _ = if compiler <> "" then () else error (Pretty.string_of (Pretty.para
- ("Environment variable " ^ env_var ^ " is not set. To test code generation with " ^
- compilerN ^ ", set this variable to your " ^ env_var_dest ^ " in the settings file.")))
+ val _ =
+ (case opt_env_var of
+ NONE => ()
+ | SOME (env_var, env_var_dest) =>
+ (case getenv env_var of
+ "" =>
+ error (Pretty.string_of (Pretty.para
+ ("Environment variable " ^ env_var ^ " is not set. To test code generation with " ^
+ compilerN ^ ", set this variable to your " ^ env_var_dest ^
+ " in the $ISABELLE_HOME_USER/etc/settings file.")))
+ | _ => ()))
fun compile NONE = ()
| compile (SOME cmd) =
- let
- val (out, ret) = Isabelle_System.bash_output cmd
- in
- if ret = 0 then () else error
- ("Compilation with " ^ compilerN ^ " failed:\n" ^ cmd ^ "\n" ^ out)
- end
+ let
+ val (out, ret) = Isabelle_System.bash_output cmd
+ in
+ if ret = 0 then ()
+ else error ("Compilation with " ^ compilerN ^ " failed:\n" ^ cmd ^ "\n" ^ out)
+ end
- fun run (path : Path.T)=
+ fun run path =
let
val modules = map fst code_files
- val {files, compile_cmd, run_cmd, mk_code_file}
- = mk_driver ctxt path modules value_name
+ val {files, compile_cmd, run_cmd, mk_code_file} = mk_driver ctxt path modules value_name
- val _ = map (fn (name, code) => File.write (mk_code_file name) code) code_files
- val _ = map (fn (name, content) => File.write name content) files
+ val _ = List.app (fn (name, code) => File.write (mk_code_file name) code) code_files
+ val _ = List.app (fn (name, content) => File.write name content) files
val _ = compile compile_cmd
val (out, res) = Isabelle_System.bash_output run_cmd
- val _ = if res = 0 then () else error
- ("Evaluation for " ^ compilerN ^ " terminated with error code " ^ Int.toString res ^
- "\nBash output:\n" ^ out)
- in
- out
- end
- in
- run
- end
+ val _ =
+ if res = 0 then ()
+ else error ("Evaluation for " ^ compilerN ^ " terminated with error code " ^
+ Int.toString res ^ "\nBash output:\n" ^ out)
+ in out end
+ in run end
-(* Driver for PolyML *)
-val ISABELLE_POLYML = "ISABELLE_POLYML"
-val polymlN = "PolyML";
+(* driver for PolyML *)
+
+val polymlN = "PolyML"
fun mk_driver_polyml _ path _ value_name =
let
@@ -327,10 +326,10 @@
val code_path = Path.append path (Path.basic generatedN)
val driver_path = Path.append path (Path.basic driverN)
- val driver =
+ val driver =
"fun main prog_name = \n" ^
" let\n" ^
- " fun format_term NONE = \"\"\n" ^
+ " fun format_term NONE = \"\"\n" ^
" | format_term (SOME t) = t ();\n" ^
" fun format (true, _) = \"" ^ successN ^ "\\n\"\n" ^
" | format (false, to) = \"" ^ failureN ^ "\" ^ format_term to ^ \"\\n\";\n" ^
@@ -342,17 +341,16 @@
" ()\n" ^
" end;\n"
val cmd =
- "echo \"use \\\"" ^ Path.implode code_path ^ "\\\"; use \\\"" ^
- Path.implode driver_path ^ "\\\"; main ();\" | " ^
- Path.implode (Path.variable ISABELLE_POLYML)
+ "echo \"use \\\"" ^ Path.implode code_path ^ "\\\"; use \\\"" ^
+ Path.implode driver_path ^ "\\\"; main ();\" | \"$ML_HOME/poly\""
in
{files = [(driver_path, driver)], compile_cmd = NONE, run_cmd = cmd, mk_code_file = K code_path}
end
-fun evaluate_in_polyml ctxt =
- gen_driver mk_driver_polyml ISABELLE_POLYML "PolyML executable" polymlN ctxt
+fun evaluate_in_polyml ctxt = evaluate mk_driver_polyml NONE polymlN ctxt
-(* Driver for mlton *)
+
+(* driver for mlton *)
val mltonN = "MLton"
val ISABELLE_MLTON = "ISABELLE_MLTON"
@@ -367,8 +365,8 @@
val code_path = Path.append path (Path.basic generatedN)
val driver_path = Path.append path (Path.basic driverN)
val ml_basis_path = Path.append path (Path.basic ml_basisN)
- val driver =
- "fun format_term NONE = \"\"\n" ^
+ val driver =
+ "fun format_term NONE = \"\"\n" ^
" | format_term (SOME t) = t ();\n" ^
"fun format (true, _) = \"" ^ successN ^ "\\n\"\n" ^
" | format (false, to) = \"" ^ failureN ^ "\" ^ format_term to ^ \"\\n\";\n" ^
@@ -391,9 +389,10 @@
end
fun evaluate_in_mlton ctxt =
- gen_driver mk_driver_mlton ISABELLE_MLTON "MLton executable" mltonN ctxt
+ evaluate mk_driver_mlton (SOME (ISABELLE_MLTON, "MLton executable")) mltonN ctxt
-(* Driver for SML/NJ *)
+
+(* driver for SML/NJ *)
val smlnjN = "SMLNJ"
val ISABELLE_SMLNJ = "ISABELLE_SMLNJ"
@@ -405,11 +404,11 @@
val code_path = Path.append path (Path.basic generatedN)
val driver_path = Path.append path (Path.basic driverN)
- val driver =
+ val driver =
"structure Test = struct\n" ^
"fun main prog_name =\n" ^
" let\n" ^
- " fun format_term NONE = \"\"\n" ^
+ " fun format_term NONE = \"\"\n" ^
" | format_term (SOME t) = t ();\n" ^
" fun format (true, _) = \"" ^ successN ^ "\\n\"\n" ^
" | format (false, to) = \"" ^ failureN ^ "\" ^ format_term to ^ \"\\n\";\n" ^
@@ -430,9 +429,10 @@
end
fun evaluate_in_smlnj ctxt =
- gen_driver mk_driver_smlnj ISABELLE_SMLNJ "SMLNJ executable" smlnjN ctxt
+ evaluate mk_driver_smlnj (SOME (ISABELLE_SMLNJ, "SMLNJ executable")) smlnjN ctxt
-(* Driver for OCaml *)
+
+(* driver for OCaml *)
val ocamlN = "OCaml"
val ISABELLE_OCAMLC = "ISABELLE_OCAMLC"
@@ -444,9 +444,9 @@
val code_path = Path.append path (Path.basic generatedN)
val driver_path = Path.append path (Path.basic driverN)
- val driver =
+ val driver =
"let format_term = function\n" ^
- " | None -> \"\"\n" ^
+ " | None -> \"\"\n" ^
" | Some t -> t ();;\n" ^
"let format = function\n" ^
" | (true, _) -> \"" ^ successN ^ "\\n\"\n" ^
@@ -471,9 +471,10 @@
end
fun evaluate_in_ocaml ctxt =
- gen_driver mk_driver_ocaml ISABELLE_OCAMLC "ocamlc executable" ocamlN ctxt
+ evaluate mk_driver_ocaml (SOME (ISABELLE_OCAMLC, "ocamlc executable")) ocamlN ctxt
-(* Driver for GHC *)
+
+(* driver for GHC *)
val ghcN = "GHC"
val ISABELLE_GHC = "ISABELLE_GHC"
@@ -486,12 +487,12 @@
fun mk_code_file name = Path.append path (Path.basic (name ^ ".hs"))
val driver_path = Path.append path (Path.basic driverN)
- val driver =
+ val driver =
"module Main where {\n" ^
String.concat (map (fn module => "import qualified " ^ module ^ ";\n") modules) ^
"main = do {\n" ^
" let {\n" ^
- " format_term Nothing = \"\";\n" ^
+ " format_term Nothing = \"\";\n" ^
" format_term (Just t) = t ();\n" ^
" format (True, _) = \"" ^ successN ^ "\\n\";\n" ^
" format (False, to) = \"" ^ failureN ^ "\" ++ format_term to ++ \"\\n\";\n" ^
@@ -516,12 +517,12 @@
end
fun evaluate_in_ghc ctxt =
- gen_driver mk_driver_ghc ISABELLE_GHC "GHC executable" ghcN ctxt
+ evaluate mk_driver_ghc (SOME (ISABELLE_GHC, "GHC executable")) ghcN ctxt
-(* Driver for Scala *)
+
+(* driver for Scala *)
val scalaN = "Scala"
-val ISABELLE_SCALA = "ISABELLE_SCALA"
fun mk_driver_scala _ path _ value_name =
let
@@ -530,7 +531,7 @@
val code_path = Path.append path (Path.basic (generatedN ^ ".scala"))
val driver_path = Path.append path (Path.basic driverN)
- val driver =
+ val driver =
"import " ^ generatedN ^ "._\n" ^
"object Test {\n" ^
" def format_term(x : Option[Unit => String]) : String = x match {\n" ^
@@ -550,37 +551,36 @@
"}\n"
val compile_cmd =
- Path.implode (Path.append (Path.variable ISABELLE_SCALA) (Path.basic "scalac")) ^
- " -d " ^ File.bash_path path ^ " -classpath " ^ File.bash_path path ^ " " ^
+ "\"$SCALA_HOME/bin/scalac\" -d " ^ File.bash_path path ^
+ " -classpath " ^ File.bash_path path ^ " " ^
File.bash_path code_path ^ " " ^ File.bash_path driver_path
- val run_cmd =
- Path.implode (Path.append (Path.variable ISABELLE_SCALA) (Path.basic "scala")) ^
- " -cp " ^ File.bash_path path ^ " Test"
+ val run_cmd = "\"$SCALA_HOME/bin/scala\" -cp " ^ File.bash_path path ^ " Test"
in
{files = [(driver_path, driver)],
compile_cmd = SOME compile_cmd, run_cmd = run_cmd, mk_code_file = K code_path}
end
-fun evaluate_in_scala ctxt =
- gen_driver mk_driver_scala ISABELLE_SCALA "Scala directory" scalaN ctxt
+fun evaluate_in_scala ctxt = evaluate mk_driver_scala NONE scalaN ctxt
+
-val test_codeP = Scan.repeat1 Parse.prop -- (@{keyword "in"} |-- Scan.repeat1 Parse.name)
+(* command setup *)
-val _ =
+val _ =
Outer_Syntax.command @{command_keyword test_code}
"compile test cases to target languages, execute them and report results"
- (test_codeP >> (fn (raw_ts, targets) => Toplevel.keep (test_code_cmd raw_ts targets)))
+ (Scan.repeat1 Parse.prop -- (@{keyword "in"} |-- Scan.repeat1 Parse.name)
+ >> (fn (ts, targets) => Toplevel.keep (test_code_cmd ts targets o Toplevel.context_of)))
-val _ = Theory.setup (fold add_driver
- [(polymlN, (evaluate_in_polyml, Code_ML.target_SML)),
- (mltonN, (evaluate_in_mlton, Code_ML.target_SML)),
- (smlnjN, (evaluate_in_smlnj, Code_ML.target_SML)),
- (ocamlN, (evaluate_in_ocaml, Code_ML.target_OCaml)),
- (ghcN, (evaluate_in_ghc, Code_Haskell.target)),
- (scalaN, (evaluate_in_scala, Code_Scala.target))]
+val _ =
+ Theory.setup (fold add_driver
+ [(polymlN, (evaluate_in_polyml, Code_ML.target_SML)),
+ (mltonN, (evaluate_in_mlton, Code_ML.target_SML)),
+ (smlnjN, (evaluate_in_smlnj, Code_ML.target_SML)),
+ (ocamlN, (evaluate_in_ocaml, Code_ML.target_OCaml)),
+ (ghcN, (evaluate_in_ghc, Code_Haskell.target)),
+ (scalaN, (evaluate_in_scala, Code_Scala.target))]
#> fold (fn target => Value_Command.add_evaluator (target, eval_term target))
- [polymlN, mltonN, smlnjN, ocamlN, ghcN, scalaN]);
+ [polymlN, mltonN, smlnjN, ocamlN, ghcN, scalaN])
end
-
--- a/src/HOL/Nonstandard_Analysis/CLim.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Nonstandard_Analysis/CLim.thy Tue Dec 20 16:18:56 2016 +0100
@@ -4,198 +4,178 @@
Conversion to Isar and new proofs by Lawrence C Paulson, 2004
*)
-section\<open>Limits, Continuity and Differentiation for Complex Functions\<close>
+section \<open>Limits, Continuity and Differentiation for Complex Functions\<close>
theory CLim
-imports CStar
+ imports CStar
begin
(*not in simpset?*)
declare hypreal_epsilon_not_zero [simp]
(*??generalize*)
-lemma lemma_complex_mult_inverse_squared [simp]:
- "x \<noteq> (0::complex) \<Longrightarrow> x * (inverse x)\<^sup>2 = inverse x"
-by (simp add: numeral_2_eq_2)
+lemma lemma_complex_mult_inverse_squared [simp]: "x \<noteq> 0 \<Longrightarrow> x * (inverse x)\<^sup>2 = inverse x"
+ for x :: complex
+ by (simp add: numeral_2_eq_2)
+
+text \<open>Changing the quantified variable. Install earlier?\<close>
+lemma all_shift: "(\<forall>x::'a::comm_ring_1. P x) \<longleftrightarrow> (\<forall>x. P (x - a))"
+ apply auto
+ apply (drule_tac x = "x + a" in spec)
+ apply (simp add: add.assoc)
+ done
-text\<open>Changing the quantified variable. Install earlier?\<close>
-lemma all_shift: "(\<forall>x::'a::comm_ring_1. P x) = (\<forall>x. P (x-a))"
-apply auto
-apply (drule_tac x="x+a" in spec)
-apply (simp add: add.assoc)
-done
+lemma complex_add_minus_iff [simp]: "x + - a = 0 \<longleftrightarrow> x = a"
+ for x a :: complex
+ by (simp add: diff_eq_eq)
-lemma complex_add_minus_iff [simp]: "(x + - a = (0::complex)) = (x=a)"
-by (simp add: diff_eq_eq)
-
-lemma complex_add_eq_0_iff [iff]: "(x+y = (0::complex)) = (y = -x)"
-apply auto
-apply (drule sym [THEN diff_eq_eq [THEN iffD2]], auto)
-done
+lemma complex_add_eq_0_iff [iff]: "x + y = 0 \<longleftrightarrow> y = - x"
+ for x y :: complex
+ apply auto
+ apply (drule sym [THEN diff_eq_eq [THEN iffD2]])
+ apply auto
+ done
-subsection\<open>Limit of Complex to Complex Function\<close>
+subsection \<open>Limit of Complex to Complex Function\<close>
+
+lemma NSLIM_Re: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L \<Longrightarrow> (\<lambda>x. Re (f x)) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S Re L"
+ by (simp add: NSLIM_def starfunC_approx_Re_Im_iff hRe_hcomplex_of_complex)
-lemma NSLIM_Re: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L ==> (%x. Re(f x)) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S Re(L)"
-by (simp add: NSLIM_def starfunC_approx_Re_Im_iff
- hRe_hcomplex_of_complex)
+lemma NSLIM_Im: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L \<Longrightarrow> (\<lambda>x. Im (f x)) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S Im L"
+ by (simp add: NSLIM_def starfunC_approx_Re_Im_iff hIm_hcomplex_of_complex)
-lemma NSLIM_Im: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L ==> (%x. Im(f x)) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S Im(L)"
-by (simp add: NSLIM_def starfunC_approx_Re_Im_iff
- hIm_hcomplex_of_complex)
+lemma LIM_Re: "f \<midarrow>a\<rightarrow> L \<Longrightarrow> (\<lambda>x. Re (f x)) \<midarrow>a\<rightarrow> Re L"
+ for f :: "'a::real_normed_vector \<Rightarrow> complex"
+ by (simp add: LIM_NSLIM_iff NSLIM_Re)
-(** get this result easily now **)
-lemma LIM_Re:
- fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
- shows "f \<midarrow>a\<rightarrow> L ==> (%x. Re(f x)) \<midarrow>a\<rightarrow> Re(L)"
-by (simp add: LIM_NSLIM_iff NSLIM_Re)
+lemma LIM_Im: "f \<midarrow>a\<rightarrow> L \<Longrightarrow> (\<lambda>x. Im (f x)) \<midarrow>a\<rightarrow> Im L"
+ for f :: "'a::real_normed_vector \<Rightarrow> complex"
+ by (simp add: LIM_NSLIM_iff NSLIM_Im)
-lemma LIM_Im:
- fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
- shows "f \<midarrow>a\<rightarrow> L ==> (%x. Im(f x)) \<midarrow>a\<rightarrow> Im(L)"
-by (simp add: LIM_NSLIM_iff NSLIM_Im)
+lemma LIM_cnj: "f \<midarrow>a\<rightarrow> L \<Longrightarrow> (\<lambda>x. cnj (f x)) \<midarrow>a\<rightarrow> cnj L"
+ for f :: "'a::real_normed_vector \<Rightarrow> complex"
+ by (simp add: LIM_eq complex_cnj_diff [symmetric] del: complex_cnj_diff)
-lemma LIM_cnj:
- fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
- shows "f \<midarrow>a\<rightarrow> L ==> (%x. cnj (f x)) \<midarrow>a\<rightarrow> cnj L"
-by (simp add: LIM_eq complex_cnj_diff [symmetric] del: complex_cnj_diff)
-
-lemma LIM_cnj_iff:
- fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
- shows "((%x. cnj (f x)) \<midarrow>a\<rightarrow> cnj L) = (f \<midarrow>a\<rightarrow> L)"
-by (simp add: LIM_eq complex_cnj_diff [symmetric] del: complex_cnj_diff)
+lemma LIM_cnj_iff: "((\<lambda>x. cnj (f x)) \<midarrow>a\<rightarrow> cnj L) \<longleftrightarrow> f \<midarrow>a\<rightarrow> L"
+ for f :: "'a::real_normed_vector \<Rightarrow> complex"
+ by (simp add: LIM_eq complex_cnj_diff [symmetric] del: complex_cnj_diff)
lemma starfun_norm: "( *f* (\<lambda>x. norm (f x))) = (\<lambda>x. hnorm (( *f* f) x))"
-by transfer (rule refl)
+ by transfer (rule refl)
lemma star_of_Re [simp]: "star_of (Re x) = hRe (star_of x)"
-by transfer (rule refl)
+ by transfer (rule refl)
lemma star_of_Im [simp]: "star_of (Im x) = hIm (star_of x)"
-by transfer (rule refl)
+ by transfer (rule refl)
-text""
-(** another equivalence result **)
-lemma NSCLIM_NSCRLIM_iff:
- "(f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S L) = ((%y. cmod(f y - L)) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S 0)"
-by (simp add: NSLIM_def starfun_norm
- approx_approx_zero_iff [symmetric] approx_minus_iff [symmetric])
+text \<open>Another equivalence result.\<close>
+lemma NSCLIM_NSCRLIM_iff: "f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S L \<longleftrightarrow> (\<lambda>y. cmod (f y - L)) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S 0"
+ by (simp add: NSLIM_def starfun_norm
+ approx_approx_zero_iff [symmetric] approx_minus_iff [symmetric])
-(** much, much easier standard proof **)
-lemma CLIM_CRLIM_iff:
- fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
- shows "(f \<midarrow>x\<rightarrow> L) = ((%y. cmod(f y - L)) \<midarrow>x\<rightarrow> 0)"
-by (simp add: LIM_eq)
+text \<open>Much, much easier standard proof.\<close>
+lemma CLIM_CRLIM_iff: "f \<midarrow>x\<rightarrow> L \<longleftrightarrow> (\<lambda>y. cmod (f y - L)) \<midarrow>x\<rightarrow> 0"
+ for f :: "'a::real_normed_vector \<Rightarrow> complex"
+ by (simp add: LIM_eq)
-(* so this is nicer nonstandard proof *)
-lemma NSCLIM_NSCRLIM_iff2:
- "(f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S L) = ((%y. cmod(f y - L)) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S 0)"
-by (simp add: LIM_NSLIM_iff [symmetric] CLIM_CRLIM_iff)
+text \<open>So this is nicer nonstandard proof.\<close>
+lemma NSCLIM_NSCRLIM_iff2: "f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S L \<longleftrightarrow> (\<lambda>y. cmod (f y - L)) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S 0"
+ by (simp add: LIM_NSLIM_iff [symmetric] CLIM_CRLIM_iff)
lemma NSLIM_NSCRLIM_Re_Im_iff:
- "(f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L) = ((%x. Re(f x)) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S Re(L) &
- (%x. Im(f x)) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S Im(L))"
-apply (auto intro: NSLIM_Re NSLIM_Im)
-apply (auto simp add: NSLIM_def starfun_Re starfun_Im)
-apply (auto dest!: spec)
-apply (simp add: hcomplex_approx_iff)
-done
+ "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L \<longleftrightarrow> (\<lambda>x. Re (f x)) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S Re L \<and> (\<lambda>x. Im (f x)) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S Im L"
+ apply (auto intro: NSLIM_Re NSLIM_Im)
+ apply (auto simp add: NSLIM_def starfun_Re starfun_Im)
+ apply (auto dest!: spec)
+ apply (simp add: hcomplex_approx_iff)
+ done
+
+lemma LIM_CRLIM_Re_Im_iff: "f \<midarrow>a\<rightarrow> L \<longleftrightarrow> (\<lambda>x. Re (f x)) \<midarrow>a\<rightarrow> Re L \<and> (\<lambda>x. Im (f x)) \<midarrow>a\<rightarrow> Im L"
+ for f :: "'a::real_normed_vector \<Rightarrow> complex"
+ by (simp add: LIM_NSLIM_iff NSLIM_NSCRLIM_Re_Im_iff)
+
+
+subsection \<open>Continuity\<close>
+
+lemma NSLIM_isContc_iff: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S f a \<longleftrightarrow> (\<lambda>h. f (a + h)) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S f a"
+ by (rule NSLIM_h_iff)
-lemma LIM_CRLIM_Re_Im_iff:
- fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
- shows "(f \<midarrow>a\<rightarrow> L) = ((%x. Re(f x)) \<midarrow>a\<rightarrow> Re(L) &
- (%x. Im(f x)) \<midarrow>a\<rightarrow> Im(L))"
-by (simp add: LIM_NSLIM_iff NSLIM_NSCRLIM_Re_Im_iff)
+
+subsection \<open>Functions from Complex to Reals\<close>
+
+lemma isNSContCR_cmod [simp]: "isNSCont cmod a"
+ by (auto intro: approx_hnorm
+ simp: starfunCR_cmod hcmod_hcomplex_of_complex [symmetric] isNSCont_def)
+
+lemma isContCR_cmod [simp]: "isCont cmod a"
+ by (simp add: isNSCont_isCont_iff [symmetric])
+
+lemma isCont_Re: "isCont f a \<Longrightarrow> isCont (\<lambda>x. Re (f x)) a"
+ for f :: "'a::real_normed_vector \<Rightarrow> complex"
+ by (simp add: isCont_def LIM_Re)
+
+lemma isCont_Im: "isCont f a \<Longrightarrow> isCont (\<lambda>x. Im (f x)) a"
+ for f :: "'a::real_normed_vector \<Rightarrow> complex"
+ by (simp add: isCont_def LIM_Im)
-subsection\<open>Continuity\<close>
-
-lemma NSLIM_isContc_iff:
- "(f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S f a) = ((%h. f(a + h)) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S f a)"
-by (rule NSLIM_h_iff)
-
-subsection\<open>Functions from Complex to Reals\<close>
-
-lemma isNSContCR_cmod [simp]: "isNSCont cmod (a)"
-by (auto intro: approx_hnorm
- simp add: starfunCR_cmod hcmod_hcomplex_of_complex [symmetric]
- isNSCont_def)
+subsection \<open>Differentiation of Natural Number Powers\<close>
-lemma isContCR_cmod [simp]: "isCont cmod (a)"
-by (simp add: isNSCont_isCont_iff [symmetric])
-
-lemma isCont_Re:
- fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
- shows "isCont f a ==> isCont (%x. Re (f x)) a"
-by (simp add: isCont_def LIM_Re)
-
-lemma isCont_Im:
- fixes f :: "'a::real_normed_vector \<Rightarrow> complex"
- shows "isCont f a ==> isCont (%x. Im (f x)) a"
-by (simp add: isCont_def LIM_Im)
+lemma CDERIV_pow [simp]: "DERIV (\<lambda>x. x ^ n) x :> complex_of_real (real n) * (x ^ (n - Suc 0))"
+ apply (induct n)
+ apply (drule_tac [2] DERIV_ident [THEN DERIV_mult])
+ apply (auto simp add: distrib_right of_nat_Suc)
+ apply (case_tac "n")
+ apply (auto simp add: ac_simps)
+ done
-subsection\<open>Differentiation of Natural Number Powers\<close>
-
-lemma CDERIV_pow [simp]:
- "DERIV (%x. x ^ n) x :> (complex_of_real (real n)) * (x ^ (n - Suc 0))"
-apply (induct n)
-apply (drule_tac [2] DERIV_ident [THEN DERIV_mult])
-apply (auto simp add: distrib_right of_nat_Suc)
-apply (case_tac "n")
-apply (auto simp add: ac_simps)
-done
+text \<open>Nonstandard version.\<close>
+lemma NSCDERIV_pow: "NSDERIV (\<lambda>x. x ^ n) x :> complex_of_real (real n) * (x ^ (n - 1))"
+ by (metis CDERIV_pow NSDERIV_DERIV_iff One_nat_def)
-text\<open>Nonstandard version\<close>
-lemma NSCDERIV_pow: "NSDERIV (%x. x ^ n) x :> complex_of_real (real n) * (x ^ (n - 1))"
- by (metis CDERIV_pow NSDERIV_DERIV_iff One_nat_def)
+text \<open>Can't relax the premise @{term "x \<noteq> 0"}: it isn't continuous at zero.\<close>
+lemma NSCDERIV_inverse: "x \<noteq> 0 \<Longrightarrow> NSDERIV (\<lambda>x. inverse x) x :> - (inverse x)\<^sup>2"
+ for x :: complex
+ unfolding numeral_2_eq_2 by (rule NSDERIV_inverse)
-text\<open>Can't relax the premise @{term "x \<noteq> 0"}: it isn't continuous at zero\<close>
-lemma NSCDERIV_inverse:
- "(x::complex) \<noteq> 0 ==> NSDERIV (%x. inverse(x)) x :> (- ((inverse x)\<^sup>2))"
-unfolding numeral_2_eq_2
-by (rule NSDERIV_inverse)
-
-lemma CDERIV_inverse:
- "(x::complex) \<noteq> 0 ==> DERIV (%x. inverse(x)) x :> (- ((inverse x)\<^sup>2))"
-unfolding numeral_2_eq_2
-by (rule DERIV_inverse)
+lemma CDERIV_inverse: "x \<noteq> 0 \<Longrightarrow> DERIV (\<lambda>x. inverse x) x :> - (inverse x)\<^sup>2"
+ for x :: complex
+ unfolding numeral_2_eq_2 by (rule DERIV_inverse)
-subsection\<open>Derivative of Reciprocals (Function @{term inverse})\<close>
+subsection \<open>Derivative of Reciprocals (Function @{term inverse})\<close>
lemma CDERIV_inverse_fun:
- "[| DERIV f x :> d; f(x) \<noteq> (0::complex) |]
- ==> DERIV (%x. inverse(f x)) x :> (- (d * inverse ((f x)\<^sup>2)))"
-unfolding numeral_2_eq_2
-by (rule DERIV_inverse_fun)
+ "DERIV f x :> d \<Longrightarrow> f x \<noteq> 0 \<Longrightarrow> DERIV (\<lambda>x. inverse (f x)) x :> - (d * inverse ((f x)\<^sup>2))"
+ for x :: complex
+ unfolding numeral_2_eq_2 by (rule DERIV_inverse_fun)
lemma NSCDERIV_inverse_fun:
- "[| NSDERIV f x :> d; f(x) \<noteq> (0::complex) |]
- ==> NSDERIV (%x. inverse(f x)) x :> (- (d * inverse ((f x)\<^sup>2)))"
-unfolding numeral_2_eq_2
-by (rule NSDERIV_inverse_fun)
+ "NSDERIV f x :> d \<Longrightarrow> f x \<noteq> 0 \<Longrightarrow> NSDERIV (\<lambda>x. inverse (f x)) x :> - (d * inverse ((f x)\<^sup>2))"
+ for x :: complex
+ unfolding numeral_2_eq_2 by (rule NSDERIV_inverse_fun)
-subsection\<open>Derivative of Quotient\<close>
+subsection \<open>Derivative of Quotient\<close>
lemma CDERIV_quotient:
- "[| DERIV f x :> d; DERIV g x :> e; g(x) \<noteq> (0::complex) |]
- ==> DERIV (%y. f(y) / (g y)) x :> (d*g(x) - (e*f(x))) / ((g x)\<^sup>2)"
-unfolding numeral_2_eq_2
-by (rule DERIV_quotient)
+ "DERIV f x :> d \<Longrightarrow> DERIV g x :> e \<Longrightarrow> g(x) \<noteq> 0 \<Longrightarrow>
+ DERIV (\<lambda>y. f y / g y) x :> (d * g x - (e * f x)) / (g x)\<^sup>2"
+ for x :: complex
+ unfolding numeral_2_eq_2 by (rule DERIV_quotient)
lemma NSCDERIV_quotient:
- "[| NSDERIV f x :> d; NSDERIV g x :> e; g(x) \<noteq> (0::complex) |]
- ==> NSDERIV (%y. f(y) / (g y)) x :> (d*g(x) - (e*f(x))) / ((g x)\<^sup>2)"
-unfolding numeral_2_eq_2
-by (rule NSDERIV_quotient)
+ "NSDERIV f x :> d \<Longrightarrow> NSDERIV g x :> e \<Longrightarrow> g x \<noteq> (0::complex) \<Longrightarrow>
+ NSDERIV (\<lambda>y. f y / g y) x :> (d * g x - (e * f x)) / (g x)\<^sup>2"
+ unfolding numeral_2_eq_2 by (rule NSDERIV_quotient)
-subsection\<open>Caratheodory Formulation of Derivative at a Point: Standard Proof\<close>
+subsection \<open>Caratheodory Formulation of Derivative at a Point: Standard Proof\<close>
lemma CARAT_CDERIVD:
- "(\<forall>z. f z - f x = g z * (z - x)) & isNSCont g x & g x = l
- ==> NSDERIV f x :> l"
-by clarify (rule CARAT_DERIVD)
+ "(\<forall>z. f z - f x = g z * (z - x)) \<and> isNSCont g x \<and> g x = l \<Longrightarrow> NSDERIV f x :> l"
+ by clarify (rule CARAT_DERIVD)
end
--- a/src/HOL/Nonstandard_Analysis/CStar.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Nonstandard_Analysis/CStar.thy Tue Dec 20 16:18:56 2016 +0100
@@ -3,37 +3,36 @@
Copyright: 2001 University of Edinburgh
*)
-section\<open>Star-transforms in NSA, Extending Sets of Complex Numbers
- and Complex Functions\<close>
+section \<open>Star-transforms in NSA, Extending Sets of Complex Numbers and Complex Functions\<close>
theory CStar
-imports NSCA
+ imports NSCA
begin
-subsection\<open>Properties of the *-Transform Applied to Sets of Reals\<close>
+subsection \<open>Properties of the \<open>*\<close>-Transform Applied to Sets of Reals\<close>
-lemma STARC_hcomplex_of_complex_Int:
- "*s* X Int SComplex = hcomplex_of_complex ` X"
-by (auto simp add: Standard_def)
+lemma STARC_hcomplex_of_complex_Int: "*s* X \<inter> SComplex = hcomplex_of_complex ` X"
+ by (auto simp: Standard_def)
-lemma lemma_not_hcomplexA:
- "x \<notin> hcomplex_of_complex ` A ==> \<forall>y \<in> A. x \<noteq> hcomplex_of_complex y"
-by auto
+lemma lemma_not_hcomplexA: "x \<notin> hcomplex_of_complex ` A \<Longrightarrow> \<forall>y \<in> A. x \<noteq> hcomplex_of_complex y"
+ by auto
+
-subsection\<open>Theorems about Nonstandard Extensions of Functions\<close>
+subsection \<open>Theorems about Nonstandard Extensions of Functions\<close>
-lemma starfunC_hcpow: "!!Z. ( *f* (%z. z ^ n)) Z = Z pow hypnat_of_nat n"
-by transfer (rule refl)
+lemma starfunC_hcpow: "\<And>Z. ( *f* (\<lambda>z. z ^ n)) Z = Z pow hypnat_of_nat n"
+ by transfer (rule refl)
lemma starfunCR_cmod: "*f* cmod = hcmod"
-by transfer (rule refl)
+ by transfer (rule refl)
-subsection\<open>Internal Functions - Some Redundancy With *f* Now\<close>
+
+subsection \<open>Internal Functions - Some Redundancy With \<open>*f*\<close> Now\<close>
(** subtraction: ( *fn) - ( *gn) = *(fn - gn) **)
(*
lemma starfun_n_diff:
- "( *fn* f) z - ( *fn* g) z = ( *fn* (%i x. f i x - g i x)) z"
+ "( *fn* f) z - ( *fn* g) z = ( *fn* (\<lambda>i x. f i x - g i x)) z"
apply (cases z)
apply (simp add: starfun_n star_n_diff)
done
@@ -41,19 +40,17 @@
(** composition: ( *fn) o ( *gn) = *(fn o gn) **)
lemma starfun_Re: "( *f* (\<lambda>x. Re (f x))) = (\<lambda>x. hRe (( *f* f) x))"
-by transfer (rule refl)
+ by transfer (rule refl)
lemma starfun_Im: "( *f* (\<lambda>x. Im (f x))) = (\<lambda>x. hIm (( *f* f) x))"
-by transfer (rule refl)
+ by transfer (rule refl)
lemma starfunC_eq_Re_Im_iff:
- "(( *f* f) x = z) = ((( *f* (%x. Re(f x))) x = hRe (z)) &
- (( *f* (%x. Im(f x))) x = hIm (z)))"
-by (simp add: hcomplex_hRe_hIm_cancel_iff starfun_Re starfun_Im)
+ "( *f* f) x = z \<longleftrightarrow> ( *f* (\<lambda>x. Re (f x))) x = hRe z \<and> ( *f* (\<lambda>x. Im (f x))) x = hIm z"
+ by (simp add: hcomplex_hRe_hIm_cancel_iff starfun_Re starfun_Im)
lemma starfunC_approx_Re_Im_iff:
- "(( *f* f) x \<approx> z) = ((( *f* (%x. Re(f x))) x \<approx> hRe (z)) &
- (( *f* (%x. Im(f x))) x \<approx> hIm (z)))"
-by (simp add: hcomplex_approx_iff starfun_Re starfun_Im)
+ "( *f* f) x \<approx> z \<longleftrightarrow> ( *f* (\<lambda>x. Re (f x))) x \<approx> hRe z \<and> ( *f* (\<lambda>x. Im (f x))) x \<approx> hIm z"
+ by (simp add: hcomplex_approx_iff starfun_Re starfun_Im)
end
--- a/src/HOL/Nonstandard_Analysis/Examples/NSPrimes.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Nonstandard_Analysis/Examples/NSPrimes.thy Tue Dec 20 16:18:56 2016 +0100
@@ -4,280 +4,287 @@
Conversion to Isar and new proofs by Lawrence C Paulson, 2004
*)
-section\<open>The Nonstandard Primes as an Extension of the Prime Numbers\<close>
+section \<open>The Nonstandard Primes as an Extension of the Prime Numbers\<close>
theory NSPrimes
-imports "~~/src/HOL/Number_Theory/Primes" "../Hyperreal"
+ imports "~~/src/HOL/Number_Theory/Primes" "../Hyperreal"
begin
-text\<open>These can be used to derive an alternative proof of the infinitude of
+text \<open>These can be used to derive an alternative proof of the infinitude of
primes by considering a property of nonstandard sets.\<close>
-definition
- starprime :: "hypnat set" where
- [transfer_unfold]: "starprime = ( *s* {p. prime p})"
+definition starprime :: "hypnat set"
+ where [transfer_unfold]: "starprime = *s* {p. prime p}"
-definition
- choicefun :: "'a set => 'a" where
- "choicefun E = (@x. \<exists>X \<in> Pow(E) -{{}}. x : X)"
+definition choicefun :: "'a set \<Rightarrow> 'a"
+ where "choicefun E = (SOME x. \<exists>X \<in> Pow E - {{}}. x \<in> X)"
-primrec injf_max :: "nat => ('a::{order} set) => 'a"
+primrec injf_max :: "nat \<Rightarrow> 'a::order set \<Rightarrow> 'a"
where
injf_max_zero: "injf_max 0 E = choicefun E"
-| injf_max_Suc: "injf_max (Suc n) E = choicefun({e. e:E & injf_max n E < e})"
-
+| injf_max_Suc: "injf_max (Suc n) E = choicefun ({e. e \<in> E \<and> injf_max n E < e})"
-lemma dvd_by_all2:
- fixes M :: nat
- shows "\<exists>N>0. \<forall>m. 0 < m \<and> m \<le> M \<longrightarrow> m dvd N"
-apply (induct M)
-apply auto
-apply (rule_tac x = "N * (Suc M) " in exI)
-apply auto
-apply (metis dvdI dvd_add_times_triv_left_iff dvd_add_triv_right_iff dvd_refl dvd_trans le_Suc_eq mult_Suc_right)
-done
+lemma dvd_by_all2: "\<exists>N>0. \<forall>m. 0 < m \<and> m \<le> M \<longrightarrow> m dvd N"
+ for M :: nat
+ apply (induct M)
+ apply auto
+ apply (rule_tac x = "N * Suc M" in exI)
+ apply auto
+ apply (metis dvdI dvd_add_times_triv_left_iff dvd_add_triv_right_iff dvd_refl dvd_trans le_Suc_eq mult_Suc_right)
+ done
-lemma dvd_by_all:
- "\<forall>M::nat. \<exists>N>0. \<forall>m. 0 < m \<and> m \<le> M \<longrightarrow> m dvd N"
+lemma dvd_by_all: "\<forall>M::nat. \<exists>N>0. \<forall>m. 0 < m \<and> m \<le> M \<longrightarrow> m dvd N"
using dvd_by_all2 by blast
-lemma hypnat_of_nat_le_zero_iff [simp]: "(hypnat_of_nat n <= 0) = (n = 0)"
-by (transfer, simp)
+lemma hypnat_of_nat_le_zero_iff [simp]: "hypnat_of_nat n \<le> 0 \<longleftrightarrow> n = 0"
+ by transfer simp
-(* Goldblatt: Exercise 5.11(2) - p. 57 *)
-lemma hdvd_by_all: "\<forall>M. \<exists>N. 0 < N & (\<forall>m. 0 < m & (m::hypnat) <= M --> m dvd N)"
-by (transfer, rule dvd_by_all)
+text \<open>Goldblatt: Exercise 5.11(2) -- p. 57.\<close>
+lemma hdvd_by_all: "\<forall>M. \<exists>N. 0 < N \<and> (\<forall>m::hypnat. 0 < m \<and> m \<le> M \<longrightarrow> m dvd N)"
+ by transfer (rule dvd_by_all)
lemmas hdvd_by_all2 = hdvd_by_all [THEN spec]
-(* Goldblatt: Exercise 5.11(2) - p. 57 *)
+text \<open>Goldblatt: Exercise 5.11(2) -- p. 57.\<close>
lemma hypnat_dvd_all_hypnat_of_nat:
- "\<exists>(N::hypnat). 0 < N & (\<forall>n \<in> -{0::nat}. hypnat_of_nat(n) dvd N)"
-apply (cut_tac hdvd_by_all)
-apply (drule_tac x = whn in spec, auto)
-apply (rule exI, auto)
-apply (drule_tac x = "hypnat_of_nat n" in spec)
-apply (auto simp add: linorder_not_less)
-done
+ "\<exists>N::hypnat. 0 < N \<and> (\<forall>n \<in> - {0::nat}. hypnat_of_nat n dvd N)"
+ apply (cut_tac hdvd_by_all)
+ apply (drule_tac x = whn in spec)
+ apply auto
+ apply (rule exI)
+ apply auto
+ apply (drule_tac x = "hypnat_of_nat n" in spec)
+ apply (auto simp add: linorder_not_less)
+ done
-text\<open>The nonstandard extension of the set prime numbers consists of precisely
-those hypernaturals exceeding 1 that have no nontrivial factors\<close>
+text \<open>The nonstandard extension of the set prime numbers consists of precisely
+ those hypernaturals exceeding 1 that have no nontrivial factors.\<close>
-(* Goldblatt: Exercise 5.11(3a) - p 57 *)
-lemma starprime:
- "starprime = {p. 1 < p & (\<forall>m. m dvd p --> m = 1 | m = p)}"
-by (transfer, auto simp add: prime_nat_iff)
+text \<open>Goldblatt: Exercise 5.11(3a) -- p 57.\<close>
+lemma starprime: "starprime = {p. 1 < p \<and> (\<forall>m. m dvd p \<longrightarrow> m = 1 \<or> m = p)}"
+ by transfer (auto simp add: prime_nat_iff)
-(* Goldblatt Exercise 5.11(3b) - p 57 *)
-lemma hyperprime_factor_exists [rule_format]:
- "!!n. 1 < n ==> (\<exists>k \<in> starprime. k dvd n)"
-by (transfer, simp add: prime_factor_nat)
+text \<open>Goldblatt Exercise 5.11(3b) -- p 57.\<close>
+lemma hyperprime_factor_exists: "\<And>n. 1 < n \<Longrightarrow> \<exists>k \<in> starprime. k dvd n"
+ by transfer (simp add: prime_factor_nat)
-(* Goldblatt Exercise 3.10(1) - p. 29 *)
-lemma NatStar_hypnat_of_nat: "finite A ==> *s* A = hypnat_of_nat ` A"
-by (rule starset_finite)
+text \<open>Goldblatt Exercise 3.10(1) -- p. 29.\<close>
+lemma NatStar_hypnat_of_nat: "finite A \<Longrightarrow> *s* A = hypnat_of_nat ` A"
+ by (rule starset_finite)
-subsection\<open>Another characterization of infinite set of natural numbers\<close>
+subsection \<open>Another characterization of infinite set of natural numbers\<close>
-lemma finite_nat_set_bounded: "finite N ==> \<exists>n. (\<forall>i \<in> N. i<(n::nat))"
-apply (erule_tac F = N in finite_induct, auto)
-apply (rule_tac x = "Suc n + x" in exI, auto)
-done
+lemma finite_nat_set_bounded: "finite N \<Longrightarrow> \<exists>n::nat. \<forall>i \<in> N. i < n"
+ apply (erule_tac F = N in finite_induct)
+ apply auto
+ apply (rule_tac x = "Suc n + x" in exI)
+ apply auto
+ done
-lemma finite_nat_set_bounded_iff: "finite N = (\<exists>n. (\<forall>i \<in> N. i<(n::nat)))"
-by (blast intro: finite_nat_set_bounded bounded_nat_set_is_finite)
+lemma finite_nat_set_bounded_iff: "finite N \<longleftrightarrow> (\<exists>n::nat. \<forall>i \<in> N. i < n)"
+ by (blast intro: finite_nat_set_bounded bounded_nat_set_is_finite)
-lemma not_finite_nat_set_iff: "(~ finite N) = (\<forall>n. \<exists>i \<in> N. n <= (i::nat))"
-by (auto simp add: finite_nat_set_bounded_iff not_less)
+lemma not_finite_nat_set_iff: "\<not> finite N \<longleftrightarrow> (\<forall>n::nat. \<exists>i \<in> N. n \<le> i)"
+ by (auto simp add: finite_nat_set_bounded_iff not_less)
-lemma bounded_nat_set_is_finite2: "(\<forall>i \<in> N. i<=(n::nat)) ==> finite N"
-apply (rule finite_subset)
- apply (rule_tac [2] finite_atMost, auto)
-done
+lemma bounded_nat_set_is_finite2: "\<forall>i::nat \<in> N. i \<le> n \<Longrightarrow> finite N"
+ apply (rule finite_subset)
+ apply (rule_tac [2] finite_atMost)
+ apply auto
+ done
-lemma finite_nat_set_bounded2: "finite N ==> \<exists>n. (\<forall>i \<in> N. i<=(n::nat))"
-apply (erule_tac F = N in finite_induct, auto)
-apply (rule_tac x = "n + x" in exI, auto)
-done
+lemma finite_nat_set_bounded2: "finite N \<Longrightarrow> \<exists>n::nat. \<forall>i \<in> N. i \<le> n"
+ apply (erule_tac F = N in finite_induct)
+ apply auto
+ apply (rule_tac x = "n + x" in exI)
+ apply auto
+ done
-lemma finite_nat_set_bounded_iff2: "finite N = (\<exists>n. (\<forall>i \<in> N. i<=(n::nat)))"
-by (blast intro: finite_nat_set_bounded2 bounded_nat_set_is_finite2)
+lemma finite_nat_set_bounded_iff2: "finite N \<longleftrightarrow> (\<exists>n::nat. \<forall>i \<in> N. i \<le> n)"
+ by (blast intro: finite_nat_set_bounded2 bounded_nat_set_is_finite2)
-lemma not_finite_nat_set_iff2: "(~ finite N) = (\<forall>n. \<exists>i \<in> N. n < (i::nat))"
-by (auto simp add: finite_nat_set_bounded_iff2 not_le)
+lemma not_finite_nat_set_iff2: "\<not> finite N \<longleftrightarrow> (\<forall>n::nat. \<exists>i \<in> N. n < i)"
+ by (auto simp add: finite_nat_set_bounded_iff2 not_le)
-subsection\<open>An injective function cannot define an embedded natural number\<close>
+subsection \<open>An injective function cannot define an embedded natural number\<close>
-lemma lemma_infinite_set_singleton: "\<forall>m n. m \<noteq> n --> f n \<noteq> f m
- ==> {n. f n = N} = {} | (\<exists>m. {n. f n = N} = {m})"
-apply auto
-apply (drule_tac x = x in spec, auto)
-apply (subgoal_tac "\<forall>n. (f n = f x) = (x = n) ")
-apply auto
-done
+lemma lemma_infinite_set_singleton:
+ "\<forall>m n. m \<noteq> n \<longrightarrow> f n \<noteq> f m \<Longrightarrow> {n. f n = N} = {} \<or> (\<exists>m. {n. f n = N} = {m})"
+ apply auto
+ apply (drule_tac x = x in spec, auto)
+ apply (subgoal_tac "\<forall>n. f n = f x \<longleftrightarrow> x = n")
+ apply auto
+ done
lemma inj_fun_not_hypnat_in_SHNat:
- assumes inj_f: "inj (f::nat=>nat)"
+ fixes f :: "nat \<Rightarrow> nat"
+ assumes inj_f: "inj f"
shows "starfun f whn \<notin> Nats"
proof
from inj_f have inj_f': "inj (starfun f)"
by (transfer inj_on_def Ball_def UNIV_def)
assume "starfun f whn \<in> Nats"
then obtain N where N: "starfun f whn = hypnat_of_nat N"
- by (auto simp add: Nats_def)
- hence "\<exists>n. starfun f n = hypnat_of_nat N" ..
- hence "\<exists>n. f n = N" by transfer
- then obtain n where n: "f n = N" ..
- hence "starfun f (hypnat_of_nat n) = hypnat_of_nat N"
+ by (auto simp: Nats_def)
+ then have "\<exists>n. starfun f n = hypnat_of_nat N" ..
+ then have "\<exists>n. f n = N" by transfer
+ then obtain n where "f n = N" ..
+ then have "starfun f (hypnat_of_nat n) = hypnat_of_nat N"
by transfer
with N have "starfun f whn = starfun f (hypnat_of_nat n)"
by simp
with inj_f' have "whn = hypnat_of_nat n"
by (rule injD)
- thus "False"
+ then show False
by (simp add: whn_neq_hypnat_of_nat)
qed
-lemma range_subset_mem_starsetNat:
- "range f <= A ==> starfun f whn \<in> *s* A"
-apply (rule_tac x="whn" in spec)
-apply (transfer, auto)
-done
+lemma range_subset_mem_starsetNat: "range f \<subseteq> A \<Longrightarrow> starfun f whn \<in> *s* A"
+ apply (rule_tac x="whn" in spec)
+ apply transfer
+ apply auto
+ done
+
+text \<open>
+ Gleason Proposition 11-5.5. pg 149, pg 155 (ex. 3) and pg. 360.
-(*--------------------------------------------------------------------------------*)
-(* Gleason Proposition 11-5.5. pg 149, pg 155 (ex. 3) and pg. 360 *)
-(* Let E be a nonvoid ordered set with no maximal elements (note: effectively an *)
-(* infinite set if we take E = N (Nats)). Then there exists an order-preserving *)
-(* injection from N to E. Of course, (as some doofus will undoubtedly point out! *)
-(* :-)) can use notion of least element in proof (i.e. no need for choice) if *)
-(* dealing with nats as we have well-ordering property *)
-(*--------------------------------------------------------------------------------*)
+ Let \<open>E\<close> be a nonvoid ordered set with no maximal elements (note: effectively an
+ infinite set if we take \<open>E = N\<close> (Nats)). Then there exists an order-preserving
+ injection from \<open>N\<close> to \<open>E\<close>. Of course, (as some doofus will undoubtedly point out!
+ :-)) can use notion of least element in proof (i.e. no need for choice) if
+ dealing with nats as we have well-ordering property.
+\<close>
-lemma lemmaPow3: "E \<noteq> {} ==> \<exists>x. \<exists>X \<in> (Pow E - {{}}). x: X"
-by auto
+lemma lemmaPow3: "E \<noteq> {} \<Longrightarrow> \<exists>x. \<exists>X \<in> Pow E - {{}}. x \<in> X"
+ by auto
-lemma choicefun_mem_set [simp]: "E \<noteq> {} ==> choicefun E \<in> E"
-apply (unfold choicefun_def)
-apply (rule lemmaPow3 [THEN someI2_ex], auto)
-done
+lemma choicefun_mem_set [simp]: "E \<noteq> {} \<Longrightarrow> choicefun E \<in> E"
+ apply (unfold choicefun_def)
+ apply (rule lemmaPow3 [THEN someI2_ex], auto)
+ done
-lemma injf_max_mem_set: "[| E \<noteq>{}; \<forall>x. \<exists>y \<in> E. x < y |] ==> injf_max n E \<in> E"
-apply (induct_tac "n", force)
-apply (simp (no_asm) add: choicefun_def)
-apply (rule lemmaPow3 [THEN someI2_ex], auto)
-done
+lemma injf_max_mem_set: "E \<noteq>{} \<Longrightarrow> \<forall>x. \<exists>y \<in> E. x < y \<Longrightarrow> injf_max n E \<in> E"
+ apply (induct n)
+ apply force
+ apply (simp add: choicefun_def)
+ apply (rule lemmaPow3 [THEN someI2_ex], auto)
+ done
-lemma injf_max_order_preserving: "\<forall>x. \<exists>y \<in> E. x < y ==> injf_max n E < injf_max (Suc n) E"
-apply (simp (no_asm) add: choicefun_def)
-apply (rule lemmaPow3 [THEN someI2_ex], auto)
-done
+lemma injf_max_order_preserving: "\<forall>x. \<exists>y \<in> E. x < y \<Longrightarrow> injf_max n E < injf_max (Suc n) E"
+ apply (simp add: choicefun_def)
+ apply (rule lemmaPow3 [THEN someI2_ex])
+ apply auto
+ done
-lemma injf_max_order_preserving2: "\<forall>x. \<exists>y \<in> E. x < y
- ==> \<forall>n m. m < n --> injf_max m E < injf_max n E"
-apply (rule allI)
-apply (induct_tac "n", auto)
-apply (simp (no_asm) add: choicefun_def)
-apply (rule lemmaPow3 [THEN someI2_ex])
-apply (auto simp add: less_Suc_eq)
-apply (drule_tac x = m in spec)
-apply (drule subsetD, auto)
-apply (drule_tac x = "injf_max m E" in order_less_trans, auto)
-done
+lemma injf_max_order_preserving2: "\<forall>x. \<exists>y \<in> E. x < y \<Longrightarrow> \<forall>n m. m < n \<longrightarrow> injf_max m E < injf_max n E"
+ apply (rule allI)
+ apply (induct_tac n)
+ apply auto
+ apply (simp add: choicefun_def)
+ apply (rule lemmaPow3 [THEN someI2_ex])
+ apply (auto simp add: less_Suc_eq)
+ apply (drule_tac x = m in spec)
+ apply (drule subsetD)
+ apply auto
+ apply (drule_tac x = "injf_max m E" in order_less_trans)
+ apply auto
+ done
-lemma inj_injf_max: "\<forall>x. \<exists>y \<in> E. x < y ==> inj (%n. injf_max n E)"
-apply (rule inj_onI)
-apply (rule ccontr, auto)
-apply (drule injf_max_order_preserving2)
-apply (metis linorder_antisym_conv3 order_less_le)
-done
+lemma inj_injf_max: "\<forall>x. \<exists>y \<in> E. x < y \<Longrightarrow> inj (\<lambda>n. injf_max n E)"
+ apply (rule inj_onI)
+ apply (rule ccontr)
+ apply auto
+ apply (drule injf_max_order_preserving2)
+ apply (metis linorder_antisym_conv3 order_less_le)
+ done
lemma infinite_set_has_order_preserving_inj:
- "[| (E::('a::{order} set)) \<noteq> {}; \<forall>x. \<exists>y \<in> E. x < y |]
- ==> \<exists>f. range f <= E & inj (f::nat => 'a) & (\<forall>m. f m < f(Suc m))"
-apply (rule_tac x = "%n. injf_max n E" in exI, safe)
-apply (rule injf_max_mem_set)
-apply (rule_tac [3] inj_injf_max)
-apply (rule_tac [4] injf_max_order_preserving, auto)
-done
+ "E \<noteq> {} \<Longrightarrow> \<forall>x. \<exists>y \<in> E. x < y \<Longrightarrow> \<exists>f. range f \<subseteq> E \<and> inj f \<and> (\<forall>m. f m < f (Suc m))"
+ for E :: "'a::order set" and f :: "nat \<Rightarrow> 'a"
+ apply (rule_tac x = "\<lambda>n. injf_max n E" in exI)
+ apply safe
+ apply (rule injf_max_mem_set)
+ apply (rule_tac [3] inj_injf_max)
+ apply (rule_tac [4] injf_max_order_preserving)
+ apply auto
+ done
-text\<open>Only need the existence of an injective function from N to A for proof\<close>
-lemma hypnat_infinite_has_nonstandard:
- "~ finite A ==> hypnat_of_nat ` A < ( *s* A)"
-apply auto
-apply (subgoal_tac "A \<noteq> {}")
-prefer 2 apply force
-apply (drule infinite_set_has_order_preserving_inj)
-apply (erule not_finite_nat_set_iff2 [THEN iffD1], auto)
-apply (drule inj_fun_not_hypnat_in_SHNat)
-apply (drule range_subset_mem_starsetNat)
-apply (auto simp add: SHNat_eq)
-done
+text \<open>Only need the existence of an injective function from \<open>N\<close> to \<open>A\<close> for proof.\<close>
-lemma starsetNat_eq_hypnat_of_nat_image_finite: "*s* A = hypnat_of_nat ` A ==> finite A"
-by (metis hypnat_infinite_has_nonstandard less_irrefl)
+lemma hypnat_infinite_has_nonstandard: "\<not> finite A \<Longrightarrow> hypnat_of_nat ` A < ( *s* A)"
+ apply auto
+ apply (subgoal_tac "A \<noteq> {}")
+ prefer 2 apply force
+ apply (drule infinite_set_has_order_preserving_inj)
+ apply (erule not_finite_nat_set_iff2 [THEN iffD1])
+ apply auto
+ apply (drule inj_fun_not_hypnat_in_SHNat)
+ apply (drule range_subset_mem_starsetNat)
+ apply (auto simp add: SHNat_eq)
+ done
-lemma finite_starsetNat_iff: "( *s* A = hypnat_of_nat ` A) = (finite A)"
-by (blast intro!: starsetNat_eq_hypnat_of_nat_image_finite NatStar_hypnat_of_nat)
+lemma starsetNat_eq_hypnat_of_nat_image_finite: "*s* A = hypnat_of_nat ` A \<Longrightarrow> finite A"
+ by (metis hypnat_infinite_has_nonstandard less_irrefl)
+
+lemma finite_starsetNat_iff: "*s* A = hypnat_of_nat ` A \<longleftrightarrow> finite A"
+ by (blast intro!: starsetNat_eq_hypnat_of_nat_image_finite NatStar_hypnat_of_nat)
-lemma hypnat_infinite_has_nonstandard_iff: "(~ finite A) = (hypnat_of_nat ` A < *s* A)"
-apply (rule iffI)
-apply (blast intro!: hypnat_infinite_has_nonstandard)
-apply (auto simp add: finite_starsetNat_iff [symmetric])
-done
+lemma hypnat_infinite_has_nonstandard_iff: "\<not> finite A \<longleftrightarrow> hypnat_of_nat ` A < *s* A"
+ apply (rule iffI)
+ apply (blast intro!: hypnat_infinite_has_nonstandard)
+ apply (auto simp add: finite_starsetNat_iff [symmetric])
+ done
-subsection\<open>Existence of Infinitely Many Primes: a Nonstandard Proof\<close>
-lemma lemma_not_dvd_hypnat_one [simp]: "~ (\<forall>n \<in> - {0}. hypnat_of_nat n dvd 1)"
-apply auto
-apply (rule_tac x = 2 in bexI)
-apply (transfer, auto)
-done
+subsection \<open>Existence of Infinitely Many Primes: a Nonstandard Proof\<close>
-lemma lemma_not_dvd_hypnat_one2 [simp]: "\<exists>n \<in> - {0}. ~ hypnat_of_nat n dvd 1"
-apply (cut_tac lemma_not_dvd_hypnat_one)
-apply (auto simp del: lemma_not_dvd_hypnat_one)
-done
+lemma lemma_not_dvd_hypnat_one [simp]: "\<not> (\<forall>n \<in> - {0}. hypnat_of_nat n dvd 1)"
+ apply auto
+ apply (rule_tac x = 2 in bexI)
+ apply transfer
+ apply auto
+ done
-lemma hypnat_add_one_gt_one:
- "!!N. 0 < N ==> 1 < (N::hypnat) + 1"
-by (transfer, simp)
+lemma lemma_not_dvd_hypnat_one2 [simp]: "\<exists>n \<in> - {0}. \<not> hypnat_of_nat n dvd 1"
+ using lemma_not_dvd_hypnat_one by (auto simp del: lemma_not_dvd_hypnat_one)
+
+lemma hypnat_add_one_gt_one: "\<And>N::hypnat. 0 < N \<Longrightarrow> 1 < N + 1"
+ by transfer simp
lemma hypnat_of_nat_zero_not_prime [simp]: "hypnat_of_nat 0 \<notin> starprime"
-by (transfer, simp)
+ by transfer simp
-lemma hypnat_zero_not_prime [simp]:
- "0 \<notin> starprime"
-by (cut_tac hypnat_of_nat_zero_not_prime, simp)
+lemma hypnat_zero_not_prime [simp]: "0 \<notin> starprime"
+ using hypnat_of_nat_zero_not_prime by simp
lemma hypnat_of_nat_one_not_prime [simp]: "hypnat_of_nat 1 \<notin> starprime"
-by (transfer, simp)
+ by transfer simp
lemma hypnat_one_not_prime [simp]: "1 \<notin> starprime"
-by (cut_tac hypnat_of_nat_one_not_prime, simp)
+ using hypnat_of_nat_one_not_prime by simp
-lemma hdvd_diff: "!!k m n :: hypnat. [| k dvd m; k dvd n |] ==> k dvd (m - n)"
-by (transfer, rule dvd_diff_nat)
+lemma hdvd_diff: "\<And>k m n :: hypnat. k dvd m \<Longrightarrow> k dvd n \<Longrightarrow> k dvd (m - n)"
+ by transfer (rule dvd_diff_nat)
-lemma hdvd_one_eq_one:
- "\<And>x::hypnat. is_unit x \<Longrightarrow> x = 1"
+lemma hdvd_one_eq_one: "\<And>x::hypnat. is_unit x \<Longrightarrow> x = 1"
by transfer simp
-text\<open>Already proved as \<open>primes_infinite\<close>, but now using non-standard naturals.\<close>
-theorem not_finite_prime: "~ finite {p::nat. prime p}"
-apply (rule hypnat_infinite_has_nonstandard_iff [THEN iffD2])
-using hypnat_dvd_all_hypnat_of_nat
-apply clarify
-apply (drule hypnat_add_one_gt_one)
-apply (drule hyperprime_factor_exists)
-apply clarify
-apply (subgoal_tac "k \<notin> hypnat_of_nat ` {p. prime p}")
-apply (force simp add: starprime_def)
- apply (metis Compl_iff add.commute dvd_add_left_iff empty_iff hdvd_one_eq_one hypnat_one_not_prime
- imageE insert_iff mem_Collect_eq not_prime_0)
-done
+text \<open>Already proved as \<open>primes_infinite\<close>, but now using non-standard naturals.\<close>
+theorem not_finite_prime: "\<not> finite {p::nat. prime p}"
+ apply (rule hypnat_infinite_has_nonstandard_iff [THEN iffD2])
+ using hypnat_dvd_all_hypnat_of_nat
+ apply clarify
+ apply (drule hypnat_add_one_gt_one)
+ apply (drule hyperprime_factor_exists)
+ apply clarify
+ apply (subgoal_tac "k \<notin> hypnat_of_nat ` {p. prime p}")
+ apply (force simp: starprime_def)
+ apply (metis Compl_iff add.commute dvd_add_left_iff empty_iff hdvd_one_eq_one hypnat_one_not_prime
+ imageE insert_iff mem_Collect_eq not_prime_0)
+ done
end
--- a/src/HOL/Nonstandard_Analysis/HDeriv.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Nonstandard_Analysis/HDeriv.thy Tue Dec 20 16:18:56 2016 +0100
@@ -53,7 +53,7 @@
apply (drule (1) bspec)+
apply (drule (1) approx_trans3)
apply simp
- apply (simp add: Infinitesimal_of_hypreal Infinitesimal_epsilon)
+ apply (simp add: Infinitesimal_of_hypreal)
apply (simp add: of_hypreal_eq_0_iff hypreal_epsilon_not_zero)
done
@@ -75,7 +75,7 @@
text \<open>While we're at it!\<close>
lemma NSDERIV_iff2:
"(NSDERIV f x :> D) \<longleftrightarrow>
- (\<forall>w. w \<noteq> star_of x & w \<approx> star_of x \<longrightarrow> ( *f* (%z. (f z - f x) / (z-x))) w \<approx> star_of D)"
+ (\<forall>w. w \<noteq> star_of x \<and> w \<approx> star_of x \<longrightarrow> ( *f* (\<lambda>z. (f z - f x) / (z - x))) w \<approx> star_of D)"
by (simp add: NSDERIV_NSLIM_iff2 NSLIM_def)
(* FIXME delete *)
@@ -91,7 +91,7 @@
apply (drule_tac x = u in spec, auto)
apply (drule_tac c = "u - hypreal_of_real x" and b = "hypreal_of_real D" in approx_mult1)
apply (drule_tac [!] hypreal_not_eq_minus_iff [THEN iffD1])
- apply (subgoal_tac [2] "( *f* (%z. z-x)) u \<noteq> (0::hypreal) ")
+ apply (subgoal_tac [2] "( *f* (\<lambda>z. z - x)) u \<noteq> (0::hypreal) ")
apply (auto simp: approx_minus_iff [THEN iffD1, THEN mem_infmal_iff [THEN iffD2]]
Infinitesimal_subset_HFinite [THEN subsetD])
done
@@ -290,7 +290,8 @@
apply (case_tac "( *f* g) (star_of (x) + xa) = star_of (g x) ")
apply (drule_tac g = g in NSDERIV_zero)
apply (auto simp add: divide_inverse)
- apply (rule_tac z1 = "( *f* g) (star_of (x) + xa) - star_of (g x) " and y1 = "inverse xa" in lemma_chain [THEN ssubst])
+ apply (rule_tac z1 = "( *f* g) (star_of (x) + xa) - star_of (g x) " and y1 = "inverse xa"
+ in lemma_chain [THEN ssubst])
apply (erule hypreal_not_eq_minus_iff [THEN iffD1])
apply (rule approx_mult_star_of)
apply (simp_all add: divide_inverse [symmetric])
--- a/src/HOL/Nonstandard_Analysis/HLim.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Nonstandard_Analysis/HLim.thy Tue Dec 20 16:18:56 2016 +0100
@@ -77,7 +77,7 @@
qed
lemma NSLIM_zero_cancel: "(\<lambda>x. f x - l) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S 0 \<Longrightarrow> f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S l"
- apply (drule_tac g = "%x. l" and m = l in NSLIM_add)
+ apply (drule_tac g = "\<lambda>x. l" and m = l in NSLIM_add)
apply (auto simp add: add.assoc)
done
@@ -205,8 +205,8 @@
lemma isCont_isNSCont: "isCont f a \<Longrightarrow> isNSCont f a"
by (erule isNSCont_isCont_iff [THEN iffD2])
-text \<open>NS continuity \<open>==>\<close> Standard continuity.\<close>
-lemma isNSCont_isCont: "isNSCont f a ==> isCont f a"
+text \<open>NS continuity \<open>\<Longrightarrow>\<close> Standard continuity.\<close>
+lemma isNSCont_isCont: "isNSCont f a \<Longrightarrow> isCont f a"
by (erule isNSCont_isCont_iff [THEN iffD1])
--- a/src/HOL/Nonstandard_Analysis/HLog.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Nonstandard_Analysis/HLog.thy Tue Dec 20 16:18:56 2016 +0100
@@ -3,154 +3,134 @@
Copyright: 2000, 2001 University of Edinburgh
*)
-section\<open>Logarithms: Non-Standard Version\<close>
+section \<open>Logarithms: Non-Standard Version\<close>
theory HLog
-imports HTranscendental
+ imports HTranscendental
begin
(* should be in NSA.ML *)
lemma epsilon_ge_zero [simp]: "0 \<le> \<epsilon>"
-by (simp add: epsilon_def star_n_zero_num star_n_le)
+ by (simp add: epsilon_def star_n_zero_num star_n_le)
-lemma hpfinite_witness: "\<epsilon> : {x. 0 \<le> x & x : HFinite}"
-by auto
+lemma hpfinite_witness: "\<epsilon> \<in> {x. 0 \<le> x \<and> x \<in> HFinite}"
+ by auto
-definition
- powhr :: "[hypreal,hypreal] => hypreal" (infixr "powhr" 80) where
- [transfer_unfold]: "x powhr a = starfun2 (op powr) x a"
-
-definition
- hlog :: "[hypreal,hypreal] => hypreal" where
- [transfer_unfold]: "hlog a x = starfun2 log a x"
+definition powhr :: "hypreal \<Rightarrow> hypreal \<Rightarrow> hypreal" (infixr "powhr" 80)
+ where [transfer_unfold]: "x powhr a = starfun2 (op powr) x a"
+
+definition hlog :: "hypreal \<Rightarrow> hypreal \<Rightarrow> hypreal"
+ where [transfer_unfold]: "hlog a x = starfun2 log a x"
+
+lemma powhr: "(star_n X) powhr (star_n Y) = star_n (\<lambda>n. (X n) powr (Y n))"
+ by (simp add: powhr_def starfun2_star_n)
-lemma powhr: "(star_n X) powhr (star_n Y) = star_n (%n. (X n) powr (Y n))"
-by (simp add: powhr_def starfun2_star_n)
-
-lemma powhr_one_eq_one [simp]: "!!a. 1 powhr a = 1"
-by (transfer, simp)
+lemma powhr_one_eq_one [simp]: "\<And>a. 1 powhr a = 1"
+ by transfer simp
-lemma powhr_mult:
- "!!a x y. [| 0 < x; 0 < y |] ==> (x * y) powhr a = (x powhr a) * (y powhr a)"
-by (transfer, simp add: powr_mult)
+lemma powhr_mult: "\<And>a x y. 0 < x \<Longrightarrow> 0 < y \<Longrightarrow> (x * y) powhr a = (x powhr a) * (y powhr a)"
+ by transfer (simp add: powr_mult)
-lemma powhr_gt_zero [simp]: "!!a x. 0 < x powhr a \<longleftrightarrow> x \<noteq> 0"
-by (transfer, simp)
+lemma powhr_gt_zero [simp]: "\<And>a x. 0 < x powhr a \<longleftrightarrow> x \<noteq> 0"
+ by transfer simp
lemma powhr_not_zero [simp]: "\<And>a x. x powhr a \<noteq> 0 \<longleftrightarrow> x \<noteq> 0"
-by transfer simp
-
-lemma powhr_divide:
- "!!a x y. [| 0 < x; 0 < y |] ==> (x / y) powhr a = (x powhr a)/(y powhr a)"
-by (transfer, rule powr_divide)
+ by transfer simp
-lemma powhr_add: "!!a b x. x powhr (a + b) = (x powhr a) * (x powhr b)"
-by (transfer, rule powr_add)
+lemma powhr_divide: "\<And>a x y. 0 < x \<Longrightarrow> 0 < y \<Longrightarrow> (x / y) powhr a = (x powhr a) / (y powhr a)"
+ by transfer (rule powr_divide)
-lemma powhr_powhr: "!!a b x. (x powhr a) powhr b = x powhr (a * b)"
-by (transfer, rule powr_powr)
+lemma powhr_add: "\<And>a b x. x powhr (a + b) = (x powhr a) * (x powhr b)"
+ by transfer (rule powr_add)
-lemma powhr_powhr_swap: "!!a b x. (x powhr a) powhr b = (x powhr b) powhr a"
-by (transfer, rule powr_powr_swap)
+lemma powhr_powhr: "\<And>a b x. (x powhr a) powhr b = x powhr (a * b)"
+ by transfer (rule powr_powr)
-lemma powhr_minus: "!!a x. x powhr (-a) = inverse (x powhr a)"
-by (transfer, rule powr_minus)
+lemma powhr_powhr_swap: "\<And>a b x. (x powhr a) powhr b = (x powhr b) powhr a"
+ by transfer (rule powr_powr_swap)
-lemma powhr_minus_divide: "x powhr (-a) = 1/(x powhr a)"
-by (simp add: divide_inverse powhr_minus)
+lemma powhr_minus: "\<And>a x. x powhr (- a) = inverse (x powhr a)"
+ by transfer (rule powr_minus)
-lemma powhr_less_mono: "!!a b x. [| a < b; 1 < x |] ==> x powhr a < x powhr b"
-by (transfer, simp)
+lemma powhr_minus_divide: "x powhr (- a) = 1 / (x powhr a)"
+ by (simp add: divide_inverse powhr_minus)
-lemma powhr_less_cancel: "!!a b x. [| x powhr a < x powhr b; 1 < x |] ==> a < b"
-by (transfer, simp)
+lemma powhr_less_mono: "\<And>a b x. a < b \<Longrightarrow> 1 < x \<Longrightarrow> x powhr a < x powhr b"
+ by transfer simp
+
+lemma powhr_less_cancel: "\<And>a b x. x powhr a < x powhr b \<Longrightarrow> 1 < x \<Longrightarrow> a < b"
+ by transfer simp
-lemma powhr_less_cancel_iff [simp]:
- "1 < x ==> (x powhr a < x powhr b) = (a < b)"
-by (blast intro: powhr_less_cancel powhr_less_mono)
+lemma powhr_less_cancel_iff [simp]: "1 < x \<Longrightarrow> x powhr a < x powhr b \<longleftrightarrow> a < b"
+ by (blast intro: powhr_less_cancel powhr_less_mono)
-lemma powhr_le_cancel_iff [simp]:
- "1 < x ==> (x powhr a \<le> x powhr b) = (a \<le> b)"
-by (simp add: linorder_not_less [symmetric])
+lemma powhr_le_cancel_iff [simp]: "1 < x \<Longrightarrow> x powhr a \<le> x powhr b \<longleftrightarrow> a \<le> b"
+ by (simp add: linorder_not_less [symmetric])
-lemma hlog:
- "hlog (star_n X) (star_n Y) =
- star_n (%n. log (X n) (Y n))"
-by (simp add: hlog_def starfun2_star_n)
+lemma hlog: "hlog (star_n X) (star_n Y) = star_n (\<lambda>n. log (X n) (Y n))"
+ by (simp add: hlog_def starfun2_star_n)
-lemma hlog_starfun_ln: "!!x. ( *f* ln) x = hlog (( *f* exp) 1) x"
-by (transfer, rule log_ln)
+lemma hlog_starfun_ln: "\<And>x. ( *f* ln) x = hlog (( *f* exp) 1) x"
+ by transfer (rule log_ln)
-lemma powhr_hlog_cancel [simp]:
- "!!a x. [| 0 < a; a \<noteq> 1; 0 < x |] ==> a powhr (hlog a x) = x"
-by (transfer, simp)
+lemma powhr_hlog_cancel [simp]: "\<And>a x. 0 < a \<Longrightarrow> a \<noteq> 1 \<Longrightarrow> 0 < x \<Longrightarrow> a powhr (hlog a x) = x"
+ by transfer simp
-lemma hlog_powhr_cancel [simp]:
- "!!a y. [| 0 < a; a \<noteq> 1 |] ==> hlog a (a powhr y) = y"
-by (transfer, simp)
+lemma hlog_powhr_cancel [simp]: "\<And>a y. 0 < a \<Longrightarrow> a \<noteq> 1 \<Longrightarrow> hlog a (a powhr y) = y"
+ by transfer simp
lemma hlog_mult:
- "!!a x y. [| 0 < a; a \<noteq> 1; 0 < x; 0 < y |]
- ==> hlog a (x * y) = hlog a x + hlog a y"
-by (transfer, rule log_mult)
+ "\<And>a x y. 0 < a \<Longrightarrow> a \<noteq> 1 \<Longrightarrow> 0 < x \<Longrightarrow> 0 < y \<Longrightarrow> hlog a (x * y) = hlog a x + hlog a y"
+ by transfer (rule log_mult)
-lemma hlog_as_starfun:
- "!!a x. [| 0 < a; a \<noteq> 1 |] ==> hlog a x = ( *f* ln) x / ( *f* ln) a"
-by (transfer, simp add: log_def)
+lemma hlog_as_starfun: "\<And>a x. 0 < a \<Longrightarrow> a \<noteq> 1 \<Longrightarrow> hlog a x = ( *f* ln) x / ( *f* ln) a"
+ by transfer (simp add: log_def)
lemma hlog_eq_div_starfun_ln_mult_hlog:
- "!!a b x. [| 0 < a; a \<noteq> 1; 0 < b; b \<noteq> 1; 0 < x |]
- ==> hlog a x = (( *f* ln) b/( *f*ln) a) * hlog b x"
-by (transfer, rule log_eq_div_ln_mult_log)
+ "\<And>a b x. 0 < a \<Longrightarrow> a \<noteq> 1 \<Longrightarrow> 0 < b \<Longrightarrow> b \<noteq> 1 \<Longrightarrow> 0 < x \<Longrightarrow>
+ hlog a x = (( *f* ln) b / ( *f* ln) a) * hlog b x"
+ by transfer (rule log_eq_div_ln_mult_log)
-lemma powhr_as_starfun: "!!a x. x powhr a = (if x=0 then 0 else ( *f* exp) (a * ( *f* real_ln) x))"
- by (transfer, simp add: powr_def)
+lemma powhr_as_starfun: "\<And>a x. x powhr a = (if x = 0 then 0 else ( *f* exp) (a * ( *f* real_ln) x))"
+ by transfer (simp add: powr_def)
lemma HInfinite_powhr:
- "[| x : HInfinite; 0 < x; a : HFinite - Infinitesimal;
- 0 < a |] ==> x powhr a : HInfinite"
-apply (auto intro!: starfun_ln_ge_zero starfun_ln_HInfinite HInfinite_HFinite_not_Infinitesimal_mult2 starfun_exp_HInfinite
- simp add: order_less_imp_le HInfinite_gt_zero_gt_one powhr_as_starfun zero_le_mult_iff)
-done
+ "x \<in> HInfinite \<Longrightarrow> 0 < x \<Longrightarrow> a \<in> HFinite - Infinitesimal \<Longrightarrow> 0 < a \<Longrightarrow> x powhr a \<in> HInfinite"
+ by (auto intro!: starfun_ln_ge_zero starfun_ln_HInfinite
+ HInfinite_HFinite_not_Infinitesimal_mult2 starfun_exp_HInfinite
+ simp add: order_less_imp_le HInfinite_gt_zero_gt_one powhr_as_starfun zero_le_mult_iff)
lemma hlog_hrabs_HInfinite_Infinitesimal:
- "[| x : HFinite - Infinitesimal; a : HInfinite; 0 < a |]
- ==> hlog a \<bar>x\<bar> : Infinitesimal"
-apply (frule HInfinite_gt_zero_gt_one)
-apply (auto intro!: starfun_ln_HFinite_not_Infinitesimal
- HInfinite_inverse_Infinitesimal Infinitesimal_HFinite_mult2
- simp add: starfun_ln_HInfinite not_Infinitesimal_not_zero
- hlog_as_starfun divide_inverse)
-done
+ "x \<in> HFinite - Infinitesimal \<Longrightarrow> a \<in> HInfinite \<Longrightarrow> 0 < a \<Longrightarrow> hlog a \<bar>x\<bar> \<in> Infinitesimal"
+ apply (frule HInfinite_gt_zero_gt_one)
+ apply (auto intro!: starfun_ln_HFinite_not_Infinitesimal
+ HInfinite_inverse_Infinitesimal Infinitesimal_HFinite_mult2
+ simp add: starfun_ln_HInfinite not_Infinitesimal_not_zero
+ hlog_as_starfun divide_inverse)
+ done
-lemma hlog_HInfinite_as_starfun:
- "[| a : HInfinite; 0 < a |] ==> hlog a x = ( *f* ln) x / ( *f* ln) a"
-by (rule hlog_as_starfun, auto)
+lemma hlog_HInfinite_as_starfun: "a \<in> HInfinite \<Longrightarrow> 0 < a \<Longrightarrow> hlog a x = ( *f* ln) x / ( *f* ln) a"
+ by (rule hlog_as_starfun) auto
-lemma hlog_one [simp]: "!!a. hlog a 1 = 0"
-by (transfer, simp)
+lemma hlog_one [simp]: "\<And>a. hlog a 1 = 0"
+ by transfer simp
-lemma hlog_eq_one [simp]: "!!a. [| 0 < a; a \<noteq> 1 |] ==> hlog a a = 1"
-by (transfer, rule log_eq_one)
+lemma hlog_eq_one [simp]: "\<And>a. 0 < a \<Longrightarrow> a \<noteq> 1 \<Longrightarrow> hlog a a = 1"
+ by transfer (rule log_eq_one)
-lemma hlog_inverse:
- "[| 0 < a; a \<noteq> 1; 0 < x |] ==> hlog a (inverse x) = - hlog a x"
-apply (rule add_left_cancel [of "hlog a x", THEN iffD1])
-apply (simp add: hlog_mult [symmetric])
-done
+lemma hlog_inverse: "0 < a \<Longrightarrow> a \<noteq> 1 \<Longrightarrow> 0 < x \<Longrightarrow> hlog a (inverse x) = - hlog a x"
+ by (rule add_left_cancel [of "hlog a x", THEN iffD1]) (simp add: hlog_mult [symmetric])
-lemma hlog_divide:
- "[| 0 < a; a \<noteq> 1; 0 < x; 0 < y|] ==> hlog a (x/y) = hlog a x - hlog a y"
-by (simp add: hlog_mult hlog_inverse divide_inverse)
+lemma hlog_divide: "0 < a \<Longrightarrow> a \<noteq> 1 \<Longrightarrow> 0 < x \<Longrightarrow> 0 < y \<Longrightarrow> hlog a (x / y) = hlog a x - hlog a y"
+ by (simp add: hlog_mult hlog_inverse divide_inverse)
lemma hlog_less_cancel_iff [simp]:
- "!!a x y. [| 1 < a; 0 < x; 0 < y |] ==> (hlog a x < hlog a y) = (x < y)"
-by (transfer, simp)
+ "\<And>a x y. 1 < a \<Longrightarrow> 0 < x \<Longrightarrow> 0 < y \<Longrightarrow> hlog a x < hlog a y \<longleftrightarrow> x < y"
+ by transfer simp
-lemma hlog_le_cancel_iff [simp]:
- "[| 1 < a; 0 < x; 0 < y |] ==> (hlog a x \<le> hlog a y) = (x \<le> y)"
-by (simp add: linorder_not_less [symmetric])
+lemma hlog_le_cancel_iff [simp]: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> 0 < y \<Longrightarrow> hlog a x \<le> hlog a y \<longleftrightarrow> x \<le> y"
+ by (simp add: linorder_not_less [symmetric])
end
--- a/src/HOL/Nonstandard_Analysis/HSEQ.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Nonstandard_Analysis/HSEQ.thy Tue Dec 20 16:18:56 2016 +0100
@@ -15,434 +15,431 @@
abbrevs "--->" = "\<longlonglongrightarrow>\<^sub>N\<^sub>S"
begin
-definition
- NSLIMSEQ :: "[nat => 'a::real_normed_vector, 'a] => bool"
+definition NSLIMSEQ :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> bool"
("((_)/ \<longlonglongrightarrow>\<^sub>N\<^sub>S (_))" [60, 60] 60) where
\<comment>\<open>Nonstandard definition of convergence of sequence\<close>
- "X \<longlonglongrightarrow>\<^sub>N\<^sub>S L = (\<forall>N \<in> HNatInfinite. ( *f* X) N \<approx> star_of L)"
+ "X \<longlonglongrightarrow>\<^sub>N\<^sub>S L \<longleftrightarrow> (\<forall>N \<in> HNatInfinite. ( *f* X) N \<approx> star_of L)"
-definition
- nslim :: "(nat => 'a::real_normed_vector) => 'a" where
- \<comment>\<open>Nonstandard definition of limit using choice operator\<close>
- "nslim X = (THE L. X \<longlonglongrightarrow>\<^sub>N\<^sub>S L)"
+definition nslim :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> 'a"
+ where "nslim X = (THE L. X \<longlonglongrightarrow>\<^sub>N\<^sub>S L)"
+ \<comment> \<open>Nonstandard definition of limit using choice operator\<close>
+
-definition
- NSconvergent :: "(nat => 'a::real_normed_vector) => bool" where
- \<comment>\<open>Nonstandard definition of convergence\<close>
- "NSconvergent X = (\<exists>L. X \<longlonglongrightarrow>\<^sub>N\<^sub>S L)"
+definition NSconvergent :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> bool"
+ where "NSconvergent X \<longleftrightarrow> (\<exists>L. X \<longlonglongrightarrow>\<^sub>N\<^sub>S L)"
+ \<comment> \<open>Nonstandard definition of convergence\<close>
-definition
- NSBseq :: "(nat => 'a::real_normed_vector) => bool" where
- \<comment>\<open>Nonstandard definition for bounded sequence\<close>
- "NSBseq X = (\<forall>N \<in> HNatInfinite. ( *f* X) N : HFinite)"
+definition NSBseq :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> bool"
+ where "NSBseq X \<longleftrightarrow> (\<forall>N \<in> HNatInfinite. ( *f* X) N \<in> HFinite)"
+ \<comment> \<open>Nonstandard definition for bounded sequence\<close>
+
-definition
- NSCauchy :: "(nat => 'a::real_normed_vector) => bool" where
- \<comment>\<open>Nonstandard definition\<close>
- "NSCauchy X = (\<forall>M \<in> HNatInfinite. \<forall>N \<in> HNatInfinite. ( *f* X) M \<approx> ( *f* X) N)"
+definition NSCauchy :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> bool"
+ where "NSCauchy X \<longleftrightarrow> (\<forall>M \<in> HNatInfinite. \<forall>N \<in> HNatInfinite. ( *f* X) M \<approx> ( *f* X) N)"
+ \<comment> \<open>Nonstandard definition\<close>
+
subsection \<open>Limits of Sequences\<close>
-lemma NSLIMSEQ_iff:
- "(X \<longlonglongrightarrow>\<^sub>N\<^sub>S L) = (\<forall>N \<in> HNatInfinite. ( *f* X) N \<approx> star_of L)"
-by (simp add: NSLIMSEQ_def)
+lemma NSLIMSEQ_iff: "(X \<longlonglongrightarrow>\<^sub>N\<^sub>S L) \<longleftrightarrow> (\<forall>N \<in> HNatInfinite. ( *f* X) N \<approx> star_of L)"
+ by (simp add: NSLIMSEQ_def)
+
+lemma NSLIMSEQ_I: "(\<And>N. N \<in> HNatInfinite \<Longrightarrow> starfun X N \<approx> star_of L) \<Longrightarrow> X \<longlonglongrightarrow>\<^sub>N\<^sub>S L"
+ by (simp add: NSLIMSEQ_def)
-lemma NSLIMSEQ_I:
- "(\<And>N. N \<in> HNatInfinite \<Longrightarrow> starfun X N \<approx> star_of L) \<Longrightarrow> X \<longlonglongrightarrow>\<^sub>N\<^sub>S L"
-by (simp add: NSLIMSEQ_def)
+lemma NSLIMSEQ_D: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S L \<Longrightarrow> N \<in> HNatInfinite \<Longrightarrow> starfun X N \<approx> star_of L"
+ by (simp add: NSLIMSEQ_def)
-lemma NSLIMSEQ_D:
- "\<lbrakk>X \<longlonglongrightarrow>\<^sub>N\<^sub>S L; N \<in> HNatInfinite\<rbrakk> \<Longrightarrow> starfun X N \<approx> star_of L"
-by (simp add: NSLIMSEQ_def)
+lemma NSLIMSEQ_const: "(\<lambda>n. k) \<longlonglongrightarrow>\<^sub>N\<^sub>S k"
+ by (simp add: NSLIMSEQ_def)
-lemma NSLIMSEQ_const: "(%n. k) \<longlonglongrightarrow>\<^sub>N\<^sub>S k"
-by (simp add: NSLIMSEQ_def)
+lemma NSLIMSEQ_add: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> Y \<longlonglongrightarrow>\<^sub>N\<^sub>S b \<Longrightarrow> (\<lambda>n. X n + Y n) \<longlonglongrightarrow>\<^sub>N\<^sub>S a + b"
+ by (auto intro: approx_add simp add: NSLIMSEQ_def)
-lemma NSLIMSEQ_add:
- "[| X \<longlonglongrightarrow>\<^sub>N\<^sub>S a; Y \<longlonglongrightarrow>\<^sub>N\<^sub>S b |] ==> (%n. X n + Y n) \<longlonglongrightarrow>\<^sub>N\<^sub>S a + b"
-by (auto intro: approx_add simp add: NSLIMSEQ_def starfun_add [symmetric])
+lemma NSLIMSEQ_add_const: "f \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> (\<lambda>n. f n + b) \<longlonglongrightarrow>\<^sub>N\<^sub>S a + b"
+ by (simp only: NSLIMSEQ_add NSLIMSEQ_const)
-lemma NSLIMSEQ_add_const: "f \<longlonglongrightarrow>\<^sub>N\<^sub>S a ==> (%n.(f n + b)) \<longlonglongrightarrow>\<^sub>N\<^sub>S a + b"
-by (simp only: NSLIMSEQ_add NSLIMSEQ_const)
+lemma NSLIMSEQ_mult: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> Y \<longlonglongrightarrow>\<^sub>N\<^sub>S b \<Longrightarrow> (\<lambda>n. X n * Y n) \<longlonglongrightarrow>\<^sub>N\<^sub>S a * b"
+ for a b :: "'a::real_normed_algebra"
+ by (auto intro!: approx_mult_HFinite simp add: NSLIMSEQ_def)
-lemma NSLIMSEQ_mult:
- fixes a b :: "'a::real_normed_algebra"
- shows "[| X \<longlonglongrightarrow>\<^sub>N\<^sub>S a; Y \<longlonglongrightarrow>\<^sub>N\<^sub>S b |] ==> (%n. X n * Y n) \<longlonglongrightarrow>\<^sub>N\<^sub>S a * b"
-by (auto intro!: approx_mult_HFinite simp add: NSLIMSEQ_def)
+lemma NSLIMSEQ_minus: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> (\<lambda>n. - X n) \<longlonglongrightarrow>\<^sub>N\<^sub>S - a"
+ by (auto simp add: NSLIMSEQ_def)
-lemma NSLIMSEQ_minus: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a ==> (%n. -(X n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S -a"
-by (auto simp add: NSLIMSEQ_def)
+lemma NSLIMSEQ_minus_cancel: "(\<lambda>n. - X n) \<longlonglongrightarrow>\<^sub>N\<^sub>S -a \<Longrightarrow> X \<longlonglongrightarrow>\<^sub>N\<^sub>S a"
+ by (drule NSLIMSEQ_minus) simp
-lemma NSLIMSEQ_minus_cancel: "(%n. -(X n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S -a ==> X \<longlonglongrightarrow>\<^sub>N\<^sub>S a"
-by (drule NSLIMSEQ_minus, simp)
-
-lemma NSLIMSEQ_diff:
- "[| X \<longlonglongrightarrow>\<^sub>N\<^sub>S a; Y \<longlonglongrightarrow>\<^sub>N\<^sub>S b |] ==> (%n. X n - Y n) \<longlonglongrightarrow>\<^sub>N\<^sub>S a - b"
+lemma NSLIMSEQ_diff: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> Y \<longlonglongrightarrow>\<^sub>N\<^sub>S b \<Longrightarrow> (\<lambda>n. X n - Y n) \<longlonglongrightarrow>\<^sub>N\<^sub>S a - b"
using NSLIMSEQ_add [of X a "- Y" "- b"] by (simp add: NSLIMSEQ_minus fun_Compl_def)
(* FIXME: delete *)
-lemma NSLIMSEQ_add_minus:
- "[| X \<longlonglongrightarrow>\<^sub>N\<^sub>S a; Y \<longlonglongrightarrow>\<^sub>N\<^sub>S b |] ==> (%n. X n + -Y n) \<longlonglongrightarrow>\<^sub>N\<^sub>S a + -b"
+lemma NSLIMSEQ_add_minus: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> Y \<longlonglongrightarrow>\<^sub>N\<^sub>S b \<Longrightarrow> (\<lambda>n. X n + - Y n) \<longlonglongrightarrow>\<^sub>N\<^sub>S a + - b"
by (simp add: NSLIMSEQ_diff)
-lemma NSLIMSEQ_diff_const: "f \<longlonglongrightarrow>\<^sub>N\<^sub>S a ==> (%n.(f n - b)) \<longlonglongrightarrow>\<^sub>N\<^sub>S a - b"
-by (simp add: NSLIMSEQ_diff NSLIMSEQ_const)
+lemma NSLIMSEQ_diff_const: "f \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> (\<lambda>n. f n - b) \<longlonglongrightarrow>\<^sub>N\<^sub>S a - b"
+ by (simp add: NSLIMSEQ_diff NSLIMSEQ_const)
-lemma NSLIMSEQ_inverse:
- fixes a :: "'a::real_normed_div_algebra"
- shows "[| X \<longlonglongrightarrow>\<^sub>N\<^sub>S a; a ~= 0 |] ==> (%n. inverse(X n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S inverse(a)"
-by (simp add: NSLIMSEQ_def star_of_approx_inverse)
+lemma NSLIMSEQ_inverse: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> (\<lambda>n. inverse (X n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S inverse a"
+ for a :: "'a::real_normed_div_algebra"
+ by (simp add: NSLIMSEQ_def star_of_approx_inverse)
-lemma NSLIMSEQ_mult_inverse:
- fixes a b :: "'a::real_normed_field"
- shows
- "[| X \<longlonglongrightarrow>\<^sub>N\<^sub>S a; Y \<longlonglongrightarrow>\<^sub>N\<^sub>S b; b ~= 0 |] ==> (%n. X n / Y n) \<longlonglongrightarrow>\<^sub>N\<^sub>S a/b"
-by (simp add: NSLIMSEQ_mult NSLIMSEQ_inverse divide_inverse)
+lemma NSLIMSEQ_mult_inverse: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> Y \<longlonglongrightarrow>\<^sub>N\<^sub>S b \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> (\<lambda>n. X n / Y n) \<longlonglongrightarrow>\<^sub>N\<^sub>S a / b"
+ for a b :: "'a::real_normed_field"
+ by (simp add: NSLIMSEQ_mult NSLIMSEQ_inverse divide_inverse)
lemma starfun_hnorm: "\<And>x. hnorm (( *f* f) x) = ( *f* (\<lambda>x. norm (f x))) x"
-by transfer simp
+ by transfer simp
lemma NSLIMSEQ_norm: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> (\<lambda>n. norm (X n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S norm a"
-by (simp add: NSLIMSEQ_def starfun_hnorm [symmetric] approx_hnorm)
-
-text\<open>Uniqueness of limit\<close>
-lemma NSLIMSEQ_unique: "[| X \<longlonglongrightarrow>\<^sub>N\<^sub>S a; X \<longlonglongrightarrow>\<^sub>N\<^sub>S b |] ==> a = b"
-apply (simp add: NSLIMSEQ_def)
-apply (drule HNatInfinite_whn [THEN [2] bspec])+
-apply (auto dest: approx_trans3)
-done
+ by (simp add: NSLIMSEQ_def starfun_hnorm [symmetric] approx_hnorm)
-lemma NSLIMSEQ_pow [rule_format]:
- fixes a :: "'a::{real_normed_algebra,power}"
- shows "(X \<longlonglongrightarrow>\<^sub>N\<^sub>S a) --> ((%n. (X n) ^ m) \<longlonglongrightarrow>\<^sub>N\<^sub>S a ^ m)"
-apply (induct "m")
-apply (auto simp add: power_Suc intro: NSLIMSEQ_mult NSLIMSEQ_const)
-done
+text \<open>Uniqueness of limit.\<close>
+lemma NSLIMSEQ_unique: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> X \<longlonglongrightarrow>\<^sub>N\<^sub>S b \<Longrightarrow> a = b"
+ apply (simp add: NSLIMSEQ_def)
+ apply (drule HNatInfinite_whn [THEN [2] bspec])+
+ apply (auto dest: approx_trans3)
+ done
-text\<open>We can now try and derive a few properties of sequences,
- starting with the limit comparison property for sequences.\<close>
+lemma NSLIMSEQ_pow [rule_format]: "(X \<longlonglongrightarrow>\<^sub>N\<^sub>S a) \<longrightarrow> ((\<lambda>n. (X n) ^ m) \<longlonglongrightarrow>\<^sub>N\<^sub>S a ^ m)"
+ for a :: "'a::{real_normed_algebra,power}"
+ by (induct m) (auto intro: NSLIMSEQ_mult NSLIMSEQ_const)
+
+text \<open>We can now try and derive a few properties of sequences,
+ starting with the limit comparison property for sequences.\<close>
-lemma NSLIMSEQ_le:
- "[| f \<longlonglongrightarrow>\<^sub>N\<^sub>S l; g \<longlonglongrightarrow>\<^sub>N\<^sub>S m;
- \<exists>N. \<forall>n \<ge> N. f(n) \<le> g(n)
- |] ==> l \<le> (m::real)"
-apply (simp add: NSLIMSEQ_def, safe)
-apply (drule starfun_le_mono)
-apply (drule HNatInfinite_whn [THEN [2] bspec])+
-apply (drule_tac x = whn in spec)
-apply (drule bex_Infinitesimal_iff2 [THEN iffD2])+
-apply clarify
-apply (auto intro: hypreal_of_real_le_add_Infininitesimal_cancel2)
-done
+lemma NSLIMSEQ_le: "f \<longlonglongrightarrow>\<^sub>N\<^sub>S l \<Longrightarrow> g \<longlonglongrightarrow>\<^sub>N\<^sub>S m \<Longrightarrow> \<exists>N. \<forall>n \<ge> N. f n \<le> g n \<Longrightarrow> l \<le> m"
+ for l m :: real
+ apply (simp add: NSLIMSEQ_def, safe)
+ apply (drule starfun_le_mono)
+ apply (drule HNatInfinite_whn [THEN [2] bspec])+
+ apply (drule_tac x = whn in spec)
+ apply (drule bex_Infinitesimal_iff2 [THEN iffD2])+
+ apply clarify
+ apply (auto intro: hypreal_of_real_le_add_Infininitesimal_cancel2)
+ done
-lemma NSLIMSEQ_le_const: "[| X \<longlonglongrightarrow>\<^sub>N\<^sub>S (r::real); \<forall>n. a \<le> X n |] ==> a \<le> r"
-by (erule NSLIMSEQ_le [OF NSLIMSEQ_const], auto)
+lemma NSLIMSEQ_le_const: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S r \<Longrightarrow> \<forall>n. a \<le> X n \<Longrightarrow> a \<le> r"
+ for a r :: real
+ by (erule NSLIMSEQ_le [OF NSLIMSEQ_const]) auto
-lemma NSLIMSEQ_le_const2: "[| X \<longlonglongrightarrow>\<^sub>N\<^sub>S (r::real); \<forall>n. X n \<le> a |] ==> r \<le> a"
-by (erule NSLIMSEQ_le [OF _ NSLIMSEQ_const], auto)
+lemma NSLIMSEQ_le_const2: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S r \<Longrightarrow> \<forall>n. X n \<le> a \<Longrightarrow> r \<le> a"
+ for a r :: real
+ by (erule NSLIMSEQ_le [OF _ NSLIMSEQ_const]) auto
-text\<open>Shift a convergent series by 1:
+text \<open>Shift a convergent series by 1:
By the equivalence between Cauchiness and convergence and because
the successor of an infinite hypernatural is also infinite.\<close>
-lemma NSLIMSEQ_Suc: "f \<longlonglongrightarrow>\<^sub>N\<^sub>S l ==> (%n. f(Suc n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S l"
-apply (unfold NSLIMSEQ_def, safe)
-apply (drule_tac x="N + 1" in bspec)
-apply (erule HNatInfinite_add)
-apply (simp add: starfun_shift_one)
-done
+lemma NSLIMSEQ_Suc: "f \<longlonglongrightarrow>\<^sub>N\<^sub>S l \<Longrightarrow> (\<lambda>n. f(Suc n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S l"
+ apply (unfold NSLIMSEQ_def)
+ apply safe
+ apply (drule_tac x="N + 1" in bspec)
+ apply (erule HNatInfinite_add)
+ apply (simp add: starfun_shift_one)
+ done
-lemma NSLIMSEQ_imp_Suc: "(%n. f(Suc n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S l ==> f \<longlonglongrightarrow>\<^sub>N\<^sub>S l"
-apply (unfold NSLIMSEQ_def, safe)
-apply (drule_tac x="N - 1" in bspec)
-apply (erule Nats_1 [THEN [2] HNatInfinite_diff])
-apply (simp add: starfun_shift_one one_le_HNatInfinite)
-done
+lemma NSLIMSEQ_imp_Suc: "(\<lambda>n. f(Suc n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S l \<Longrightarrow> f \<longlonglongrightarrow>\<^sub>N\<^sub>S l"
+ apply (unfold NSLIMSEQ_def)
+ apply safe
+ apply (drule_tac x="N - 1" in bspec)
+ apply (erule Nats_1 [THEN [2] HNatInfinite_diff])
+ apply (simp add: starfun_shift_one one_le_HNatInfinite)
+ done
-lemma NSLIMSEQ_Suc_iff: "((%n. f(Suc n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S l) = (f \<longlonglongrightarrow>\<^sub>N\<^sub>S l)"
-by (blast intro: NSLIMSEQ_imp_Suc NSLIMSEQ_Suc)
+lemma NSLIMSEQ_Suc_iff: "(\<lambda>n. f (Suc n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S l \<longleftrightarrow> f \<longlonglongrightarrow>\<^sub>N\<^sub>S l"
+ by (blast intro: NSLIMSEQ_imp_Suc NSLIMSEQ_Suc)
+
subsubsection \<open>Equivalence of @{term LIMSEQ} and @{term NSLIMSEQ}\<close>
lemma LIMSEQ_NSLIMSEQ:
- assumes X: "X \<longlonglongrightarrow> L" shows "X \<longlonglongrightarrow>\<^sub>N\<^sub>S L"
+ assumes X: "X \<longlonglongrightarrow> L"
+ shows "X \<longlonglongrightarrow>\<^sub>N\<^sub>S L"
proof (rule NSLIMSEQ_I)
- fix N assume N: "N \<in> HNatInfinite"
+ fix N
+ assume N: "N \<in> HNatInfinite"
have "starfun X N - star_of L \<in> Infinitesimal"
proof (rule InfinitesimalI2)
- fix r::real assume r: "0 < r"
- from LIMSEQ_D [OF X r]
- obtain no where "\<forall>n\<ge>no. norm (X n - L) < r" ..
- hence "\<forall>n\<ge>star_of no. hnorm (starfun X n - star_of L) < star_of r"
+ fix r :: real
+ assume r: "0 < r"
+ from LIMSEQ_D [OF X r] obtain no where "\<forall>n\<ge>no. norm (X n - L) < r" ..
+ then have "\<forall>n\<ge>star_of no. hnorm (starfun X n - star_of L) < star_of r"
by transfer
- thus "hnorm (starfun X N - star_of L) < star_of r"
+ then show "hnorm (starfun X N - star_of L) < star_of r"
using N by (simp add: star_of_le_HNatInfinite)
qed
- thus "starfun X N \<approx> star_of L"
- by (unfold approx_def)
+ then show "starfun X N \<approx> star_of L"
+ by (simp only: approx_def)
qed
lemma NSLIMSEQ_LIMSEQ:
- assumes X: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S L" shows "X \<longlonglongrightarrow> L"
+ assumes X: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S L"
+ shows "X \<longlonglongrightarrow> L"
proof (rule LIMSEQ_I)
- fix r::real assume r: "0 < r"
+ fix r :: real
+ assume r: "0 < r"
have "\<exists>no. \<forall>n\<ge>no. hnorm (starfun X n - star_of L) < star_of r"
proof (intro exI allI impI)
- fix n assume "whn \<le> n"
+ fix n
+ assume "whn \<le> n"
with HNatInfinite_whn have "n \<in> HNatInfinite"
by (rule HNatInfinite_upward_closed)
with X have "starfun X n \<approx> star_of L"
by (rule NSLIMSEQ_D)
- hence "starfun X n - star_of L \<in> Infinitesimal"
- by (unfold approx_def)
- thus "hnorm (starfun X n - star_of L) < star_of r"
+ then have "starfun X n - star_of L \<in> Infinitesimal"
+ by (simp only: approx_def)
+ then show "hnorm (starfun X n - star_of L) < star_of r"
using r by (rule InfinitesimalD2)
qed
- thus "\<exists>no. \<forall>n\<ge>no. norm (X n - L) < r"
+ then show "\<exists>no. \<forall>n\<ge>no. norm (X n - L) < r"
by transfer
qed
-theorem LIMSEQ_NSLIMSEQ_iff: "(f \<longlonglongrightarrow> L) = (f \<longlonglongrightarrow>\<^sub>N\<^sub>S L)"
-by (blast intro: LIMSEQ_NSLIMSEQ NSLIMSEQ_LIMSEQ)
+theorem LIMSEQ_NSLIMSEQ_iff: "f \<longlonglongrightarrow> L \<longleftrightarrow> f \<longlonglongrightarrow>\<^sub>N\<^sub>S L"
+ by (blast intro: LIMSEQ_NSLIMSEQ NSLIMSEQ_LIMSEQ)
+
subsubsection \<open>Derived theorems about @{term NSLIMSEQ}\<close>
-text\<open>We prove the NS version from the standard one, since the NS proof
- seems more complicated than the standard one above!\<close>
-lemma NSLIMSEQ_norm_zero: "((\<lambda>n. norm (X n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0) = (X \<longlonglongrightarrow>\<^sub>N\<^sub>S 0)"
-by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] tendsto_norm_zero_iff)
+text \<open>We prove the NS version from the standard one, since the NS proof
+ seems more complicated than the standard one above!\<close>
+lemma NSLIMSEQ_norm_zero: "(\<lambda>n. norm (X n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0 \<longleftrightarrow> X \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"
+ by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] tendsto_norm_zero_iff)
-lemma NSLIMSEQ_rabs_zero: "((%n. \<bar>f n\<bar>) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0) = (f \<longlonglongrightarrow>\<^sub>N\<^sub>S (0::real))"
-by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] tendsto_rabs_zero_iff)
-
-text\<open>Generalization to other limits\<close>
-lemma NSLIMSEQ_imp_rabs: "f \<longlonglongrightarrow>\<^sub>N\<^sub>S (l::real) ==> (%n. \<bar>f n\<bar>) \<longlonglongrightarrow>\<^sub>N\<^sub>S \<bar>l\<bar>"
-apply (simp add: NSLIMSEQ_def)
-apply (auto intro: approx_hrabs
- simp add: starfun_abs)
-done
+lemma NSLIMSEQ_rabs_zero: "(\<lambda>n. \<bar>f n\<bar>) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0 \<longleftrightarrow> f \<longlonglongrightarrow>\<^sub>N\<^sub>S (0::real)"
+ by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] tendsto_rabs_zero_iff)
-lemma NSLIMSEQ_inverse_zero:
- "\<forall>y::real. \<exists>N. \<forall>n \<ge> N. y < f(n)
- ==> (%n. inverse(f n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"
-by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_zero)
+text \<open>Generalization to other limits.\<close>
+lemma NSLIMSEQ_imp_rabs: "f \<longlonglongrightarrow>\<^sub>N\<^sub>S l \<Longrightarrow> (\<lambda>n. \<bar>f n\<bar>) \<longlonglongrightarrow>\<^sub>N\<^sub>S \<bar>l\<bar>"
+ for l :: real
+ by (simp add: NSLIMSEQ_def) (auto intro: approx_hrabs simp add: starfun_abs)
+
+lemma NSLIMSEQ_inverse_zero: "\<forall>y::real. \<exists>N. \<forall>n \<ge> N. y < f n \<Longrightarrow> (\<lambda>n. inverse (f n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"
+ by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_zero)
-lemma NSLIMSEQ_inverse_real_of_nat: "(%n. inverse(real(Suc n))) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"
-by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_real_of_nat del: of_nat_Suc)
+lemma NSLIMSEQ_inverse_real_of_nat: "(\<lambda>n. inverse (real (Suc n))) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"
+ by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_real_of_nat del: of_nat_Suc)
-lemma NSLIMSEQ_inverse_real_of_nat_add:
- "(%n. r + inverse(real(Suc n))) \<longlonglongrightarrow>\<^sub>N\<^sub>S r"
-by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_real_of_nat_add del: of_nat_Suc)
+lemma NSLIMSEQ_inverse_real_of_nat_add: "(\<lambda>n. r + inverse (real (Suc n))) \<longlonglongrightarrow>\<^sub>N\<^sub>S r"
+ by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_real_of_nat_add del: of_nat_Suc)
-lemma NSLIMSEQ_inverse_real_of_nat_add_minus:
- "(%n. r + -inverse(real(Suc n))) \<longlonglongrightarrow>\<^sub>N\<^sub>S r"
+lemma NSLIMSEQ_inverse_real_of_nat_add_minus: "(\<lambda>n. r + - inverse (real (Suc n))) \<longlonglongrightarrow>\<^sub>N\<^sub>S r"
using LIMSEQ_inverse_real_of_nat_add_minus by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric])
lemma NSLIMSEQ_inverse_real_of_nat_add_minus_mult:
- "(%n. r*( 1 + -inverse(real(Suc n)))) \<longlonglongrightarrow>\<^sub>N\<^sub>S r"
- using LIMSEQ_inverse_real_of_nat_add_minus_mult by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric])
+ "(\<lambda>n. r * (1 + - inverse (real (Suc n)))) \<longlonglongrightarrow>\<^sub>N\<^sub>S r"
+ using LIMSEQ_inverse_real_of_nat_add_minus_mult
+ by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric])
subsection \<open>Convergence\<close>
-lemma nslimI: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S L ==> nslim X = L"
-apply (simp add: nslim_def)
-apply (blast intro: NSLIMSEQ_unique)
-done
+lemma nslimI: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S L \<Longrightarrow> nslim X = L"
+ by (simp add: nslim_def) (blast intro: NSLIMSEQ_unique)
lemma lim_nslim_iff: "lim X = nslim X"
-by (simp add: lim_def nslim_def LIMSEQ_NSLIMSEQ_iff)
+ by (simp add: lim_def nslim_def LIMSEQ_NSLIMSEQ_iff)
-lemma NSconvergentD: "NSconvergent X ==> \<exists>L. (X \<longlonglongrightarrow>\<^sub>N\<^sub>S L)"
-by (simp add: NSconvergent_def)
+lemma NSconvergentD: "NSconvergent X \<Longrightarrow> \<exists>L. X \<longlonglongrightarrow>\<^sub>N\<^sub>S L"
+ by (simp add: NSconvergent_def)
-lemma NSconvergentI: "(X \<longlonglongrightarrow>\<^sub>N\<^sub>S L) ==> NSconvergent X"
-by (auto simp add: NSconvergent_def)
+lemma NSconvergentI: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S L \<Longrightarrow> NSconvergent X"
+ by (auto simp add: NSconvergent_def)
lemma convergent_NSconvergent_iff: "convergent X = NSconvergent X"
-by (simp add: convergent_def NSconvergent_def LIMSEQ_NSLIMSEQ_iff)
+ by (simp add: convergent_def NSconvergent_def LIMSEQ_NSLIMSEQ_iff)
-lemma NSconvergent_NSLIMSEQ_iff: "NSconvergent X = (X \<longlonglongrightarrow>\<^sub>N\<^sub>S nslim X)"
-by (auto intro: theI NSLIMSEQ_unique simp add: NSconvergent_def nslim_def)
+lemma NSconvergent_NSLIMSEQ_iff: "NSconvergent X \<longleftrightarrow> X \<longlonglongrightarrow>\<^sub>N\<^sub>S nslim X"
+ by (auto intro: theI NSLIMSEQ_unique simp add: NSconvergent_def nslim_def)
subsection \<open>Bounded Monotonic Sequences\<close>
-lemma NSBseqD: "[| NSBseq X; N: HNatInfinite |] ==> ( *f* X) N : HFinite"
-by (simp add: NSBseq_def)
+lemma NSBseqD: "NSBseq X \<Longrightarrow> N \<in> HNatInfinite \<Longrightarrow> ( *f* X) N \<in> HFinite"
+ by (simp add: NSBseq_def)
lemma Standard_subset_HFinite: "Standard \<subseteq> HFinite"
-unfolding Standard_def by auto
+ by (auto simp: Standard_def)
lemma NSBseqD2: "NSBseq X \<Longrightarrow> ( *f* X) N \<in> HFinite"
-apply (cases "N \<in> HNatInfinite")
-apply (erule (1) NSBseqD)
-apply (rule subsetD [OF Standard_subset_HFinite])
-apply (simp add: HNatInfinite_def Nats_eq_Standard)
-done
+ apply (cases "N \<in> HNatInfinite")
+ apply (erule (1) NSBseqD)
+ apply (rule subsetD [OF Standard_subset_HFinite])
+ apply (simp add: HNatInfinite_def Nats_eq_Standard)
+ done
-lemma NSBseqI: "\<forall>N \<in> HNatInfinite. ( *f* X) N : HFinite ==> NSBseq X"
-by (simp add: NSBseq_def)
-
-text\<open>The standard definition implies the nonstandard definition\<close>
+lemma NSBseqI: "\<forall>N \<in> HNatInfinite. ( *f* X) N \<in> HFinite \<Longrightarrow> NSBseq X"
+ by (simp add: NSBseq_def)
-lemma Bseq_NSBseq: "Bseq X ==> NSBseq X"
-proof (unfold NSBseq_def, safe)
+text \<open>The standard definition implies the nonstandard definition.\<close>
+lemma Bseq_NSBseq: "Bseq X \<Longrightarrow> NSBseq X"
+ unfolding NSBseq_def
+proof safe
assume X: "Bseq X"
- fix N assume N: "N \<in> HNatInfinite"
- from BseqD [OF X] obtain K where "\<forall>n. norm (X n) \<le> K" by fast
- hence "\<forall>N. hnorm (starfun X N) \<le> star_of K" by transfer
- hence "hnorm (starfun X N) \<le> star_of K" by simp
- also have "star_of K < star_of (K + 1)" by simp
- finally have "\<exists>x\<in>Reals. hnorm (starfun X N) < x" by (rule bexI, simp)
- thus "starfun X N \<in> HFinite" by (simp add: HFinite_def)
+ fix N
+ assume N: "N \<in> HNatInfinite"
+ from BseqD [OF X] obtain K where "\<forall>n. norm (X n) \<le> K"
+ by fast
+ then have "\<forall>N. hnorm (starfun X N) \<le> star_of K"
+ by transfer
+ then have "hnorm (starfun X N) \<le> star_of K"
+ by simp
+ also have "star_of K < star_of (K + 1)"
+ by simp
+ finally have "\<exists>x\<in>Reals. hnorm (starfun X N) < x"
+ by (rule bexI) simp
+ then show "starfun X N \<in> HFinite"
+ by (simp add: HFinite_def)
qed
-text\<open>The nonstandard definition implies the standard definition\<close>
-
+text \<open>The nonstandard definition implies the standard definition.\<close>
lemma SReal_less_omega: "r \<in> \<real> \<Longrightarrow> r < \<omega>"
-apply (insert HInfinite_omega)
-apply (simp add: HInfinite_def)
-apply (simp add: order_less_imp_le)
-done
+ using HInfinite_omega
+ by (simp add: HInfinite_def) (simp add: order_less_imp_le)
lemma NSBseq_Bseq: "NSBseq X \<Longrightarrow> Bseq X"
proof (rule ccontr)
let ?n = "\<lambda>K. LEAST n. K < norm (X n)"
assume "NSBseq X"
- hence finite: "( *f* X) (( *f* ?n) \<omega>) \<in> HFinite"
+ then have finite: "( *f* X) (( *f* ?n) \<omega>) \<in> HFinite"
by (rule NSBseqD2)
assume "\<not> Bseq X"
- hence "\<forall>K>0. \<exists>n. K < norm (X n)"
+ then have "\<forall>K>0. \<exists>n. K < norm (X n)"
by (simp add: Bseq_def linorder_not_le)
- hence "\<forall>K>0. K < norm (X (?n K))"
+ then have "\<forall>K>0. K < norm (X (?n K))"
by (auto intro: LeastI_ex)
- hence "\<forall>K>0. K < hnorm (( *f* X) (( *f* ?n) K))"
+ then have "\<forall>K>0. K < hnorm (( *f* X) (( *f* ?n) K))"
by transfer
- hence "\<omega> < hnorm (( *f* X) (( *f* ?n) \<omega>))"
+ then have "\<omega> < hnorm (( *f* X) (( *f* ?n) \<omega>))"
by simp
- hence "\<forall>r\<in>\<real>. r < hnorm (( *f* X) (( *f* ?n) \<omega>))"
+ then have "\<forall>r\<in>\<real>. r < hnorm (( *f* X) (( *f* ?n) \<omega>))"
by (simp add: order_less_trans [OF SReal_less_omega])
- hence "( *f* X) (( *f* ?n) \<omega>) \<in> HInfinite"
+ then have "( *f* X) (( *f* ?n) \<omega>) \<in> HInfinite"
by (simp add: HInfinite_def)
with finite show "False"
by (simp add: HFinite_HInfinite_iff)
qed
-text\<open>Equivalence of nonstandard and standard definitions
- for a bounded sequence\<close>
-lemma Bseq_NSBseq_iff: "(Bseq X) = (NSBseq X)"
-by (blast intro!: NSBseq_Bseq Bseq_NSBseq)
+text \<open>Equivalence of nonstandard and standard definitions for a bounded sequence.\<close>
+lemma Bseq_NSBseq_iff: "Bseq X = NSBseq X"
+ by (blast intro!: NSBseq_Bseq Bseq_NSBseq)
-text\<open>A convergent sequence is bounded:
- Boundedness as a necessary condition for convergence.
- The nonstandard version has no existential, as usual\<close>
+text \<open>A convergent sequence is bounded:
+ Boundedness as a necessary condition for convergence.
+ The nonstandard version has no existential, as usual.\<close>
+lemma NSconvergent_NSBseq: "NSconvergent X \<Longrightarrow> NSBseq X"
+ by (simp add: NSconvergent_def NSBseq_def NSLIMSEQ_def)
+ (blast intro: HFinite_star_of approx_sym approx_HFinite)
-lemma NSconvergent_NSBseq: "NSconvergent X ==> NSBseq X"
-apply (simp add: NSconvergent_def NSBseq_def NSLIMSEQ_def)
-apply (blast intro: HFinite_star_of approx_sym approx_HFinite)
-done
+text \<open>Standard Version: easily now proved using equivalence of NS and
+ standard definitions.\<close>
-text\<open>Standard Version: easily now proved using equivalence of NS and
- standard definitions\<close>
+lemma convergent_Bseq: "convergent X \<Longrightarrow> Bseq X"
+ for X :: "nat \<Rightarrow> 'b::real_normed_vector"
+ by (simp add: NSconvergent_NSBseq convergent_NSconvergent_iff Bseq_NSBseq_iff)
-lemma convergent_Bseq: "convergent X ==> Bseq (X::nat \<Rightarrow> _::real_normed_vector)"
-by (simp add: NSconvergent_NSBseq convergent_NSconvergent_iff Bseq_NSBseq_iff)
-subsubsection\<open>Upper Bounds and Lubs of Bounded Sequences\<close>
+subsubsection \<open>Upper Bounds and Lubs of Bounded Sequences\<close>
-lemma NSBseq_isUb: "NSBseq X ==> \<exists>U::real. isUb UNIV {x. \<exists>n. X n = x} U"
-by (simp add: Bseq_NSBseq_iff [symmetric] Bseq_isUb)
+lemma NSBseq_isUb: "NSBseq X \<Longrightarrow> \<exists>U::real. isUb UNIV {x. \<exists>n. X n = x} U"
+ by (simp add: Bseq_NSBseq_iff [symmetric] Bseq_isUb)
-lemma NSBseq_isLub: "NSBseq X ==> \<exists>U::real. isLub UNIV {x. \<exists>n. X n = x} U"
-by (simp add: Bseq_NSBseq_iff [symmetric] Bseq_isLub)
+lemma NSBseq_isLub: "NSBseq X \<Longrightarrow> \<exists>U::real. isLub UNIV {x. \<exists>n. X n = x} U"
+ by (simp add: Bseq_NSBseq_iff [symmetric] Bseq_isLub)
+
-subsubsection\<open>A Bounded and Monotonic Sequence Converges\<close>
+subsubsection \<open>A Bounded and Monotonic Sequence Converges\<close>
-text\<open>The best of both worlds: Easier to prove this result as a standard
+text \<open>The best of both worlds: Easier to prove this result as a standard
theorem and then use equivalence to "transfer" it into the
equivalent nonstandard form if needed!\<close>
-lemma Bmonoseq_NSLIMSEQ: "\<forall>n \<ge> m. X n = X m ==> \<exists>L. (X \<longlonglongrightarrow>\<^sub>N\<^sub>S L)"
-by (auto dest!: Bmonoseq_LIMSEQ simp add: LIMSEQ_NSLIMSEQ_iff)
+lemma Bmonoseq_NSLIMSEQ: "\<forall>n \<ge> m. X n = X m \<Longrightarrow> \<exists>L. X \<longlonglongrightarrow>\<^sub>N\<^sub>S L"
+ by (auto dest!: Bmonoseq_LIMSEQ simp add: LIMSEQ_NSLIMSEQ_iff)
-lemma NSBseq_mono_NSconvergent:
- "[| NSBseq X; \<forall>m. \<forall>n \<ge> m. X m \<le> X n |] ==> NSconvergent (X::nat=>real)"
-by (auto intro: Bseq_mono_convergent
- simp add: convergent_NSconvergent_iff [symmetric]
- Bseq_NSBseq_iff [symmetric])
+lemma NSBseq_mono_NSconvergent: "NSBseq X \<Longrightarrow> \<forall>m. \<forall>n \<ge> m. X m \<le> X n \<Longrightarrow> NSconvergent X"
+ for X :: "nat \<Rightarrow> real"
+ by (auto intro: Bseq_mono_convergent
+ simp: convergent_NSconvergent_iff [symmetric] Bseq_NSBseq_iff [symmetric])
subsection \<open>Cauchy Sequences\<close>
lemma NSCauchyI:
- "(\<And>M N. \<lbrakk>M \<in> HNatInfinite; N \<in> HNatInfinite\<rbrakk> \<Longrightarrow> starfun X M \<approx> starfun X N)
- \<Longrightarrow> NSCauchy X"
-by (simp add: NSCauchy_def)
+ "(\<And>M N. M \<in> HNatInfinite \<Longrightarrow> N \<in> HNatInfinite \<Longrightarrow> starfun X M \<approx> starfun X N) \<Longrightarrow> NSCauchy X"
+ by (simp add: NSCauchy_def)
lemma NSCauchyD:
- "\<lbrakk>NSCauchy X; M \<in> HNatInfinite; N \<in> HNatInfinite\<rbrakk>
- \<Longrightarrow> starfun X M \<approx> starfun X N"
-by (simp add: NSCauchy_def)
+ "NSCauchy X \<Longrightarrow> M \<in> HNatInfinite \<Longrightarrow> N \<in> HNatInfinite \<Longrightarrow> starfun X M \<approx> starfun X N"
+ by (simp add: NSCauchy_def)
-subsubsection\<open>Equivalence Between NS and Standard\<close>
+
+subsubsection \<open>Equivalence Between NS and Standard\<close>
lemma Cauchy_NSCauchy:
- assumes X: "Cauchy X" shows "NSCauchy X"
+ assumes X: "Cauchy X"
+ shows "NSCauchy X"
proof (rule NSCauchyI)
- fix M assume M: "M \<in> HNatInfinite"
- fix N assume N: "N \<in> HNatInfinite"
+ fix M
+ assume M: "M \<in> HNatInfinite"
+ fix N
+ assume N: "N \<in> HNatInfinite"
have "starfun X M - starfun X N \<in> Infinitesimal"
proof (rule InfinitesimalI2)
- fix r :: real assume r: "0 < r"
- from CauchyD [OF X r]
- obtain k where "\<forall>m\<ge>k. \<forall>n\<ge>k. norm (X m - X n) < r" ..
- hence "\<forall>m\<ge>star_of k. \<forall>n\<ge>star_of k.
- hnorm (starfun X m - starfun X n) < star_of r"
+ fix r :: real
+ assume r: "0 < r"
+ from CauchyD [OF X r] obtain k where "\<forall>m\<ge>k. \<forall>n\<ge>k. norm (X m - X n) < r" ..
+ then have "\<forall>m\<ge>star_of k. \<forall>n\<ge>star_of k. hnorm (starfun X m - starfun X n) < star_of r"
by transfer
- thus "hnorm (starfun X M - starfun X N) < star_of r"
+ then show "hnorm (starfun X M - starfun X N) < star_of r"
using M N by (simp add: star_of_le_HNatInfinite)
qed
- thus "starfun X M \<approx> starfun X N"
- by (unfold approx_def)
+ then show "starfun X M \<approx> starfun X N"
+ by (simp only: approx_def)
qed
lemma NSCauchy_Cauchy:
- assumes X: "NSCauchy X" shows "Cauchy X"
+ assumes X: "NSCauchy X"
+ shows "Cauchy X"
proof (rule CauchyI)
- fix r::real assume r: "0 < r"
+ fix r :: real
+ assume r: "0 < r"
have "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. hnorm (starfun X m - starfun X n) < star_of r"
proof (intro exI allI impI)
- fix M assume "whn \<le> M"
+ fix M
+ assume "whn \<le> M"
with HNatInfinite_whn have M: "M \<in> HNatInfinite"
by (rule HNatInfinite_upward_closed)
- fix N assume "whn \<le> N"
+ fix N
+ assume "whn \<le> N"
with HNatInfinite_whn have N: "N \<in> HNatInfinite"
by (rule HNatInfinite_upward_closed)
from X M N have "starfun X M \<approx> starfun X N"
by (rule NSCauchyD)
- hence "starfun X M - starfun X N \<in> Infinitesimal"
- by (unfold approx_def)
- thus "hnorm (starfun X M - starfun X N) < star_of r"
+ then have "starfun X M - starfun X N \<in> Infinitesimal"
+ by (simp only: approx_def)
+ then show "hnorm (starfun X M - starfun X N) < star_of r"
using r by (rule InfinitesimalD2)
qed
- thus "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. norm (X m - X n) < r"
+ then show "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. norm (X m - X n) < r"
by transfer
qed
theorem NSCauchy_Cauchy_iff: "NSCauchy X = Cauchy X"
-by (blast intro!: NSCauchy_Cauchy Cauchy_NSCauchy)
+ by (blast intro!: NSCauchy_Cauchy Cauchy_NSCauchy)
+
subsubsection \<open>Cauchy Sequences are Bounded\<close>
-text\<open>A Cauchy sequence is bounded -- nonstandard version\<close>
+text \<open>A Cauchy sequence is bounded -- nonstandard version.\<close>
-lemma NSCauchy_NSBseq: "NSCauchy X ==> NSBseq X"
-by (simp add: Cauchy_Bseq Bseq_NSBseq_iff [symmetric] NSCauchy_Cauchy_iff)
+lemma NSCauchy_NSBseq: "NSCauchy X \<Longrightarrow> NSBseq X"
+ by (simp add: Cauchy_Bseq Bseq_NSBseq_iff [symmetric] NSCauchy_Cauchy_iff)
+
subsubsection \<open>Cauchy Sequences are Convergent\<close>
-text\<open>Equivalence of Cauchy criterion and convergence:
+text \<open>Equivalence of Cauchy criterion and convergence:
We will prove this using our NS formulation which provides a
much easier proof than using the standard definition. We do not
need to use properties of subsequences such as boundedness,
@@ -453,64 +450,60 @@
since the NS formulations do not involve existential quantifiers.\<close>
lemma NSconvergent_NSCauchy: "NSconvergent X \<Longrightarrow> NSCauchy X"
-apply (simp add: NSconvergent_def NSLIMSEQ_def NSCauchy_def, safe)
-apply (auto intro: approx_trans2)
-done
+ by (simp add: NSconvergent_def NSLIMSEQ_def NSCauchy_def) (auto intro: approx_trans2)
-lemma real_NSCauchy_NSconvergent:
- fixes X :: "nat \<Rightarrow> real"
- shows "NSCauchy X \<Longrightarrow> NSconvergent X"
-apply (simp add: NSconvergent_def NSLIMSEQ_def)
-apply (frule NSCauchy_NSBseq)
-apply (simp add: NSBseq_def NSCauchy_def)
-apply (drule HNatInfinite_whn [THEN [2] bspec])
-apply (drule HNatInfinite_whn [THEN [2] bspec])
-apply (auto dest!: st_part_Ex simp add: SReal_iff)
-apply (blast intro: approx_trans3)
-done
+lemma real_NSCauchy_NSconvergent: "NSCauchy X \<Longrightarrow> NSconvergent X"
+ for X :: "nat \<Rightarrow> real"
+ apply (simp add: NSconvergent_def NSLIMSEQ_def)
+ apply (frule NSCauchy_NSBseq)
+ apply (simp add: NSBseq_def NSCauchy_def)
+ apply (drule HNatInfinite_whn [THEN [2] bspec])
+ apply (drule HNatInfinite_whn [THEN [2] bspec])
+ apply (auto dest!: st_part_Ex simp add: SReal_iff)
+ apply (blast intro: approx_trans3)
+ done
-lemma NSCauchy_NSconvergent:
- fixes X :: "nat \<Rightarrow> 'a::banach"
- shows "NSCauchy X \<Longrightarrow> NSconvergent X"
-apply (drule NSCauchy_Cauchy [THEN Cauchy_convergent])
-apply (erule convergent_NSconvergent_iff [THEN iffD1])
-done
+lemma NSCauchy_NSconvergent: "NSCauchy X \<Longrightarrow> NSconvergent X"
+ for X :: "nat \<Rightarrow> 'a::banach"
+ apply (drule NSCauchy_Cauchy [THEN Cauchy_convergent])
+ apply (erule convergent_NSconvergent_iff [THEN iffD1])
+ done
-lemma NSCauchy_NSconvergent_iff:
- fixes X :: "nat \<Rightarrow> 'a::banach"
- shows "NSCauchy X = NSconvergent X"
-by (fast intro: NSCauchy_NSconvergent NSconvergent_NSCauchy)
+lemma NSCauchy_NSconvergent_iff: "NSCauchy X = NSconvergent X"
+ for X :: "nat \<Rightarrow> 'a::banach"
+ by (fast intro: NSCauchy_NSconvergent NSconvergent_NSCauchy)
subsection \<open>Power Sequences\<close>
-text\<open>The sequence @{term "x^n"} tends to 0 if @{term "0\<le>x"} and @{term
-"x<1"}. Proof will use (NS) Cauchy equivalence for convergence and
+text \<open>The sequence @{term "x^n"} tends to 0 if @{term "0\<le>x"} and @{term
+ "x<1"}. Proof will use (NS) Cauchy equivalence for convergence and
also fact that bounded and monotonic sequence converges.\<close>
-text\<open>We now use NS criterion to bring proof of theorem through\<close>
+text \<open>We now use NS criterion to bring proof of theorem through.\<close>
+lemma NSLIMSEQ_realpow_zero: "0 \<le> x \<Longrightarrow> x < 1 \<Longrightarrow> (\<lambda>n. x ^ n) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"
+ for x :: real
+ apply (simp add: NSLIMSEQ_def)
+ apply (auto dest!: convergent_realpow simp add: convergent_NSconvergent_iff)
+ apply (frule NSconvergentD)
+ apply (auto simp add: NSLIMSEQ_def NSCauchy_NSconvergent_iff [symmetric] NSCauchy_def starfun_pow)
+ apply (frule HNatInfinite_add_one)
+ apply (drule bspec, assumption)
+ apply (drule bspec, assumption)
+ apply (drule_tac x = "N + 1" in bspec, assumption)
+ apply (simp add: hyperpow_add)
+ apply (drule approx_mult_subst_star_of, assumption)
+ apply (drule approx_trans3, assumption)
+ apply (auto simp del: star_of_mult simp add: star_of_mult [symmetric])
+ done
-lemma NSLIMSEQ_realpow_zero:
- "[| 0 \<le> (x::real); x < 1 |] ==> (%n. x ^ n) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"
-apply (simp add: NSLIMSEQ_def)
-apply (auto dest!: convergent_realpow simp add: convergent_NSconvergent_iff)
-apply (frule NSconvergentD)
-apply (auto simp add: NSLIMSEQ_def NSCauchy_NSconvergent_iff [symmetric] NSCauchy_def starfun_pow)
-apply (frule HNatInfinite_add_one)
-apply (drule bspec, assumption)
-apply (drule bspec, assumption)
-apply (drule_tac x = "N + (1::hypnat) " in bspec, assumption)
-apply (simp add: hyperpow_add)
-apply (drule approx_mult_subst_star_of, assumption)
-apply (drule approx_trans3, assumption)
-apply (auto simp del: star_of_mult simp add: star_of_mult [symmetric])
-done
+lemma NSLIMSEQ_rabs_realpow_zero: "\<bar>c\<bar> < 1 \<Longrightarrow> (\<lambda>n. \<bar>c\<bar> ^ n) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"
+ for c :: real
+ by (simp add: LIMSEQ_rabs_realpow_zero LIMSEQ_NSLIMSEQ_iff [symmetric])
-lemma NSLIMSEQ_rabs_realpow_zero: "\<bar>c\<bar> < (1::real) ==> (%n. \<bar>c\<bar> ^ n) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"
-by (simp add: LIMSEQ_rabs_realpow_zero LIMSEQ_NSLIMSEQ_iff [symmetric])
-
-lemma NSLIMSEQ_rabs_realpow_zero2: "\<bar>c\<bar> < (1::real) ==> (%n. c ^ n) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"
-by (simp add: LIMSEQ_rabs_realpow_zero2 LIMSEQ_NSLIMSEQ_iff [symmetric])
+lemma NSLIMSEQ_rabs_realpow_zero2: "\<bar>c\<bar> < 1 \<Longrightarrow> (\<lambda>n. c ^ n) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"
+ for c :: real
+ by (simp add: LIMSEQ_rabs_realpow_zero2 LIMSEQ_NSLIMSEQ_iff [symmetric])
(***---------------------------------------------------------------
Theorems proved by Harrison in HOL that we do not need
--- a/src/HOL/Nonstandard_Analysis/HSeries.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Nonstandard_Analysis/HSeries.thy Tue Dec 20 16:18:56 2016 +0100
@@ -5,200 +5,177 @@
Converted to Isar and polished by lcp
*)
-section\<open>Finite Summation and Infinite Series for Hyperreals\<close>
+section \<open>Finite Summation and Infinite Series for Hyperreals\<close>
theory HSeries
-imports HSEQ
+ imports HSEQ
begin
-definition
- sumhr :: "(hypnat * hypnat * (nat=>real)) => hypreal" where
- "sumhr =
- (%(M,N,f). starfun2 (%m n. sum f {m..<n}) M N)"
+definition sumhr :: "hypnat \<times> hypnat \<times> (nat \<Rightarrow> real) \<Rightarrow> hypreal"
+ where "sumhr = (\<lambda>(M,N,f). starfun2 (\<lambda>m n. sum f {m..<n}) M N)"
+
+definition NSsums :: "(nat \<Rightarrow> real) \<Rightarrow> real \<Rightarrow> bool" (infixr "NSsums" 80)
+ where "f NSsums s = (\<lambda>n. sum f {..<n}) \<longlonglongrightarrow>\<^sub>N\<^sub>S s"
-definition
- NSsums :: "[nat=>real,real] => bool" (infixr "NSsums" 80) where
- "f NSsums s = (%n. sum f {..<n}) \<longlonglongrightarrow>\<^sub>N\<^sub>S s"
+definition NSsummable :: "(nat \<Rightarrow> real) \<Rightarrow> bool"
+ where "NSsummable f \<longleftrightarrow> (\<exists>s. f NSsums s)"
-definition
- NSsummable :: "(nat=>real) => bool" where
- "NSsummable f = (\<exists>s. f NSsums s)"
+definition NSsuminf :: "(nat \<Rightarrow> real) \<Rightarrow> real"
+ where "NSsuminf f = (THE s. f NSsums s)"
-definition
- NSsuminf :: "(nat=>real) => real" where
- "NSsuminf f = (THE s. f NSsums s)"
+lemma sumhr_app: "sumhr (M, N, f) = ( *f2* (\<lambda>m n. sum f {m..<n})) M N"
+ by (simp add: sumhr_def)
-lemma sumhr_app: "sumhr(M,N,f) = ( *f2* (\<lambda>m n. sum f {m..<n})) M N"
-by (simp add: sumhr_def)
+text \<open>Base case in definition of @{term sumr}.\<close>
+lemma sumhr_zero [simp]: "\<And>m. sumhr (m, 0, f) = 0"
+ unfolding sumhr_app by transfer simp
-text\<open>Base case in definition of @{term sumr}\<close>
-lemma sumhr_zero [simp]: "!!m. sumhr (m,0,f) = 0"
-unfolding sumhr_app by transfer simp
-
-text\<open>Recursive case in definition of @{term sumr}\<close>
+text \<open>Recursive case in definition of @{term sumr}.\<close>
lemma sumhr_if:
- "!!m n. sumhr(m,n+1,f) =
- (if n + 1 \<le> m then 0 else sumhr(m,n,f) + ( *f* f) n)"
-unfolding sumhr_app by transfer simp
+ "\<And>m n. sumhr (m, n + 1, f) = (if n + 1 \<le> m then 0 else sumhr (m, n, f) + ( *f* f) n)"
+ unfolding sumhr_app by transfer simp
+
+lemma sumhr_Suc_zero [simp]: "\<And>n. sumhr (n + 1, n, f) = 0"
+ unfolding sumhr_app by transfer simp
-lemma sumhr_Suc_zero [simp]: "!!n. sumhr (n + 1, n, f) = 0"
-unfolding sumhr_app by transfer simp
+lemma sumhr_eq_bounds [simp]: "\<And>n. sumhr (n, n, f) = 0"
+ unfolding sumhr_app by transfer simp
-lemma sumhr_eq_bounds [simp]: "!!n. sumhr (n,n,f) = 0"
-unfolding sumhr_app by transfer simp
+lemma sumhr_Suc [simp]: "\<And>m. sumhr (m, m + 1, f) = ( *f* f) m"
+ unfolding sumhr_app by transfer simp
-lemma sumhr_Suc [simp]: "!!m. sumhr (m,m + 1,f) = ( *f* f) m"
-unfolding sumhr_app by transfer simp
+lemma sumhr_add_lbound_zero [simp]: "\<And>k m. sumhr (m + k, k, f) = 0"
+ unfolding sumhr_app by transfer simp
-lemma sumhr_add_lbound_zero [simp]: "!!k m. sumhr(m+k,k,f) = 0"
-unfolding sumhr_app by transfer simp
+lemma sumhr_add: "\<And>m n. sumhr (m, n, f) + sumhr (m, n, g) = sumhr (m, n, \<lambda>i. f i + g i)"
+ unfolding sumhr_app by transfer (rule sum.distrib [symmetric])
-lemma sumhr_add:
- "!!m n. sumhr (m,n,f) + sumhr(m,n,g) = sumhr(m,n,%i. f i + g i)"
-unfolding sumhr_app by transfer (rule sum.distrib [symmetric])
+lemma sumhr_mult: "\<And>m n. hypreal_of_real r * sumhr (m, n, f) = sumhr (m, n, \<lambda>n. r * f n)"
+ unfolding sumhr_app by transfer (rule sum_distrib_left)
-lemma sumhr_mult:
- "!!m n. hypreal_of_real r * sumhr(m,n,f) = sumhr(m,n,%n. r * f n)"
-unfolding sumhr_app by transfer (rule sum_distrib_left)
+lemma sumhr_split_add: "\<And>n p. n < p \<Longrightarrow> sumhr (0, n, f) + sumhr (n, p, f) = sumhr (0, p, f)"
+ unfolding sumhr_app by transfer (simp add: sum_add_nat_ivl)
-lemma sumhr_split_add:
- "!!n p. n < p ==> sumhr(0,n,f) + sumhr(n,p,f) = sumhr(0,p,f)"
-unfolding sumhr_app by transfer (simp add: sum_add_nat_ivl)
+lemma sumhr_split_diff: "n < p \<Longrightarrow> sumhr (0, p, f) - sumhr (0, n, f) = sumhr (n, p, f)"
+ by (drule sumhr_split_add [symmetric, where f = f]) simp
-lemma sumhr_split_diff: "n<p ==> sumhr(0,p,f) - sumhr(0,n,f) = sumhr(n,p,f)"
-by (drule_tac f = f in sumhr_split_add [symmetric], simp)
+lemma sumhr_hrabs: "\<And>m n. \<bar>sumhr (m, n, f)\<bar> \<le> sumhr (m, n, \<lambda>i. \<bar>f i\<bar>)"
+ unfolding sumhr_app by transfer (rule sum_abs)
-lemma sumhr_hrabs: "!!m n. \<bar>sumhr(m,n,f)\<bar> \<le> sumhr(m,n,%i. \<bar>f i\<bar>)"
-unfolding sumhr_app by transfer (rule sum_abs)
-
-text\<open>other general version also needed\<close>
+text \<open>Other general version also needed.\<close>
lemma sumhr_fun_hypnat_eq:
- "(\<forall>r. m \<le> r & r < n --> f r = g r) -->
- sumhr(hypnat_of_nat m, hypnat_of_nat n, f) =
- sumhr(hypnat_of_nat m, hypnat_of_nat n, g)"
-unfolding sumhr_app by transfer simp
+ "(\<forall>r. m \<le> r \<and> r < n \<longrightarrow> f r = g r) \<longrightarrow>
+ sumhr (hypnat_of_nat m, hypnat_of_nat n, f) =
+ sumhr (hypnat_of_nat m, hypnat_of_nat n, g)"
+ unfolding sumhr_app by transfer simp
-lemma sumhr_const:
- "!!n. sumhr(0, n, %i. r) = hypreal_of_hypnat n * hypreal_of_real r"
-unfolding sumhr_app by transfer simp
+lemma sumhr_const: "\<And>n. sumhr (0, n, \<lambda>i. r) = hypreal_of_hypnat n * hypreal_of_real r"
+ unfolding sumhr_app by transfer simp
-lemma sumhr_less_bounds_zero [simp]: "!!m n. n < m ==> sumhr(m,n,f) = 0"
-unfolding sumhr_app by transfer simp
+lemma sumhr_less_bounds_zero [simp]: "\<And>m n. n < m \<Longrightarrow> sumhr (m, n, f) = 0"
+ unfolding sumhr_app by transfer simp
-lemma sumhr_minus: "!!m n. sumhr(m, n, %i. - f i) = - sumhr(m, n, f)"
-unfolding sumhr_app by transfer (rule sum_negf)
+lemma sumhr_minus: "\<And>m n. sumhr (m, n, \<lambda>i. - f i) = - sumhr (m, n, f)"
+ unfolding sumhr_app by transfer (rule sum_negf)
lemma sumhr_shift_bounds:
- "!!m n. sumhr(m+hypnat_of_nat k,n+hypnat_of_nat k,f) =
- sumhr(m,n,%i. f(i + k))"
-unfolding sumhr_app by transfer (rule sum_shift_bounds_nat_ivl)
+ "\<And>m n. sumhr (m + hypnat_of_nat k, n + hypnat_of_nat k, f) =
+ sumhr (m, n, \<lambda>i. f (i + k))"
+ unfolding sumhr_app by transfer (rule sum_shift_bounds_nat_ivl)
-subsection\<open>Nonstandard Sums\<close>
+subsection \<open>Nonstandard Sums\<close>
-text\<open>Infinite sums are obtained by summing to some infinite hypernatural
- (such as @{term whn})\<close>
-lemma sumhr_hypreal_of_hypnat_omega:
- "sumhr(0,whn,%i. 1) = hypreal_of_hypnat whn"
-by (simp add: sumhr_const)
+text \<open>Infinite sums are obtained by summing to some infinite hypernatural
+ (such as @{term whn}).\<close>
+lemma sumhr_hypreal_of_hypnat_omega: "sumhr (0, whn, \<lambda>i. 1) = hypreal_of_hypnat whn"
+ by (simp add: sumhr_const)
-lemma sumhr_hypreal_omega_minus_one: "sumhr(0, whn, %i. 1) = \<omega> - 1"
-apply (simp add: sumhr_const)
-(* FIXME: need lemma: hypreal_of_hypnat whn = \<omega> - 1 *)
-(* maybe define \<omega> = hypreal_of_hypnat whn + 1 *)
-apply (unfold star_class_defs omega_def hypnat_omega_def
- of_hypnat_def star_of_def)
-apply (simp add: starfun_star_n starfun2_star_n)
-done
+lemma sumhr_hypreal_omega_minus_one: "sumhr(0, whn, \<lambda>i. 1) = \<omega> - 1"
+ apply (simp add: sumhr_const)
+ (* FIXME: need lemma: hypreal_of_hypnat whn = \<omega> - 1 *)
+ (* maybe define \<omega> = hypreal_of_hypnat whn + 1 *)
+ apply (unfold star_class_defs omega_def hypnat_omega_def of_hypnat_def star_of_def)
+ apply (simp add: starfun_star_n starfun2_star_n)
+ done
-lemma sumhr_minus_one_realpow_zero [simp]:
- "!!N. sumhr(0, N + N, %i. (-1) ^ (i+1)) = 0"
-unfolding sumhr_app
-apply transfer
-apply (simp del: power_Suc add: mult_2 [symmetric])
-apply (induct_tac N)
-apply simp_all
-done
+lemma sumhr_minus_one_realpow_zero [simp]: "\<And>N. sumhr (0, N + N, \<lambda>i. (-1) ^ (i + 1)) = 0"
+ unfolding sumhr_app
+ apply transfer
+ apply (simp del: power_Suc add: mult_2 [symmetric])
+ apply (induct_tac N)
+ apply simp_all
+ done
lemma sumhr_interval_const:
- "(\<forall>n. m \<le> Suc n --> f n = r) & m \<le> na
- ==> sumhr(hypnat_of_nat m,hypnat_of_nat na,f) =
- (hypreal_of_nat (na - m) * hypreal_of_real r)"
-unfolding sumhr_app by transfer simp
+ "(\<forall>n. m \<le> Suc n \<longrightarrow> f n = r) \<and> m \<le> na \<Longrightarrow>
+ sumhr (hypnat_of_nat m, hypnat_of_nat na, f) = hypreal_of_nat (na - m) * hypreal_of_real r"
+ unfolding sumhr_app by transfer simp
-lemma starfunNat_sumr: "!!N. ( *f* (%n. sum f {0..<n})) N = sumhr(0,N,f)"
-unfolding sumhr_app by transfer (rule refl)
+lemma starfunNat_sumr: "\<And>N. ( *f* (\<lambda>n. sum f {0..<n})) N = sumhr (0, N, f)"
+ unfolding sumhr_app by transfer (rule refl)
-lemma sumhr_hrabs_approx [simp]: "sumhr(0, M, f) \<approx> sumhr(0, N, f)
- ==> \<bar>sumhr(M, N, f)\<bar> \<approx> 0"
-apply (cut_tac x = M and y = N in linorder_less_linear)
-apply (auto simp add: approx_refl)
-apply (drule approx_sym [THEN approx_minus_iff [THEN iffD1]])
-apply (auto dest: approx_hrabs
- simp add: sumhr_split_diff)
-done
+lemma sumhr_hrabs_approx [simp]: "sumhr (0, M, f) \<approx> sumhr (0, N, f) \<Longrightarrow> \<bar>sumhr (M, N, f)\<bar> \<approx> 0"
+ using linorder_less_linear [where x = M and y = N]
+ apply auto
+ apply (drule approx_sym [THEN approx_minus_iff [THEN iffD1]])
+ apply (auto dest: approx_hrabs simp add: sumhr_split_diff)
+ done
+
+
+subsection \<open>Infinite sums: Standard and NS theorems\<close>
-(*----------------------------------------------------------------
- infinite sums: Standard and NS theorems
- ----------------------------------------------------------------*)
-lemma sums_NSsums_iff: "(f sums l) = (f NSsums l)"
-by (simp add: sums_def NSsums_def LIMSEQ_NSLIMSEQ_iff)
+lemma sums_NSsums_iff: "f sums l \<longleftrightarrow> f NSsums l"
+ by (simp add: sums_def NSsums_def LIMSEQ_NSLIMSEQ_iff)
-lemma summable_NSsummable_iff: "(summable f) = (NSsummable f)"
-by (simp add: summable_def NSsummable_def sums_NSsums_iff)
+lemma summable_NSsummable_iff: "summable f \<longleftrightarrow> NSsummable f"
+ by (simp add: summable_def NSsummable_def sums_NSsums_iff)
-lemma suminf_NSsuminf_iff: "(suminf f) = (NSsuminf f)"
-by (simp add: suminf_def NSsuminf_def sums_NSsums_iff)
+lemma suminf_NSsuminf_iff: "suminf f = NSsuminf f"
+ by (simp add: suminf_def NSsuminf_def sums_NSsums_iff)
-lemma NSsums_NSsummable: "f NSsums l ==> NSsummable f"
-by (simp add: NSsums_def NSsummable_def, blast)
+lemma NSsums_NSsummable: "f NSsums l \<Longrightarrow> NSsummable f"
+ unfolding NSsums_def NSsummable_def by blast
-lemma NSsummable_NSsums: "NSsummable f ==> f NSsums (NSsuminf f)"
-apply (simp add: NSsummable_def NSsuminf_def NSsums_def)
-apply (blast intro: theI NSLIMSEQ_unique)
-done
+lemma NSsummable_NSsums: "NSsummable f \<Longrightarrow> f NSsums (NSsuminf f)"
+ unfolding NSsummable_def NSsuminf_def NSsums_def
+ by (blast intro: theI NSLIMSEQ_unique)
-lemma NSsums_unique: "f NSsums s ==> (s = NSsuminf f)"
-by (simp add: suminf_NSsuminf_iff [symmetric] sums_NSsums_iff sums_unique)
+lemma NSsums_unique: "f NSsums s \<Longrightarrow> s = NSsuminf f"
+ by (simp add: suminf_NSsuminf_iff [symmetric] sums_NSsums_iff sums_unique)
-lemma NSseries_zero:
- "\<forall>m. n \<le> Suc m --> f(m) = 0 ==> f NSsums (sum f {..<n})"
-by (auto simp add: sums_NSsums_iff [symmetric] not_le[symmetric] intro!: sums_finite)
+lemma NSseries_zero: "\<forall>m. n \<le> Suc m \<longrightarrow> f m = 0 \<Longrightarrow> f NSsums (sum f {..<n})"
+ by (auto simp add: sums_NSsums_iff [symmetric] not_le[symmetric] intro!: sums_finite)
lemma NSsummable_NSCauchy:
- "NSsummable f =
- (\<forall>M \<in> HNatInfinite. \<forall>N \<in> HNatInfinite. \<bar>sumhr(M,N,f)\<bar> \<approx> 0)"
-apply (auto simp add: summable_NSsummable_iff [symmetric]
- summable_iff_convergent convergent_NSconvergent_iff atLeast0LessThan[symmetric]
- NSCauchy_NSconvergent_iff [symmetric] NSCauchy_def starfunNat_sumr)
-apply (cut_tac x = M and y = N in linorder_less_linear)
-apply auto
-apply (rule approx_minus_iff [THEN iffD2, THEN approx_sym])
-apply (rule_tac [2] approx_minus_iff [THEN iffD2])
-apply (auto dest: approx_hrabs_zero_cancel
- simp add: sumhr_split_diff atLeast0LessThan[symmetric])
-done
+ "NSsummable f \<longleftrightarrow> (\<forall>M \<in> HNatInfinite. \<forall>N \<in> HNatInfinite. \<bar>sumhr (M, N, f)\<bar> \<approx> 0)"
+ apply (auto simp add: summable_NSsummable_iff [symmetric]
+ summable_iff_convergent convergent_NSconvergent_iff atLeast0LessThan[symmetric]
+ NSCauchy_NSconvergent_iff [symmetric] NSCauchy_def starfunNat_sumr)
+ apply (cut_tac x = M and y = N in linorder_less_linear)
+ apply auto
+ apply (rule approx_minus_iff [THEN iffD2, THEN approx_sym])
+ apply (rule_tac [2] approx_minus_iff [THEN iffD2])
+ apply (auto dest: approx_hrabs_zero_cancel simp: sumhr_split_diff atLeast0LessThan[symmetric])
+ done
-text\<open>Terms of a convergent series tend to zero\<close>
-lemma NSsummable_NSLIMSEQ_zero: "NSsummable f ==> f \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"
-apply (auto simp add: NSLIMSEQ_def NSsummable_NSCauchy)
-apply (drule bspec, auto)
-apply (drule_tac x = "N + 1 " in bspec)
-apply (auto intro: HNatInfinite_add_one approx_hrabs_zero_cancel)
-done
+text \<open>Terms of a convergent series tend to zero.\<close>
+lemma NSsummable_NSLIMSEQ_zero: "NSsummable f \<Longrightarrow> f \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"
+ apply (auto simp add: NSLIMSEQ_def NSsummable_NSCauchy)
+ apply (drule bspec)
+ apply auto
+ apply (drule_tac x = "N + 1 " in bspec)
+ apply (auto intro: HNatInfinite_add_one approx_hrabs_zero_cancel)
+ done
-text\<open>Nonstandard comparison test\<close>
-lemma NSsummable_comparison_test:
- "[| \<exists>N. \<forall>n. N \<le> n --> \<bar>f n\<bar> \<le> g n; NSsummable g |] ==> NSsummable f"
-apply (fold summable_NSsummable_iff)
-apply (rule summable_comparison_test, simp, assumption)
-done
+text \<open>Nonstandard comparison test.\<close>
+lemma NSsummable_comparison_test: "\<exists>N. \<forall>n. N \<le> n \<longrightarrow> \<bar>f n\<bar> \<le> g n \<Longrightarrow> NSsummable g \<Longrightarrow> NSsummable f"
+ apply (fold summable_NSsummable_iff)
+ apply (rule summable_comparison_test, simp, assumption)
+ done
lemma NSsummable_rabs_comparison_test:
- "[| \<exists>N. \<forall>n. N \<le> n --> \<bar>f n\<bar> \<le> g n; NSsummable g |]
- ==> NSsummable (%k. \<bar>f k\<bar>)"
-apply (rule NSsummable_comparison_test)
-apply (auto)
-done
+ "\<exists>N. \<forall>n. N \<le> n \<longrightarrow> \<bar>f n\<bar> \<le> g n \<Longrightarrow> NSsummable g \<Longrightarrow> NSsummable (\<lambda>k. \<bar>f k\<bar>)"
+ by (rule NSsummable_comparison_test) auto
end
--- a/src/HOL/Nonstandard_Analysis/StarDef.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Nonstandard_Analysis/StarDef.thy Tue Dec 20 16:18:56 2016 +0100
@@ -81,7 +81,7 @@
by (simp add: FreeUltrafilterNat.proper)
text \<open>Standard principles that play a central role in the transfer tactic.\<close>
-definition Ifun :: "('a \<Rightarrow> 'b) star \<Rightarrow> 'a star \<Rightarrow> 'b star" ("_ \<star> _" [300,301] 300)
+definition Ifun :: "('a \<Rightarrow> 'b) star \<Rightarrow> 'a star \<Rightarrow> 'b star" ("(_ \<star>/ _)" [300, 301] 300)
where "Ifun f \<equiv>
\<lambda>x. Abs_star (\<Union>F\<in>Rep_star f. \<Union>X\<in>Rep_star x. starrel``{\<lambda>n. F n (X n)})"
--- a/src/HOL/Number_Theory/Cong.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Number_Theory/Cong.thy Tue Dec 20 16:18:56 2016 +0100
@@ -251,7 +251,7 @@
done
lemma cong_altdef_int: "[(a::int) = b] (mod m) = (m dvd (a - b))"
- by (metis cong_int_def zmod_eq_dvd_iff)
+ by (metis cong_int_def mod_eq_dvd_iff)
lemma cong_abs_int: "[(x::int) = y] (mod abs m) = [x = y] (mod m)"
by (simp add: cong_altdef_int)
@@ -429,7 +429,7 @@
by (simp add: cong_nat_def mod_mult2_eq mod_add_left_eq)
lemma neg_cong_int: "([(a::int) = b] (mod m)) = ([-a = -b] (mod m))"
- by (metis cong_int_def minus_minus zminus_zmod)
+ by (metis cong_int_def minus_minus mod_minus_cong)
lemma cong_modulus_neg_int: "([(a::int) = b] (mod m)) = ([a = b] (mod -m))"
by (auto simp add: cong_altdef_int)
--- a/src/HOL/Number_Theory/Euclidean_Algorithm.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Number_Theory/Euclidean_Algorithm.thy Tue Dec 20 16:18:56 2016 +0100
@@ -17,7 +17,7 @@
The existence of these functions makes it possible to derive gcd and lcm functions
for any Euclidean semiring.
\<close>
-class euclidean_semiring = semiring_modulo + normalization_semidom +
+class euclidean_semiring = semidom_modulo + normalization_semidom +
fixes euclidean_size :: "'a \<Rightarrow> nat"
assumes size_0 [simp]: "euclidean_size 0 = 0"
assumes mod_size_less:
@@ -26,30 +26,6 @@
"b \<noteq> 0 \<Longrightarrow> euclidean_size a \<le> euclidean_size (a * b)"
begin
-lemma mod_0 [simp]: "0 mod a = 0"
- using div_mult_mod_eq [of 0 a] by simp
-
-lemma dvd_mod_iff:
- assumes "k dvd n"
- shows "(k dvd m mod n) = (k dvd m)"
-proof -
- from assms have "(k dvd m mod n) \<longleftrightarrow> (k dvd ((m div n) * n + m mod n))"
- by (simp add: dvd_add_right_iff)
- also have "(m div n) * n + m mod n = m"
- using div_mult_mod_eq [of m n] by simp
- finally show ?thesis .
-qed
-
-lemma mod_0_imp_dvd:
- assumes "a mod b = 0"
- shows "b dvd a"
-proof -
- have "b dvd ((a div b) * b)" by simp
- also have "(a div b) * b = a"
- using div_mult_mod_eq [of a b] by (simp add: assms)
- finally show ?thesis .
-qed
-
lemma euclidean_size_normalize [simp]:
"euclidean_size (normalize a) = euclidean_size a"
proof (cases "a = 0")
--- a/src/HOL/Number_Theory/Pocklington.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Number_Theory/Pocklington.thy Tue Dec 20 16:18:56 2016 +0100
@@ -369,7 +369,7 @@
hence th: "[a^?r = 1] (mod n)"
using eqo mod_mult_left_eq[of "(a^?o)^?q" "a^?r" n]
apply (simp add: cong_nat_def del: One_nat_def)
- by (simp add: mod_mult_left_eq[symmetric])
+ by (metis mod_mult_left_eq nat_mult_1)
{assume r: "?r = 0" hence ?rhs by (simp add: dvd_eq_mod_eq_0)}
moreover
{assume r: "?r \<noteq> 0"
--- a/src/HOL/Number_Theory/Primes.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Number_Theory/Primes.thy Tue Dec 20 16:18:56 2016 +0100
@@ -241,12 +241,18 @@
"prime (p::int) \<longleftrightarrow> p > 1 \<and> (\<forall>n \<in> {2..<p}. \<not> n dvd p)"
by (auto simp add: prime_int_code)
-lemmas prime_nat_simp_numeral [simp] = prime_nat_simp [of "numeral m"] for m
+lemma prime_int_numeral_eq [simp]:
+ "prime (numeral m :: int) \<longleftrightarrow> prime (numeral m :: nat)"
+ by (simp add: prime_int_nat_transfer)
lemma two_is_prime_nat [simp]: "prime (2::nat)"
- by simp
+ by (simp add: prime_nat_simp)
-declare prime_int_nat_transfer[of "numeral m" for m, simp]
+lemma prime_nat_numeral_eq [simp]:
+ "prime (numeral m :: nat) \<longleftrightarrow>
+ (1::nat) < numeral m \<and>
+ (\<forall>n::nat\<in>set [2..<numeral m]. \<not> n dvd numeral m)"
+ by (fact prime_nat_simp) -- \<open>TODO Sieve Of Erathosthenes might speed this up\<close>
text\<open>A bit of regression testing:\<close>
--- a/src/HOL/Number_Theory/Quadratic_Reciprocity.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Number_Theory/Quadratic_Reciprocity.thy Tue Dec 20 16:18:56 2016 +0100
@@ -167,7 +167,7 @@
fix a b
assume a: "P_1 res a" "P_1 res b"
hence "int p * int q dvd a - b"
- using divides_mult[of "int p" "a - b" "int q"] pq_coprime_int zmod_eq_dvd_iff[of a _ b]
+ using divides_mult[of "int p" "a - b" "int q"] pq_coprime_int mod_eq_dvd_iff [of a _ b]
unfolding P_1_def by force
hence "a = b" using dvd_imp_le_int[of "a - b"] a unfolding P_1_def by fastforce
}
@@ -187,7 +187,7 @@
assume a: "x \<in> BuC" "y \<in> BuC" "f_1 x = f_1 y"
hence "int p * int q dvd x - y"
using f_1_def pq_coprime_int divides_mult[of "int p" "x - y" "int q"]
- zmod_eq_dvd_iff[of x _ y] by auto
+ mod_eq_dvd_iff[of x _ y] by auto
hence "x = y"
using dvd_imp_le_int[of "x - y" "int p * int q"] a unfolding BuC_def by force
}
--- a/src/HOL/Number_Theory/Residues.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Number_Theory/Residues.thy Tue Dec 20 16:18:56 2016 +0100
@@ -40,7 +40,7 @@
apply (insert m_gt_one)
apply (rule abelian_groupI)
apply (unfold R_def residue_ring_def)
- apply (auto simp add: mod_add_right_eq [symmetric] ac_simps)
+ apply (auto simp add: mod_add_right_eq ac_simps)
apply (case_tac "x = 0")
apply force
apply (subgoal_tac "(x + (m - x)) mod m = 0")
@@ -55,7 +55,7 @@
apply auto
apply (subgoal_tac "x * y mod m * z mod m = z * (x * y mod m) mod m")
apply (erule ssubst)
- apply (subst mod_mult_right_eq [symmetric])+
+ apply (subst mod_mult_right_eq)+
apply (simp_all only: ac_simps)
done
@@ -64,9 +64,9 @@
apply (rule abelian_group)
apply (rule comm_monoid)
apply (unfold R_def residue_ring_def, auto)
- apply (subst mod_add_eq [symmetric])
+ apply (subst mod_add_eq)
apply (subst mult.commute)
- apply (subst mod_mult_right_eq [symmetric])
+ apply (subst mod_mult_right_eq)
apply (simp add: field_simps)
done
@@ -116,9 +116,9 @@
apply auto
apply (rule the_equality)
apply auto
- apply (subst mod_add_right_eq [symmetric])
+ apply (subst mod_add_right_eq)
apply auto
- apply (subst mod_add_left_eq [symmetric])
+ apply (subst mod_add_left_eq)
apply auto
apply (subgoal_tac "y mod m = - x mod m")
apply simp
@@ -143,13 +143,11 @@
lemma add_cong: "(x mod m) \<oplus> (y mod m) = (x + y) mod m"
unfolding R_def residue_ring_def
- apply auto
- apply presburger
- done
+ by (auto simp add: mod_simps)
lemma mult_cong: "(x mod m) \<otimes> (y mod m) = (x * y) mod m"
unfolding R_def residue_ring_def
- by auto (metis mod_mult_eq)
+ by (auto simp add: mod_simps)
lemma zero_cong: "\<zero> = 0"
unfolding R_def residue_ring_def by auto
--- a/src/HOL/ROOT Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/ROOT Tue Dec 20 16:18:56 2016 +0100
@@ -31,18 +31,15 @@
*}
theories
Library
- Nonpos_Ints
- Periodic_Fun
+ (*conflicting type class instantiations and dependent applications*)
+ Field_as_Ring
+ Finite_Lattice
+ List_lexord
Polynomial_Factorial
- Predicate_Compile_Quickcheck
Prefix_Order
- Rewrite
- (*conflicting type class instantiations*)
- List_lexord
- Sublist_Order
Product_Lexorder
Product_Order
- Finite_Lattice
+ Sublist_Order
(*data refinements and dependent applications*)
AList_Mapping
Code_Binary_Nat
@@ -54,11 +51,13 @@
DAList_Multiset
RBT_Mapping
RBT_Set
+ (*prototypic tools*)
+ Predicate_Compile_Quickcheck
(*legacy tools*)
- Refute
Old_Datatype
Old_Recdef
Old_SMT
+ Refute
document_files "root.bib" "root.tex"
session "HOL-Hahn_Banach" in Hahn_Banach = HOL +
@@ -238,16 +237,14 @@
Generate_Target_Nat
Generate_Efficient_Datastructures
Generate_Pretty_Char
+ Code_Test_PolyML
+ Code_Test_Scala
theories [condition = "ISABELLE_GHC"]
Code_Test_GHC
theories [condition = "ISABELLE_MLTON"]
Code_Test_MLton
theories [condition = "ISABELLE_OCAMLC"]
Code_Test_OCaml
- theories [condition = "ISABELLE_POLYML"]
- Code_Test_PolyML
- theories [condition = "ISABELLE_SCALA"]
- Code_Test_Scala
theories [condition = "ISABELLE_SMLNJ"]
Code_Test_SMLNJ
@@ -394,7 +391,7 @@
theories Decision_Procs
session "HOL-Proofs-ex" in "Proofs/ex" = "HOL-Proofs" +
- options [document = false, parallel_proofs = 0]
+ options [document = false]
theories
Hilbert_Classical
Proof_Terms
--- a/src/HOL/Relation.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Relation.thy Tue Dec 20 16:18:56 2016 +0100
@@ -389,8 +389,9 @@
lemma trans_INTER: "\<forall>x\<in>S. trans (r x) \<Longrightarrow> trans (INTER S r)"
by (fast intro: transI elim: transD)
-(* FIXME thm trans_INTER [to_pred] *)
-
+lemma transp_INF: "\<forall>x\<in>S. transp (r x) \<Longrightarrow> transp (INFIMUM S r)"
+ by (fact trans_INTER [to_pred])
+
lemma trans_join [code]: "trans r \<longleftrightarrow> (\<forall>(x, y1) \<in> r. \<forall>(y2, z) \<in> r. y1 = y2 \<longrightarrow> (x, z) \<in> r)"
by (auto simp add: trans_def)
@@ -620,13 +621,15 @@
lemma relcomp_UNION_distrib: "s O UNION I r = (\<Union>i\<in>I. s O r i) "
by auto
-(* FIXME thm relcomp_UNION_distrib [to_pred] *)
-
+lemma relcompp_SUP_distrib: "s OO SUPREMUM I r = (\<Squnion>i\<in>I. s OO r i)"
+ by (fact relcomp_UNION_distrib [to_pred])
+
lemma relcomp_UNION_distrib2: "UNION I r O s = (\<Union>i\<in>I. r i O s) "
by auto
-(* FIXME thm relcomp_UNION_distrib2 [to_pred] *)
-
+lemma relcompp_SUP_distrib2: "SUPREMUM I r OO s = (\<Squnion>i\<in>I. r i OO s)"
+ by (fact relcomp_UNION_distrib2 [to_pred])
+
lemma single_valued_relcomp: "single_valued r \<Longrightarrow> single_valued s \<Longrightarrow> single_valued (r O s)"
unfolding single_valued_def by blast
--- a/src/HOL/Rings.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Rings.thy Tue Dec 20 16:18:56 2016 +0100
@@ -713,9 +713,61 @@
lemma div_by_1 [simp]: "a div 1 = a"
using nonzero_mult_div_cancel_left [of 1 a] by simp
+lemma dvd_div_eq_0_iff:
+ assumes "b dvd a"
+ shows "a div b = 0 \<longleftrightarrow> a = 0"
+ using assms by (elim dvdE, cases "b = 0") simp_all
+
+lemma dvd_div_eq_cancel:
+ "a div c = b div c \<Longrightarrow> c dvd a \<Longrightarrow> c dvd b \<Longrightarrow> a = b"
+ by (elim dvdE, cases "c = 0") simp_all
+
+lemma dvd_div_eq_iff:
+ "c dvd a \<Longrightarrow> c dvd b \<Longrightarrow> a div c = b div c \<longleftrightarrow> a = b"
+ by (elim dvdE, cases "c = 0") simp_all
+
end
class idom_divide = idom + semidom_divide
+begin
+
+lemma dvd_neg_div:
+ assumes "b dvd a"
+ shows "- a div b = - (a div b)"
+proof (cases "b = 0")
+ case True
+ then show ?thesis by simp
+next
+ case False
+ from assms obtain c where "a = b * c" ..
+ then have "- a div b = (b * - c) div b"
+ by simp
+ from False also have "\<dots> = - c"
+ by (rule nonzero_mult_div_cancel_left)
+ with False \<open>a = b * c\<close> show ?thesis
+ by simp
+qed
+
+lemma dvd_div_neg:
+ assumes "b dvd a"
+ shows "a div - b = - (a div b)"
+proof (cases "b = 0")
+ case True
+ then show ?thesis by simp
+next
+ case False
+ then have "- b \<noteq> 0"
+ by simp
+ from assms obtain c where "a = b * c" ..
+ then have "a div - b = (- b * - c) div - b"
+ by simp
+ from \<open>- b \<noteq> 0\<close> also have "\<dots> = - c"
+ by (rule nonzero_mult_div_cancel_left)
+ with False \<open>a = b * c\<close> show ?thesis
+ by simp
+qed
+
+end
class algebraic_semidom = semidom_divide
begin
@@ -884,6 +936,39 @@
by (simp add: mult.commute [of a] mult.assoc)
qed
+lemma div_div_eq_right:
+ assumes "c dvd b" "b dvd a"
+ shows "a div (b div c) = a div b * c"
+proof (cases "c = 0 \<or> b = 0")
+ case True
+ then show ?thesis
+ by auto
+next
+ case False
+ from assms obtain r s where "b = c * r" and "a = c * r * s"
+ by (blast elim: dvdE)
+ moreover with False have "r \<noteq> 0"
+ by auto
+ ultimately show ?thesis using False
+ by simp (simp add: mult.commute [of _ r] mult.assoc mult.commute [of c])
+qed
+
+lemma div_div_div_same:
+ assumes "d dvd b" "b dvd a"
+ shows "(a div d) div (b div d) = a div b"
+proof (cases "b = 0 \<or> d = 0")
+ case True
+ with assms show ?thesis
+ by auto
+next
+ case False
+ from assms obtain r s
+ where "a = d * r * s" and "b = d * r"
+ by (blast elim: dvdE)
+ with False show ?thesis
+ by simp (simp add: ac_simps)
+qed
+
text \<open>Units: invertible elements in a ring\<close>
@@ -1060,6 +1145,15 @@
shows "a div (b * a) = 1 div b"
using assms is_unit_div_mult_cancel_left [of a b] by (simp add: ac_simps)
+lemma unit_div_eq_0_iff:
+ assumes "is_unit b"
+ shows "a div b = 0 \<longleftrightarrow> a = 0"
+ by (rule dvd_div_eq_0_iff) (insert assms, auto)
+
+lemma div_mult_unit2:
+ "is_unit c \<Longrightarrow> b dvd a \<Longrightarrow> a div (b * c) = a div b div c"
+ by (rule dvd_div_mult2_eq) (simp_all add: mult_unit_dvd_iff)
+
end
class normalization_semidom = algebraic_semidom +
@@ -1373,6 +1467,17 @@
by simp
qed
+lemma normalize_unit_factor_eqI:
+ assumes "normalize a = normalize b"
+ and "unit_factor a = unit_factor b"
+ shows "a = b"
+proof -
+ from assms have "unit_factor a * normalize a = unit_factor b * normalize b"
+ by simp
+ then show ?thesis
+ by simp
+qed
+
end
--- a/src/HOL/SPARK/Examples/RIPEMD-160/Round.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/Round.thy Tue Dec 20 16:18:56 2016 +0100
@@ -456,7 +456,7 @@
unfolding round_def
unfolding steps_to_steps'
unfolding H1[symmetric]
- by (simp add: uint_word_ariths(1) rdmods
+ by (simp add: uint_word_ariths(1) mod_simps
uint_word_of_int_id)
qed
--- a/src/HOL/SPARK/Manual/Proc1.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/SPARK/Manual/Proc1.thy Tue Dec 20 16:18:56 2016 +0100
@@ -5,10 +5,10 @@
spark_open "loop_invariant/proc1"
spark_vc procedure_proc1_5
- by (simp add: ring_distribs pull_mods)
+ by (simp add: ring_distribs mod_simps)
spark_vc procedure_proc1_8
- by (simp add: ring_distribs pull_mods)
+ by (simp add: ring_distribs mod_simps)
spark_end
--- a/src/HOL/SPARK/Manual/Proc2.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/SPARK/Manual/Proc2.thy Tue Dec 20 16:18:56 2016 +0100
@@ -5,7 +5,7 @@
spark_open "loop_invariant/proc2"
spark_vc procedure_proc2_7
- by (simp add: ring_distribs pull_mods)
+ by (simp add: ring_distribs mod_simps)
spark_end
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML Tue Dec 20 16:18:56 2016 +0100
@@ -141,8 +141,6 @@
| flatten' (t as Abs _) (names, prems) = [(t, (names, prems))]
| flatten' (t as _ $ _) (names, prems) =
if is_constrt ctxt t orelse keep_functions thy t then
- (* FIXME: constructor terms are supposed to be seen in the way the code generator
- sees constructors?*)
[(t, (names, prems))]
else
case (fst (strip_comb t)) of
--- a/src/HOL/Tools/Qelim/cooper.ML Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Tools/Qelim/cooper.ML Tue Dec 20 16:18:56 2016 +0100
@@ -823,15 +823,13 @@
|> fold Simplifier.add_cong [@{thm "conj_le_cong"}, @{thm "imp_le_cong"}])
val div_mod_ss =
simpset_of (put_simpset HOL_basic_ss @{context}
- addsimps @{thms simp_thms}
- @ map (Thm.symmetric o mk_meta_eq)
- [@{thm "dvd_eq_mod_eq_0"},
- @{thm "mod_add_left_eq"}, @{thm "mod_add_right_eq"},
- @{thm "mod_add_eq"}, @{thm "div_add1_eq"}, @{thm "zdiv_zadd1_eq"}]
- @ [@{thm "mod_self"}, @{thm "mod_by_0"}, @{thm "div_by_0"},
- @{thm "div_0"}, @{thm "mod_0"}, @{thm "div_by_1"}, @{thm "mod_by_1"},
- @{thm "div_by_Suc_0"}, @{thm "mod_by_Suc_0"}, @{thm "Suc_eq_plus1"}]
- @ @{thms ac_simps}
+ addsimps @{thms simp_thms
+ mod_eq_0_iff_dvd mod_add_left_eq mod_add_right_eq
+ mod_add_eq div_add1_eq [symmetric] zdiv_zadd1_eq [symmetric]
+ mod_self mod_by_0 div_by_0
+ div_0 mod_0 div_by_1 mod_by_1
+ div_by_Suc_0 mod_by_Suc_0 Suc_eq_plus1
+ ac_simps}
addsimprocs [@{simproc cancel_div_mod_nat}, @{simproc cancel_div_mod_int}])
val splits_ss =
simpset_of (put_simpset comp_ss @{context}
--- a/src/HOL/Word/Bit_Representation.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Word/Bit_Representation.thy Tue Dec 20 16:18:56 2016 +0100
@@ -308,17 +308,12 @@
lemma sbintrunc_mod2p: "sbintrunc n w = (w + 2 ^ n) mod 2 ^ (Suc n) - 2 ^ n"
apply (induct n arbitrary: w)
- apply simp
- apply (subst mod_add_left_eq)
- apply (simp add: bin_last_def)
- apply arith
- apply (simp add: bin_last_def bin_rest_def Bit_def)
- apply (clarsimp simp: mod_mult_mult1 [symmetric]
- mult_div_mod_eq [symmetric, THEN diff_eq_eq [THEN iffD2 [THEN sym]]])
- apply (rule trans [symmetric, OF _ emep1])
- apply auto
+ apply (auto simp add: bin_last_odd bin_rest_def Bit_def elim!: evenE oddE)
+ apply (smt pos_zmod_mult_2 zle2p)
+ apply (simp add: mult_mod_right)
done
+
subsection "Simplifications for (s)bintrunc"
lemma bintrunc_n_0 [simp]: "bintrunc n 0 = 0"
@@ -647,28 +642,6 @@
"x >= (2 ^ n) ==> x - 2 ^ (Suc n) >= sbintrunc n x"
unfolding no_sbintr_alt2 by (drule sb_dec_lem') simp
-lemmas zmod_uminus' = zminus_zmod [where m=c] for c
-lemmas zpower_zmod' = power_mod [where b=c and n=k] for c k
-
-lemmas brdmod1s' [symmetric] =
- mod_add_left_eq mod_add_right_eq
- mod_diff_left_eq mod_diff_right_eq
- mod_mult_left_eq mod_mult_right_eq
-
-lemmas brdmods' [symmetric] =
- zpower_zmod' [symmetric]
- trans [OF mod_add_left_eq mod_add_right_eq]
- trans [OF mod_diff_left_eq mod_diff_right_eq]
- trans [OF mod_mult_right_eq mod_mult_left_eq]
- zmod_uminus' [symmetric]
- mod_add_left_eq [where b = "1::int"]
- mod_diff_left_eq [where b = "1::int"]
-
-lemmas bintr_arith1s =
- brdmod1s' [where c="2^n::int", folded bintrunc_mod2p] for n
-lemmas bintr_ariths =
- brdmods' [where c="2^n::int", folded bintrunc_mod2p] for n
-
lemmas m2pths = pos_mod_sign pos_mod_bound [OF zless2p]
lemma bintr_ge0: "0 \<le> bintrunc n w"
--- a/src/HOL/Word/Bits_Int.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Word/Bits_Int.thy Tue Dec 20 16:18:56 2016 +0100
@@ -643,14 +643,14 @@
lemma mod_BIT:
"bin BIT bit mod 2 ^ Suc n = (bin mod 2 ^ n) BIT bit"
proof -
- have "bin mod 2 ^ n < 2 ^ n" by simp
- then have "bin mod 2 ^ n \<le> 2 ^ n - 1" by simp
- then have "2 * (bin mod 2 ^ n) \<le> 2 * (2 ^ n - 1)"
- by (rule mult_left_mono) simp
- then have "2 * (bin mod 2 ^ n) + 1 < 2 * 2 ^ n" by simp
- then show ?thesis
- by (auto simp add: Bit_def mod_mult_mult1 mod_add_left_eq [of "2 * bin"]
- mod_pos_pos_trivial)
+ have "2 * (bin mod 2 ^ n) + 1 = (2 * bin mod 2 ^ Suc n) + 1"
+ by (simp add: mod_mult_mult1)
+ also have "\<dots> = ((2 * bin mod 2 ^ Suc n) + 1) mod 2 ^ Suc n"
+ by (simp add: ac_simps p1mod22k')
+ also have "\<dots> = (2 * bin + 1) mod 2 ^ Suc n"
+ by (simp only: mod_simps)
+ finally show ?thesis
+ by (auto simp add: Bit_def)
qed
lemma AND_mod:
--- a/src/HOL/Word/Misc_Numeric.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Word/Misc_Numeric.thy Tue Dec 20 16:18:56 2016 +0100
@@ -8,6 +8,10 @@
imports Main
begin
+lemma one_mod_exp_eq_one [simp]:
+ "1 mod (2 * 2 ^ n) = (1::int)"
+ by (smt mod_pos_pos_trivial zero_less_power)
+
lemma mod_2_neq_1_eq_eq_0:
fixes k :: int
shows "k mod 2 \<noteq> 1 \<longleftrightarrow> k mod 2 = 0"
--- a/src/HOL/Word/Word.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Word/Word.thy Tue Dec 20 16:18:56 2016 +0100
@@ -282,10 +282,10 @@
subsection \<open>Arithmetic operations\<close>
lift_definition word_succ :: "'a::len0 word \<Rightarrow> 'a word" is "\<lambda>x. x + 1"
- by (metis bintr_ariths(6))
+ by (auto simp add: bintrunc_mod2p intro: mod_add_cong)
lift_definition word_pred :: "'a::len0 word \<Rightarrow> 'a word" is "\<lambda>x. x - 1"
- by (metis bintr_ariths(7))
+ by (auto simp add: bintrunc_mod2p intro: mod_diff_cong)
instantiation word :: (len0) "{neg_numeral, modulo, comm_monoid_mult, comm_ring}"
begin
@@ -295,16 +295,16 @@
lift_definition one_word :: "'a word" is "1" .
lift_definition plus_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word" is "op +"
- by (metis bintr_ariths(2))
+ by (auto simp add: bintrunc_mod2p intro: mod_add_cong)
lift_definition minus_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word" is "op -"
- by (metis bintr_ariths(3))
+ by (auto simp add: bintrunc_mod2p intro: mod_diff_cong)
lift_definition uminus_word :: "'a word \<Rightarrow> 'a word" is uminus
- by (metis bintr_ariths(5))
+ by (auto simp add: bintrunc_mod2p intro: mod_minus_cong)
lift_definition times_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word" is "op *"
- by (metis bintr_ariths(4))
+ by (auto simp add: bintrunc_mod2p intro: mod_mult_cong)
definition
word_div_def: "a div b = word_of_int (uint a div uint b)"
@@ -332,9 +332,6 @@
unfolding word_succ_def word_pred_def zero_word_def one_word_def
by simp_all
-lemmas arths =
- bintr_ariths [THEN word_ubin.norm_eq_iff [THEN iffD1], folded word_ubin.eq_norm]
-
lemma wi_homs:
shows
wi_hom_add: "word_of_int a + word_of_int b = word_of_int (a + b)" and
@@ -1340,10 +1337,11 @@
and "sint (word_pred a) = sbintrunc (len_of TYPE('a) - 1) (sint a - 1)"
and "sint (0 :: 'a word) = sbintrunc (len_of TYPE('a) - 1) 0"
and "sint (1 :: 'a word) = sbintrunc (len_of TYPE('a) - 1) 1"
- by (simp_all add: uint_word_arith_bintrs
- [THEN uint_sint [symmetric, THEN trans],
- unfolded uint_sint bintr_arith1s bintr_ariths
- len_gt_0 [THEN bin_sbin_eq_iff'] word_sbin.norm_Rep])
+ apply (simp_all only: word_sbin.inverse_norm [symmetric])
+ apply (simp_all add: wi_hom_syms)
+ apply transfer apply simp
+ apply transfer apply simp
+ done
lemmas uint_div_alt = word_div_def [THEN trans [OF uint_cong int_word_uint]]
lemmas uint_mod_alt = word_mod_def [THEN trans [OF uint_cong int_word_uint]]
@@ -1443,7 +1441,7 @@
with \<open>1 \<le> uint w\<close> have "nat ((uint w - 1) mod 2 ^ len_of TYPE('a)) = nat (uint w) - 1"
by auto
then show ?thesis
- by (simp only: unat_def int_word_uint word_arith_wis mod_diff_right_eq [symmetric])
+ by (simp only: unat_def int_word_uint word_arith_wis mod_diff_right_eq)
qed
lemma measure_unat: "p ~= 0 \<Longrightarrow> unat (p - 1) < unat p"
@@ -2699,7 +2697,7 @@
lemma nth_w2p:
"((2::'a::len word) ^ n) !! m \<longleftrightarrow> m = n \<and> m < len_of TYPE('a::len)"
unfolding test_bit_2p [symmetric] word_of_int [symmetric]
- by (simp add: of_int_power)
+ by simp
lemma uint_2p:
"(0::'a::len word) < 2 ^ n \<Longrightarrow> uint (2 ^ n::'a::len word) = 2 ^ n"
@@ -2723,16 +2721,7 @@
done
lemma word_of_int_2p: "(word_of_int (2 ^ n) :: 'a :: len word) = 2 ^ n"
- apply (unfold word_arith_power_alt)
- apply (case_tac "len_of TYPE ('a)")
- apply clarsimp
- apply (case_tac "nat")
- apply (rule word_ubin.norm_eq_iff [THEN iffD1])
- apply (rule box_equals)
- apply (rule_tac [2] bintr_ariths (1))+
- apply simp
- apply simp
- done
+ by (induct n) (simp_all add: wi_hom_syms)
lemma bang_is_le: "x !! m \<Longrightarrow> 2 ^ m <= (x :: 'a :: len word)"
apply (rule xtr3)
@@ -3244,6 +3233,9 @@
lemma and_mask_wi: "word_of_int i AND mask n = word_of_int (bintrunc n i)"
by (auto simp add: nth_bintr word_size word_ops_nth_size word_eq_iff)
+lemma and_mask_wi': "word_of_int i AND mask n = (word_of_int (bintrunc (min LENGTH('a) n) i) :: 'a::len word)"
+ by (auto simp add: nth_bintr word_size word_ops_nth_size word_eq_iff)
+
lemma and_mask_no: "numeral i AND mask n = word_of_int (bintrunc n (numeral i))"
unfolding word_numeral_alt by (rule and_mask_wi)
@@ -3342,12 +3334,12 @@
"word_succ (a AND mask n) AND mask n = word_succ a AND mask n"
"word_pred (a AND mask n) AND mask n = word_pred a AND mask n"
using word_of_int_Ex [where x=a] word_of_int_Ex [where x=b]
- by (auto simp: and_mask_wi bintr_ariths bintr_arith1s word_of_int_homs)
+ by (auto simp add: and_mask_wi' word_of_int_homs word.abs_eq_iff bintrunc_mod2p mod_simps)
lemma mask_power_eq:
"(x AND mask n) ^ k AND mask n = x ^ k AND mask n"
using word_of_int_Ex [where x=x]
- by (clarsimp simp: and_mask_wi word_of_int_power_hom bintr_ariths)
+ by (auto simp add: and_mask_wi' word_of_int_power_hom word.abs_eq_iff bintrunc_mod2p mod_simps)
subsubsection \<open>Revcast\<close>
@@ -4242,7 +4234,7 @@
apply (simp add: word_size nat_mod_distrib)
apply (rule of_nat_eq_0_iff [THEN iffD1])
apply (auto simp add: not_le mod_eq_0_iff_dvd zdvd_int nat_add_distrib [symmetric])
- using mod_mod_trivial zmod_eq_dvd_iff
+ using mod_mod_trivial mod_eq_dvd_iff
apply blast
done
@@ -4579,9 +4571,9 @@
shows "(x + y) mod b = z' mod b'"
proof -
from 1 2[symmetric] 3[symmetric] have "(x + y) mod b = (x' mod b' + y' mod b') mod b'"
- by (simp add: mod_add_eq[symmetric])
+ by (simp add: mod_add_eq)
also have "\<dots> = (x' + y') mod b'"
- by (simp add: mod_add_eq[symmetric])
+ by (simp add: mod_add_eq)
finally show ?thesis by (simp add: 4)
qed
@@ -4591,11 +4583,8 @@
and 3: "y mod b' = y' mod b'"
and 4: "x' - y' = z'"
shows "(x - y) mod b = z' mod b'"
- using assms
- apply (subst mod_diff_left_eq)
- apply (subst mod_diff_right_eq)
- apply (simp add: mod_diff_left_eq [symmetric] mod_diff_right_eq [symmetric])
- done
+ using 1 2 3 4 [symmetric]
+ by (auto intro: mod_diff_cong)
lemma word_induct_less:
"\<lbrakk>P (0::'a::len word); \<And>n. \<lbrakk>n < m; P n\<rbrakk> \<Longrightarrow> P (1 + n)\<rbrakk> \<Longrightarrow> P m"
--- a/src/HOL/Word/Word_Miscellaneous.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/Word/Word_Miscellaneous.thy Tue Dec 20 16:18:56 2016 +0100
@@ -201,10 +201,6 @@
lemmas push_mods = push_mods' [THEN eq_reflection]
lemmas pull_mods = push_mods [symmetric] rdmods [THEN eq_reflection]
-lemmas mod_simps =
- mod_mult_self2_is_0 [THEN eq_reflection]
- mod_mult_self1_is_0 [THEN eq_reflection]
- mod_mod_trivial [THEN eq_reflection]
lemma nat_mod_eq:
"!!b. b < n ==> a mod n = b mod n ==> a mod n = (b :: nat)"
--- a/src/HOL/ex/Word_Type.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/HOL/ex/Word_Type.thy Tue Dec 20 16:18:56 2016 +0100
@@ -57,7 +57,7 @@
lemma bitrunc_plus:
"bitrunc n (bitrunc n a + bitrunc n b) = bitrunc n (a + b)"
- by (simp add: bitrunc_eq_mod mod_add_eq [symmetric])
+ by (simp add: bitrunc_eq_mod mod_add_eq)
lemma bitrunc_of_1_eq_0_iff [simp]:
"bitrunc n 1 = 0 \<longleftrightarrow> n = 0"
@@ -74,12 +74,12 @@
lemma bitrunc_uminus:
fixes k :: int
shows "bitrunc n (- (bitrunc n k)) = bitrunc n (- k)"
- by (simp add: bitrunc_eq_mod mod_minus_eq [symmetric])
+ by (simp add: bitrunc_eq_mod mod_minus_eq)
lemma bitrunc_minus:
fixes k l :: int
shows "bitrunc n (bitrunc n k - bitrunc n l) = bitrunc n (k - l)"
- by (simp add: bitrunc_eq_mod mod_diff_eq [symmetric])
+ by (simp add: bitrunc_eq_mod mod_diff_eq)
lemma bitrunc_nonnegative [simp]:
fixes k :: int
@@ -279,7 +279,7 @@
lemma of_int_signed [simp]:
"of_int (signed a) = a"
- by transfer (simp add: signed_bitrunc_eq_bitrunc bitrunc_eq_mod zdiff_zmod_left)
+ by transfer (simp add: signed_bitrunc_eq_bitrunc bitrunc_eq_mod mod_simps)
subsubsection \<open>Properties\<close>
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/General/logger.scala Tue Dec 20 16:18:56 2016 +0100
@@ -0,0 +1,31 @@
+/* Title: Pure/General/logger.scala
+ Author: Makarius
+
+Minimal logging support.
+*/
+
+package isabelle
+
+
+object Logger
+{
+ def make(log_file: Option[Path]): Logger =
+ log_file match { case Some(file) => new File_Logger(file) case None => No_Logger }
+}
+
+trait Logger
+{
+ def apply(msg: => String): Unit
+}
+
+object No_Logger extends Logger
+{
+ def apply(msg: => String) { }
+}
+
+class File_Logger(path: Path) extends Logger
+{
+ def apply(msg: => String) { synchronized { File.append(path, msg + "\n") } }
+
+ override def toString: String = path.toString
+}
--- a/src/Pure/Isar/overloading.ML Tue Dec 20 16:17:13 2016 +0100
+++ b/src/Pure/Isar/overloading.ML Tue Dec 20 16:18:56 2016 +0100
@@ -139,7 +139,7 @@
| NONE => NONE);
val unchecks =
map (fn (c_ty as (_, ty), (v, _)) => (Free (v, ty), Const c_ty)) overloading;
- in
+ in
ctxt
|> map_improvable_syntax (K {primary_constraints = [],
secondary_constraints = [], improve = K NONE, subst = subst,
@@ -168,7 +168,7 @@
val overloading = get_overloading lthy;
fun pr_operation ((c, ty), (v, _)) =
Pretty.block (Pretty.breaks
- [Pretty.str v, Pretty.str "==", Proof_Context.pretty_const lthy c,
+ [Pretty.str v, Pretty.str "\<equiv>", Proof_Context.pretty_const lthy c,
Pretty.str "::", Syntax.pretty_typ lthy ty]);
in
[Pretty.block
--- a/src/Pure/Isar/proof_display.ML Tue Dec 20 16:17:13 2016 +0100
+++ b/src/Pure/Isar/proof_display.ML Tue Dec 20 16:18:56 2016 +0100
@@ -95,7 +95,7 @@
val pretty_arities = maps (fn (t, ars) => map (prt_arity t) ars);
fun pretty_abbrev (c, (ty, t)) = Pretty.block
- (prt_const (c, ty) @ [Pretty.str " ==", Pretty.brk 1, prt_term_no_vars t]);
+ (prt_const (c, ty) @ [Pretty.str " \<equiv>", Pretty.brk 1, prt_term_no_vars t]);
fun pretty_axm (a, t) =
Pretty.block [Pretty.mark_str a, Pretty.str ":", Pretty.brk 1, prt_term_no_vars t];
--- a/src/Pure/ML/ml_antiquotations.ML Tue Dec 20 16:17:13 2016 +0100
+++ b/src/Pure/ML/ml_antiquotations.ML Tue Dec 20 16:18:56 2016 +0100
@@ -260,10 +260,12 @@
val _ = Theory.setup
(ML_Antiquotation.value @{binding keyword}
- (Args.theory -- Scan.lift (Parse.position Parse.string) >> (fn (thy, (name, pos)) =>
- if Keyword.is_keyword (Thy_Header.get_keywords thy) name then
- "Parse.$$$ " ^ ML_Syntax.print_string name
- else error ("Bad outer syntax keyword " ^ quote name ^ Position.here pos))) #>
+ (Args.context -- Scan.lift (Parse.position (Parse.name || Parse.keyword_with (K true)))
+ >> (fn (ctxt, (name, pos)) =>
+ if Keyword.is_keyword (Thy_Header.get_keywords' ctxt) name then
+ (Context_Position.report ctxt pos (Token.keyword_markup (true, Markup.keyword2) name);
+ "Parse.$$$ " ^ ML_Syntax.print_string name)
+ else error ("Bad outer syntax keyword " ^ quote name ^ Position.here pos))) #>
ML_Antiquotation.value @{binding command_keyword}
(Args.context -- Scan.lift (Parse.position Parse.name) >> (fn (ctxt, (name, pos)) =>
(case Keyword.command_markup (Thy_Header.get_keywords' ctxt) name of
--- a/src/Pure/Pure.thy Tue Dec 20 16:17:13 2016 +0100
+++ b/src/Pure/Pure.thy Tue Dec 20 16:18:56 2016 +0100
@@ -193,19 +193,19 @@
val _ =
Outer_Syntax.command @{command_keyword oracle} "declare oracle"
- (Parse.range Parse.name -- (@{keyword "="} |-- Parse.ML_source) >>
+ (Parse.range Parse.name -- (@{keyword =} |-- Parse.ML_source) >>
(fn (x, y) => Toplevel.theory (Isar_Cmd.oracle x y)));
val _ =
Outer_Syntax.local_theory @{command_keyword attribute_setup} "define attribute in ML"
(Parse.position Parse.name --
- Parse.!!! (@{keyword "="} |-- Parse.ML_source -- Scan.optional Parse.text "")
+ Parse.!!! (@{keyword =} |-- Parse.ML_source -- Scan.optional Parse.text "")
>> (fn (name, (txt, cmt)) => Attrib.attribute_setup name txt cmt));
val _ =
Outer_Syntax.local_theory @{command_keyword method_setup} "define proof method in ML"
(Parse.position Parse.name --
- Parse.!!! (@{keyword "="} |-- Parse.ML_source -- Scan.optional Parse.text "")
+ Parse.!!! (@{keyword =} |-- Parse.ML_source -- Scan.optional Parse.text "")
>> (fn (name, (txt, cmt)) => Method.method_setup name txt cmt));
val _ =
@@ -221,7 +221,7 @@
val _ =
Outer_Syntax.local_theory @{command_keyword simproc_setup} "define simproc in ML"
(Parse.position Parse.name --
- (@{keyword "("} |-- Parse.enum1 "|" Parse.term --| @{keyword ")"} --| @{keyword "="}) --
+ (@{keyword "("} |-- Parse.enum1 "|" Parse.term --| @{keyword ")"} --| @{keyword =}) --
Parse.ML_source >> (fn ((a, b), c) => Isar_Cmd.simproc_setup a b c));
in end\<close>
@@ -248,7 +248,7 @@
val _ =
Outer_Syntax.local_theory @{command_keyword type_synonym} "declare type abbreviation"
(Parse.type_args -- Parse.binding --
- (@{keyword "="} |-- Parse.!!! (Parse.typ -- Parse.opt_mixfix'))
+ (@{keyword =} |-- Parse.!!! (Parse.typ -- Parse.opt_mixfix'))
>> (fn ((args, a), (rhs, mx)) => snd o Typedecl.abbrev_cmd (a, args, mx) rhs));
in end\<close>
@@ -296,9 +296,9 @@
-- Parse.inner_syntax Parse.string;
fun trans_arrow toks =
- ((@{keyword "\<rightharpoonup>"} || @{keyword "=>"}) >> K Syntax.Parse_Rule ||
- (@{keyword "\<leftharpoondown>"} || @{keyword "<="}) >> K Syntax.Print_Rule ||
- (@{keyword "\<rightleftharpoons>"} || @{keyword "=="}) >> K Syntax.Parse_Print_Rule) toks;
+ ((@{keyword \<rightharpoonup>} || @{keyword =>}) >> K Syntax.Parse_Rule ||
+ (@{keyword \<leftharpoondown>} || @{keyword <=}) >> K Syntax.Print_Rule ||
+ (@{keyword \<rightleftharpoons>} || @{keyword ==}) >> K Syntax.Parse_Print_Rule) toks;
val trans_line =
trans_pat -- Parse.!!! (trans_arrow -- trans_pat)
@@ -505,7 +505,7 @@
val _ =
Outer_Syntax.maybe_begin_local_theory @{command_keyword bundle}
"define bundle of declarations"
- ((Parse.binding --| @{keyword "="}) -- Parse.thms1 -- Parse.for_fixes
+ ((Parse.binding --| @{keyword =}) -- Parse.thms1 -- Parse.for_fixes
>> (uncurry Bundle.bundle_cmd))
(Parse.binding --| Parse.begin >> Bundle.init);
@@ -563,13 +563,13 @@
val locale_val =
Parse_Spec.locale_expression --
- Scan.optional (@{keyword "+"} |-- Parse.!!! (Scan.repeat1 Parse_Spec.context_element)) [] ||
+ Scan.optional (@{keyword +} |-- Parse.!!! (Scan.repeat1 Parse_Spec.context_element)) [] ||
Scan.repeat1 Parse_Spec.context_element >> pair ([], []);
val _ =
Outer_Syntax.command @{command_keyword locale} "define named specification context"
(Parse.binding --
- Scan.optional (@{keyword "="} |-- Parse.!!! locale_val) (([], []), []) -- Parse.opt_begin
+ Scan.optional (@{keyword =} |-- Parse.!!! locale_val) (([], []), []) -- Parse.opt_begin
>> (fn ((name, (expr, elems)), begin) =>
Toplevel.begin_local_theory begin
(Expression.add_locale_cmd name Binding.empty expr elems #> snd)));
@@ -583,7 +583,7 @@
val interpretation_args =
Parse.!!! Parse_Spec.locale_expression --
Scan.optional
- (@{keyword "rewrites"} |-- Parse.and_list1 (Parse_Spec.opt_thm_name ":" -- Parse.prop)) [];
+ (@{keyword rewrites} |-- Parse.and_list1 (Parse_Spec.opt_thm_name ":" -- Parse.prop)) [];
val _ =
Outer_Syntax.command @{command_keyword interpret}
@@ -593,10 +593,10 @@
val interpretation_args_with_defs =
Parse.!!! Parse_Spec.locale_expression --
- (Scan.optional (@{keyword "defines"} |-- Parse.and_list1 (Parse_Spec.opt_thm_name ":"
- -- ((Parse.binding -- Parse.opt_mixfix') --| @{keyword "="} -- Parse.term))) [] --
+ (Scan.optional (@{keyword defines} |-- Parse.and_list1 (Parse_Spec.opt_thm_name ":"
+ -- ((Parse.binding -- Parse.opt_mixfix') --| @{keyword =} -- Parse.term))) [] --
Scan.optional
- (@{keyword "rewrites"} |-- Parse.and_list1 (Parse_Spec.opt_thm_name ":" -- Parse.prop)) []);
+ (@{keyword rewrites} |-- Parse.and_list1 (Parse_Spec.opt_thm_name ":" -- Parse.prop)) []);
val _ =
Outer_Syntax.local_theory_to_proof @{command_keyword global_interpretation}
@@ -607,7 +607,7 @@
val _ =
Outer_Syntax.command @{command_keyword sublocale}
"prove sublocale relation between a locale and a locale expression"
- ((Parse.position Parse.name --| (@{keyword "\<subseteq>"} || @{keyword "<"}) --
+ ((Parse.position Parse.name --| (@{keyword \<subseteq>} || @{keyword <}) --
interpretation_args_with_defs >> (fn (loc, (expr, (defs, equations))) =>
Toplevel.theory_to_proof (Interpretation.global_sublocale_cmd loc expr defs equations)))
|| interpretation_args_with_defs >> (fn (expr, (defs, equations)) =>
@@ -630,12 +630,12 @@
val class_val =
Parse_Spec.class_expression --
- Scan.optional (@{keyword "+"} |-- Parse.!!! (Scan.repeat1 Parse_Spec.context_element)) [] ||
+ Scan.optional (@{keyword +} |-- Parse.!!! (Scan.repeat1 Parse_Spec.context_element)) [] ||
Scan.repeat1 Parse_Spec.context_element >> pair [];
val _ =
Outer_Syntax.command @{command_keyword class} "define type class"
- (Parse.binding -- Scan.optional (@{keyword "="} |-- class_val) ([], []) -- Parse.opt_begin
+ (Parse.binding -- Scan.optional (@{keyword =} |-- class_val) ([], []) -- Parse.opt_begin
>> (fn ((name, (supclasses, elems)), begin) =>
Toplevel.begin_local_theory begin
(Class_Declaration.class_cmd name supclasses elems #> snd)));
@@ -652,7 +652,7 @@
val _ =
Outer_Syntax.command @{command_keyword instance} "prove type arity or subclass relation"
((Parse.class --
- ((@{keyword "\<subseteq>"} || @{keyword "<"}) |-- Parse.!!! Parse.class) >> Class.classrel_cmd ||
+ ((@{keyword \<subseteq>} || @{keyword <}) |-- Parse.!!! Parse.class) >> Class.classrel_cmd ||
Parse.multi_arity >> Class.instance_arity_cmd) >> Toplevel.theory_to_proof ||
Scan.succeed (Toplevel.local_theory_to_proof NONE NONE (Class.instantiation_instance I)));
@@ -666,8 +666,8 @@
val _ =
Outer_Syntax.command @{command_keyword overloading} "overloaded definitions"
- (Scan.repeat1 (Parse.name --| (@{keyword "=="} || @{keyword "\<equiv>"}) -- Parse.term --
- Scan.optional (@{keyword "("} |-- (@{keyword "unchecked"} >> K false) --| @{keyword ")"}) true
+ (Scan.repeat1 (Parse.name --| (@{keyword ==} || @{keyword \<equiv>}) -- Parse.term --
+ Scan.optional (@{keyword "("} |-- (@{keyword unchecked} >> K false) --| @{keyword ")"}) true
>> Scan.triple1) --| Parse.begin
>> (fn operations => Toplevel.begin_local_theory true (Overloading.overloading_cmd operations)));
@@ -787,7 +787,7 @@
(Parse.and_list1
(Parse_Spec.opt_thm_name ":" --
((Parse.binding -- Parse.opt_mixfix) --
- ((@{keyword "\<equiv>"} || @{keyword "=="}) |-- Parse.!!! Parse.termp)))
+ ((@{keyword \<equiv>} || @{keyword ==}) |-- Parse.!!! Parse.termp)))
>> (fn args => Toplevel.proof (fn state =>
(legacy_feature "Old 'def' command -- use 'define' instead"; Proof.def_cmd args state))));
@@ -806,7 +806,7 @@
val _ =
Outer_Syntax.command @{command_keyword let} "bind text variables"
- (Parse.and_list1 (Parse.and_list1 Parse.term -- (@{keyword "="} |-- Parse.term))
+ (Parse.and_list1 (Parse.and_list1 Parse.term -- (@{keyword =} |-- Parse.term))
>> (Toplevel.proof o Proof.let_bind_cmd));
val _ =
@@ -937,7 +937,7 @@
val for_params =
Scan.optional
- (@{keyword "for"} |--
+ (@{keyword for} |--
Parse.!!! ((Scan.option Parse.dots >> is_some) --
(Scan.repeat1 (Parse.position (Parse.maybe Parse.name)))))
(false, []);
@@ -945,7 +945,7 @@
val _ =
Outer_Syntax.command @{command_keyword subgoal}
"focus on first subgoal within backward refinement"
- (opt_fact_binding -- (Scan.option (@{keyword "premises"} |-- Parse.!!! opt_fact_binding)) --
+ (opt_fact_binding -- (Scan.option (@{keyword premises} |-- Parse.!!! opt_fact_binding)) --
for_params >> (fn ((a, b), c) =>
Toplevel.proofs (Seq.make_results o Seq.single o #2 o Subgoal.subgoal_cmd a b c)));
@@ -1176,7 +1176,7 @@
val _ =
Outer_Syntax.command @{command_keyword typ} "read and print type"
- (opt_modes -- (Parse.typ -- Scan.option (@{keyword "::"} |-- Parse.!!! Parse.sort))
+ (opt_modes -- (Parse.typ -- Scan.option (@{keyword ::} |-- Parse.!!! Parse.sort))
>> Isar_Cmd.print_type);
val _ =
--- a/src/Pure/System/isabelle_tool.scala Tue Dec 20 16:17:13 2016 +0100
+++ b/src/Pure/System/isabelle_tool.scala Tue Dec 20 16:18:56 2016 +0100
@@ -113,7 +113,8 @@
Update_Cartouches.isabelle_tool,
Update_Header.isabelle_tool,
Update_Then.isabelle_tool,
- Update_Theorems.isabelle_tool)
+ Update_Theorems.isabelle_tool,
+ isabelle.vscode.Server.isabelle_tool)
private def list_internal(): List[(String, String)] =
for (tool <- internal_tools.toList if tool.accessible)
--- a/src/Pure/Tools/build.ML Tue Dec 20 16:17:13 2016 +0100
+++ b/src/Pure/Tools/build.ML Tue Dec 20 16:18:56 2016 +0100
@@ -109,6 +109,7 @@
(if Options.bool options "checkpoint" then ML_Heap.share_common_data () else ();
Options.set_default options;
Isabelle_Process.init_options ();
+ Future.fork I;
(Thy_Info.use_theories {
document = Present.document_enabled (Options.string options "document"),
symbols = symbols,
--- a/src/Pure/Tools/build.scala Tue Dec 20 16:17:13 2016 +0100
+++ b/src/Pure/Tools/build.scala Tue Dec 20 16:18:56 2016 +0100
@@ -96,12 +96,12 @@
/* source dependencies and static content */
sealed case class Session_Content(
- loaded_theories: Set[String],
- known_theories: Map[String, Document.Node.Name],
- keywords: Thy_Header.Keywords,
- syntax: Outer_Syntax,
- sources: List[(Path, SHA1.Digest)],
- session_graph: Graph_Display.Graph)
+ loaded_theories: Set[String] = Set.empty,
+ known_theories: Map[String, Document.Node.Name] = Map.empty,
+ keywords: Thy_Header.Keywords = Nil,
+ syntax: Outer_Syntax = Outer_Syntax.empty,
+ sources: List[(Path, SHA1.Digest)] = Nil,
+ session_graph: Graph_Display.Graph = Graph_Display.empty_graph)
sealed case class Deps(deps: Map[String, Session_Content])
{
--- a/src/Pure/Tools/ml_console.scala Tue Dec 20 16:17:13 2016 +0100
+++ b/src/Pure/Tools/ml_console.scala Tue Dec 20 16:18:56 2016 +0100
@@ -70,7 +70,7 @@
// process loop
val process =
- ML_Process(options, logic = logic, args = List("-i"), dirs = dirs,
+ ML_Process(options, logic = logic, args = List("-i"), dirs = dirs, redirect = true,
modes = if (raw_ml_system) Nil else modes ::: List("ASCII"),
raw_ml_system = raw_ml_system, store = Sessions.store(system_mode))
val process_output = Future.thread[Unit]("process_output") {
--- a/src/Pure/build-jars Tue Dec 20 16:17:13 2016 +0100
+++ b/src/Pure/build-jars Tue Dec 20 16:18:56 2016 +0100
@@ -49,6 +49,7 @@
General/http_server.scala
General/json.scala
General/linear_set.scala
+ General/logger.scala
General/long_name.scala
General/mercurial.scala
General/multi_map.scala
@@ -153,6 +154,12 @@
"../Tools/Graphview/popups.scala"
"../Tools/Graphview/shapes.scala"
"../Tools/Graphview/tree_panel.scala"
+ "../Tools/VSCode/src/channel.scala"
+ "../Tools/VSCode/src/document_model.scala"
+ "../Tools/VSCode/src/line.scala"
+ "../Tools/VSCode/src/protocol.scala"
+ "../Tools/VSCode/src/server.scala"
+ "../Tools/VSCode/src/uri_resources.scala"
)
--- a/src/Pure/primitive_defs.ML Tue Dec 20 16:17:13 2016 +0100
+++ b/src/Pure/primitive_defs.ML Tue Dec 20 16:18:56 2016 +0100
@@ -34,7 +34,7 @@
commas_quote o map (Syntax.string_of_term ctxt o Syntax_Trans.bound_vars eq_vars);
val display_types = commas_quote o map (Syntax.string_of_typ ctxt);
- val (raw_lhs, rhs) = Logic.dest_equals eq_body handle TERM _ => err "Not a meta-equality (==)";
+ val (raw_lhs, rhs) = Logic.dest_equals eq_body handle TERM _ => err "Not a meta-equality (\<equiv>)";
val lhs = Envir.beta_eta_contract raw_lhs;
val (head, args) = Term.strip_comb lhs;
val head_tfrees = Term.add_tfrees head [];
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/README.md Tue Dec 20 16:18:56 2016 +0100
@@ -0,0 +1,58 @@
+# Isabelle/PIDE for Visual Studio Code editor #
+
+* Extension for the editor ([TypeScript](extension/src/extension.ts))
+* Language Server protocol implementation ([Isabelle/Scala](src/server.scala))
+
+
+## Build and run ##
+
+* shell> `isabelle build -b HOL`
+
+* shell> `cd src/Tools/VSCode/extension; vsce package`
+
+* Preferences / User settings / edit settings.json: e.g.
+ `"isabelle.home": "/home/makarius/isabelle/repos"`
+
+* Extensions / ... / Install from VSIX: `src/Tools/VSCode/extension/isabelle-0.0.1.vsix`
+
+* File / Open Folder: e.g. `src/HOL/Isar_Examples/` then open .thy files
+
+
+## Debug
+
+* shell> `code src/Tools/VSCode/extension`
+
+* View / Debug / Launch Extension
+
+* File / Open Folder: e.g. `src/HOL/Isar_Examples/` then open .thy files
+
+
+## Relevant links ##
+
+### VSCode editor ###
+
+* https://code.visualstudio.com
+* https://code.visualstudio.com/docs/extensionAPI/extension-points
+* https://code.visualstudio.com/docs/extensions/example-language-server
+* https://github.com/Microsoft/vscode-languageserver-node-example
+
+
+### Protocol ###
+
+* https://code.visualstudio.com/blogs/2016/06/27/common-language-protocol
+* https://github.com/Microsoft/vscode-languageserver-node
+* https://github.com/Microsoft/language-server-protocol
+* https://github.com/Microsoft/language-server-protocol/blob/master/protocol.md
+* http://www.jsonrpc.org/specification
+* http://www.json.org
+
+
+### Similar projects ###
+
+* Coq: https://github.com/siegebell/vscoq
+* OCaml: https://github.com/freebroccolo/vscode-reasonml
+* Scala: https://github.com/dragos/dragos-vscode-scala
+* Rust:
+ * https://github.com/jonathandturner/rls
+ * https://github.com/jonathandturner/rls_vscode
+ * https://github.com/RustDT/RustLSP
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/extension/.vscode/launch.json Tue Dec 20 16:18:56 2016 +0100
@@ -0,0 +1,28 @@
+// A launch configuration that compiles the extension and then opens it inside a new window
+{
+ "version": "0.1.0",
+ "configurations": [
+ {
+ "name": "Launch Extension",
+ "type": "extensionHost",
+ "request": "launch",
+ "runtimeExecutable": "${execPath}",
+ "args": ["--extensionDevelopmentPath=${workspaceRoot}" ],
+ "stopOnEntry": false,
+ "sourceMaps": true,
+ "outDir": "${workspaceRoot}/out/src",
+ "preLaunchTask": "npm"
+ },
+ {
+ "name": "Launch Tests",
+ "type": "extensionHost",
+ "request": "launch",
+ "runtimeExecutable": "${execPath}",
+ "args": ["--extensionDevelopmentPath=${workspaceRoot}", "--extensionTestsPath=${workspaceRoot}/out/test" ],
+ "stopOnEntry": false,
+ "sourceMaps": true,
+ "outDir": "${workspaceRoot}/out/test",
+ "preLaunchTask": "npm"
+ }
+ ]
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/extension/.vscode/settings.json Tue Dec 20 16:18:56 2016 +0100
@@ -0,0 +1,10 @@
+// Place your settings in this file to overwrite default and user settings.
+{
+ "files.exclude": {
+ "out": false // set this to true to hide the "out" folder with the compiled JS files
+ },
+ "search.exclude": {
+ "out": true // set this to false to include "out" folder in search results
+ },
+ "typescript.tsdk": "./node_modules/typescript/lib" // we want to use the TS server from our node_modules folder to control its version
+}
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/extension/.vscode/tasks.json Tue Dec 20 16:18:56 2016 +0100
@@ -0,0 +1,30 @@
+// Available variables which can be used inside of strings.
+// ${workspaceRoot}: the root folder of the team
+// ${file}: the current opened file
+// ${fileBasename}: the current opened file's basename
+// ${fileDirname}: the current opened file's dirname
+// ${fileExtname}: the current opened file's extension
+// ${cwd}: the current working directory of the spawned process
+
+// A task runner that calls a custom npm script that compiles the extension.
+{
+ "version": "0.1.0",
+
+ // we want to run npm
+ "command": "npm",
+
+ // the command is a shell script
+ "isShellCommand": true,
+
+ // show the output window only if unrecognized errors occur.
+ "showOutput": "silent",
+
+ // we run the custom script "compile" as defined in package.json
+ "args": ["run", "compile", "--loglevel", "silent"],
+
+ // The tsc compiler is started in watching mode
+ "isWatching": true,
+
+ // use the standard tsc in watch mode problem matcher to find compile problems in the output.
+ "problemMatcher": "$tsc-watch"
+}
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/extension/.vscodeignore Tue Dec 20 16:18:56 2016 +0100
@@ -0,0 +1,9 @@
+.vscode/**
+.vscode-test/**
+out/test/**
+test/**
+src/**
+**/*.map
+.gitignore
+tsconfig.json
+vsc-extension-quickstart.md
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/extension/README.md Tue Dec 20 16:18:56 2016 +0100
@@ -0,0 +1,6 @@
+# Isabelle language support
+
+This extension provides language support for Isabelle.
+
+Make sure that User Settings `isabelle.home` points to the ISABELLE_HOME
+directory.
Binary file src/Tools/VSCode/extension/isabelle.png has changed
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/extension/language-configuration.json Tue Dec 20 16:18:56 2016 +0100
@@ -0,0 +1,16 @@
+{
+ "brackets": [
+ ["(", ")"],
+ ["[", "]"],
+ ["{", "}"],
+ ["<", ">"],
+ ["«", "»"],
+ ["‹", "›"],
+ ["⟨", "⟩"],
+ ["⌈", "⌉"],
+ ["⌊", "⌋"],
+ ["⦇", "⦈"],
+ ["⟦", "⟧"],
+ ["⦃", "⦄"]
+ ]
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/extension/package.json Tue Dec 20 16:18:56 2016 +0100
@@ -0,0 +1,58 @@
+{
+ "name": "isabelle",
+ "displayName": "Isabelle",
+ "description": "Isabelle Theorem Prover",
+ "keywords": [
+ "theorem prover",
+ "formalized mathematics",
+ "mathematical logic",
+ "functional programming",
+ "document preparation"
+ ],
+ "icon": "isabelle.png",
+ "version": "0.0.1",
+ "publisher": "Makarius",
+ "license": "BSD-3-Clause",
+ "repository": { "url": "http://isabelle.in.tum.de/repos/isabelle" },
+ "engines": { "vscode": "^1.5.0" },
+ "categories": ["Languages"],
+ "activationEvents": [
+ "onLanguage:isabelle"
+ ],
+ "main": "./out/src/extension",
+ "contributes": {
+ "languages": [
+ {
+ "id": "isabelle",
+ "aliases": ["Isabelle"],
+ "extensions": [".thy"],
+ "configuration": "./language-configuration.json"
+ }
+ ],
+ "configuration": {
+ "title": "Isabelle",
+ "properties": {
+ "isabelle.home": {
+ "type": "string",
+ "default": "",
+ "description": "ISABELLE_HOME directory"
+ }
+ }
+ }
+ },
+ "scripts": {
+ "vscode:prepublish": "tsc -p ./",
+ "compile": "tsc -watch -p ./",
+ "postinstall": "node ./node_modules/vscode/bin/install"
+ },
+ "devDependencies": {
+ "typescript": "^2.0.3",
+ "vscode": "^1.0.0",
+ "mocha": "^2.3.3",
+ "@types/node": "^6.0.40",
+ "@types/mocha": "^2.2.32"
+ },
+ "dependencies": {
+ "vscode-languageclient": "^2.6.3"
+ }
+}
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/extension/src/extension.ts Tue Dec 20 16:18:56 2016 +0100
@@ -0,0 +1,32 @@
+'use strict';
+
+import * as vscode from 'vscode';
+import * as path from 'path';
+
+import { LanguageClient, LanguageClientOptions, SettingMonitor, ServerOptions, TransportKind }
+ from 'vscode-languageclient';
+
+
+export function activate(context: vscode.ExtensionContext)
+{
+ let isabelle_home = vscode.workspace.getConfiguration("isabelle").get("home");
+
+ let run = {
+ command: path.join(isabelle_home, "bin", "isabelle"),
+ args: ["vscode_server"]
+ };
+ let server_options: ServerOptions =
+ {
+ run: run,
+ debug: {
+ command: run.command,
+ args: run.args.concat(["-L", path.join(context.extensionPath, "protocol.log")]) }
+ };
+ let client_options: LanguageClientOptions = { documentSelector: "isabelle" };
+
+ let disposable =
+ new LanguageClient("Isabelle Language Service", server_options, client_options, false).start();
+ context.subscriptions.push(disposable);
+}
+
+export function deactivate() { }
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/extension/test/extension.test.ts Tue Dec 20 16:18:56 2016 +0100
@@ -0,0 +1,22 @@
+//
+// Note: This example test is leveraging the Mocha test framework.
+// Please refer to their documentation on https://mochajs.org/ for help.
+//
+
+// The module 'assert' provides assertion methods from node
+import * as assert from 'assert';
+
+// You can import and use all API from the 'vscode' module
+// as well as import your extension to test it
+import * as vscode from 'vscode';
+import * as myExtension from '../src/extension';
+
+// Defines a Mocha test suite to group tests of similar kind together
+suite("Extension Tests", () => {
+
+ // Defines a Mocha unit test
+ test("Something 1", () => {
+ assert.equal(-1, [1, 2, 3].indexOf(5));
+ assert.equal(-1, [1, 2, 3].indexOf(0));
+ });
+});
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/extension/test/index.ts Tue Dec 20 16:18:56 2016 +0100
@@ -0,0 +1,22 @@
+//
+// PLEASE DO NOT MODIFY / DELETE UNLESS YOU KNOW WHAT YOU ARE DOING
+//
+// This file is providing the test runner to use when running extension tests.
+// By default the test runner in use is Mocha based.
+//
+// You can provide your own test runner if you want to override it by exporting
+// a function run(testRoot: string, clb: (error:Error) => void) that the extension
+// host can call to run the tests. The test runner is expected to use console.log
+// to report the results back to the caller. When the tests are finished, return
+// a possible error to the callback or null if none.
+
+var testRunner = require('vscode/lib/testrunner');
+
+// You can directly control Mocha options by uncommenting the following lines
+// See https://github.com/mochajs/mocha/wiki/Using-mocha-programmatically#set-options for more info
+testRunner.configure({
+ ui: 'tdd', // the TDD UI is being used in extension.test.ts (suite, test, etc.)
+ useColors: true // colored output from test results
+});
+
+module.exports = testRunner;
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/extension/tsconfig.json Tue Dec 20 16:18:56 2016 +0100
@@ -0,0 +1,16 @@
+{
+ "compilerOptions": {
+ "module": "commonjs",
+ "target": "es6",
+ "outDir": "out",
+ "lib": [
+ "es6"
+ ],
+ "sourceMap": true,
+ "rootDir": "."
+ },
+ "exclude": [
+ "node_modules",
+ ".vscode-test"
+ ]
+}
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/extension/vsc-extension-quickstart.md Tue Dec 20 16:18:56 2016 +0100
@@ -0,0 +1,33 @@
+# Welcome to your first VS Code Extension
+
+## What's in the folder
+* This folder contains all of the files necessary for your extension
+* `package.json` - this is the manifest file in which you declare your extension and command.
+The sample plugin registers a command and defines its title and command name. With this information
+VS Code can show the command in the command palette. It doesn’t yet need to load the plugin.
+* `src/extension.ts` - this is the main file where you will provide the implementation of your command.
+The file exports one function, `activate`, which is called the very first time your extension is
+activated (in this case by executing the command). Inside the `activate` function we call `registerCommand`.
+We pass the function containing the implementation of the command as the second parameter to
+`registerCommand`.
+
+## Get up and running straight away
+* press `F5` to open a new window with your extension loaded
+* run your command from the command palette by pressing (`Ctrl+Shift+P` or `Cmd+Shift+P` on Mac) and typing `Hello World`
+* set breakpoints in your code inside `src/extension.ts` to debug your extension
+* find output from your extension in the debug console
+
+## Make changes
+* you can relaunch the extension from the debug toolbar after changing code in `src/extension.ts`
+* you can also reload (`Ctrl+R` or `Cmd+R` on Mac) the VS Code window with your extension to load your changes
+
+## Explore the API
+* you can open the full set of our API when you open the file `node_modules/vscode/vscode.d.ts`
+
+## Run tests
+* open the debug viewlet (`Ctrl+Shift+D` or `Cmd+Shift+D` on Mac) and from the launch configuration dropdown pick `Launch Tests`
+* press `F5` to run the tests in a new window with your extension loaded
+* see the output of the test result in the debug console
+* make changes to `test/extension.test.ts` or create new test files inside the `test` folder
+ * by convention, the test runner will only consider files matching the name pattern `**.test.ts`
+ * you can create folders inside the `test` folder to structure your tests any way you want
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/src/channel.scala Tue Dec 20 16:18:56 2016 +0100
@@ -0,0 +1,89 @@
+/* Title: Tools/VSCode/src/channel.scala
+ Author: Makarius
+
+Channel with JSON RPC protocol.
+*/
+
+package isabelle.vscode
+
+
+import isabelle._
+
+import java.io.{InputStream, OutputStream, FileOutputStream, ByteArrayOutputStream}
+
+import scala.collection.mutable
+
+
+class Channel(in: InputStream, out: OutputStream, log_file: Option[Path] = None)
+{
+ val log: Logger = Logger.make(log_file)
+
+
+ /* read message */
+
+ private val Content_Length = """^\s*Content-Length:\s*(\d+)\s*$""".r
+
+ private def read_line(): String =
+ {
+ val buffer = new ByteArrayOutputStream(100)
+ var c = 0
+ while ({ c = in.read; c != -1 && c != 10 }) buffer.write(c)
+ Library.trim_line(buffer.toString(UTF8.charset_name))
+ }
+
+ private def read_header(): List[String] =
+ {
+ val header = new mutable.ListBuffer[String]
+ var line = ""
+ while ({ line = read_line(); line != "" }) header += line
+ header.toList
+ }
+
+ private def read_content(n: Int): String =
+ {
+ val buffer = new Array[Byte](n)
+ var i = 0
+ var m = 0
+ do {
+ m = in.read(buffer, i, n - i)
+ if (m != -1) i += m
+ }
+ while (m != -1 && n > i)
+
+ if (i == n) new String(buffer, UTF8.charset)
+ else error("Bad message content (unexpected EOF after " + i + " of " + n + " bytes)")
+ }
+
+ def read(): Option[JSON.T] =
+ {
+ read_header() match {
+ case Nil => None
+ case Content_Length(s) :: _ =>
+ s match {
+ case Value.Int(n) if n >= 0 =>
+ val msg = read_content(n)
+ log("IN:\n" + msg)
+ Some(JSON.parse(msg))
+ case _ => error("Bad Content-Length: " + s)
+ }
+ case header => error(cat_lines("Malformed header:" :: header))
+ }
+ }
+
+
+ /* write message */
+
+ def write(json: JSON.T)
+ {
+ val msg = JSON.Format(json)
+ log("OUT:\n" + msg)
+
+ val content = UTF8.bytes(msg)
+ val header = UTF8.bytes("Content-Length: " + content.length + "\r\n\r\n")
+ out.synchronized {
+ out.write(header)
+ out.write(content)
+ out.flush
+ }
+ }
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/src/document_model.scala Tue Dec 20 16:18:56 2016 +0100
@@ -0,0 +1,58 @@
+/* Title: Tools/VSCode/src/document_model.scala
+ Author: Makarius
+
+Document model for line-oriented text.
+*/
+
+package isabelle.vscode
+
+
+import isabelle._
+
+import scala.util.parsing.input.CharSequenceReader
+
+
+case class Document_Model(
+ session: Session, doc: Line.Document, node_name: Document.Node.Name,
+ changed: Boolean = true)
+{
+ /* header */
+
+ def is_theory: Boolean = node_name.is_theory
+
+ lazy val node_header: Document.Node.Header =
+ if (is_theory) {
+ val toks = Token.explode(Thy_Header.bootstrap_syntax.keywords, doc.text)
+ val toks1 = toks.dropWhile(tok => !tok.is_command(Thy_Header.THEORY))
+ toks1.iterator.zipWithIndex.collectFirst({ case (tok, i) if tok.is_begin => i }) match {
+ case Some(i) =>
+ session.resources.check_thy_reader("", node_name,
+ new CharSequenceReader(toks1.take(i + 1).map(_.source).mkString), Token.Pos.command)
+ case None => Document.Node.no_header
+ }
+ }
+ else Document.Node.no_header
+
+
+ /* edits */
+
+ def text_edits: List[Text.Edit] =
+ if (changed) List(Text.Edit.insert(0, doc.text)) else Nil
+
+ def node_edits: List[Document.Edit_Text] =
+ if (changed) {
+ List(session.header_edit(node_name, node_header),
+ node_name -> Document.Node.Clear(),
+ node_name -> Document.Node.Edits(text_edits),
+ node_name ->
+ Document.Node.Perspective(true, Text.Perspective.full, Document.Node.Overlays.empty))
+ }
+ else Nil
+
+ def unchanged: Document_Model = if (changed) copy(changed = false) else this
+
+
+ /* snapshot */
+
+ def snapshot(): Document.Snapshot = session.snapshot(node_name, text_edits)
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/src/line.scala Tue Dec 20 16:18:56 2016 +0100
@@ -0,0 +1,151 @@
+/* Title: Tools/VSCode/src/line.scala
+ Author: Makarius
+
+Line-oriented text.
+*/
+
+package isabelle.vscode
+
+
+import isabelle._
+
+import scala.annotation.tailrec
+
+
+object Line
+{
+ /* position */
+
+ object Position
+ {
+ val zero: Position = Position(0, 0)
+ }
+
+ sealed case class Position(line: Int, column: Int)
+ {
+ def line1: Int = line + 1
+ def column1: Int = column + 1
+ def print: String = line1.toString + ":" + column1.toString
+
+ def compare(that: Position): Int =
+ line compare that.line match {
+ case 0 => column compare that.column
+ case i => i
+ }
+
+ private def advance[A](iterator: Iterator[A], is_newline: A => Boolean): Position =
+ {
+ var l = line
+ var c = column
+ for (x <- iterator) {
+ if (is_newline(x)) { l += 1; c = 0 } else c += 1
+ }
+ Position(l, c)
+ }
+
+ def advance(text: String): Position =
+ if (text.isEmpty) this else advance[Char](text.iterator, _ == '\n')
+
+ def advance_symbols(text: String): Position =
+ if (text.isEmpty) this else advance[Symbol.Symbol](Symbol.iterator(text), Symbol.is_newline _)
+
+ def advance_codepoints(text: String): Position =
+ if (text.isEmpty) this else advance[Int](Word.codepoint_iterator(text), _ == 10)
+ }
+
+
+ /* range (right-open interval) */
+
+ sealed case class Range(start: Position, stop: Position)
+ {
+ if (start.compare(stop) > 0)
+ error("Bad line range: " + start.print + ".." + stop.print)
+
+ def print: String = start.print + ".." + stop.print
+ }
+
+
+ /* document with newline as separator (not terminator) */
+
+ object Document
+ {
+ val empty: Document = new Document("", Nil)
+
+ def apply(lines: List[Line]): Document =
+ if (lines.isEmpty) empty
+ else new Document(lines.mkString("", "\n", ""), lines)
+
+ def apply(text: String): Document =
+ if (text.contains('\r'))
+ apply(Library.cat_lines(Library.split_lines(text).map(Library.trim_line(_))))
+ else if (text == "") Document.empty
+ else new Document(text, Library.split_lines(text).map(Line(_)))
+ }
+
+ final class Document private(val text: String, val lines: List[Line])
+ {
+ override def toString: String = text
+
+ override def equals(that: Any): Boolean =
+ that match {
+ case other: Document => lines == other.lines
+ case _ => false
+ }
+ override def hashCode(): Int = lines.hashCode
+
+ private def position(advance: (Position, String) => Position, offset: Text.Offset): Position =
+ {
+ @tailrec def move(i: Text.Offset, lines_count: Int, lines_rest: List[Line]): Position =
+ {
+ lines_rest match {
+ case Nil => require(i == 0); Position(lines_count, 0)
+ case line :: ls =>
+ val n = line.text.length
+ if (ls.isEmpty || i <= n) advance(Position(lines_count, 0), line.text.substring(n - i))
+ else move(i - (n + 1), lines_count + 1, ls)
+ }
+ }
+ move(offset, 0, lines)
+ }
+
+ def position(i: Text.Offset): Position = position(_.advance(_), i)
+ def position_symbols(i: Text.Offset): Position = position(_.advance_symbols(_), i)
+ def position_codepoints(i: Text.Offset): Position = position(_.advance_codepoints(_), i)
+
+ // FIXME symbols, codepoints
+ def offset(pos: Position): Option[Text.Offset] =
+ {
+ val l = pos.line
+ val c = pos.column
+ if (0 <= l && l < lines.length) {
+ val line_offset =
+ if (l == 0) 0
+ else (0 /: lines.take(l - 1))({ case (n, line) => n + line.text.length + 1 })
+ if (c <= lines(l).text.length) Some(line_offset + c) else None
+ }
+ else None
+ }
+ }
+
+
+ /* line text */
+
+ val empty: Line = new Line("")
+ def apply(text: String): Line = if (text == "") empty else new Line(text)
+}
+
+final class Line private(val text: String)
+{
+ require(text.forall(c => c != '\r' && c != '\n'))
+
+ lazy val length_symbols: Int = Symbol.iterator(text).length
+ lazy val length_codepoints: Int = Word.codepoint_iterator(text).length
+
+ override def equals(that: Any): Boolean =
+ that match {
+ case other: Line => text == other.text
+ case _ => false
+ }
+ override def hashCode(): Int = text.hashCode
+ override def toString: String = text
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/src/protocol.scala Tue Dec 20 16:18:56 2016 +0100
@@ -0,0 +1,299 @@
+/* Title: Tools/VSCode/src/protocol.scala
+ Author: Makarius
+
+Message formats for Language Server Protocol 2.0, see
+https://github.com/Microsoft/language-server-protocol/blob/master/protocol.md
+*/
+
+package isabelle.vscode
+
+
+import isabelle._
+
+
+object Protocol
+{
+ /* abstract message */
+
+ object Message
+ {
+ val empty: Map[String, JSON.T] = Map("jsonrpc" -> "2.0")
+ }
+
+
+ /* notification */
+
+ object Notification
+ {
+ def apply(method: String, params: JSON.T): JSON.T =
+ Message.empty + ("method" -> method) + ("params" -> params)
+
+ def unapply(json: JSON.T): Option[(String, Option[JSON.T])] =
+ for {
+ method <- JSON.string(json, "method")
+ params = JSON.value(json, "params")
+ } yield (method, params)
+ }
+
+ class Notification0(name: String)
+ {
+ def unapply(json: JSON.T): Option[Unit] =
+ json match {
+ case Notification(method, _) if method == name => Some(())
+ case _ => None
+ }
+ }
+
+
+ /* request message */
+
+ sealed case class Id(id: Any)
+ {
+ require(
+ id.isInstanceOf[Int] ||
+ id.isInstanceOf[Long] ||
+ id.isInstanceOf[Double] ||
+ id.isInstanceOf[String])
+
+ override def toString: String = id.toString
+ }
+
+ object RequestMessage
+ {
+ def apply(id: Id, method: String, params: JSON.T): JSON.T =
+ Message.empty + ("id" -> id.id) + ("method" -> method) + ("params" -> params)
+
+ def unapply(json: JSON.T): Option[(Id, String, Option[JSON.T])] =
+ for {
+ id <- JSON.long(json, "id") orElse JSON.double(json, "id") orElse JSON.string(json, "id")
+ method <- JSON.string(json, "method")
+ params = JSON.value(json, "params")
+ } yield (Id(id), method, params)
+ }
+
+ class Request0(name: String)
+ {
+ def unapply(json: JSON.T): Option[Id] =
+ json match {
+ case RequestMessage(id, method, _) if method == name => Some(id)
+ case _ => None
+ }
+ }
+
+
+ /* response message */
+
+ object ResponseMessage
+ {
+ def apply(id: Id, result: Option[JSON.T] = None, error: Option[ResponseError] = None): JSON.T =
+ Message.empty + ("id" -> id.id) ++
+ (result match { case Some(x) => Map("result" -> x) case None => Map.empty }) ++
+ (error match { case Some(x) => Map("error" -> x.json) case None => Map.empty })
+
+ def strict(id: Id, result: Option[JSON.T] = None, error: String = ""): JSON.T =
+ if (error == "") apply(id, result = result)
+ else apply(id, error = Some(ResponseError(code = ErrorCodes.serverErrorEnd, message = error)))
+ }
+
+ sealed case class ResponseError(code: Int, message: String, data: Option[JSON.T] = None)
+ {
+ def json: JSON.T =
+ Map("code" -> code, "message" -> message) ++
+ (data match { case Some(x) => Map("data" -> x) case None => Map.empty })
+ }
+
+ object ErrorCodes
+ {
+ val ParseError = -32700
+ val InvalidRequest = -32600
+ val MethodNotFound = -32601
+ val InvalidParams = -32602
+ val InternalError = -32603
+ val serverErrorStart = -32099
+ val serverErrorEnd = -32000
+ }
+
+
+ /* init and exit */
+
+ object Initialize extends Request0("initialize")
+ {
+ def reply(id: Id, error: String): JSON.T =
+ ResponseMessage.strict(id, Some(Map("capabilities" -> ServerCapabilities.json)), error)
+ }
+
+ object ServerCapabilities
+ {
+ val json: JSON.T =
+ Map("textDocumentSync" -> 1, "hoverProvider" -> true)
+ }
+
+ object Shutdown extends Request0("shutdown")
+ {
+ def reply(id: Id, error: String): JSON.T =
+ ResponseMessage.strict(id, Some("OK"), error)
+ }
+
+ object Exit extends Notification0("exit")
+
+
+ /* document positions */
+
+ object Position
+ {
+ def apply(pos: Line.Position): JSON.T =
+ Map("line" -> pos.line, "character" -> pos.column)
+
+ def unapply(json: JSON.T): Option[Line.Position] =
+ for {
+ line <- JSON.int(json, "line")
+ column <- JSON.int(json, "character")
+ } yield Line.Position(line, column)
+ }
+
+ object Range
+ {
+ def apply(range: Line.Range): JSON.T =
+ Map("start" -> Position(range.start), "end" -> Position(range.stop))
+
+ def unapply(json: JSON.T): Option[Line.Range] =
+ (JSON.value(json, "start"), JSON.value(json, "end")) match {
+ case (Some(Position(start)), Some(Position(stop))) => Some(Line.Range(start, stop))
+ case _ => None
+ }
+ }
+
+ object Location
+ {
+ def apply(uri: String, range: Line.Range): JSON.T =
+ Map("uri" -> uri, "range" -> Range(range))
+
+ def unapply(json: JSON.T): Option[(String, Line.Range)] =
+ for {
+ uri <- JSON.string(json, "uri")
+ range_json <- JSON.value(json, "range")
+ range <- Range.unapply(range_json)
+ } yield (uri, range)
+ }
+
+ object TextDocumentPosition
+ {
+ def unapply(json: JSON.T): Option[(String, Line.Position)] =
+ for {
+ doc <- JSON.value(json, "textDocument")
+ uri <- JSON.string(doc, "uri")
+ pos_json <- JSON.value(json, "position")
+ pos <- Position.unapply(pos_json)
+ } yield (uri, pos)
+ }
+
+
+ /* diagnostic messages */
+
+ object MessageType
+ {
+ val Error = 1
+ val Warning = 2
+ val Info = 3
+ val Log = 4
+ }
+
+ object DisplayMessage
+ {
+ def apply(message_type: Int, message: String, show: Boolean = true): JSON.T =
+ Notification(if (show) "window/showMessage" else "window/logMessage",
+ Map("type" -> message_type, "message" -> message))
+ }
+
+
+ /* document edits */
+
+ object DidOpenTextDocument
+ {
+ def unapply(json: JSON.T): Option[(String, String, Long, String)] =
+ json match {
+ case Notification("textDocument/didOpen", Some(params)) =>
+ for {
+ doc <- JSON.value(params, "textDocument")
+ uri <- JSON.string(doc, "uri")
+ lang <- JSON.string(doc, "languageId")
+ version <- JSON.long(doc, "version")
+ text <- JSON.string(doc, "text")
+ } yield (uri, lang, version, text)
+ case _ => None
+ }
+ }
+
+
+ sealed abstract class TextDocumentContentChange
+ case class TextDocumentContent(text: String) extends TextDocumentContentChange
+ case class TextDocumentChange(range: Line.Range, range_length: Int, text: String)
+ extends TextDocumentContentChange
+
+ object TextDocumentContentChange
+ {
+ def unapply(json: JSON.T): Option[TextDocumentContentChange] =
+ for { text <- JSON.string(json, "text") }
+ yield {
+ (JSON.value(json, "range"), JSON.int(json, "rangeLength")) match {
+ case (Some(Range(range)), Some(range_length)) =>
+ TextDocumentChange(range, range_length, text)
+ case _ => TextDocumentContent(text)
+ }
+ }
+ }
+
+ object DidChangeTextDocument
+ {
+ def unapply(json: JSON.T): Option[(String, Long, List[TextDocumentContentChange])] =
+ json match {
+ case Notification("textDocument/didChange", Some(params)) =>
+ for {
+ doc <- JSON.value(params, "textDocument")
+ uri <- JSON.string(doc, "uri")
+ version <- JSON.long(doc, "version")
+ changes <- JSON.array(params, "contentChanges", TextDocumentContentChange.unapply _)
+ } yield (uri, version, changes)
+ case _ => None
+ }
+ }
+
+ class TextDocumentNotification(name: String)
+ {
+ def unapply(json: JSON.T): Option[String] =
+ json match {
+ case Notification(method, Some(params)) if method == name =>
+ for {
+ doc <- JSON.value(params, "textDocument")
+ uri <- JSON.string(doc, "uri")
+ } yield uri
+ case _ => None
+ }
+ }
+
+ object DidCloseTextDocument extends TextDocumentNotification("textDocument/didClose")
+ object DidSaveTextDocument extends TextDocumentNotification("textDocument/didSave")
+
+
+ /* hover request */
+
+ object Hover
+ {
+ def unapply(json: JSON.T): Option[(Id, String, Line.Position)] =
+ json match {
+ case RequestMessage(id, "textDocument/hover", Some(TextDocumentPosition(uri, pos))) =>
+ Some((id, uri, pos))
+ case _ => None
+ }
+
+ def reply(id: Id, result: Option[(Line.Range, List[String])]): JSON.T =
+ {
+ val res =
+ result match {
+ case Some((range, contents)) => Map("contents" -> contents, "range" -> Range(range))
+ case None => Map("contents" -> Nil)
+ }
+ ResponseMessage(id, Some(res))
+ }
+ }
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/src/server.scala Tue Dec 20 16:18:56 2016 +0100
@@ -0,0 +1,354 @@
+/* Title: Tools/VSCode/src/server.scala
+ Author: Makarius
+
+Server for VS Code Language Server Protocol 2.0, see also
+https://github.com/Microsoft/language-server-protocol
+https://github.com/Microsoft/language-server-protocol/blob/master/protocol.md
+*/
+
+package isabelle.vscode
+
+
+import isabelle._
+
+import java.io.{PrintStream, OutputStream}
+
+import scala.annotation.tailrec
+
+
+object Server
+{
+ /* Isabelle tool wrapper */
+
+ private lazy val default_logic = Isabelle_System.getenv("ISABELLE_LOGIC")
+
+ val isabelle_tool = Isabelle_Tool("vscode_server", "VSCode Language Server for PIDE", args =>
+ {
+ var log_file: Option[Path] = None
+ var dirs: List[Path] = Nil
+ var logic = default_logic
+ var modes: List[String] = Nil
+ var options = Options.init()
+
+ val getopts = Getopts("""
+Usage: isabelle vscode_server [OPTIONS]
+
+ Options are:
+ -L FILE enable logging on FILE
+ -d DIR include session directory
+ -l NAME logic session name (default ISABELLE_LOGIC=""" + quote(logic) + """)
+ -m MODE add print mode for output
+ -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME)
+
+ Run the VSCode Language Server protocol (JSON RPC) over stdin/stdout.
+""",
+ "L:" -> (arg => log_file = Some(Path.explode(arg))),
+ "d:" -> (arg => dirs = dirs ::: List(Path.explode(arg))),
+ "l:" -> (arg => logic = arg),
+ "m:" -> (arg => modes = arg :: modes),
+ "o:" -> (arg => options = options + arg))
+
+ val more_args = getopts(args)
+ if (more_args.nonEmpty) getopts.usage()
+
+ if (!Build.build(options = options, build_heap = true, no_build = true,
+ dirs = dirs, sessions = List(logic)).ok)
+ error("Missing logic image " + quote(logic))
+
+ val channel = new Channel(System.in, System.out, log_file)
+ val server = new Server(channel, options, logic, dirs, modes)
+
+ // prevent spurious garbage on the main protocol channel
+ val orig_out = System.out
+ try {
+ System.setOut(new PrintStream(new OutputStream { def write(n: Int) {} }))
+ server.start()
+ }
+ finally { System.setOut(orig_out) }
+ })
+
+
+ /* server state */
+
+ sealed case class State(
+ session: Option[Session] = None,
+ models: Map[String, Document_Model] = Map.empty)
+}
+
+class Server(
+ channel: Channel,
+ options: Options,
+ session_name: String = Server.default_logic,
+ session_dirs: List[Path] = Nil,
+ modes: List[String] = Nil)
+{
+ /* server state */
+
+ private val state = Synchronized(Server.State())
+
+ def session: Session = state.value.session getOrElse error("Session inactive")
+ def resources: URI_Resources = session.resources.asInstanceOf[URI_Resources]
+
+
+ /* init and exit */
+
+ def initialize(reply: String => Unit)
+ {
+ val content = Build.session_content(options, false, session_dirs, session_name)
+ val resources =
+ new URI_Resources(content.loaded_theories, content.known_theories, content.syntax)
+
+ val session =
+ new Session(resources) {
+ override def output_delay = options.seconds("editor_output_delay")
+ override def prune_delay = options.seconds("editor_prune_delay")
+ override def syslog_limit = options.int("editor_syslog_limit")
+ override def reparse_limit = options.int("editor_reparse_limit")
+ }
+
+ var session_phase: Session.Consumer[Session.Phase] = null
+ session_phase =
+ Session.Consumer(getClass.getName) {
+ case Session.Ready =>
+ session.phase_changed -= session_phase
+ session.update_options(options)
+ reply("")
+ case Session.Failed =>
+ session.phase_changed -= session_phase
+ reply("Prover startup failed")
+ case _ =>
+ }
+ session.phase_changed += session_phase
+
+ session.start(receiver =>
+ Isabelle_Process(options = options, logic = session_name, dirs = session_dirs,
+ modes = modes, receiver = receiver))
+
+ state.change(_.copy(session = Some(session)))
+ }
+
+ def shutdown(reply: String => Unit)
+ {
+ state.change(st =>
+ st.session match {
+ case None => reply("Prover inactive"); st
+ case Some(session) =>
+ var session_phase: Session.Consumer[Session.Phase] = null
+ session_phase =
+ Session.Consumer(getClass.getName) {
+ case Session.Inactive =>
+ session.phase_changed -= session_phase
+ reply("")
+ case _ =>
+ }
+ session.phase_changed += session_phase
+ session.stop()
+ st.copy(session = None)
+ })
+ }
+
+ def exit() { sys.exit(if (state.value.session.isDefined) 1 else 0) }
+
+
+ /* document management */
+
+ private val delay_flush =
+ Standard_Thread.delay_last(options.seconds("editor_input_delay")) {
+ state.change(st =>
+ {
+ val models = st.models
+ val changed = (for { entry <- models.iterator if entry._2.changed } yield entry).toList
+ val edits = for { (_, model) <- changed; edit <- model.node_edits } yield edit
+ val models1 =
+ (models /: changed)({ case (m, (uri, model)) => m + (uri -> model.unchanged) })
+
+ session.update(Document.Blobs.empty, edits)
+ st.copy(models = models1)
+ })
+ }
+
+ def update_document(uri: String, text: String)
+ {
+ state.change(st =>
+ {
+ val node_name = resources.node_name(uri)
+ val model = Document_Model(session, Line.Document(text), node_name)
+ st.copy(models = st.models + (uri -> model))
+ })
+ delay_flush.invoke()
+ }
+
+
+ /* tooltips -- see $ISABELLE_HOME/src/Tools/jEdit/rendering.scala */
+
+ def hover(uri: String, pos: Line.Position): Option[(Line.Range, List[String])] =
+ for {
+ model <- state.value.models.get(uri)
+ snapshot = model.snapshot()
+ offset <- model.doc.offset(pos)
+ info <- tooltip(snapshot, Text.Range(offset, offset + 1))
+ } yield {
+ val r = Line.Range(model.doc.position(info.range.start), model.doc.position(info.range.stop))
+ val s = Pretty.string_of(info.info, margin = tooltip_margin.toDouble)
+ (r, List("```\n" + s + "\n```"))
+ }
+
+ private val tooltip_descriptions =
+ Map(
+ Markup.CITATION -> "citation",
+ Markup.TOKEN_RANGE -> "inner syntax token",
+ Markup.FREE -> "free variable",
+ Markup.SKOLEM -> "skolem variable",
+ Markup.BOUND -> "bound variable",
+ Markup.VAR -> "schematic variable",
+ Markup.TFREE -> "free type variable",
+ Markup.TVAR -> "schematic type variable")
+
+ private val tooltip_elements =
+ Markup.Elements(Markup.LANGUAGE, Markup.EXPRESSION, Markup.TIMING, Markup.ENTITY,
+ Markup.SORTING, Markup.TYPING, Markup.CLASS_PARAMETER, Markup.ML_TYPING,
+ Markup.ML_BREAKPOINT, Markup.PATH, Markup.DOC, Markup.URL, Markup.MARKDOWN_PARAGRAPH,
+ Markup.Markdown_List.name) ++ Markup.Elements(tooltip_descriptions.keySet)
+
+ private def pretty_typing(kind: String, body: XML.Body): XML.Tree =
+ Pretty.block(XML.Text(kind) :: Pretty.brk(1) :: body)
+
+ private def timing_threshold: Double = options.real("jedit_timing_threshold")
+
+ def tooltip(snapshot: Document.Snapshot, range: Text.Range): Option[Text.Info[XML.Body]] =
+ {
+ def add(prev: Text.Info[(Timing, Vector[(Boolean, XML.Tree)])],
+ r0: Text.Range, p: (Boolean, XML.Tree)): Text.Info[(Timing, Vector[(Boolean, XML.Tree)])] =
+ {
+ val r = snapshot.convert(r0)
+ val (t, info) = prev.info
+ if (prev.range == r)
+ Text.Info(r, (t, info :+ p))
+ else Text.Info(r, (t, Vector(p)))
+ }
+
+ val results =
+ snapshot.cumulate[Text.Info[(Timing, Vector[(Boolean, XML.Tree)])]](
+ range, Text.Info(range, (Timing.zero, Vector.empty)), tooltip_elements, _ =>
+ {
+ case (Text.Info(r, (t1, info)), Text.Info(_, XML.Elem(Markup.Timing(t2), _))) =>
+ Some(Text.Info(r, (t1 + t2, info)))
+
+ case (prev, Text.Info(r, XML.Elem(Markup.Entity(kind, name), _)))
+ if kind != "" &&
+ kind != Markup.ML_DEF &&
+ kind != Markup.ML_OPEN &&
+ kind != Markup.ML_STRUCTURE =>
+ val kind1 = Word.implode(Word.explode('_', kind))
+ val txt1 =
+ if (name == "") kind1
+ else if (kind1 == "") quote(name)
+ else kind1 + " " + quote(name)
+ val t = prev.info._1
+ val txt2 =
+ if (kind == Markup.COMMAND && t.elapsed.seconds >= timing_threshold)
+ "\n" + t.message
+ else ""
+ Some(add(prev, r, (true, XML.Text(txt1 + txt2))))
+
+ case (prev, Text.Info(r, XML.Elem(Markup.Doc(name), _))) =>
+ val text = "doc " + quote(name)
+ Some(add(prev, r, (true, XML.Text(text))))
+
+ case (prev, Text.Info(r, XML.Elem(Markup.Url(name), _))) =>
+ Some(add(prev, r, (true, XML.Text("URL " + quote(name)))))
+
+ case (prev, Text.Info(r, XML.Elem(Markup(name, _), body)))
+ if name == Markup.SORTING || name == Markup.TYPING =>
+ Some(add(prev, r, (true, pretty_typing("::", body))))
+
+ case (prev, Text.Info(r, XML.Elem(Markup(Markup.CLASS_PARAMETER, _), body))) =>
+ Some(add(prev, r, (true, Pretty.block(0, body))))
+
+ case (prev, Text.Info(r, XML.Elem(Markup(Markup.ML_TYPING, _), body))) =>
+ Some(add(prev, r, (false, pretty_typing("ML:", body))))
+
+ case (prev, Text.Info(r, XML.Elem(Markup.Language(language, _, _, _), _))) =>
+ val lang = Word.implode(Word.explode('_', language))
+ Some(add(prev, r, (true, XML.Text("language: " + lang))))
+
+ case (prev, Text.Info(r, XML.Elem(Markup.Expression(kind), _))) =>
+ val descr = if (kind == "") "expression" else "expression: " + kind
+ Some(add(prev, r, (true, XML.Text(descr))))
+
+ case (prev, Text.Info(r, XML.Elem(Markup(Markup.MARKDOWN_PARAGRAPH, _), _))) =>
+ Some(add(prev, r, (true, XML.Text("Markdown: paragraph"))))
+ case (prev, Text.Info(r, XML.Elem(Markup.Markdown_List(kind), _))) =>
+ Some(add(prev, r, (true, XML.Text("Markdown: " + kind))))
+
+ case (prev, Text.Info(r, XML.Elem(Markup(name, _), _))) =>
+ tooltip_descriptions.get(name).
+ map(descr => add(prev, r, (true, XML.Text(descr))))
+ }).map(_.info)
+
+ results.map(_.info).flatMap(res => res._2.toList) match {
+ case Nil => None
+ case tips =>
+ val r = Text.Range(results.head.range.start, results.last.range.stop)
+ val all_tips = (tips.filter(_._1) ++ tips.filter(!_._1).lastOption.toList).map(_._2)
+ Some(Text.Info(r, Library.separate(Pretty.fbrk, all_tips)))
+ }
+ }
+
+ def tooltip_margin: Int = options.int("jedit_tooltip_margin")
+
+
+ /* main loop */
+
+ def start()
+ {
+ channel.log("\nServer started " + Date.now())
+
+ def handle(json: JSON.T)
+ {
+ try {
+ json match {
+ case Protocol.Initialize(id) =>
+ def initialize_reply(err: String)
+ {
+ channel.write(Protocol.Initialize.reply(id, err))
+ if (err == "") {
+ channel.write(Protocol.DisplayMessage(Protocol.MessageType.Info,
+ "Welcome to Isabelle/" + session_name + " (" + Distribution.version + ")"))
+ }
+ else channel.write(Protocol.DisplayMessage(Protocol.MessageType.Error, err))
+ }
+ initialize(initialize_reply _)
+ case Protocol.Shutdown(id) =>
+ shutdown((error: String) => channel.write(Protocol.Shutdown.reply(id, error)))
+ case Protocol.Exit =>
+ exit()
+ case Protocol.DidOpenTextDocument(uri, lang, version, text) =>
+ update_document(uri, text)
+ case Protocol.DidChangeTextDocument(uri, version, List(Protocol.TextDocumentContent(text))) =>
+ update_document(uri, text)
+ case Protocol.DidCloseTextDocument(uri) => channel.log("CLOSE " + uri)
+ case Protocol.DidSaveTextDocument(uri) => channel.log("SAVE " + uri)
+ case Protocol.Hover(id, uri, pos) =>
+ channel.write(Protocol.Hover.reply(id, hover(uri, pos)))
+ case _ => channel.log("### IGNORED")
+ }
+ }
+ catch { case exn: Throwable => channel.log("*** ERROR: " + Exn.message(exn)) }
+ }
+
+ @tailrec def loop()
+ {
+ channel.read() match {
+ case Some(json) =>
+ json match {
+ case bulk: List[_] => bulk.foreach(handle(_))
+ case _ => handle(json)
+ }
+ loop()
+ case None => channel.log("### TERMINATE")
+ }
+ }
+ loop()
+ }
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/src/uri_resources.scala Tue Dec 20 16:18:56 2016 +0100
@@ -0,0 +1,43 @@
+/* Title: Tools/VSCode/src/uri_resources.scala
+ Author: Makarius
+
+Resources based on file-system URIs.
+*/
+
+package isabelle.vscode
+
+
+import isabelle._
+
+import java.net.{URI, URISyntaxException}
+import java.io.{File => JFile}
+
+
+object URI_Resources
+{
+ def is_wellformed(uri: String): Boolean =
+ try { new JFile(new URI(uri)); true }
+ catch { case _: URISyntaxException | _: IllegalArgumentException => false }
+
+ def canonical_file(uri: String): JFile =
+ new JFile(new URI(uri)).getCanonicalFile
+
+ val empty: URI_Resources =
+ new URI_Resources(Set.empty, Map.empty, Outer_Syntax.empty)
+}
+
+class URI_Resources(
+ loaded_theories: Set[String],
+ known_theories: Map[String, Document.Node.Name],
+ base_syntax: Outer_Syntax)
+ extends Resources(loaded_theories, known_theories, base_syntax)
+{
+ def node_name(uri: String): Document.Node.Name =
+ {
+ val file = URI_Resources.canonical_file(uri) // FIXME wellformed!?
+ val node = file.getPath
+ val theory = Thy_Header.thy_name_bootstrap(node).getOrElse("")
+ val master_dir = if (theory == "") "" else file.getParent
+ Document.Node.Name(node, master_dir, theory)
+ }
+}
--- a/src/Tools/jEdit/lib/Tools/jedit Tue Dec 20 16:17:13 2016 +0100
+++ b/src/Tools/jEdit/lib/Tools/jedit Tue Dec 20 16:18:56 2016 +0100
@@ -98,6 +98,7 @@
echo " Options are:"
echo " -D NAME=X set JVM system property"
echo " -J OPTION add JVM runtime option (default $JEDIT_JAVA_SYSTEM_OPTIONS $JEDIT_JAVA_OPTIONS)"
+ echo " -R open ROOT entry of logic session and use its parent"
echo " -b build only"
echo " -d DIR include session directory"
echo " -f fresh build"
@@ -135,6 +136,7 @@
BUILD_JARS="jars"
ML_PROCESS_POLICY=""
JEDIT_SESSION_DIRS=""
+JEDIT_LOGIC_ROOT=""
JEDIT_LOGIC=""
JEDIT_PRINT_MODE=""
JEDIT_BUILD_MODE="normal"
@@ -142,7 +144,7 @@
function getoptions()
{
OPTIND=1
- while getopts "D:J:bd:fj:l:m:np:s" OPT
+ while getopts "D:J:Rbd:fj:l:m:np:s" OPT
do
case "$OPT" in
D)
@@ -151,6 +153,9 @@
J)
JAVA_ARGS["${#JAVA_ARGS[@]}"]="$OPTARG"
;;
+ R)
+ JEDIT_LOGIC_ROOT="true"
+ ;;
b)
BUILD_ONLY=true
;;
@@ -365,7 +370,7 @@
if [ "$BUILD_ONLY" = false ]
then
- export JEDIT_SESSION_DIRS JEDIT_LOGIC JEDIT_PRINT_MODE JEDIT_BUILD_MODE
+ export JEDIT_SESSION_DIRS JEDIT_LOGIC JEDIT_LOGIC_ROOT JEDIT_PRINT_MODE JEDIT_BUILD_MODE
export JEDIT_ML_PROCESS_POLICY="$ML_PROCESS_POLICY"
classpath "$JEDIT_HOME/dist/jedit.jar"
exec isabelle java "${JAVA_ARGS[@]}" isabelle.Main "${ARGS[@]}"
--- a/src/Tools/jEdit/src/jedit_sessions.scala Tue Dec 20 16:17:13 2016 +0100
+++ b/src/Tools/jEdit/src/jedit_sessions.scala Tue Dec 20 16:18:56 2016 +0100
@@ -19,10 +19,7 @@
private val option_name = "jedit_logic"
- def session_name(): String =
- Isabelle_System.default_logic(
- Isabelle_System.getenv("JEDIT_LOGIC"),
- PIDE.options.string(option_name))
+ sealed case class Info(name: String, open_root: Position.T)
def session_dirs(): List[Path] = Path.split(Isabelle_System.getenv("JEDIT_SESSION_DIRS"))
@@ -32,6 +29,25 @@
case s => PIDE.options.value.string("ML_process_policy") = s
}
+ def session_info(): Info =
+ {
+ val logic =
+ Isabelle_System.default_logic(
+ Isabelle_System.getenv("JEDIT_LOGIC"),
+ PIDE.options.string(option_name))
+
+ (for {
+ tree <-
+ try { Some(Sessions.load(session_options(), dirs = session_dirs())) }
+ catch { case ERROR(_) => None }
+ info <- tree.lift(logic)
+ parent <- info.parent
+ if Isabelle_System.getenv("JEDIT_LOGIC_ROOT") == "true"
+ } yield Info(parent, info.pos)) getOrElse Info(logic, Position.none)
+ }
+
+ def session_name(): String = session_info().name
+
def session_build_mode(): String = Isabelle_System.getenv("JEDIT_BUILD_MODE")
def session_build(progress: Progress = Ignore_Progress, no_build: Boolean = false): Int =
@@ -66,7 +82,10 @@
def session_content(inlined_files: Boolean): Build.Session_Content =
{
val content =
- Build.session_content(PIDE.options.value, inlined_files, session_dirs(), session_name())
+ try {
+ Build.session_content(PIDE.options.value, inlined_files, session_dirs(), session_name())
+ }
+ catch { case ERROR(_) => Build.Session_Content() }
content.copy(known_theories =
content.known_theories.mapValues(name => name.map(File.platform_path(_))))
}
--- a/src/Tools/jEdit/src/plugin.scala Tue Dec 20 16:17:13 2016 +0100
+++ b/src/Tools/jEdit/src/plugin.scala Tue Dec 20 16:18:56 2016 +0100
@@ -333,9 +333,14 @@
"It is for testing only, not for production use.")
}
- Session_Build.session_build(jEdit.getActiveView())
+ val view = jEdit.getActiveView()
+
+ Session_Build.session_build(view)
- Keymap_Merge.check_dialog(jEdit.getActiveView())
+ Keymap_Merge.check_dialog(view)
+
+ PIDE.editor.hyperlink_position(true, Document.Snapshot.init,
+ JEdit_Sessions.session_info().open_root).foreach(_.follow(view))
case msg: BufferUpdate
if msg.getWhat == BufferUpdate.LOADED ||